--- a/ESOP-Paper/Appendix.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,135 +0,0 @@
-(*<*)
-theory Appendix
-imports "../Nominal/Nominal2" "~~/src/HOL/Library/LaTeXsugar"
-begin
-
-consts
- fv :: "'a \<Rightarrow> 'b"
- abs_set :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
- alpha_bn :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
- abs_set2 :: "'a \<Rightarrow> perm \<Rightarrow> 'b \<Rightarrow> 'c"
- Abs_dist :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
- Abs_print :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
-
-definition
- "equal \<equiv> (op =)"
-
-notation (latex output)
- swap ("'(_ _')" [1000, 1000] 1000) and
- fresh ("_ # _" [51, 51] 50) and
- fresh_star ("_ #\<^sup>* _" [51, 51] 50) and
- supp ("supp _" [78] 73) and
- uminus ("-_" [78] 73) and
- If ("if _ then _ else _" 10) and
- alpha_set ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{set}}$}}>\<^bsup>_, _, _\<^esup> _") and
- alpha_lst ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{list}}$}}>\<^bsup>_, _, _\<^esup> _") and
- alpha_res ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{res}}$}}>\<^bsup>_, _, _\<^esup> _") and
- abs_set ("_ \<approx>\<^raw:{$\,_{\textit{abs\_set}}$}> _") and
- abs_set2 ("_ \<approx>\<^raw:\raisebox{-1pt}{\makebox[0mm][l]{$\,_{\textit{list}}$}}>\<^bsup>_\<^esup> _") and
- fv ("fa'(_')" [100] 100) and
- equal ("=") and
- alpha_abs_set ("_ \<approx>\<^raw:{$\,_{\textit{abs\_set}}$}> _") and
- Abs_set ("[_]\<^bsub>set\<^esub>._" [20, 101] 999) and
- Abs_lst ("[_]\<^bsub>list\<^esub>._") and
- Abs_dist ("[_]\<^bsub>#list\<^esub>._") and
- Abs_res ("[_]\<^bsub>res\<^esub>._") and
- Abs_print ("_\<^bsub>set\<^esub>._") and
- Cons ("_::_" [78,77] 73) and
- supp_set ("aux _" [1000] 10) and
- alpha_bn ("_ \<approx>bn _")
-
-consts alpha_trm ::'a
-consts fa_trm :: 'a
-consts alpha_trm2 ::'a
-consts fa_trm2 :: 'a
-consts ast :: 'a
-consts ast' :: 'a
-notation (latex output)
- alpha_trm ("\<approx>\<^bsub>trm\<^esub>") and
- fa_trm ("fa\<^bsub>trm\<^esub>") and
- alpha_trm2 ("'(\<approx>\<^bsub>assn\<^esub>, \<approx>\<^bsub>trm\<^esub>')") and
- fa_trm2 ("'(fa\<^bsub>assn\<^esub>, fa\<^bsub>trm\<^esub>')") and
- ast ("'(as, t')") and
- ast' ("'(as', t\<PRIME> ')")
-
-(*>*)
-
-text {*
-\appendix
-\section*{Appendix}
-
- Details for one case in Theorem \ref{suppabs}, which the reader might like to ignore.
- By definition of the abstraction type @{text "abs_set"}
- we have
- %
- \begin{equation}\label{abseqiff}
- @{thm (lhs) Abs_eq_iff(1)[where bs="as" and cs="bs", no_vars]} \;\;\text{if and only if}\;\;
- @{thm (rhs) Abs_eq_iff(1)[where bs="as" and cs="bs", no_vars]}
- \end{equation}
-
- \noindent
- and also
-
- \begin{equation}\label{absperm}
- @{thm permute_Abs(1)[no_vars]}%
- \end{equation}
-
- \noindent
- The second fact derives from the definition of permutations acting on pairs
- and $\alpha$-equivalence being equivariant. With these two facts at our disposal, we can show
- the following lemma about swapping two atoms in an abstraction.
-
- \begin{lemma}
- @{thm[mode=IfThen] Abs_swap1(1)[where bs="as", no_vars]}
- \end{lemma}
-
- \begin{proof}
- This lemma is straightforward using \eqref{abseqiff} and observing that
- the assumptions give us @{term "(a \<rightleftharpoons> b) \<bullet> (supp x - as) = (supp x - as)"}.
- Moreover @{text supp} and set difference are equivariant (see \cite{HuffmanUrban10}).
- \end{proof}
-
- \noindent
- Assuming that @{text "x"} has finite support, this lemma together
- with \eqref{absperm} allows us to show
-
- \begin{equation}\label{halfone}
- @{thm Abs_supports(1)[no_vars]}
- \end{equation}
-
- \noindent
- which gives us ``one half'' of
- Theorem~\ref{suppabs} (the notion of supports is defined in \cite{HuffmanUrban10}).
- The ``other half'' is a bit more involved. To establish
- it, we use a trick from \cite{Pitts04} and first define an auxiliary
- function @{text aux}, taking an abstraction as argument:
- @{thm supp_set.simps[THEN eq_reflection, no_vars]}.
-
- We can show that
- @{text "aux"} is equivariant (since @{term "p \<bullet> (supp x - as) = (supp (p \<bullet> x)) - (p \<bullet> as)"})
- and therefore has empty support.
- This in turn means
-
- \begin{center}
- @{text "supp (aux ([as]\<^bsub>set\<^esub>. x)) \<subseteq> supp ([as]\<^bsub>set\<^esub> x)"}
- \end{center}
-
- \noindent
- Assuming @{term "supp x - as"} is a finite set,
- we further obtain
-
- \begin{equation}\label{halftwo}
- @{thm (concl) Abs_supp_subset1(1)[no_vars]}
- \end{equation}
-
- \noindent
- since for finite sets of atoms, @{text "bs"}, we have
- @{thm (concl) supp_finite_atom_set[where S="bs", no_vars]}.
- Finally, taking \eqref{halfone} and \eqref{halftwo} together establishes
- Theorem~\ref{suppabs}.
-
-*}
-
-(*<*)
-end
-(*>*)
--- a/ESOP-Paper/Paper.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2393 +0,0 @@
-(*<*)
-theory Paper
-imports "../Nominal/Nominal2"
- "~~/src/HOL/Library/LaTeXsugar"
-begin
-
-consts
- fv :: "'a \<Rightarrow> 'b"
- abs_set :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
- alpha_bn :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
- abs_set2 :: "'a \<Rightarrow> perm \<Rightarrow> 'b \<Rightarrow> 'c"
- Abs_dist :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
- Abs_print :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
-
-definition
- "equal \<equiv> (op =)"
-
-notation (latex output)
- swap ("'(_ _')" [1000, 1000] 1000) and
- fresh ("_ # _" [51, 51] 50) and
- fresh_star ("_ #\<^sup>* _" [51, 51] 50) and
- supp ("supp _" [78] 73) and
- uminus ("-_" [78] 73) and
- If ("if _ then _ else _" 10) and
- alpha_set ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{set}}$}}>\<^bsup>_, _, _\<^esup> _") and
- alpha_lst ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{list}}$}}>\<^bsup>_, _, _\<^esup> _") and
- alpha_res ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{set+}}$}}>\<^bsup>_, _, _\<^esup> _") and
- abs_set ("_ \<approx>\<^raw:{$\,_{\textit{abs\_set}}$}> _") and
- abs_set2 ("_ \<approx>\<^raw:\raisebox{-1pt}{\makebox[0mm][l]{$\,_{\textit{list}}$}}>\<^bsup>_\<^esup> _") and
- fv ("fa'(_')" [100] 100) and
- equal ("=") and
- alpha_abs_set ("_ \<approx>\<^raw:{$\,_{\textit{abs\_set}}$}> _") and
- Abs_set ("[_]\<^bsub>set\<^esub>._" [20, 101] 999) and
- Abs_lst ("[_]\<^bsub>list\<^esub>._") and
- Abs_dist ("[_]\<^bsub>#list\<^esub>._") and
- Abs_res ("[_]\<^bsub>set+\<^esub>._") and
- Abs_print ("_\<^bsub>set\<^esub>._") and
- Cons ("_::_" [78,77] 73) and
- supp_set ("aux _" [1000] 10) and
- alpha_bn ("_ \<approx>bn _")
-
-consts alpha_trm ::'a
-consts fa_trm :: 'a
-consts alpha_trm2 ::'a
-consts fa_trm2 :: 'a
-consts ast :: 'a
-consts ast' :: 'a
-notation (latex output)
- alpha_trm ("\<approx>\<^bsub>trm\<^esub>") and
- fa_trm ("fa\<^bsub>trm\<^esub>") and
- alpha_trm2 ("'(\<approx>\<^bsub>assn\<^esub>, \<approx>\<^bsub>trm\<^esub>')") and
- fa_trm2 ("'(fa\<^bsub>assn\<^esub>, fa\<^bsub>trm\<^esub>')") and
- ast ("'(as, t')") and
- ast' ("'(as', t\<PRIME> ')")
-
-(*>*)
-
-
-section {* Introduction *}
-
-text {*
-
- So far, Nominal Isabelle provided a mechanism for constructing
- $\alpha$-equated terms, for example lambda-terms,
- @{text "t ::= x | t t | \<lambda>x. t"},
- where free and bound variables have names. For such $\alpha$-equated terms,
- Nominal Isabelle derives automatically a reasoning infrastructure that has
- been used successfully in formalisations of an equivalence checking
- algorithm for LF \cite{UrbanCheneyBerghofer08}, Typed
- Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for concurrency
- \cite{BengtsonParow09} and a strong normalisation result for cut-elimination
- in classical logic \cite{UrbanZhu08}. It has also been used by Pollack for
- formalisations in the locally-nameless approach to binding
- \cite{SatoPollack10}.
-
- However, Nominal Isabelle has fared less well in a formalisation of
- the algorithm W \cite{UrbanNipkow09}, where types and type-schemes are,
- respectively, of the form
- %
- \begin{equation}\label{tysch}
- \begin{array}{l}
- @{text "T ::= x | T \<rightarrow> T"}\hspace{9mm}
- @{text "S ::= \<forall>{x\<^isub>1,\<dots>, x\<^isub>n}. T"}
- \end{array}
- \end{equation}
- %
- \noindent
- and the @{text "\<forall>"}-quantification binds a finite (possibly empty) set of
- type-variables. While it is possible to implement this kind of more general
- binders by iterating single binders, this leads to a rather clumsy
- formalisation of W.
- %The need of iterating single binders is also one reason
- %why Nominal Isabelle
- % and similar theorem provers that only provide
- %mechanisms for binding single variables
- %has not fared extremely well with the
- %more advanced tasks in the POPLmark challenge \cite{challenge05}, because
- %also there one would like to bind multiple variables at once.
-
- Binding multiple variables has interesting properties that cannot be captured
- easily by iterating single binders. For example in the case of type-schemes we do not
- want to make a distinction about the order of the bound variables. Therefore
- we would like to regard the first pair of type-schemes as $\alpha$-equivalent,
- but assuming that @{text x}, @{text y} and @{text z} are distinct variables,
- the second pair should \emph{not} be $\alpha$-equivalent:
- %
- \begin{equation}\label{ex1}
- @{text "\<forall>{x, y}. x \<rightarrow> y \<approx>\<^isub>\<alpha> \<forall>{y, x}. y \<rightarrow> x"}\hspace{10mm}
- @{text "\<forall>{x, y}. x \<rightarrow> y \<notapprox>\<^isub>\<alpha> \<forall>{z}. z \<rightarrow> z"}
- \end{equation}
- %
- \noindent
- Moreover, we like to regard type-schemes as $\alpha$-equivalent, if they differ
- only on \emph{vacuous} binders, such as
- %
- \begin{equation}\label{ex3}
- @{text "\<forall>{x}. x \<rightarrow> y \<approx>\<^isub>\<alpha> \<forall>{x, z}. x \<rightarrow> y"}
- \end{equation}
- %
- \noindent
- where @{text z} does not occur freely in the type. In this paper we will
- give a general binding mechanism and associated notion of $\alpha$-equivalence
- that can be used to faithfully represent this kind of binding in Nominal
- Isabelle.
- %The difficulty of finding the right notion for $\alpha$-equivalence
- %can be appreciated in this case by considering that the definition given by
- %Leroy in \cite{Leroy92} is incorrect (it omits a side-condition).
-
- However, the notion of $\alpha$-equivalence that is preserved by vacuous
- binders is not always wanted. For example in terms like
- %
- \begin{equation}\label{one}
- @{text "\<LET> x = 3 \<AND> y = 2 \<IN> x - y \<END>"}
- \end{equation}
-
- \noindent
- we might not care in which order the assignments @{text "x = 3"} and
- \mbox{@{text "y = 2"}} are given, but it would be often unusual to regard
- \eqref{one} as $\alpha$-equivalent with
- %
- \begin{center}
- @{text "\<LET> x = 3 \<AND> y = 2 \<AND> z = foo \<IN> x - y \<END>"}
- \end{center}
- %
- \noindent
- Therefore we will also provide a separate binding mechanism for cases in
- which the order of binders does not matter, but the ``cardinality'' of the
- binders has to agree.
-
- However, we found that this is still not sufficient for dealing with
- language constructs frequently occurring in programming language
- research. For example in @{text "\<LET>"}s containing patterns like
- %
- \begin{equation}\label{two}
- @{text "\<LET> (x, y) = (3, 2) \<IN> x - y \<END>"}
- \end{equation}
- %
- \noindent
- we want to bind all variables from the pattern inside the body of the
- $\mathtt{let}$, but we also care about the order of these variables, since
- we do not want to regard \eqref{two} as $\alpha$-equivalent with
- %
- \begin{center}
- @{text "\<LET> (y, x) = (3, 2) \<IN> x - y \<END>"}
- \end{center}
- %
- \noindent
- As a result, we provide three general binding mechanisms each of which binds
- multiple variables at once, and let the user chose which one is intended
- in a formalisation.
- %%when formalising a term-calculus.
-
- By providing these general binding mechanisms, however, we have to work
- around a problem that has been pointed out by Pottier \cite{Pottier06} and
- Cheney \cite{Cheney05}: in @{text "\<LET>"}-constructs of the form
- %
- \begin{center}
- @{text "\<LET> x\<^isub>1 = t\<^isub>1 \<AND> \<dots> \<AND> x\<^isub>n = t\<^isub>n \<IN> s \<END>"}
- \end{center}
- %
- \noindent
- we care about the
- information that there are as many bound variables @{text
- "x\<^isub>i"} as there are @{text "t\<^isub>i"}. We lose this information if
- we represent the @{text "\<LET>"}-constructor by something like
- %
- \begin{center}
- @{text "\<LET> (\<lambda>x\<^isub>1\<dots>x\<^isub>n . s) [t\<^isub>1,\<dots>,t\<^isub>n]"}
- \end{center}
- %
- \noindent
- where the notation @{text "\<lambda>_ . _"} indicates that the list of @{text
- "x\<^isub>i"} becomes bound in @{text s}. In this representation the term
- \mbox{@{text "\<LET> (\<lambda>x . s) [t\<^isub>1, t\<^isub>2]"}} is a perfectly legal
- instance, but the lengths of the two lists do not agree. To exclude such
- terms, additional predicates about well-formed terms are needed in order to
- ensure that the two lists are of equal length. This can result in very messy
- reasoning (see for example~\cite{BengtsonParow09}). To avoid this, we will
- allow type specifications for @{text "\<LET>"}s as follows
- %
- \begin{center}
- \begin{tabular}{r@ {\hspace{2mm}}r@ {\hspace{2mm}}cl}
- @{text trm} & @{text "::="} & @{text "\<dots>"}
- & @{text "|"} @{text "\<LET> as::assn s::trm"}\hspace{2mm}
- \isacommand{bind} @{text "bn(as)"} \isacommand{in} @{text "s"}\\%%%[1mm]
- @{text assn} & @{text "::="} & @{text "\<ANIL>"}
- & @{text "|"} @{text "\<ACONS> name trm assn"}
- \end{tabular}
- \end{center}
- %
- \noindent
- where @{text assn} is an auxiliary type representing a list of assignments
- and @{text bn} an auxiliary function identifying the variables to be bound
- by the @{text "\<LET>"}. This function can be defined by recursion over @{text
- assn} as follows
- %
- \begin{center}
- @{text "bn(\<ANIL>) ="} @{term "{}"} \hspace{5mm}
- @{text "bn(\<ACONS> x t as) = {x} \<union> bn(as)"}
- \end{center}
- %
- \noindent
- The scope of the binding is indicated by labels given to the types, for
- example @{text "s::trm"}, and a binding clause, in this case
- \isacommand{bind} @{text "bn(as)"} \isacommand{in} @{text "s"}. This binding
- clause states that all the names the function @{text
- "bn(as)"} returns should be bound in @{text s}. This style of specifying terms and bindings is heavily
- inspired by the syntax of the Ott-tool \cite{ott-jfp}.
-
- %Though, Ott
- %has only one binding mode, namely the one where the order of
- %binders matters. Consequently, type-schemes with binding sets
- %of names cannot be modelled in Ott.
-
- However, we will not be able to cope with all specifications that are
- allowed by Ott. One reason is that Ott lets the user specify ``empty''
- types like @{text "t ::= t t | \<lambda>x. t"}
- where no clause for variables is given. Arguably, such specifications make
- some sense in the context of Coq's type theory (which Ott supports), but not
- at all in a HOL-based environment where every datatype must have a non-empty
- set-theoretic model. % \cite{Berghofer99}.
-
- Another reason is that we establish the reasoning infrastructure
- for $\alpha$-\emph{equated} terms. In contrast, Ott produces a reasoning
- infrastructure in Isabelle/HOL for
- \emph{non}-$\alpha$-equated, or ``raw'', terms. While our $\alpha$-equated terms
- and the raw terms produced by Ott use names for bound variables,
- there is a key difference: working with $\alpha$-equated terms means, for example,
- that the two type-schemes
-
- \begin{center}
- @{text "\<forall>{x}. x \<rightarrow> y = \<forall>{x, z}. x \<rightarrow> y"}
- \end{center}
-
- \noindent
- are not just $\alpha$-equal, but actually \emph{equal}! As a result, we can
- only support specifications that make sense on the level of $\alpha$-equated
- terms (offending specifications, which for example bind a variable according
- to a variable bound somewhere else, are not excluded by Ott, but we have
- to).
-
- %Our insistence on reasoning with $\alpha$-equated terms comes from the
- %wealth of experience we gained with the older version of Nominal Isabelle:
- %for non-trivial properties, reasoning with $\alpha$-equated terms is much
- %easier than reasoning with raw terms. The fundamental reason for this is
- %that the HOL-logic underlying Nominal Isabelle allows us to replace
- %``equals-by-equals''. In contrast, replacing
- %``$\alpha$-equals-by-$\alpha$-equals'' in a representation based on raw terms
- %requires a lot of extra reasoning work.
-
- Although in informal settings a reasoning infrastructure for $\alpha$-equated
- terms is nearly always taken for granted, establishing it automatically in
- Isabelle/HOL is a rather non-trivial task. For every
- specification we will need to construct type(s) containing as elements the
- $\alpha$-equated terms. To do so, we use the standard HOL-technique of defining
- a new type by identifying a non-empty subset of an existing type. The
- construction we perform in Isabelle/HOL can be illustrated by the following picture:
- %
- \begin{center}
- \begin{tikzpicture}[scale=0.89]
- %\draw[step=2mm] (-4,-1) grid (4,1);
-
- \draw[very thick] (0.7,0.4) circle (4.25mm);
- \draw[rounded corners=1mm, very thick] ( 0.0,-0.8) rectangle ( 1.8, 0.9);
- \draw[rounded corners=1mm, very thick] (-1.95,0.85) rectangle (-2.85,-0.05);
-
- \draw (-2.0, 0.845) -- (0.7,0.845);
- \draw (-2.0,-0.045) -- (0.7,-0.045);
-
- \draw ( 0.7, 0.4) node {\footnotesize\begin{tabular}{@ {}c@ {}}$\alpha$-\\[-1mm]clas.\end{tabular}};
- \draw (-2.4, 0.4) node {\footnotesize\begin{tabular}{@ {}c@ {}}$\alpha$-eq.\\[-1mm]terms\end{tabular}};
- \draw (1.8, 0.48) node[right=-0.1mm]
- {\footnotesize\begin{tabular}{@ {}l@ {}}existing\\[-1mm] type\\ (sets of raw terms)\end{tabular}};
- \draw (0.9, -0.35) node {\footnotesize\begin{tabular}{@ {}l@ {}}non-empty\\[-1mm]subset\end{tabular}};
- \draw (-3.25, 0.55) node {\footnotesize\begin{tabular}{@ {}l@ {}}new\\[-1mm]type\end{tabular}};
-
- \draw[<->, very thick] (-1.8, 0.3) -- (-0.1,0.3);
- \draw (-0.95, 0.3) node[above=0mm] {\footnotesize{}isomorphism};
-
- \end{tikzpicture}
- \end{center}
- %
- \noindent
- We take as the starting point a definition of raw terms (defined as a
- datatype in Isabelle/HOL); then identify the $\alpha$-equivalence classes in
- the type of sets of raw terms according to our $\alpha$-equivalence relation,
- and finally define the new type as these $\alpha$-equivalence classes
- (non-emptiness is satisfied whenever the raw terms are definable as datatype
- in Isabelle/HOL and our relation for $\alpha$-equivalence is
- an equivalence relation).
-
- %The fact that we obtain an isomorphism between the new type and the
- %non-empty subset shows that the new type is a faithful representation of
- %$\alpha$-equated terms. That is not the case for example for terms using the
- %locally nameless representation of binders \cite{McKinnaPollack99}: in this
- %representation there are ``junk'' terms that need to be excluded by
- %reasoning about a well-formedness predicate.
-
- The problem with introducing a new type in Isabelle/HOL is that in order to
- be useful, a reasoning infrastructure needs to be ``lifted'' from the
- underlying subset to the new type. This is usually a tricky and arduous
- task. To ease it, we re-implemented in Isabelle/HOL \cite{KaliszykUrban11} the quotient package
- described by Homeier \cite{Homeier05} for the HOL4 system. This package
- allows us to lift definitions and theorems involving raw terms to
- definitions and theorems involving $\alpha$-equated terms. For example if we
- define the free-variable function over raw lambda-terms
-
- \begin{center}
- @{text "fv(x) = {x}"}\hspace{8mm}
- @{text "fv(t\<^isub>1 t\<^isub>2) = fv(t\<^isub>1) \<union> fv(t\<^isub>2)"}\hspace{8mm}
- @{text "fv(\<lambda>x.t) = fv(t) - {x}"}
- \end{center}
-
- \noindent
- then with the help of the quotient package we can obtain a function @{text "fv\<^sup>\<alpha>"}
- operating on quotients, or $\alpha$-equivalence classes of lambda-terms. This
- lifted function is characterised by the equations
-
- \begin{center}
- @{text "fv\<^sup>\<alpha>(x) = {x}"}\hspace{8mm}
- @{text "fv\<^sup>\<alpha>(t\<^isub>1 t\<^isub>2) = fv\<^sup>\<alpha>(t\<^isub>1) \<union> fv\<^sup>\<alpha>(t\<^isub>2)"}\hspace{8mm}
- @{text "fv\<^sup>\<alpha>(\<lambda>x.t) = fv\<^sup>\<alpha>(t) - {x}"}
- \end{center}
-
- \noindent
- (Note that this means also the term-constructors for variables, applications
- and lambda are lifted to the quotient level.) This construction, of course,
- only works if $\alpha$-equivalence is indeed an equivalence relation, and the
- ``raw'' definitions and theorems are respectful w.r.t.~$\alpha$-equivalence.
- %For example, we will not be able to lift a bound-variable function. Although
- %this function can be defined for raw terms, it does not respect
- %$\alpha$-equivalence and therefore cannot be lifted.
- To sum up, every lifting
- of theorems to the quotient level needs proofs of some respectfulness
- properties (see \cite{Homeier05}). In the paper we show that we are able to
- automate these proofs and as a result can automatically establish a reasoning
- infrastructure for $\alpha$-equated terms.\smallskip
-
- %The examples we have in mind where our reasoning infrastructure will be
- %helpful includes the term language of Core-Haskell. This term language
- %involves patterns that have lists of type-, coercion- and term-variables,
- %all of which are bound in @{text "\<CASE>"}-expressions. In these
- %patterns we do not know in advance how many variables need to
- %be bound. Another example is the specification of SML, which includes
- %includes bindings as in type-schemes.\medskip
-
- \noindent
- {\bf Contributions:} We provide three new definitions for when terms
- involving general binders are $\alpha$-equivalent. These definitions are
- inspired by earlier work of Pitts \cite{Pitts04}. By means of automatic
- proofs, we establish a reasoning infrastructure for $\alpha$-equated
- terms, including properties about support, freshness and equality
- conditions for $\alpha$-equated terms. We are also able to derive strong
- induction principles that have the variable convention already built in.
- The method behind our specification of general binders is taken
- from the Ott-tool, but we introduce crucial restrictions, and also extensions, so
- that our specifications make sense for reasoning about $\alpha$-equated terms.
- The main improvement over Ott is that we introduce three binding modes
- (only one is present in Ott), provide formalised definitions for $\alpha$-equivalence and
- for free variables of our terms, and also derive a reasoning infrastructure
- for our specifications from ``first principles''.
-
-
- %\begin{figure}
- %\begin{boxedminipage}{\linewidth}
- %%\begin{center}
- %\begin{tabular}{r@ {\hspace{1mm}}r@ {\hspace{2mm}}l}
- %\multicolumn{3}{@ {}l}{Type Kinds}\\
- %@{text "\<kappa>"} & @{text "::="} & @{text "\<star> | \<kappa>\<^isub>1 \<rightarrow> \<kappa>\<^isub>2"}\smallskip\\
- %\multicolumn{3}{@ {}l}{Coercion Kinds}\\
- %@{text "\<iota>"} & @{text "::="} & @{text "\<sigma>\<^isub>1 \<sim> \<sigma>\<^isub>2"}\smallskip\\
- %\multicolumn{3}{@ {}l}{Types}\\
- %@{text "\<sigma>"} & @{text "::="} & @{text "a | T | \<sigma>\<^isub>1 \<sigma>\<^isub>2 | S\<^isub>n"}$\;\overline{@{text "\<sigma>"}}$@{text "\<^sup>n"}
- %@{text "| \<forall>a:\<kappa>. \<sigma> | \<iota> \<Rightarrow> \<sigma>"}\smallskip\\
- %\multicolumn{3}{@ {}l}{Coercion Types}\\
- %@{text "\<gamma>"} & @{text "::="} & @{text "c | C | \<gamma>\<^isub>1 \<gamma>\<^isub>2 | S\<^isub>n"}$\;\overline{@{text "\<gamma>"}}$@{text "\<^sup>n"}
- %@{text "| \<forall>c:\<iota>. \<gamma> | \<iota> \<Rightarrow> \<gamma> "}\\
- %& @{text "|"} & @{text "refl \<sigma> | sym \<gamma> | \<gamma>\<^isub>1 \<circ> \<gamma>\<^isub>2 | \<gamma> @ \<sigma> | left \<gamma> | right \<gamma>"}\\
- %& @{text "|"} & @{text "\<gamma>\<^isub>1 \<sim> \<gamma>\<^isub>2 | rightc \<gamma> | leftc \<gamma> | \<gamma>\<^isub>1 \<triangleright> \<gamma>\<^isub>2"}\smallskip\\
- %\multicolumn{3}{@ {}l}{Terms}\\
- %@{text "e"} & @{text "::="} & @{text "x | K | \<Lambda>a:\<kappa>. e | \<Lambda>c:\<iota>. e | e \<sigma> | e \<gamma>"}\\
- %& @{text "|"} & @{text "\<lambda>x:\<sigma>. e | e\<^isub>1 e\<^isub>2 | \<LET> x:\<sigma> = e\<^isub>1 \<IN> e\<^isub>2"}\\
- %& @{text "|"} & @{text "\<CASE> e\<^isub>1 \<OF>"}$\;\overline{@{text "p \<rightarrow> e\<^isub>2"}}$ @{text "| e \<triangleright> \<gamma>"}\smallskip\\
- %\multicolumn{3}{@ {}l}{Patterns}\\
- %@{text "p"} & @{text "::="} & @{text "K"}$\;\overline{@{text "a:\<kappa>"}}\;\overline{@{text "c:\<iota>"}}\;\overline{@{text "x:\<sigma>"}}$\smallskip\\
- %\multicolumn{3}{@ {}l}{Constants}\\
- %& @{text C} & coercion constants\\
- %& @{text T} & value type constructors\\
- %& @{text "S\<^isub>n"} & n-ary type functions (which need to be fully applied)\\
- %& @{text K} & data constructors\smallskip\\
- %\multicolumn{3}{@ {}l}{Variables}\\
- %& @{text a} & type variables\\
- %& @{text c} & coercion variables\\
- %& @{text x} & term variables\\
- %\end{tabular}
- %\end{center}
- %\end{boxedminipage}
- %\caption{The System @{text "F\<^isub>C"}
- %\cite{CoreHaskell}, also often referred to as \emph{Core-Haskell}. In this
- %version of @{text "F\<^isub>C"} we made a modification by separating the
- %grammars for type kinds and coercion kinds, as well as for types and coercion
- %types. For this paper the interesting term-constructor is @{text "\<CASE>"},
- %which binds multiple type-, coercion- and term-variables.\label{corehas}}
- %\end{figure}
-*}
-
-section {* A Short Review of the Nominal Logic Work *}
-
-text {*
- At its core, Nominal Isabelle is an adaption of the nominal logic work by
- Pitts \cite{Pitts03}. This adaptation for Isabelle/HOL is described in
- \cite{HuffmanUrban10} (including proofs). We shall briefly review this work
- to aid the description of what follows.
-
- Two central notions in the nominal logic work are sorted atoms and
- sort-respecting permutations of atoms. We will use the letters @{text "a,
- b, c, \<dots>"} to stand for atoms and @{text "p, q, \<dots>"} to stand for
- permutations. The purpose of atoms is to represent variables, be they bound or free.
- %The sorts of atoms can be used to represent different kinds of
- %variables, such as the term-, coercion- and type-variables in Core-Haskell.
- It is assumed that there is an infinite supply of atoms for each
- sort. In the interest of brevity, we shall restrict ourselves
- in what follows to only one sort of atoms.
-
- Permutations are bijective functions from atoms to atoms that are
- the identity everywhere except on a finite number of atoms. There is a
- two-place permutation operation written
- @{text "_ \<bullet> _ :: perm \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
- where the generic type @{text "\<beta>"} is the type of the object
- over which the permutation
- acts. In Nominal Isabelle, the identity permutation is written as @{term "0::perm"},
- the composition of two permutations @{term p} and @{term q} as \mbox{@{term "p + q"}},
- and the inverse permutation of @{term p} as @{text "- p"}. The permutation
- operation is defined over the type-hierarchy \cite{HuffmanUrban10};
- for example permutations acting on products, lists, sets, functions and booleans are
- given by:
- %
- %\begin{equation}\label{permute}
- %\mbox{\begin{tabular}{@ {}c@ {\hspace{10mm}}c@ {}}
- %\begin{tabular}{@ {}l@ {}}
- %@{thm permute_prod.simps[no_vars, THEN eq_reflection]}\\[2mm]
- %@{thm permute_list.simps(1)[no_vars, THEN eq_reflection]}\\
- %@{thm permute_list.simps(2)[no_vars, THEN eq_reflection]}\\
- %\end{tabular} &
- %\begin{tabular}{@ {}l@ {}}
- %@{thm permute_set_eq[no_vars, THEN eq_reflection]}\\
- %@{text "p \<bullet> f \<equiv> \<lambda>x. p \<bullet> (f (- p \<bullet> x))"}\\
- %@{thm permute_bool_def[no_vars, THEN eq_reflection]}
- %\end{tabular}
- %\end{tabular}}
- %\end{equation}
- %
- \begin{center}
- \mbox{\begin{tabular}{@ {}c@ {\hspace{4mm}}c@ {\hspace{4mm}}c@ {}}
- \begin{tabular}{@ {}l@ {}}
- @{thm permute_prod.simps[no_vars, THEN eq_reflection]}\\
- @{thm permute_bool_def[no_vars, THEN eq_reflection]}
- \end{tabular} &
- \begin{tabular}{@ {}l@ {}}
- @{thm permute_list.simps(1)[no_vars, THEN eq_reflection]}\\
- @{thm permute_list.simps(2)[no_vars, THEN eq_reflection]}\\
- \end{tabular} &
- \begin{tabular}{@ {}l@ {}}
- @{thm permute_set_eq[no_vars, THEN eq_reflection]}\\
- @{text "p \<bullet> f \<equiv> \<lambda>x. p \<bullet> (f (- p \<bullet> x))"}\\
- \end{tabular}
- \end{tabular}}
- \end{center}
-
- \noindent
- Concrete permutations in Nominal Isabelle are built up from swappings,
- written as \mbox{@{text "(a b)"}}, which are permutations that behave
- as follows:
- %
- \begin{center}
- @{text "(a b) = \<lambda>c. if a = c then b else if b = c then a else c"}
- \end{center}
-
- The most original aspect of the nominal logic work of Pitts is a general
- definition for the notion of the ``set of free variables of an object @{text
- "x"}''. This notion, written @{term "supp x"}, is general in the sense that
- it applies not only to lambda-terms ($\alpha$-equated or not), but also to lists,
- products, sets and even functions. The definition depends only on the
- permutation operation and on the notion of equality defined for the type of
- @{text x}, namely:
- %
- \begin{equation}\label{suppdef}
- @{thm supp_def[no_vars, THEN eq_reflection]}
- \end{equation}
-
- \noindent
- There is also the derived notion for when an atom @{text a} is \emph{fresh}
- for an @{text x}, defined as @{thm fresh_def[no_vars]}.
- We use for sets of atoms the abbreviation
- @{thm (lhs) fresh_star_def[no_vars]}, defined as
- @{thm (rhs) fresh_star_def[no_vars]}.
- A striking consequence of these definitions is that we can prove
- without knowing anything about the structure of @{term x} that
- swapping two fresh atoms, say @{text a} and @{text b}, leaves
- @{text x} unchanged, namely if @{text "a \<FRESH> x"} and @{text "b \<FRESH> x"}
- then @{term "(a \<rightleftharpoons> b) \<bullet> x = x"}.
- %
- %\begin{myproperty}\label{swapfreshfresh}
- %@{thm[mode=IfThen] swap_fresh_fresh[no_vars]}
- %\end{myproperty}
- %
- %While often the support of an object can be relatively easily
- %described, for example for atoms, products, lists, function applications,
- %booleans and permutations as follows
- %%
- %\begin{center}
- %\begin{tabular}{c@ {\hspace{10mm}}c}
- %\begin{tabular}{rcl}
- %@{term "supp a"} & $=$ & @{term "{a}"}\\
- %@{term "supp (x, y)"} & $=$ & @{term "supp x \<union> supp y"}\\
- %@{term "supp []"} & $=$ & @{term "{}"}\\
- %@{term "supp (x#xs)"} & $=$ & @{term "supp x \<union> supp xs"}\\
- %\end{tabular}
- %&
- %\begin{tabular}{rcl}
- %@{text "supp (f x)"} & @{text "\<subseteq>"} & @{term "supp f \<union> supp x"}\\
- %@{term "supp b"} & $=$ & @{term "{}"}\\
- %@{term "supp p"} & $=$ & @{term "{a. p \<bullet> a \<noteq> a}"}
- %\end{tabular}
- %\end{tabular}
- %\end{center}
- %
- %\noindent
- %in some cases it can be difficult to characterise the support precisely, and
- %only an approximation can be established (as for functions above).
- %
- %Reasoning about
- %such approximations can be simplified with the notion \emph{supports}, defined
- %as follows:
- %
- %\begin{definition}
- %A set @{text S} \emph{supports} @{text x} if for all atoms @{text a} and @{text b}
- %not in @{text S} we have @{term "(a \<rightleftharpoons> b) \<bullet> x = x"}.
- %\end{definition}
- %
- %\noindent
- %The main point of @{text supports} is that we can establish the following
- %two properties.
- %
- %\begin{myproperty}\label{supportsprop}
- %Given a set @{text "as"} of atoms.
- %{\it (i)} @{thm[mode=IfThen] supp_is_subset[where S="as", no_vars]}
- %{\it (ii)} @{thm supp_supports[no_vars]}.
- %\end{myproperty}
- %
- %Another important notion in the nominal logic work is \emph{equivariance}.
- %For a function @{text f}, say of type @{text "\<alpha> \<Rightarrow> \<beta>"}, to be equivariant
- %it is required that every permutation leaves @{text f} unchanged, that is
- %%
- %\begin{equation}\label{equivariancedef}
- %@{term "\<forall>p. p \<bullet> f = f"}
- %\end{equation}
- %
- %\noindent or equivalently that a permutation applied to the application
- %@{text "f x"} can be moved to the argument @{text x}. That means for equivariant
- %functions @{text f}, we have for all permutations @{text p}:
- %%
- %\begin{equation}\label{equivariance}
- %@{text "p \<bullet> f = f"} \;\;\;\textit{if and only if}\;\;\;
- %@{text "p \<bullet> (f x) = f (p \<bullet> x)"}
- %\end{equation}
- %
- %\noindent
- %From property \eqref{equivariancedef} and the definition of @{text supp}, we
- %can easily deduce that equivariant functions have empty support. There is
- %also a similar notion for equivariant relations, say @{text R}, namely the property
- %that
- %%
- %\begin{center}
- %@{text "x R y"} \;\;\textit{implies}\;\; @{text "(p \<bullet> x) R (p \<bullet> y)"}
- %\end{center}
- %
- %Using freshness, the nominal logic work provides us with general means for renaming
- %binders.
- %
- %\noindent
- While in the older version of Nominal Isabelle, we used extensively
- %Property~\ref{swapfreshfresh}
- this property to rename single binders, it %%this property
- proved too unwieldy for dealing with multiple binders. For such binders the
- following generalisations turned out to be easier to use.
-
- \begin{myproperty}\label{supppermeq}
- @{thm[mode=IfThen] supp_perm_eq[no_vars]}
- \end{myproperty}
-
- \begin{myproperty}\label{avoiding}
- For a finite set @{text as} and a finitely supported @{text x} with
- @{term "as \<sharp>* x"} and also a finitely supported @{text c}, there
- exists a permutation @{text p} such that @{term "(p \<bullet> as) \<sharp>* c"} and
- @{term "supp x \<sharp>* p"}.
- \end{myproperty}
-
- \noindent
- The idea behind the second property is that given a finite set @{text as}
- of binders (being bound, or fresh, in @{text x} is ensured by the
- assumption @{term "as \<sharp>* x"}), then there exists a permutation @{text p} such that
- the renamed binders @{term "p \<bullet> as"} avoid @{text c} (which can be arbitrarily chosen
- as long as it is finitely supported) and also @{text "p"} does not affect anything
- in the support of @{text x} (that is @{term "supp x \<sharp>* p"}). The last
- fact and Property~\ref{supppermeq} allow us to ``rename'' just the binders
- @{text as} in @{text x}, because @{term "p \<bullet> x = x"}.
-
- Most properties given in this section are described in detail in \cite{HuffmanUrban10}
- and all are formalised in Isabelle/HOL. In the next sections we will make
- extensive use of these properties in order to define $\alpha$-equivalence in
- the presence of multiple binders.
-*}
-
-
-section {* General Bindings\label{sec:binders} *}
-
-text {*
- In Nominal Isabelle, the user is expected to write down a specification of a
- term-calculus and then a reasoning infrastructure is automatically derived
- from this specification (remember that Nominal Isabelle is a definitional
- extension of Isabelle/HOL, which does not introduce any new axioms).
-
- In order to keep our work with deriving the reasoning infrastructure
- manageable, we will wherever possible state definitions and perform proofs
- on the ``user-level'' of Isabelle/HOL, as opposed to write custom ML-code. % that
- %generates them anew for each specification.
- To that end, we will consider
- first pairs @{text "(as, x)"} of type @{text "(atom set) \<times> \<beta>"}. These pairs
- are intended to represent the abstraction, or binding, of the set of atoms @{text
- "as"} in the body @{text "x"}.
-
- The first question we have to answer is when two pairs @{text "(as, x)"} and
- @{text "(bs, y)"} are $\alpha$-equivalent? (For the moment we are interested in
- the notion of $\alpha$-equivalence that is \emph{not} preserved by adding
- vacuous binders.) To answer this question, we identify four conditions: {\it (i)}
- given a free-atom function @{text "fa"} of type \mbox{@{text "\<beta> \<Rightarrow> atom
- set"}}, then @{text x} and @{text y} need to have the same set of free
- atoms; moreover there must be a permutation @{text p} such that {\it
- (ii)} @{text p} leaves the free atoms of @{text x} and @{text y} unchanged, but
- {\it (iii)} ``moves'' their bound names so that we obtain modulo a relation,
- say \mbox{@{text "_ R _"}}, two equivalent terms. We also require that {\it (iv)}
- @{text p} makes the sets of abstracted atoms @{text as} and @{text bs} equal. The
- requirements {\it (i)} to {\it (iv)} can be stated formally as the conjunction of:
- %
- \begin{equation}\label{alphaset}
- \begin{array}{@ {\hspace{10mm}}l@ {\hspace{5mm}}l@ {\hspace{10mm}}l@ {\hspace{5mm}}l}
- \multicolumn{4}{l}{@{term "(as, x) \<approx>set R fa p (bs, y)"}\hspace{2mm}@{text "\<equiv>"}}\\[1mm]
- \mbox{\it (i)} & @{term "fa(x) - as = fa(y) - bs"} &
- \mbox{\it (iii)} & @{text "(p \<bullet> x) R y"} \\
- \mbox{\it (ii)} & @{term "(fa(x) - as) \<sharp>* p"} &
- \mbox{\it (iv)} & @{term "(p \<bullet> as) = bs"} \\
- \end{array}
- \end{equation}
- %
- \noindent
- Note that this relation depends on the permutation @{text
- "p"}; $\alpha$-equivalence between two pairs is then the relation where we
- existentially quantify over this @{text "p"}. Also note that the relation is
- dependent on a free-atom function @{text "fa"} and a relation @{text
- "R"}. The reason for this extra generality is that we will use
- $\approx_{\,\textit{set}}$ for both ``raw'' terms and $\alpha$-equated terms. In
- the latter case, @{text R} will be replaced by equality @{text "="} and we
- will prove that @{text "fa"} is equal to @{text "supp"}.
-
- The definition in \eqref{alphaset} does not make any distinction between the
- order of abstracted atoms. If we want this, then we can define $\alpha$-equivalence
- for pairs of the form \mbox{@{text "(as, x)"}} with type @{text "(atom list) \<times> \<beta>"}
- as follows
- %
- \begin{equation}\label{alphalist}
- \begin{array}{@ {\hspace{10mm}}l@ {\hspace{5mm}}l@ {\hspace{10mm}}l@ {\hspace{5mm}}l}
- \multicolumn{4}{l}{@{term "(as, x) \<approx>lst R fa p (bs, y)"}\hspace{2mm}@{text "\<equiv>"}}\\[1mm]
- \mbox{\it (i)} & @{term "fa(x) - (set as) = fa(y) - (set bs)"} &
- \mbox{\it (iii)} & @{text "(p \<bullet> x) R y"}\\
- \mbox{\it (ii)} & @{term "(fa(x) - set as) \<sharp>* p"} &
- \mbox{\it (iv)} & @{term "(p \<bullet> as) = bs"}\\
- \end{array}
- \end{equation}
- %
- \noindent
- where @{term set} is the function that coerces a list of atoms into a set of atoms.
- Now the last clause ensures that the order of the binders matters (since @{text as}
- and @{text bs} are lists of atoms).
-
- If we do not want to make any difference between the order of binders \emph{and}
- also allow vacuous binders, that means \emph{restrict} names, then we keep sets of binders, but drop
- condition {\it (iv)} in \eqref{alphaset}:
- %
- \begin{equation}\label{alphares}
- \begin{array}{@ {\hspace{10mm}}l@ {\hspace{5mm}}l@ {\hspace{10mm}}l@ {\hspace{5mm}}l}
- \multicolumn{2}{l}{@{term "(as, x) \<approx>res R fa p (bs, y)"}\hspace{2mm}@{text "\<equiv>"}}\\[1mm]
- \mbox{\it (i)} & @{term "fa(x) - as = fa(y) - bs"} &
- \mbox{\it (iii)} & @{text "(p \<bullet> x) R y"}\\
- \mbox{\it (ii)} & @{term "(fa(x) - as) \<sharp>* p"}\\
- \end{array}
- \end{equation}
-
- It might be useful to consider first some examples how these definitions
- of $\alpha$-equivalence pan out in practice. For this consider the case of
- abstracting a set of atoms over types (as in type-schemes). We set
- @{text R} to be the usual equality @{text "="} and for @{text "fa(T)"} we
- define
- %
- \begin{center}
- @{text "fa(x) = {x}"} \hspace{5mm} @{text "fa(T\<^isub>1 \<rightarrow> T\<^isub>2) = fa(T\<^isub>1) \<union> fa(T\<^isub>2)"}
- \end{center}
-
- \noindent
- Now recall the examples shown in \eqref{ex1} and
- \eqref{ex3}. It can be easily checked that @{text "({x, y}, x \<rightarrow> y)"} and
- @{text "({y, x}, y \<rightarrow> x)"} are $\alpha$-equivalent according to
- $\approx_{\,\textit{set}}$ and $\approx_{\,\textit{set+}}$ by taking @{text p} to
- be the swapping @{term "(x \<rightleftharpoons> y)"}. In case of @{text "x \<noteq> y"}, then @{text
- "([x, y], x \<rightarrow> y)"} $\not\approx_{\,\textit{list}}$ @{text "([y, x], x \<rightarrow> y)"}
- since there is no permutation that makes the lists @{text "[x, y]"} and
- @{text "[y, x]"} equal, and also leaves the type \mbox{@{text "x \<rightarrow> y"}}
- unchanged. Another example is @{text "({x}, x)"} $\approx_{\,\textit{set+}}$
- @{text "({x, y}, x)"} which holds by taking @{text p} to be the identity
- permutation. However, if @{text "x \<noteq> y"}, then @{text "({x}, x)"}
- $\not\approx_{\,\textit{set}}$ @{text "({x, y}, x)"} since there is no
- permutation that makes the sets @{text "{x}"} and @{text "{x, y}"} equal
- (similarly for $\approx_{\,\textit{list}}$). It can also relatively easily be
- shown that all three notions of $\alpha$-equivalence coincide, if we only
- abstract a single atom.
-
- In the rest of this section we are going to introduce three abstraction
- types. For this we define
- %
- \begin{equation}
- @{term "abs_set (as, x) (bs, x) \<equiv> \<exists>p. alpha_set (as, x) equal supp p (bs, x)"}
- \end{equation}
-
- \noindent
- (similarly for $\approx_{\,\textit{abs\_set+}}$
- and $\approx_{\,\textit{abs\_list}}$). We can show that these relations are equivalence
- relations. %% and equivariant.
-
- \begin{lemma}\label{alphaeq}
- The relations $\approx_{\,\textit{abs\_set}}$, $\approx_{\,\textit{abs\_list}}$
- and $\approx_{\,\textit{abs\_set+}}$ are equivalence relations. %, and if
- %@{term "abs_set (as, x) (bs, y)"} then also
- %@{term "abs_set (p \<bullet> as, p \<bullet> x) (p \<bullet> bs, p \<bullet> y)"} (similarly for the other two relations).
- \end{lemma}
-
- \begin{proof}
- Reflexivity is by taking @{text "p"} to be @{text "0"}. For symmetry we have
- a permutation @{text p} and for the proof obligation take @{term "-p"}. In case
- of transitivity, we have two permutations @{text p} and @{text q}, and for the
- proof obligation use @{text "q + p"}. All conditions are then by simple
- calculations.
- \end{proof}
-
- \noindent
- This lemma allows us to use our quotient package for introducing
- new types @{text "\<beta> abs_set"}, @{text "\<beta> abs_set+"} and @{text "\<beta> abs_list"}
- representing $\alpha$-equivalence classes of pairs of type
- @{text "(atom set) \<times> \<beta>"} (in the first two cases) and of type @{text "(atom list) \<times> \<beta>"}
- (in the third case).
- The elements in these types will be, respectively, written as
- %
- %\begin{center}
- @{term "Abs_set as x"}, %\hspace{5mm}
- @{term "Abs_res as x"} and %\hspace{5mm}
- @{term "Abs_lst as x"},
- %\end{center}
- %
- %\noindent
- indicating that a set (or list) of atoms @{text as} is abstracted in @{text x}. We will
- call the types \emph{abstraction types} and their elements
- \emph{abstractions}. The important property we need to derive is the support of
- abstractions, namely:
-
- \begin{theorem}[Support of Abstractions]\label{suppabs}
- Assuming @{text x} has finite support, then
-
- \begin{center}
- \begin{tabular}{l}
- @{thm (lhs) supp_Abs(1)[no_vars]} $\;\;=\;\;$
- @{thm (lhs) supp_Abs(2)[no_vars]} $\;\;=\;\;$ @{thm (rhs) supp_Abs(2)[no_vars]}, and\\
- @{thm (lhs) supp_Abs(3)[where bs="bs", no_vars]} $\;\;=\;\;$
- @{thm (rhs) supp_Abs(3)[where bs="bs", no_vars]}
- \end{tabular}
- \end{center}
- \end{theorem}
-
- \noindent
- This theorem states that the bound names do not appear in the support.
- For brevity we omit the proof and again refer the reader to
- our formalisation in Isabelle/HOL.
-
- %\noindent
- %Below we will show the first equation. The others
- %follow by similar arguments. By definition of the abstraction type @{text "abs_set"}
- %we have
- %%
- %\begin{equation}\label{abseqiff}
- %@{thm (lhs) Abs_eq_iff(1)[where bs="as" and cs="bs", no_vars]} \;\;\text{if and only if}\;\;
- %@{thm (rhs) Abs_eq_iff(1)[where bs="as" and cs="bs", no_vars]}
- %\end{equation}
- %
- %\noindent
- %and also
- %
- %\begin{equation}\label{absperm}
- %%@%{%thm %permute_Abs[no_vars]}%
- %\end{equation}
-
- %\noindent
- %The second fact derives from the definition of permutations acting on pairs
- %\eqref{permute} and $\alpha$-equivalence being equivariant
- %(see Lemma~\ref{alphaeq}). With these two facts at our disposal, we can show
- %the following lemma about swapping two atoms in an abstraction.
- %
- %\begin{lemma}
- %@{thm[mode=IfThen] Abs_swap1(1)[where bs="as", no_vars]}
- %\end{lemma}
- %
- %\begin{proof}
- %This lemma is straightforward using \eqref{abseqiff} and observing that
- %the assumptions give us @{term "(a \<rightleftharpoons> b) \<bullet> (supp x - as) = (supp x - as)"}.
- %Moreover @{text supp} and set difference are equivariant (see \cite{HuffmanUrban10}).
- %\end{proof}
- %
- %\noindent
- %Assuming that @{text "x"} has finite support, this lemma together
- %with \eqref{absperm} allows us to show
- %
- %\begin{equation}\label{halfone}
- %@{thm Abs_supports(1)[no_vars]}
- %\end{equation}
- %
- %\noindent
- %which by Property~\ref{supportsprop} gives us ``one half'' of
- %Theorem~\ref{suppabs}. The ``other half'' is a bit more involved. To establish
- %it, we use a trick from \cite{Pitts04} and first define an auxiliary
- %function @{text aux}, taking an abstraction as argument:
- %@{thm supp_set.simps[THEN eq_reflection, no_vars]}.
- %
- %Using the second equation in \eqref{equivariance}, we can show that
- %@{text "aux"} is equivariant (since @{term "p \<bullet> (supp x - as) = (supp (p \<bullet> x)) - (p \<bullet> as)"})
- %and therefore has empty support.
- %This in turn means
- %
- %\begin{center}
- %@{term "supp (supp_gen (Abs_set as x)) \<subseteq> supp (Abs_set as x)"}
- %\end{center}
- %
- %\noindent
- %using \eqref{suppfun}. Assuming @{term "supp x - as"} is a finite set,
- %we further obtain
- %
- %\begin{equation}\label{halftwo}
- %@{thm (concl) Abs_supp_subset1(1)[no_vars]}
- %\end{equation}
- %
- %\noindent
- %since for finite sets of atoms, @{text "bs"}, we have
- %@{thm (concl) supp_finite_atom_set[where S="bs", no_vars]}.
- %Finally, taking \eqref{halfone} and \eqref{halftwo} together establishes
- %Theorem~\ref{suppabs}.
-
- The method of first considering abstractions of the
- form @{term "Abs_set as x"} etc is motivated by the fact that
- we can conveniently establish at the Isabelle/HOL level
- properties about them. It would be
- laborious to write custom ML-code that derives automatically such properties
- for every term-constructor that binds some atoms. Also the generality of
- the definitions for $\alpha$-equivalence will help us in the next sections.
-*}
-
-section {* Specifying General Bindings\label{sec:spec} *}
-
-text {*
- Our choice of syntax for specifications is influenced by the existing
- datatype package of Isabelle/HOL %\cite{Berghofer99}
- and by the syntax of the
- Ott-tool \cite{ott-jfp}. For us a specification of a term-calculus is a
- collection of (possibly mutual recursive) type declarations, say @{text
- "ty\<AL>\<^isub>1, \<dots>, ty\<AL>\<^isub>n"}, and an associated collection of
- binding functions, say @{text "bn\<AL>\<^isub>1, \<dots>, bn\<AL>\<^isub>m"}. The
- syntax in Nominal Isabelle for such specifications is roughly as follows:
- %
- \begin{equation}\label{scheme}
- \mbox{\begin{tabular}{@ {}p{2.5cm}l}
- type \mbox{declaration part} &
- $\begin{cases}
- \mbox{\small\begin{tabular}{l}
- \isacommand{nominal\_datatype} @{text "ty\<AL>\<^isub>1 = \<dots>"}\\
- \isacommand{and} @{text "ty\<AL>\<^isub>2 = \<dots>"}\\
- \raisebox{2mm}{$\ldots$}\\[-2mm]
- \isacommand{and} @{text "ty\<AL>\<^isub>n = \<dots>"}\\
- \end{tabular}}
- \end{cases}$\\
- binding \mbox{function part} &
- $\begin{cases}
- \mbox{\small\begin{tabular}{l}
- \isacommand{binder} @{text "bn\<AL>\<^isub>1"} \isacommand{and} \ldots \isacommand{and} @{text "bn\<AL>\<^isub>m"}\\
- \isacommand{where}\\
- \raisebox{2mm}{$\ldots$}\\[-2mm]
- \end{tabular}}
- \end{cases}$\\
- \end{tabular}}
- \end{equation}
-
- \noindent
- Every type declaration @{text ty}$^\alpha_{1..n}$ consists of a collection of
- term-constructors, each of which comes with a list of labelled
- types that stand for the types of the arguments of the term-constructor.
- For example a term-constructor @{text "C\<^sup>\<alpha>"} might be specified with
-
- \begin{center}
- @{text "C\<^sup>\<alpha> label\<^isub>1::ty"}$'_1$ @{text "\<dots> label\<^isub>l::ty"}$'_l\;\;$ @{text "binding_clauses"}
- \end{center}
-
- \noindent
- whereby some of the @{text ty}$'_{1..l}$ %%(or their components)
- can be contained
- in the collection of @{text ty}$^\alpha_{1..n}$ declared in
- \eqref{scheme}.
- In this case we will call the corresponding argument a
- \emph{recursive argument} of @{text "C\<^sup>\<alpha>"}.
- %The types of such recursive
- %arguments need to satisfy a ``positivity''
- %restriction, which ensures that the type has a set-theoretic semantics
- %\cite{Berghofer99}.
- The labels
- annotated on the types are optional. Their purpose is to be used in the
- (possibly empty) list of \emph{binding clauses}, which indicate the binders
- and their scope in a term-constructor. They come in three \emph{modes}:
- %
- \begin{center}
- \begin{tabular}{@ {}l@ {}}
- \isacommand{bind} {\it binders} \isacommand{in} {\it bodies}\;\;\;\,
- \isacommand{bind (set)} {\it binders} \isacommand{in} {\it bodies}\;\;\;\,
- \isacommand{bind (set+)} {\it binders} \isacommand{in} {\it bodies}
- \end{tabular}
- \end{center}
- %
- \noindent
- The first mode is for binding lists of atoms (the order of binders matters);
- the second is for sets of binders (the order does not matter, but the
- cardinality does) and the last is for sets of binders (with vacuous binders
- preserving $\alpha$-equivalence). As indicated, the labels in the ``\isacommand{in}-part'' of a binding
- clause will be called \emph{bodies}; the
- ``\isacommand{bind}-part'' will be called \emph{binders}. In contrast to
- Ott, we allow multiple labels in binders and bodies.
-
- %For example we allow
- %binding clauses of the form:
- %
- %\begin{center}
- %\begin{tabular}{@ {}ll@ {}}
- %@{text "Foo\<^isub>1 x::name y::name t::trm s::trm"} &
- % \isacommand{bind} @{text "x y"} \isacommand{in} @{text "t s"}\\
- %@{text "Foo\<^isub>2 x::name y::name t::trm s::trm"} &
- % \isacommand{bind} @{text "x y"} \isacommand{in} @{text "t"},
- % \isacommand{bind} @{text "x y"} \isacommand{in} @{text "s"}\\
- %\end{tabular}
- %\end{center}
-
- \noindent
- %Similarly for the other binding modes.
- %Interestingly, in case of \isacommand{bind (set)}
- %and \isacommand{bind (set+)} the binding clauses above will make a difference to the semantics
- %of the specifications (the corresponding $\alpha$-equivalence will differ). We will
- %show this later with an example.
-
- There are also some restrictions we need to impose on our binding clauses in comparison to
- the ones of Ott. The
- main idea behind these restrictions is that we obtain a sensible notion of
- $\alpha$-equivalence where it is ensured that within a given scope an
- atom occurrence cannot be both bound and free at the same time. The first
- restriction is that a body can only occur in
- \emph{one} binding clause of a term constructor (this ensures that the bound
- atoms of a body cannot be free at the same time by specifying an
- alternative binder for the same body).
-
- For binders we distinguish between
- \emph{shallow} and \emph{deep} binders. Shallow binders are just
- labels. The restriction we need to impose on them is that in case of
- \isacommand{bind (set)} and \isacommand{bind (set+)} the labels must either
- refer to atom types or to sets of atom types; in case of \isacommand{bind}
- the labels must refer to atom types or lists of atom types. Two examples for
- the use of shallow binders are the specification of lambda-terms, where a
- single name is bound, and type-schemes, where a finite set of names is
- bound:
-
- \begin{center}\small
- \begin{tabular}{@ {}c@ {\hspace{7mm}}c@ {}}
- \begin{tabular}{@ {}l}
- \isacommand{nominal\_datatype} @{text lam} $=$\\
- \hspace{2mm}\phantom{$\mid$}~@{text "Var name"}\\
- \hspace{2mm}$\mid$~@{text "App lam lam"}\\
- \hspace{2mm}$\mid$~@{text "Lam x::name t::lam"}~~\isacommand{bind} @{text x} \isacommand{in} @{text t}\\
- \end{tabular} &
- \begin{tabular}{@ {}l@ {}}
- \isacommand{nominal\_datatype}~@{text ty} $=$\\
- \hspace{5mm}\phantom{$\mid$}~@{text "TVar name"}\\
- \hspace{5mm}$\mid$~@{text "TFun ty ty"}\\
- \isacommand{and}~@{text "tsc = All xs::(name fset) T::ty"}~~%
- \isacommand{bind (set+)} @{text xs} \isacommand{in} @{text T}\\
- \end{tabular}
- \end{tabular}
- \end{center}
-
- \noindent
- In these specifications @{text "name"} refers to an atom type, and @{text
- "fset"} to the type of finite sets.
- Note that for @{text lam} it does not matter which binding mode we use. The
- reason is that we bind only a single @{text name}. However, having
- \isacommand{bind (set)} or \isacommand{bind} in the second case makes a
- difference to the semantics of the specification (which we will define in the next section).
-
-
- A \emph{deep} binder uses an auxiliary binding function that ``picks'' out
- the atoms in one argument of the term-constructor, which can be bound in
- other arguments and also in the same argument (we will call such binders
- \emph{recursive}, see below). The binding functions are
- expected to return either a set of atoms (for \isacommand{bind (set)} and
- \isacommand{bind (set+)}) or a list of atoms (for \isacommand{bind}). They can
- be defined by recursion over the corresponding type; the equations
- must be given in the binding function part of the scheme shown in
- \eqref{scheme}. For example a term-calculus containing @{text "Let"}s with
- tuple patterns might be specified as:
- %
- \begin{equation}\label{letpat}
- \mbox{\small%
- \begin{tabular}{l}
- \isacommand{nominal\_datatype} @{text trm} $=$\\
- \hspace{5mm}\phantom{$\mid$}~@{term "Var name"}\\
- \hspace{5mm}$\mid$~@{term "App trm trm"}\\
- \hspace{5mm}$\mid$~@{text "Lam x::name t::trm"}
- \;\;\isacommand{bind} @{text x} \isacommand{in} @{text t}\\
- \hspace{5mm}$\mid$~@{text "Let p::pat trm t::trm"}
- \;\;\isacommand{bind} @{text "bn(p)"} \isacommand{in} @{text t}\\
- \isacommand{and} @{text pat} $=$
- @{text PNil}
- $\mid$~@{text "PVar name"}
- $\mid$~@{text "PTup pat pat"}\\
- \isacommand{binder}~@{text "bn::pat \<Rightarrow> atom list"}\\
- \isacommand{where}~@{text "bn(PNil) = []"}\\
- \hspace{5mm}$\mid$~@{text "bn(PVar x) = [atom x]"}\\
- \hspace{5mm}$\mid$~@{text "bn(PTup p\<^isub>1 p\<^isub>2) = bn(p\<^isub>1) @ bn(p\<^isub>2)"}\smallskip\\
- \end{tabular}}
- \end{equation}
- %
- \noindent
- In this specification the function @{text "bn"} determines which atoms of
- the pattern @{text p} are bound in the argument @{text "t"}. Note that in the
- second-last @{text bn}-clause the function @{text "atom"} coerces a name
- into the generic atom type of Nominal Isabelle \cite{HuffmanUrban10}. This
- allows us to treat binders of different atom type uniformly.
-
- As said above, for deep binders we allow binding clauses such as
- %
- %\begin{center}
- %\begin{tabular}{ll}
- @{text "Bar p::pat t::trm"} %%%&
- \isacommand{bind} @{text "bn(p)"} \isacommand{in} @{text "p t"} %%\\
- %\end{tabular}
- %\end{center}
- %
- %\noindent
- where the argument of the deep binder also occurs in the body. We call such
- binders \emph{recursive}. To see the purpose of such recursive binders,
- compare ``plain'' @{text "Let"}s and @{text "Let_rec"}s in the following
- specification:
- %
- \begin{equation}\label{letrecs}
- \mbox{\small%
- \begin{tabular}{@ {}l@ {}}
- \isacommand{nominal\_datatype}~@{text "trm ="}~\ldots\\
- \hspace{5mm}$\mid$~@{text "Let as::assn t::trm"}
- \;\;\isacommand{bind} @{text "bn(as)"} \isacommand{in} @{text t}\\
- \hspace{5mm}$\mid$~@{text "Let_rec as::assn t::trm"}
- \;\;\isacommand{bind} @{text "bn(as)"} \isacommand{in} @{text "as t"}\\
- \isacommand{and} @{text "assn"} $=$
- @{text "ANil"}
- $\mid$~@{text "ACons name trm assn"}\\
- \isacommand{binder} @{text "bn::assn \<Rightarrow> atom list"}\\
- \isacommand{where}~@{text "bn(ANil) = []"}\\
- \hspace{5mm}$\mid$~@{text "bn(ACons a t as) = [atom a] @ bn(as)"}\\
- \end{tabular}}
- \end{equation}
- %
- \noindent
- The difference is that with @{text Let} we only want to bind the atoms @{text
- "bn(as)"} in the term @{text t}, but with @{text "Let_rec"} we also want to bind the atoms
- inside the assignment. This difference has consequences for the associated
- notions of free-atoms and $\alpha$-equivalence.
-
- To make sure that atoms bound by deep binders cannot be free at the
- same time, we cannot have more than one binding function for a deep binder.
- Consequently we exclude specifications such as
- %
- \begin{center}\small
- \begin{tabular}{@ {}l@ {\hspace{2mm}}l@ {}}
- @{text "Baz\<^isub>1 p::pat t::trm"} &
- \isacommand{bind} @{text "bn\<^isub>1(p) bn\<^isub>2(p)"} \isacommand{in} @{text t}\\
- @{text "Baz\<^isub>2 p::pat t\<^isub>1::trm t\<^isub>2::trm"} &
- \isacommand{bind} @{text "bn\<^isub>1(p)"} \isacommand{in} @{text "t\<^isub>1"},
- \isacommand{bind} @{text "bn\<^isub>2(p)"} \isacommand{in} @{text "t\<^isub>2"}\\
- \end{tabular}
- \end{center}
-
- \noindent
- Otherwise it is possible that @{text "bn\<^isub>1"} and @{text "bn\<^isub>2"} pick
- out different atoms to become bound, respectively be free, in @{text "p"}.
- (Since the Ott-tool does not derive a reasoning infrastructure for
- $\alpha$-equated terms with deep binders, it can permit such specifications.)
-
- We also need to restrict the form of the binding functions in order
- to ensure the @{text "bn"}-functions can be defined for $\alpha$-equated
- terms. The main restriction is that we cannot return an atom in a binding function that is also
- bound in the corresponding term-constructor. That means in \eqref{letpat}
- that the term-constructors @{text PVar} and @{text PTup} may
- not have a binding clause (all arguments are used to define @{text "bn"}).
- In contrast, in case of \eqref{letrecs} the term-constructor @{text ACons}
- may have a binding clause involving the argument @{text trm} (the only one that
- is \emph{not} used in the definition of the binding function). This restriction
- is sufficient for lifting the binding function to $\alpha$-equated terms.
-
- In the version of
- Nominal Isabelle described here, we also adopted the restriction from the
- Ott-tool that binding functions can only return: the empty set or empty list
- (as in case @{text PNil}), a singleton set or singleton list containing an
- atom (case @{text PVar}), or unions of atom sets or appended atom lists
- (case @{text PTup}). This restriction will simplify some automatic definitions and proofs
- later on.
-
- In order to simplify our definitions of free atoms and $\alpha$-equivalence,
- we shall assume specifications
- of term-calculi are implicitly \emph{completed}. By this we mean that
- for every argument of a term-constructor that is \emph{not}
- already part of a binding clause given by the user, we add implicitly a special \emph{empty} binding
- clause, written \isacommand{bind}~@{term "{}"}~\isacommand{in}~@{text "labels"}. In case
- of the lambda-terms, the completion produces
-
- \begin{center}\small
- \begin{tabular}{@ {}l@ {\hspace{-1mm}}}
- \isacommand{nominal\_datatype} @{text lam} =\\
- \hspace{5mm}\phantom{$\mid$}~@{text "Var x::name"}
- \;\;\isacommand{bind}~@{term "{}"}~\isacommand{in}~@{text "x"}\\
- \hspace{5mm}$\mid$~@{text "App t\<^isub>1::lam t\<^isub>2::lam"}
- \;\;\isacommand{bind}~@{term "{}"}~\isacommand{in}~@{text "t\<^isub>1 t\<^isub>2"}\\
- \hspace{5mm}$\mid$~@{text "Lam x::name t::lam"}
- \;\;\isacommand{bind}~@{text x} \isacommand{in} @{text t}\\
- \end{tabular}
- \end{center}
-
- \noindent
- The point of completion is that we can make definitions over the binding
- clauses and be sure to have captured all arguments of a term constructor.
-*}
-
-section {* Alpha-Equivalence and Free Atoms\label{sec:alpha} *}
-
-text {*
- Having dealt with all syntax matters, the problem now is how we can turn
- specifications into actual type definitions in Isabelle/HOL and then
- establish a reasoning infrastructure for them. As
- Pottier and Cheney pointed out \cite{Pottier06,Cheney05}, just
- re-arranging the arguments of
- term-constructors so that binders and their bodies are next to each other will
- result in inadequate representations in cases like @{text "Let x\<^isub>1 = t\<^isub>1\<dots>x\<^isub>n = t\<^isub>n in s"}.
- Therefore we will first
- extract ``raw'' datatype definitions from the specification and then define
- explicitly an $\alpha$-equivalence relation over them. We subsequently
- construct the quotient of the datatypes according to our $\alpha$-equivalence.
-
- The ``raw'' datatype definition can be obtained by stripping off the
- binding clauses and the labels from the types. We also have to invent
- new names for the types @{text "ty\<^sup>\<alpha>"} and term-constructors @{text "C\<^sup>\<alpha>"}
- given by the user. In our implementation we just use the affix ``@{text "_raw"}''.
- But for the purpose of this paper, we use the superscript @{text "_\<^sup>\<alpha>"} to indicate
- that a notion is given for $\alpha$-equivalence classes and leave it out
- for the corresponding notion given on the ``raw'' level. So for example
- we have @{text "ty\<^sup>\<alpha> \<mapsto> ty"} and @{text "C\<^sup>\<alpha> \<mapsto> C"}
- where @{term ty} is the type used in the quotient construction for
- @{text "ty\<^sup>\<alpha>"} and @{text "C"} is the term-constructor on the ``raw'' type @{text "ty"}.
-
- %The resulting datatype definition is legal in Isabelle/HOL provided the datatypes are
- %non-empty and the types in the constructors only occur in positive
- %position (see \cite{Berghofer99} for an in-depth description of the datatype package
- %in Isabelle/HOL).
- We subsequently define each of the user-specified binding
- functions @{term "bn"}$_{1..m}$ by recursion over the corresponding
- raw datatype. We can also easily define permutation operations by
- recursion so that for each term constructor @{text "C"} we have that
- %
- \begin{equation}\label{ceqvt}
- @{text "p \<bullet> (C z\<^isub>1 \<dots> z\<^isub>n) = C (p \<bullet> z\<^isub>1) \<dots> (p \<bullet> z\<^isub>n)"}
- \end{equation}
-
- The first non-trivial step we have to perform is the generation of
- free-atom functions from the specification. For the
- \emph{raw} types @{text "ty"}$_{1..n}$ we define the free-atom functions
- %
- %\begin{equation}\label{fvars}
- @{text "fa_ty\<^isub>"}$_{1..n}$
- %\end{equation}
- %
- %\noindent
- by recursion.
- We define these functions together with auxiliary free-atom functions for
- the binding functions. Given raw binding functions @{text "bn"}$_{1..m}$
- we define
- %
- %\begin{center}
- @{text "fa_bn\<^isub>"}$_{1..m}$.
- %\end{center}
- %
- %\noindent
- The reason for this setup is that in a deep binder not all atoms have to be
- bound, as we saw in the example with ``plain'' @{text Let}s. We need therefore a function
- that calculates those free atoms in a deep binder.
-
- While the idea behind these free-atom functions is clear (they just
- collect all atoms that are not bound), because of our rather complicated
- binding mechanisms their definitions are somewhat involved. Given
- a term-constructor @{text "C"} of type @{text ty} and some associated
- binding clauses @{text "bc\<^isub>1\<dots>bc\<^isub>k"}, the result of @{text
- "fa_ty (C z\<^isub>1 \<dots> z\<^isub>n)"} will be the union @{text
- "fa(bc\<^isub>1) \<union> \<dots> \<union> fa(bc\<^isub>k)"} where we will define below what @{text "fa"} for a binding
- clause means. We only show the details for the mode \isacommand{bind (set)} (the other modes are similar).
- Suppose the binding clause @{text bc\<^isub>i} is of the form
- %
- %\begin{center}
- \mbox{\isacommand{bind (set)} @{text "b\<^isub>1\<dots>b\<^isub>p"} \isacommand{in} @{text "d\<^isub>1\<dots>d\<^isub>q"}}
- %\end{center}
- %
- %\noindent
- in which the body-labels @{text "d"}$_{1..q}$ refer to types @{text ty}$_{1..q}$,
- and the binders @{text b}$_{1..p}$
- either refer to labels of atom types (in case of shallow binders) or to binding
- functions taking a single label as argument (in case of deep binders). Assuming
- @{text "D"} stands for the set of free atoms of the bodies, @{text B} for the
- set of binding atoms in the binders and @{text "B'"} for the set of free atoms in
- non-recursive deep binders,
- then the free atoms of the binding clause @{text bc\<^isub>i} are\\[-2mm]
- %
- \begin{equation}\label{fadef}
- \mbox{@{text "fa(bc\<^isub>i) \<equiv> (D - B) \<union> B'"}}.
- \end{equation}
- %
- \noindent
- The set @{text D} is formally defined as
- %
- %\begin{center}
- @{text "D \<equiv> fa_ty\<^isub>1 d\<^isub>1 \<union> ... \<union> fa_ty\<^isub>q d\<^isub>q"}
- %\end{center}
- %
- %\noindent
- where in case @{text "d\<^isub>i"} refers to one of the raw types @{text "ty"}$_{1..n}$ from the
- specification, the function @{text "fa_ty\<^isub>i"} is the corresponding free-atom function
- we are defining by recursion;
- %(see \eqref{fvars});
- otherwise we set @{text "fa_ty\<^isub>i d\<^isub>i = supp d\<^isub>i"}.
-
- In order to formally define the set @{text B} we use the following auxiliary @{text "bn"}-functions
- for atom types to which shallow binders may refer\\[-4mm]
- %
- %\begin{center}
- %\begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
- %@{text "bn_atom a"} & @{text "\<equiv>"} & @{text "{atom a}"}\\
- %@{text "bn_atom_set as"} & @{text "\<equiv>"} & @{text "atoms as"}\\
- %@{text "bn_atom_list as"} & @{text "\<equiv>"} & @{text "atoms (set as)"}
- %\end{tabular}
- %\end{center}
- %
- \begin{center}
- @{text "bn\<^bsub>atom\<^esub> a \<equiv> {atom a}"}\hfill
- @{text "bn\<^bsub>atom_set\<^esub> as \<equiv> atoms as"}\hfill
- @{text "bn\<^bsub>atom_list\<^esub> as \<equiv> atoms (set as)"}
- \end{center}
- %
- \noindent
- Like the function @{text atom}, the function @{text "atoms"} coerces
- a set of atoms to a set of the generic atom type.
- %It is defined as @{text "atoms as \<equiv> {atom a | a \<in> as}"}.
- The set @{text B} is then formally defined as\\[-4mm]
- %
- \begin{center}
- @{text "B \<equiv> bn_ty\<^isub>1 b\<^isub>1 \<union> ... \<union> bn_ty\<^isub>p b\<^isub>p"}
- \end{center}
- %
- \noindent
- where we use the auxiliary binding functions for shallow binders.
- The set @{text "B'"} collects all free atoms in non-recursive deep
- binders. Let us assume these binders in @{text "bc\<^isub>i"} are
- %
- %\begin{center}
- \mbox{@{text "bn\<^isub>1 l\<^isub>1, \<dots>, bn\<^isub>r l\<^isub>r"}}
- %\end{center}
- %
- %\noindent
- with @{text "l"}$_{1..r}$ $\subseteq$ @{text "b"}$_{1..p}$ and none of the
- @{text "l"}$_{1..r}$ being among the bodies @{text
- "d"}$_{1..q}$. The set @{text "B'"} is defined as\\[-5mm]
- %
- \begin{center}
- @{text "B' \<equiv> fa_bn\<^isub>1 l\<^isub>1 \<union> ... \<union> fa_bn\<^isub>r l\<^isub>r"}\\[-9mm]
- \end{center}
- %
- \noindent
- This completes the definition of the free-atom functions @{text "fa_ty"}$_{1..n}$.
-
- Note that for non-recursive deep binders, we have to add in \eqref{fadef}
- the set of atoms that are left unbound by the binding functions @{text
- "bn"}$_{1..m}$. We used for the definition of
- this set the functions @{text "fa_bn"}$_{1..m}$, which are also defined by mutual
- recursion. Assume the user specified a @{text bn}-clause of the form
- %
- %\begin{center}
- @{text "bn (C z\<^isub>1 \<dots> z\<^isub>s) = rhs"}
- %\end{center}
- %
- %\noindent
- where the @{text "z"}$_{1..s}$ are of types @{text "ty"}$_{1..s}$. For each of
- the arguments we calculate the free atoms as follows:
- %
- \begin{center}
- \begin{tabular}{c@ {\hspace{2mm}}p{0.9\textwidth}}
- $\bullet$ & @{term "fa_ty\<^isub>i z\<^isub>i"} provided @{text "z\<^isub>i"} does not occur in @{text "rhs"}
- (that means nothing is bound in @{text "z\<^isub>i"} by the binding function),\\
- $\bullet$ & @{term "fa_bn\<^isub>i z\<^isub>i"} provided @{text "z\<^isub>i"} occurs in @{text "rhs"}
- with the recursive call @{text "bn\<^isub>i z\<^isub>i"}, and\\
- $\bullet$ & @{term "{}"} provided @{text "z\<^isub>i"} occurs in @{text "rhs"},
- but without a recursive call.
- \end{tabular}
- \end{center}
- %
- \noindent
- For defining @{text "fa_bn (C z\<^isub>1 \<dots> z\<^isub>n)"} we just union up all these sets.
-
- To see how these definitions work in practice, let us reconsider the
- term-constructors @{text "Let"} and @{text "Let_rec"} shown in
- \eqref{letrecs} together with the term-constructors for assignments @{text
- "ANil"} and @{text "ACons"}. Since there is a binding function defined for
- assignments, we have three free-atom functions, namely @{text
- "fa\<^bsub>trm\<^esub>"}, @{text "fa\<^bsub>assn\<^esub>"} and @{text
- "fa\<^bsub>bn\<^esub>"} as follows:
- %
- \begin{center}\small
- \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
- @{text "fa\<^bsub>trm\<^esub> (Let as t)"} & @{text "="} & @{text "(fa\<^bsub>trm\<^esub> t - set (bn as)) \<union> fa\<^bsub>bn\<^esub> as"}\\
- @{text "fa\<^bsub>trm\<^esub> (Let_rec as t)"} & @{text "="} & @{text "(fa\<^bsub>assn\<^esub> as \<union> fa\<^bsub>trm\<^esub> t) - set (bn as)"}\\[1mm]
-
- @{text "fa\<^bsub>assn\<^esub> (ANil)"} & @{text "="} & @{term "{}"}\\
- @{text "fa\<^bsub>assn\<^esub> (ACons a t as)"} & @{text "="} & @{text "(supp a) \<union> (fa\<^bsub>trm\<^esub> t) \<union> (fa\<^bsub>assn\<^esub> as)"}\\[1mm]
-
- @{text "fa\<^bsub>bn\<^esub> (ANil)"} & @{text "="} & @{term "{}"}\\
- @{text "fa\<^bsub>bn\<^esub> (ACons a t as)"} & @{text "="} & @{text "(fa\<^bsub>trm\<^esub> t) \<union> (fa\<^bsub>bn\<^esub> as)"}
- \end{tabular}
- \end{center}
-
- \noindent
- Recall that @{text ANil} and @{text "ACons"} have no
- binding clause in the specification. The corresponding free-atom
- function @{text "fa\<^bsub>assn\<^esub>"} therefore returns all free atoms
- of an assignment (in case of @{text "ACons"}, they are given in
- terms of @{text supp}, @{text "fa\<^bsub>trm\<^esub>"} and @{text "fa\<^bsub>assn\<^esub>"}).
- The binding only takes place in @{text Let} and
- @{text "Let_rec"}. In case of @{text "Let"}, the binding clause specifies
- that all atoms given by @{text "set (bn as)"} have to be bound in @{text
- t}. Therefore we have to subtract @{text "set (bn as)"} from @{text
- "fa\<^bsub>trm\<^esub> t"}. However, we also need to add all atoms that are
- free in @{text "as"}. This is
- in contrast with @{text "Let_rec"} where we have a recursive
- binder to bind all occurrences of the atoms in @{text
- "set (bn as)"} also inside @{text "as"}. Therefore we have to subtract
- @{text "set (bn as)"} from both @{text "fa\<^bsub>trm\<^esub> t"} and @{text "fa\<^bsub>assn\<^esub> as"}.
- %Like the function @{text "bn"}, the function @{text "fa\<^bsub>bn\<^esub>"} traverses the
- %list of assignments, but instead returns the free atoms, which means in this
- %example the free atoms in the argument @{text "t"}.
-
- An interesting point in this
- example is that a ``naked'' assignment (@{text "ANil"} or @{text "ACons"}) does not bind any
- atoms, even if the binding function is specified over assignments.
- Only in the context of a @{text Let} or @{text "Let_rec"}, where the binding clauses are given, will
- some atoms actually become bound. This is a phenomenon that has also been pointed
- out in \cite{ott-jfp}. For us this observation is crucial, because we would
- not be able to lift the @{text "bn"}-functions to $\alpha$-equated terms if they act on
- atoms that are bound. In that case, these functions would \emph{not} respect
- $\alpha$-equivalence.
-
- Next we define the $\alpha$-equivalence relations for the raw types @{text
- "ty"}$_{1..n}$ from the specification. We write them as
- %
- %\begin{center}
- @{text "\<approx>ty"}$_{1..n}$.
- %\end{center}
- %
- %\noindent
- Like with the free-atom functions, we also need to
- define auxiliary $\alpha$-equivalence relations
- %
- %\begin{center}
- @{text "\<approx>bn\<^isub>"}$_{1..m}$
- %\end{center}
- %
- %\noindent
- for the binding functions @{text "bn"}$_{1..m}$,
- To simplify our definitions we will use the following abbreviations for
- \emph{compound equivalence relations} and \emph{compound free-atom functions} acting on tuples.
- %
- \begin{center}
- \begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
- @{text "(x\<^isub>1,\<dots>, x\<^isub>n) (R\<^isub>1,\<dots>, R\<^isub>n) (x\<PRIME>\<^isub>1,\<dots>, x\<PRIME>\<^isub>n)"} & @{text "\<equiv>"} &
- @{text "x\<^isub>1 R\<^isub>1 x\<PRIME>\<^isub>1 \<and> \<dots> \<and> x\<^isub>n R\<^isub>n x\<PRIME>\<^isub>n"}\\
- @{text "(fa\<^isub>1,\<dots>, fa\<^isub>n) (x\<^isub>1,\<dots>, x\<^isub>n)"} & @{text "\<equiv>"} & @{text "fa\<^isub>1 x\<^isub>1 \<union> \<dots> \<union> fa\<^isub>n x\<^isub>n"}\\
- \end{tabular}
- \end{center}
-
-
- The $\alpha$-equivalence relations are defined as inductive predicates
- having a single clause for each term-constructor. Assuming a
- term-constructor @{text C} is of type @{text ty} and has the binding clauses
- @{term "bc"}$_{1..k}$, then the $\alpha$-equivalence clause has the form
- %
- \begin{center}
- \mbox{\infer{@{text "C z\<^isub>1 \<dots> z\<^isub>n \<approx>ty C z\<PRIME>\<^isub>1 \<dots> z\<PRIME>\<^isub>n"}}
- {@{text "prems(bc\<^isub>1) \<dots> prems(bc\<^isub>k)"}}}
- \end{center}
-
- \noindent
- The task below is to specify what the premises of a binding clause are. As a
- special instance, we first treat the case where @{text "bc\<^isub>i"} is the
- empty binding clause of the form
- %
- \begin{center}
- \mbox{\isacommand{bind (set)} @{term "{}"} \isacommand{in} @{text "d\<^isub>1\<dots>d\<^isub>q"}.}
- \end{center}
-
- \noindent
- In this binding clause no atom is bound and we only have to $\alpha$-relate the bodies. For this
- we build first the tuples @{text "D \<equiv> (d\<^isub>1,\<dots>, d\<^isub>q)"} and @{text "D' \<equiv> (d\<PRIME>\<^isub>1,\<dots>, d\<PRIME>\<^isub>q)"}
- whereby the labels @{text "d"}$_{1..q}$ refer to arguments @{text "z"}$_{1..n}$ and
- respectively @{text "d\<PRIME>"}$_{1..q}$ to @{text "z\<PRIME>"}$_{1..n}$. In order to relate
- two such tuples we define the compound $\alpha$-equivalence relation @{text "R"} as follows
- %
- \begin{equation}\label{rempty}
- \mbox{@{text "R \<equiv> (R\<^isub>1,\<dots>, R\<^isub>q)"}}
- \end{equation}
-
- \noindent
- with @{text "R\<^isub>i"} being @{text "\<approx>ty\<^isub>i"} if the corresponding labels @{text "d\<^isub>i"} and
- @{text "d\<PRIME>\<^isub>i"} refer
- to a recursive argument of @{text C} with type @{text "ty\<^isub>i"}; otherwise
- we take @{text "R\<^isub>i"} to be the equality @{text "="}. This lets us define
- the premise for an empty binding clause succinctly as @{text "prems(bc\<^isub>i) \<equiv> D R D'"},
- which can be unfolded to the series of premises
- %
- %\begin{center}
- @{text "d\<^isub>1 R\<^isub>1 d\<PRIME>\<^isub>1 \<dots> d\<^isub>q R\<^isub>q d\<PRIME>\<^isub>q"}.
- %\end{center}
- %
- %\noindent
- We will use the unfolded version in the examples below.
-
- Now suppose the binding clause @{text "bc\<^isub>i"} is of the general form
- %
- \begin{equation}\label{nonempty}
- \mbox{\isacommand{bind (set)} @{text "b\<^isub>1\<dots>b\<^isub>p"} \isacommand{in} @{text "d\<^isub>1\<dots>d\<^isub>q"}.}
- \end{equation}
-
- \noindent
- In this case we define a premise @{text P} using the relation
- $\approx_{\,\textit{set}}$ given in Section~\ref{sec:binders} (similarly
- $\approx_{\,\textit{set+}}$ and $\approx_{\,\textit{list}}$ for the other
- binding modes). This premise defines $\alpha$-equivalence of two abstractions
- involving multiple binders. As above, we first build the tuples @{text "D"} and
- @{text "D'"} for the bodies @{text "d"}$_{1..q}$, and the corresponding
- compound $\alpha$-relation @{text "R"} (shown in \eqref{rempty}).
- For $\approx_{\,\textit{set}}$ we also need
- a compound free-atom function for the bodies defined as
- %
- \begin{center}
- \mbox{@{text "fa \<equiv> (fa_ty\<^isub>1,\<dots>, fa_ty\<^isub>q)"}}
- \end{center}
-
- \noindent
- with the assumption that the @{text "d"}$_{1..q}$ refer to arguments of types @{text "ty"}$_{1..q}$.
- The last ingredient we need are the sets of atoms bound in the bodies.
- For this we take
-
- \begin{center}
- @{text "B \<equiv> bn_ty\<^isub>1 b\<^isub>1 \<union> \<dots> \<union> bn_ty\<^isub>p b\<^isub>p"}\;.\\
- \end{center}
-
- \noindent
- Similarly for @{text "B'"} using the labels @{text "b\<PRIME>"}$_{1..p}$. This
- lets us formally define the premise @{text P} for a non-empty binding clause as:
- %
- \begin{center}
- \mbox{@{term "P \<equiv> \<exists>p. (B, D) \<approx>set R fa p (B', D')"}}\;.
- \end{center}
-
- \noindent
- This premise accounts for $\alpha$-equivalence of the bodies of the binding
- clause.
- However, in case the binders have non-recursive deep binders, this premise
- is not enough:
- we also have to ``propagate'' $\alpha$-equivalence inside the structure of
- these binders. An example is @{text "Let"} where we have to make sure the
- right-hand sides of assignments are $\alpha$-equivalent. For this we use
- relations @{text "\<approx>bn"}$_{1..m}$ (which we will formally define shortly).
- Let us assume the non-recursive deep binders in @{text "bc\<^isub>i"} are
- %
- %\begin{center}
- @{text "bn\<^isub>1 l\<^isub>1, \<dots>, bn\<^isub>r l\<^isub>r"}.
- %\end{center}
- %
- %\noindent
- The tuple @{text L} is then @{text "(l\<^isub>1,\<dots>,l\<^isub>r)"} (similarly @{text "L'"})
- and the compound equivalence relation @{text "R'"} is @{text "(\<approx>bn\<^isub>1,\<dots>,\<approx>bn\<^isub>r)"}.
- All premises for @{text "bc\<^isub>i"} are then given by
- %
- \begin{center}
- @{text "prems(bc\<^isub>i) \<equiv> P \<and> L R' L'"}
- \end{center}
-
- \noindent
- The auxiliary $\alpha$-equivalence relations @{text "\<approx>bn"}$_{1..m}$
- in @{text "R'"} are defined as follows: assuming a @{text bn}-clause is of the form
- %
- %\begin{center}
- @{text "bn (C z\<^isub>1 \<dots> z\<^isub>s) = rhs"}
- %\end{center}
- %
- %\noindent
- where the @{text "z"}$_{1..s}$ are of types @{text "ty"}$_{1..s}$,
- then the corresponding $\alpha$-equivalence clause for @{text "\<approx>bn"} has the form
- %
- \begin{center}
- \mbox{\infer{@{text "C z\<^isub>1 \<dots> z\<^isub>s \<approx>bn C z\<PRIME>\<^isub>1 \<dots> z\<PRIME>\<^isub>s"}}
- {@{text "z\<^isub>1 R\<^isub>1 z\<PRIME>\<^isub>1 \<dots> z\<^isub>s R\<^isub>s z\<PRIME>\<^isub>s"}}}
- \end{center}
-
- \noindent
- In this clause the relations @{text "R"}$_{1..s}$ are given by
-
- \begin{center}
- \begin{tabular}{c@ {\hspace{2mm}}p{0.9\textwidth}}
- $\bullet$ & @{text "z\<^isub>i \<approx>ty z\<PRIME>\<^isub>i"} provided @{text "z\<^isub>i"} does not occur in @{text rhs} and
- is a recursive argument of @{text C},\\
- $\bullet$ & @{text "z\<^isub>i = z\<PRIME>\<^isub>i"} provided @{text "z\<^isub>i"} does not occur in @{text rhs}
- and is a non-recursive argument of @{text C},\\
- $\bullet$ & @{text "z\<^isub>i \<approx>bn\<^isub>i z\<PRIME>\<^isub>i"} provided @{text "z\<^isub>i"} occurs in @{text rhs}
- with the recursive call @{text "bn\<^isub>i x\<^isub>i"} and\\
- $\bullet$ & @{text True} provided @{text "z\<^isub>i"} occurs in @{text rhs} but without a
- recursive call.
- \end{tabular}
- \end{center}
-
- \noindent
- This completes the definition of $\alpha$-equivalence. As a sanity check, we can show
- that the premises of empty binding clauses are a special case of the clauses for
- non-empty ones (we just have to unfold the definition of $\approx_{\,\textit{set}}$ and take @{text "0"}
- for the existentially quantified permutation).
-
- Again let us take a look at a concrete example for these definitions. For \eqref{letrecs}
- we have three relations $\approx_{\textit{trm}}$, $\approx_{\textit{assn}}$ and
- $\approx_{\textit{bn}}$ with the following clauses:
-
- \begin{center}\small
- \begin{tabular}{@ {}c @ {}}
- \infer{@{text "Let as t \<approx>\<^bsub>trm\<^esub> Let as' t'"}}
- {@{term "\<exists>p. (bn as, t) \<approx>lst alpha_trm fa_trm p (bn as', t')"} & @{text "as \<approx>\<^bsub>bn\<^esub> as'"}}\smallskip\\
- \makebox[0mm]{\infer{@{text "Let_rec as t \<approx>\<^bsub>trm\<^esub> Let_rec as' t'"}}
- {@{term "\<exists>p. (bn as, ast) \<approx>lst alpha_trm2 fa_trm2 p (bn as', ast')"}}}
- \end{tabular}
- \end{center}
-
- \begin{center}\small
- \begin{tabular}{@ {}c @ {}}
- \infer{@{text "ANil \<approx>\<^bsub>assn\<^esub> ANil"}}{}\hspace{9mm}
- \infer{@{text "ACons a t as \<approx>\<^bsub>assn\<^esub> ACons a' t' as"}}
- {@{text "a = a'"} & @{text "t \<approx>\<^bsub>trm\<^esub> t'"} & @{text "as \<approx>\<^bsub>assn\<^esub> as'"}}
- \end{tabular}
- \end{center}
-
- \begin{center}\small
- \begin{tabular}{@ {}c @ {}}
- \infer{@{text "ANil \<approx>\<^bsub>bn\<^esub> ANil"}}{}\hspace{9mm}
- \infer{@{text "ACons a t as \<approx>\<^bsub>bn\<^esub> ACons a' t' as"}}
- {@{text "t \<approx>\<^bsub>trm\<^esub> t'"} & @{text "as \<approx>\<^bsub>bn\<^esub> as'"}}
- \end{tabular}
- \end{center}
-
- \noindent
- Note the difference between $\approx_{\textit{assn}}$ and
- $\approx_{\textit{bn}}$: the latter only ``tracks'' $\alpha$-equivalence of
- the components in an assignment that are \emph{not} bound. This is needed in the
- clause for @{text "Let"} (which has
- a non-recursive binder).
- %The underlying reason is that the terms inside an assignment are not meant
- %to be ``under'' the binder. Such a premise is \emph{not} needed in @{text "Let_rec"},
- %because there all components of an assignment are ``under'' the binder.
-*}
-
-section {* Establishing the Reasoning Infrastructure *}
-
-text {*
- Having made all necessary definitions for raw terms, we can start
- with establishing the reasoning infrastructure for the $\alpha$-equated types
- @{text "ty\<AL>"}$_{1..n}$, that is the types the user originally specified. We sketch
- in this section the proofs we need for establishing this infrastructure. One
- main point of our work is that we have completely automated these proofs in Isabelle/HOL.
-
- First we establish that the
- $\alpha$-equivalence relations defined in the previous section are
- equivalence relations.
-
- \begin{lemma}\label{equiv}
- Given the raw types @{text "ty"}$_{1..n}$ and binding functions
- @{text "bn"}$_{1..m}$, the relations @{text "\<approx>ty"}$_{1..n}$ and
- @{text "\<approx>bn"}$_{1..m}$ are equivalence relations.%% and equivariant.
- \end{lemma}
-
- \begin{proof}
- The proof is by mutual induction over the definitions. The non-trivial
- cases involve premises built up by $\approx_{\textit{set}}$,
- $\approx_{\textit{set+}}$ and $\approx_{\textit{list}}$. They
- can be dealt with as in Lemma~\ref{alphaeq}.
- \end{proof}
-
- \noindent
- We can feed this lemma into our quotient package and obtain new types @{text
- "ty"}$^\alpha_{1..n}$ representing $\alpha$-equated terms of types @{text "ty"}$_{1..n}$.
- We also obtain definitions for the term-constructors @{text
- "C"}$^\alpha_{1..k}$ from the raw term-constructors @{text
- "C"}$_{1..k}$, and similar definitions for the free-atom functions @{text
- "fa_ty"}$^\alpha_{1..n}$ and @{text "fa_bn"}$^\alpha_{1..m}$ as well as the binding functions @{text
- "bn"}$^\alpha_{1..m}$. However, these definitions are not really useful to the
- user, since they are given in terms of the isomorphisms we obtained by
- creating new types in Isabelle/HOL (recall the picture shown in the
- Introduction).
-
- The first useful property for the user is the fact that distinct
- term-constructors are not
- equal, that is
- %
- \begin{equation}\label{distinctalpha}
- \mbox{@{text "C"}$^\alpha$~@{text "x\<^isub>1 \<dots> x\<^isub>r"}~@{text "\<noteq>"}~%
- @{text "D"}$^\alpha$~@{text "y\<^isub>1 \<dots> y\<^isub>s"}}
- \end{equation}
-
- \noindent
- whenever @{text "C"}$^\alpha$~@{text "\<noteq>"}~@{text "D"}$^\alpha$.
- In order to derive this fact, we use the definition of $\alpha$-equivalence
- and establish that
- %
- \begin{equation}\label{distinctraw}
- \mbox{@{text "C x\<^isub>1 \<dots> x\<^isub>r"}\;$\not\approx$@{text ty}\;@{text "D y\<^isub>1 \<dots> y\<^isub>s"}}
- \end{equation}
-
- \noindent
- holds for the corresponding raw term-constructors.
- In order to deduce \eqref{distinctalpha} from \eqref{distinctraw}, our quotient
- package needs to know that the raw term-constructors @{text "C"} and @{text "D"}
- are \emph{respectful} w.r.t.~the $\alpha$-equivalence relations (see \cite{Homeier05}).
- Assuming, for example, @{text "C"} is of type @{text "ty"} with argument types
- @{text "ty"}$_{1..r}$, respectfulness amounts to showing that
- %
- \begin{center}
- @{text "C x\<^isub>1 \<dots> x\<^isub>r \<approx>ty C x\<PRIME>\<^isub>1 \<dots> x\<PRIME>\<^isub>r"}
- \end{center}
-
- \noindent
- holds under the assumptions that we have \mbox{@{text
- "x\<^isub>i \<approx>ty\<^isub>i x\<PRIME>\<^isub>i"}} whenever @{text "x\<^isub>i"}
- and @{text "x\<PRIME>\<^isub>i"} are recursive arguments of @{text C} and
- @{text "x\<^isub>i = x\<PRIME>\<^isub>i"} whenever they are non-recursive arguments. We can prove this
- implication by applying the corresponding rule in our $\alpha$-equivalence
- definition and by establishing the following auxiliary implications %facts
- %
- \begin{equation}\label{fnresp}
- \mbox{%
- \begin{tabular}{ll@ {\hspace{7mm}}ll}
- \mbox{\it (i)} & @{text "x \<approx>ty\<^isub>i x\<PRIME>"}~~@{text "\<Rightarrow>"}~~@{text "fa_ty\<^isub>i x = fa_ty\<^isub>i x\<PRIME>"} &
- \mbox{\it (iii)} & @{text "x \<approx>ty\<^isub>j x\<PRIME>"}~~@{text "\<Rightarrow>"}~~@{text "bn\<^isub>j x = bn\<^isub>j x\<PRIME>"}\\
-
- \mbox{\it (ii)} & @{text "x \<approx>ty\<^isub>j x\<PRIME>"}~~@{text "\<Rightarrow>"}~~@{text "fa_bn\<^isub>j x = fa_bn\<^isub>j x\<PRIME>"} &
- \mbox{\it (iv)} & @{text "x \<approx>ty\<^isub>j x\<PRIME>"}~~@{text "\<Rightarrow>"}~~@{text "x \<approx>bn\<^isub>j x\<PRIME>"}\\
- \end{tabular}}
- \end{equation}
-
- \noindent
- They can be established by induction on @{text "\<approx>ty"}$_{1..n}$. Whereas the first,
- second and last implication are true by how we stated our definitions, the
- third \emph{only} holds because of our restriction
- imposed on the form of the binding functions---namely \emph{not} returning
- any bound atoms. In Ott, in contrast, the user may
- define @{text "bn"}$_{1..m}$ so that they return bound
- atoms and in this case the third implication is \emph{not} true. A
- result is that the lifting of the corresponding binding functions in Ott to $\alpha$-equated
- terms is impossible.
-
- Having established respectfulness for the raw term-constructors, the
- quotient package is able to automatically deduce \eqref{distinctalpha} from
- \eqref{distinctraw}. Having the facts \eqref{fnresp} at our disposal, we can
- also lift properties that characterise when two raw terms of the form
- %
- \begin{center}
- @{text "C x\<^isub>1 \<dots> x\<^isub>r \<approx>ty C x\<PRIME>\<^isub>1 \<dots> x\<PRIME>\<^isub>r"}
- \end{center}
-
- \noindent
- are $\alpha$-equivalent. This gives us conditions when the corresponding
- $\alpha$-equated terms are \emph{equal}, namely
- %
- %\begin{center}
- @{text "C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>r = C\<^sup>\<alpha> x\<PRIME>\<^isub>1 \<dots> x\<PRIME>\<^isub>r"}.
- %\end{center}
- %
- %\noindent
- We call these conditions as \emph{quasi-injectivity}. They correspond to
- the premises in our $\alpha$-equivalence relations.
-
- Next we can lift the permutation
- operations defined in \eqref{ceqvt}. In order to make this
- lifting to go through, we have to show that the permutation operations are respectful.
- This amounts to showing that the
- $\alpha$-equivalence relations are equivariant \cite{HuffmanUrban10}.
- %, which we already established
- %in Lemma~\ref{equiv}.
- As a result we can add the equations
- %
- \begin{equation}\label{calphaeqvt}
- @{text "p \<bullet> (C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>r) = C\<^sup>\<alpha> (p \<bullet> x\<^isub>1) \<dots> (p \<bullet> x\<^isub>r)"}
- \end{equation}
-
- \noindent
- to our infrastructure. In a similar fashion we can lift the defining equations
- of the free-atom functions @{text "fn_ty\<AL>"}$_{1..n}$ and
- @{text "fa_bn\<AL>"}$_{1..m}$ as well as of the binding functions @{text
- "bn\<AL>"}$_{1..m}$ and the size functions @{text "size_ty\<AL>"}$_{1..n}$.
- The latter are defined automatically for the raw types @{text "ty"}$_{1..n}$
- by the datatype package of Isabelle/HOL.
-
- Finally we can add to our infrastructure a cases lemma (explained in the next section)
- and a structural induction principle
- for the types @{text "ty\<AL>"}$_{1..n}$. The conclusion of the induction principle is
- of the form
- %
- %\begin{equation}\label{weakinduct}
- \mbox{@{text "P\<^isub>1 x\<^isub>1 \<and> \<dots> \<and> P\<^isub>n x\<^isub>n "}}
- %\end{equation}
- %
- %\noindent
- whereby the @{text P}$_{1..n}$ are predicates and the @{text x}$_{1..n}$
- have types @{text "ty\<AL>"}$_{1..n}$. This induction principle has for each
- term constructor @{text "C"}$^\alpha$ a premise of the form
- %
- \begin{equation}\label{weakprem}
- \mbox{@{text "\<forall>x\<^isub>1\<dots>x\<^isub>r. P\<^isub>i x\<^isub>i \<and> \<dots> \<and> P\<^isub>j x\<^isub>j \<Rightarrow> P (C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>r)"}}
- \end{equation}
-
- \noindent
- in which the @{text "x"}$_{i..j}$ @{text "\<subseteq>"} @{text "x"}$_{1..r}$ are
- the recursive arguments of @{text "C\<AL>"}.
-
- By working now completely on the $\alpha$-equated level, we
- can first show that the free-atom functions and binding functions are
- equivariant, namely
- %
- \begin{center}
- \begin{tabular}{rcl@ {\hspace{10mm}}rcl}
- @{text "p \<bullet> (fa_ty\<AL>\<^isub>i x)"} & $=$ & @{text "fa_ty\<AL>\<^isub>i (p \<bullet> x)"} &
- @{text "p \<bullet> (bn\<AL>\<^isub>j x)"} & $=$ & @{text "bn\<AL>\<^isub>j (p \<bullet> x)"}\\
- @{text "p \<bullet> (fa_bn\<AL>\<^isub>j x)"} & $=$ & @{text "fa_bn\<AL>\<^isub>j (p \<bullet> x)"}\\
- \end{tabular}
- \end{center}
- %
- \noindent
- These properties can be established using the induction principle for the types @{text "ty\<AL>"}$_{1..n}$.
- %%in \eqref{weakinduct}.
- Having these equivariant properties established, we can
- show that the support of term-constructors @{text "C\<^sup>\<alpha>"} is included in
- the support of its arguments, that means
-
- \begin{center}
- @{text "supp (C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>r) \<subseteq> (supp x\<^isub>1 \<union> \<dots> \<union> supp x\<^isub>r)"}
- \end{center}
-
- \noindent
- holds. This allows us to prove by induction that
- every @{text x} of type @{text "ty\<AL>"}$_{1..n}$ is finitely supported.
- %This can be again shown by induction
- %over @{text "ty\<AL>"}$_{1..n}$.
- Lastly, we can show that the support of
- elements in @{text "ty\<AL>"}$_{1..n}$ is the same as @{text "fa_ty\<AL>"}$_{1..n}$.
- This fact is important in a nominal setting, but also provides evidence
- that our notions of free-atoms and $\alpha$-equivalence are correct.
-
- \begin{theorem}
- For @{text "x"}$_{1..n}$ with type @{text "ty\<AL>"}$_{1..n}$, we have
- @{text "supp x\<^isub>i = fa_ty\<AL>\<^isub>i x\<^isub>i"}.
- \end{theorem}
-
- \begin{proof}
- The proof is by induction. In each case
- we unfold the definition of @{text "supp"}, move the swapping inside the
- term-constructors and then use the quasi-injectivity lemmas in order to complete the
- proof. For the abstraction cases we use the facts derived in Theorem~\ref{suppabs}.
- \end{proof}
-
- \noindent
- To sum up this section, we can establish automatically a reasoning infrastructure
- for the types @{text "ty\<AL>"}$_{1..n}$
- by first lifting definitions from the raw level to the quotient level and
- then by establishing facts about these lifted definitions. All necessary proofs
- are generated automatically by custom ML-code.
-
- %This code can deal with
- %specifications such as the one shown in Figure~\ref{nominalcorehas} for Core-Haskell.
-
- %\begin{figure}[t!]
- %\begin{boxedminipage}{\linewidth}
- %\small
- %\begin{tabular}{l}
- %\isacommand{atom\_decl}~@{text "var cvar tvar"}\\[1mm]
- %\isacommand{nominal\_datatype}~@{text "tkind ="}\\
- %\phantom{$|$}~@{text "KStar"}~$|$~@{text "KFun tkind tkind"}\\
- %\isacommand{and}~@{text "ckind ="}\\
- %\phantom{$|$}~@{text "CKSim ty ty"}\\
- %\isacommand{and}~@{text "ty ="}\\
- %\phantom{$|$}~@{text "TVar tvar"}~$|$~@{text "T string"}~$|$~@{text "TApp ty ty"}\\
- %$|$~@{text "TFun string ty_list"}~%
- %$|$~@{text "TAll tv::tvar tkind ty::ty"} \isacommand{bind}~@{text "tv"}~\isacommand{in}~@{text ty}\\
- %$|$~@{text "TArr ckind ty"}\\
- %\isacommand{and}~@{text "ty_lst ="}\\
- %\phantom{$|$}~@{text "TNil"}~$|$~@{text "TCons ty ty_lst"}\\
- %\isacommand{and}~@{text "cty ="}\\
- %\phantom{$|$}~@{text "CVar cvar"}~%
- %$|$~@{text "C string"}~$|$~@{text "CApp cty cty"}~$|$~@{text "CFun string co_lst"}\\
- %$|$~@{text "CAll cv::cvar ckind cty::cty"} \isacommand{bind}~@{text "cv"}~\isacommand{in}~@{text cty}\\
- %$|$~@{text "CArr ckind cty"}~$|$~@{text "CRefl ty"}~$|$~@{text "CSym cty"}~$|$~@{text "CCirc cty cty"}\\
- %$|$~@{text "CAt cty ty"}~$|$~@{text "CLeft cty"}~$|$~@{text "CRight cty"}~$|$~@{text "CSim cty cty"}\\
- %$|$~@{text "CRightc cty"}~$|$~@{text "CLeftc cty"}~$|$~@{text "Coerce cty cty"}\\
- %\isacommand{and}~@{text "co_lst ="}\\
- %\phantom{$|$}@{text "CNil"}~$|$~@{text "CCons cty co_lst"}\\
- %\isacommand{and}~@{text "trm ="}\\
- %\phantom{$|$}~@{text "Var var"}~$|$~@{text "K string"}\\
- %$|$~@{text "LAM_ty tv::tvar tkind t::trm"} \isacommand{bind}~@{text "tv"}~\isacommand{in}~@{text t}\\
- %$|$~@{text "LAM_cty cv::cvar ckind t::trm"} \isacommand{bind}~@{text "cv"}~\isacommand{in}~@{text t}\\
- %$|$~@{text "App_ty trm ty"}~$|$~@{text "App_cty trm cty"}~$|$~@{text "App trm trm"}\\
- %$|$~@{text "Lam v::var ty t::trm"} \isacommand{bind}~@{text "v"}~\isacommand{in}~@{text t}\\
- %$|$~@{text "Let x::var ty trm t::trm"} \isacommand{bind}~@{text x}~\isacommand{in}~@{text t}\\
- %$|$~@{text "Case trm assoc_lst"}~$|$~@{text "Cast trm co"}\\
- %\isacommand{and}~@{text "assoc_lst ="}\\
- %\phantom{$|$}~@{text ANil}~%
- %$|$~@{text "ACons p::pat t::trm assoc_lst"} \isacommand{bind}~@{text "bv p"}~\isacommand{in}~@{text t}\\
- %\isacommand{and}~@{text "pat ="}\\
- %\phantom{$|$}~@{text "Kpat string tvtk_lst tvck_lst vt_lst"}\\
- %\isacommand{and}~@{text "vt_lst ="}\\
- %\phantom{$|$}~@{text VTNil}~$|$~@{text "VTCons var ty vt_lst"}\\
- %\isacommand{and}~@{text "tvtk_lst ="}\\
- %\phantom{$|$}~@{text TVTKNil}~$|$~@{text "TVTKCons tvar tkind tvtk_lst"}\\
- %\isacommand{and}~@{text "tvck_lst ="}\\
- %\phantom{$|$}~@{text TVCKNil}~$|$ @{text "TVCKCons cvar ckind tvck_lst"}\\
- %\isacommand{binder}\\
- %@{text "bv :: pat \<Rightarrow> atom list"}~\isacommand{and}~%
- %@{text "bv1 :: vt_lst \<Rightarrow> atom list"}~\isacommand{and}\\
- %@{text "bv2 :: tvtk_lst \<Rightarrow> atom list"}~\isacommand{and}~%
- %@{text "bv3 :: tvck_lst \<Rightarrow> atom list"}\\
- %\isacommand{where}\\
- %\phantom{$|$}~@{text "bv (K s tvts tvcs vs) = (bv3 tvts) @ (bv2 tvcs) @ (bv1 vs)"}\\
- %$|$~@{text "bv1 VTNil = []"}\\
- %$|$~@{text "bv1 (VTCons x ty tl) = (atom x)::(bv1 tl)"}\\
- %$|$~@{text "bv2 TVTKNil = []"}\\
- %$|$~@{text "bv2 (TVTKCons a ty tl) = (atom a)::(bv2 tl)"}\\
- %$|$~@{text "bv3 TVCKNil = []"}\\
- %$|$~@{text "bv3 (TVCKCons c cty tl) = (atom c)::(bv3 tl)"}\\
- %\end{tabular}
- %\end{boxedminipage}
- %\caption{The nominal datatype declaration for Core-Haskell. For the moment we
- %do not support nested types; therefore we explicitly have to unfold the
- %lists @{text "co_lst"}, @{text "assoc_lst"} and so on. This will be improved
- %in a future version of Nominal Isabelle. Apart from that, the
- %declaration follows closely the original in Figure~\ref{corehas}. The
- %point of our work is that having made such a declaration in Nominal Isabelle,
- %one obtains automatically a reasoning infrastructure for Core-Haskell.
- %\label{nominalcorehas}}
- %\end{figure}
-*}
-
-
-section {* Strong Induction Principles *}
-
-text {*
- In the previous section we derived induction principles for $\alpha$-equated terms.
- We call such induction principles \emph{weak}, because for a
- term-constructor \mbox{@{text "C\<^sup>\<alpha> x\<^isub>1\<dots>x\<^isub>r"}}
- the induction hypothesis requires us to establish the implications \eqref{weakprem}.
- The problem with these implications is that in general they are difficult to establish.
- The reason is that we cannot make any assumption about the bound atoms that might be in @{text "C\<^sup>\<alpha>"}.
- %%(for example we cannot assume the variable convention for them).
-
- In \cite{UrbanTasson05} we introduced a method for automatically
- strengthening weak induction principles for terms containing single
- binders. These stronger induction principles allow the user to make additional
- assumptions about bound atoms.
- %These additional assumptions amount to a formal
- %version of the informal variable convention for binders.
- To sketch how this strengthening extends to the case of multiple binders, we use as
- running example the term-constructors @{text "Lam"} and @{text "Let"}
- from example \eqref{letpat}. Instead of establishing @{text " P\<^bsub>trm\<^esub> t \<and> P\<^bsub>pat\<^esub> p"},
- the stronger induction principle for \eqref{letpat} establishes properties @{text " P\<^bsub>trm\<^esub> c t \<and> P\<^bsub>pat\<^esub> c p"}
- where the additional parameter @{text c} controls
- which freshness assumptions the binders should satisfy. For the two term constructors
- this means that the user has to establish in inductions the implications
- %
- \begin{center}
- \begin{tabular}{l}
- @{text "\<forall>a t c. {atom a} \<FRESH>\<^sup>* c \<and> (\<forall>d. P\<^bsub>trm\<^esub> d t) \<Rightarrow> P\<^bsub>trm\<^esub> c (Lam a t)"}\\
- @{text "\<forall>p t c. (set (bn p)) \<FRESH>\<^sup>* c \<and> (\<forall>d. P\<^bsub>pat\<^esub> d p) \<and> (\<forall>d. P\<^bsub>trm\<^esub> d t) \<and> \<Rightarrow> P\<^bsub>trm\<^esub> c (Let p t)"}\\%[-0mm]
- \end{tabular}
- \end{center}
-
- In \cite{UrbanTasson05} we showed how the weaker induction principles imply
- the stronger ones. This was done by some quite complicated, nevertheless automated,
- induction proof. In this paper we simplify this work by leveraging the automated proof
- methods from the function package of Isabelle/HOL.
- The reasoning principle these methods employ is well-founded induction.
- To use them in our setting, we have to discharge
- two proof obligations: one is that we have
- well-founded measures (for each type @{text "ty"}$^\alpha_{1..n}$) that decrease in
- every induction step and the other is that we have covered all cases.
- As measures we use the size functions
- @{text "size_ty"}$^\alpha_{1..n}$, which we lifted in the previous section and which are
- all well-founded. %It is straightforward to establish that these measures decrease
- %in every induction step.
-
- What is left to show is that we covered all cases. To do so, we use
- a \emph{cases lemma} derived for each type. For the terms in \eqref{letpat}
- this lemma is of the form
- %
- \begin{equation}\label{weakcases}
- \infer{@{text "P\<^bsub>trm\<^esub>"}}
- {\begin{array}{l@ {\hspace{9mm}}l}
- @{text "\<forall>x. t = Var x \<Rightarrow> P\<^bsub>trm\<^esub>"} & @{text "\<forall>a t'. t = Lam a t' \<Rightarrow> P\<^bsub>trm\<^esub>"}\\
- @{text "\<forall>t\<^isub>1 t\<^isub>2. t = App t\<^isub>1 t\<^isub>2 \<Rightarrow> P\<^bsub>trm\<^esub>"} & @{text "\<forall>p t'. t = Let p t' \<Rightarrow> P\<^bsub>trm\<^esub>"}\\
- \end{array}}\\[-1mm]
- \end{equation}
- %
- where we have a premise for each term-constructor.
- The idea behind such cases lemmas is that we can conclude with a property @{text "P\<^bsub>trm\<^esub>"},
- provided we can show that this property holds if we substitute for @{text "t"} all
- possible term-constructors.
-
- The only remaining difficulty is that in order to derive the stronger induction
- principles conveniently, the cases lemma in \eqref{weakcases} is too weak. For this note that
- in order to apply this lemma, we have to establish @{text "P\<^bsub>trm\<^esub>"} for \emph{all} @{text Lam}- and
- \emph{all} @{text Let}-terms.
- What we need instead is a cases lemma where we only have to consider terms that have
- binders that are fresh w.r.t.~a context @{text "c"}. This gives the implications
- %
- \begin{center}
- \begin{tabular}{l}
- @{text "\<forall>a t'. t = Lam a t' \<and> {atom a} \<FRESH>\<^sup>* c \<Rightarrow> P\<^bsub>trm\<^esub>"}\\
- @{text "\<forall>p t'. t = Let p t' \<and> (set (bn p)) \<FRESH>\<^sup>* c \<Rightarrow> P\<^bsub>trm\<^esub>"}\\%[-2mm]
- \end{tabular}
- \end{center}
- %
- \noindent
- which however can be relatively easily be derived from the implications in \eqref{weakcases}
- by a renaming using Properties \ref{supppermeq} and \ref{avoiding}. In the first case we know
- that @{text "{atom a} \<FRESH>\<^sup>* Lam a t"}. Property \eqref{avoiding} provides us therefore with
- a permutation @{text q}, such that @{text "{atom (q \<bullet> a)} \<FRESH>\<^sup>* c"} and
- @{text "supp (Lam a t) \<FRESH>\<^sup>* q"} hold.
- By using Property \ref{supppermeq}, we can infer from the latter
- that @{text "Lam (q \<bullet> a) (q \<bullet> t) = Lam a t"}
- and we are done with this case.
-
- The @{text Let}-case involving a (non-recursive) deep binder is a bit more complicated.
- The reason is that the we cannot apply Property \ref{avoiding} to the whole term @{text "Let p t"},
- because @{text p} might contain names bound by @{text bn}, but also some that are
- free. To solve this problem we have to introduce a permutation function that only
- permutes names bound by @{text bn} and leaves the other names unchanged. We do this again
- by lifting. For a
- clause @{text "bn (C x\<^isub>1 \<dots> x\<^isub>r) = rhs"}, we define
- %
- \begin{center}
- @{text "p \<bullet>\<^bsub>bn\<^esub> (C x\<^isub>1 \<dots> x\<^isub>r) \<equiv> C y\<^isub>1 \<dots> y\<^isub>r"} with
- $\begin{cases}
- \text{@{text "y\<^isub>i \<equiv> x\<^isub>i"} provided @{text "x\<^isub>i"} does not occur in @{text "rhs"}}\\
- \text{@{text "y\<^isub>i \<equiv> p \<bullet>\<^bsub>bn'\<^esub> x\<^isub>i"} provided @{text "bn' x\<^isub>i"} is in @{text "rhs"}}\\
- \text{@{text "y\<^isub>i \<equiv> p \<bullet> x\<^isub>i"} otherwise}
- \end{cases}$
- \end{center}
- %
- %\noindent
- %with @{text "y\<^isub>i"} determined as follows:
- %
- %\begin{center}
- %\begin{tabular}{c@ {\hspace{2mm}}p{0.9\textwidth}}
- %$\bullet$ & @{text "y\<^isub>i \<equiv> x\<^isub>i"} provided @{text "x\<^isub>i"} does not occur in @{text "rhs"}\\
- %$\bullet$ & @{text "y\<^isub>i \<equiv> p \<bullet>\<^bsub>bn'\<^esub> x\<^isub>i"} provided @{text "bn' x\<^isub>i"} is in @{text "rhs"}\\
- %$\bullet$ & @{text "y\<^isub>i \<equiv> p \<bullet> x\<^isub>i"} otherwise
- %\end{tabular}
- %\end{center}
- %
- \noindent
- Now Properties \ref{supppermeq} and \ref{avoiding} give us a permutation @{text q} such that
- @{text "(set (bn (q \<bullet>\<^bsub>bn\<^esub> p)) \<FRESH>\<^sup>* c"} holds and such that @{text "[q \<bullet>\<^bsub>bn\<^esub> p]\<^bsub>list\<^esub>.(q \<bullet> t)"}
- is equal to @{text "[p]\<^bsub>list\<^esub>. t"}. We can also show that @{text "(q \<bullet>\<^bsub>bn\<^esub> p) \<approx>\<^bsub>bn\<^esub> p"}.
- These facts establish that @{text "Let (q \<bullet>\<^bsub>bn\<^esub> p) (p \<bullet> t) = Let p t"}, as we need. This
- completes the non-trivial cases in \eqref{letpat} for strengthening the corresponding induction
- principle.
-
-
-
- %A natural question is
- %whether we can also strengthen the weak induction principles involving
- %the general binders presented here. We will indeed be able to so, but for this we need an
- %additional notion for permuting deep binders.
-
- %Given a binding function @{text "bn"} we define an auxiliary permutation
- %operation @{text "_ \<bullet>\<^bsub>bn\<^esub> _"} which permutes only bound arguments in a deep binder.
- %Assuming a clause of @{text bn} is given as
- %
- %\begin{center}
- %@{text "bn (C x\<^isub>1 \<dots> x\<^isub>r) = rhs"},
- %\end{center}
-
- %\noindent
- %then we define
- %
- %\begin{center}
- %@{text "p \<bullet>\<^bsub>bn\<^esub> (C x\<^isub>1 \<dots> x\<^isub>r) \<equiv> C y\<^isub>1 \<dots> y\<^isub>r"}
- %\end{center}
-
- %\noindent
- %with @{text "y\<^isub>i"} determined as follows:
- %
- %\begin{center}
- %\begin{tabular}{c@ {\hspace{2mm}}p{7cm}}
- %$\bullet$ & @{text "y\<^isub>i \<equiv> x\<^isub>i"} provided @{text "x\<^isub>i"} does not occur in @{text "rhs"}\\
- %$\bullet$ & @{text "y\<^isub>i \<equiv> p \<bullet>\<^bsub>bn'\<^esub> x\<^isub>i"} provided @{text "bn' x\<^isub>i"} is in @{text "rhs"}\\
- %$\bullet$ & @{text "y\<^isub>i \<equiv> p \<bullet> x\<^isub>i"} otherwise
- %\end{tabular}
- %\end{center}
-
- %\noindent
- %Using again the quotient package we can lift the @{text "_ \<bullet>\<^bsub>bn\<^esub> _"} function to
- %$\alpha$-equated terms. We can then prove the following two facts
-
- %\begin{lemma}\label{permutebn}
- %Given a binding function @{text "bn\<^sup>\<alpha>"} then for all @{text p}
- %{\it (i)} @{text "p \<bullet> (bn\<^sup>\<alpha> x) = bn\<^sup>\<alpha> (p \<bullet>\<AL>\<^bsub>bn\<^esub> x)"} and {\it (ii)}
- % @{text "fa_bn\<^isup>\<alpha> x = fa_bn\<^isup>\<alpha> (p \<bullet>\<AL>\<^bsub>bn\<^esub> x)"}.
- %\end{lemma}
-
- %\begin{proof}
- %By induction on @{text x}. The equations follow by simple unfolding
- %of the definitions.
- %\end{proof}
-
- %\noindent
- %The first property states that a permutation applied to a binding function is
- %equivalent to first permuting the binders and then calculating the bound
- %atoms. The second amounts to the fact that permuting the binders has no
- %effect on the free-atom function. The main point of this permutation
- %function, however, is that if we have a permutation that is fresh
- %for the support of an object @{text x}, then we can use this permutation
- %to rename the binders in @{text x}, without ``changing'' @{text x}. In case of the
- %@{text "Let"} term-constructor from the example shown
- %in \eqref{letpat} this means for a permutation @{text "r"}
- %%
- %\begin{equation}\label{renaming}
- %\begin{array}{l}
- %\mbox{if @{term "supp (Abs_lst (bn p) t\<^isub>2) \<sharp>* r"}}\\
- %\qquad\mbox{then @{text "Let p t\<^isub>1 t\<^isub>2 = Let (r \<bullet>\<AL>\<^bsub>bn_pat\<^esub> p) t\<^isub>1 (r \<bullet> t\<^isub>2)"}}
- %\end{array}
- %\end{equation}
-
- %\noindent
- %This fact will be crucial when establishing the strong induction principles below.
-
-
- %In our running example about @{text "Let"}, the strong induction
- %principle means that instead
- %of establishing the implication
- %
- %\begin{center}
- %@{text "\<forall>p t\<^isub>1 t\<^isub>2. P\<^bsub>pat\<^esub> p \<and> P\<^bsub>trm\<^esub> t\<^isub>1 \<and> P\<^bsub>trm\<^esub> t\<^isub>2 \<Rightarrow> P\<^bsub>trm\<^esub> (Let p t\<^isub>1 t\<^isub>2)"}
- %\end{center}
- %
- %\noindent
- %it is sufficient to establish the following implication
- %
- %\begin{equation}\label{strong}
- %\mbox{\begin{tabular}{l}
- %@{text "\<forall>p t\<^isub>1 t\<^isub>2 c."}\\
- %\hspace{5mm}@{text "set (bn p) #\<^sup>* c \<and>"}\\
- %\hspace{5mm}@{text "(\<forall>d. P\<^bsub>pat\<^esub> d p) \<and> (\<forall>d. P\<^bsub>trm\<^esub> d t\<^isub>1) \<and> (\<forall>d. P\<^bsub>trm\<^esub> d t\<^isub>2)"}\\
- %\hspace{15mm}@{text "\<Rightarrow> P\<^bsub>trm\<^esub> c (Let p t\<^isub>1 t\<^isub>2)"}
- %\end{tabular}}
- %\end{equation}
- %
- %\noindent
- %While this implication contains an additional argument, namely @{text c}, and
- %also additional universal quantifications, it is usually easier to establish.
- %The reason is that we have the freshness
- %assumption @{text "set (bn\<^sup>\<alpha> p) #\<^sup>* c"}, whereby @{text c} can be arbitrarily
- %chosen by the user as long as it has finite support.
- %
- %Let us now show how we derive the strong induction principles from the
- %weak ones. In case of the @{text "Let"}-example we derive by the weak
- %induction the following two properties
- %
- %\begin{equation}\label{hyps}
- %@{text "\<forall>q c. P\<^bsub>trm\<^esub> c (q \<bullet> t)"} \hspace{4mm}
- %@{text "\<forall>q\<^isub>1 q\<^isub>2 c. P\<^bsub>pat\<^esub> (q\<^isub>1 \<bullet>\<AL>\<^bsub>bn\<^esub> (q\<^isub>2 \<bullet> p))"}
- %\end{equation}
- %
- %\noindent
- %For the @{text Let} term-constructor we therefore have to establish @{text "P\<^bsub>trm\<^esub> c (q \<bullet> Let p t\<^isub>1 t\<^isub>2)"}
- %assuming \eqref{hyps} as induction hypotheses (the first for @{text t\<^isub>1} and @{text t\<^isub>2}).
- %By Property~\ref{avoiding} we
- %obtain a permutation @{text "r"} such that
- %
- %\begin{equation}\label{rprops}
- %@{term "(r \<bullet> set (bn (q \<bullet> p))) \<sharp>* c "}\hspace{4mm}
- %@{term "supp (Abs_lst (bn (q \<bullet> p)) (q \<bullet> t\<^isub>2)) \<sharp>* r"}
- %\end{equation}
- %
- %\noindent
- %hold. The latter fact and \eqref{renaming} give us
- %%
- %\begin{center}
- %\begin{tabular}{l}
- %@{text "Let (q \<bullet> p) (q \<bullet> t\<^isub>1) (q \<bullet> t\<^isub>2) ="} \\
- %\hspace{15mm}@{text "Let (r \<bullet>\<AL>\<^bsub>bn\<^esub> (q \<bullet> p)) (q \<bullet> t\<^isub>1) (r \<bullet> (q \<bullet> t\<^isub>2))"}
- %\end{tabular}
- %\end{center}
- %
- %\noindent
- %So instead of proving @{text "P\<^bsub>trm\<^esub> c (q \<bullet> Let p t\<^isub>1 t\<^isub>2)"}, we can equally
- %establish @{text "P\<^bsub>trm\<^esub> c (Let (r \<bullet>\<AL>\<^bsub>bn\<^esub> (q \<bullet> p)) (q \<bullet> t\<^isub>1) (r \<bullet> (q \<bullet> t\<^isub>2)))"}.
- %To do so, we will use the implication \eqref{strong} of the strong induction
- %principle, which requires us to discharge
- %the following four proof obligations:
- %%
- %\begin{center}
- %\begin{tabular}{rl}
- %{\it (i)} & @{text "set (bn (r \<bullet>\<AL>\<^bsub>bn\<^esub> (q \<bullet> p))) #\<^sup>* c"}\\
- %{\it (ii)} & @{text "\<forall>d. P\<^bsub>pat\<^esub> d (r \<bullet>\<AL>\<^bsub>bn\<^esub> (q \<bullet> p))"}\\
- %{\it (iii)} & @{text "\<forall>d. P\<^bsub>trm\<^esub> d (q \<bullet> t\<^isub>1)"}\\
- %{\it (iv)} & @{text "\<forall>d. P\<^bsub>trm\<^esub> d (r \<bullet> (q \<bullet> t\<^isub>2))"}\\
- %\end{tabular}
- %\end{center}
- %
- %\noindent
- %The first follows from \eqref{rprops} and Lemma~\ref{permutebn}.{\it (i)}; the
- %others from the induction hypotheses in \eqref{hyps} (in the fourth case
- %we have to use the fact that @{term "(r \<bullet> (q \<bullet> t\<^isub>2)) = (r + q) \<bullet> t\<^isub>2"}).
- %
- %Taking now the identity permutation @{text 0} for the permutations in \eqref{hyps},
- %we can establish our original goals, namely @{text "P\<^bsub>trm\<^esub> c t"} and \mbox{@{text "P\<^bsub>pat\<^esub> c p"}}.
- %This completes the proof showing that the weak induction principles imply
- %the strong induction principles.
-*}
-
-
-section {* Related Work\label{related} *}
-
-text {*
- To our knowledge the earliest usage of general binders in a theorem prover
- is described in \cite{NaraschewskiNipkow99} about a formalisation of the
- algorithm W. This formalisation implements binding in type-schemes using a
- de-Bruijn indices representation. Since type-schemes in W contain only a single
- place where variables are bound, different indices do not refer to different binders (as in the usual
- de-Bruijn representation), but to different bound variables. A similar idea
- has been recently explored for general binders in the locally nameless
- approach to binding \cite{chargueraud09}. There, de-Bruijn indices consist
- of two numbers, one referring to the place where a variable is bound, and the
- other to which variable is bound. The reasoning infrastructure for both
- representations of bindings comes for free in theorem provers like Isabelle/HOL or
- Coq, since the corresponding term-calculi can be implemented as ``normal''
- datatypes. However, in both approaches it seems difficult to achieve our
- fine-grained control over the ``semantics'' of bindings (i.e.~whether the
- order of binders should matter, or vacuous binders should be taken into
- account). %To do so, one would require additional predicates that filter out
- %unwanted terms. Our guess is that such predicates result in rather
- %intricate formal reasoning.
-
- Another technique for representing binding is higher-order abstract syntax
- (HOAS). %, which for example is implemented in the Twelf system.
- This %%representation
- technique supports very elegantly many aspects of \emph{single} binding, and
- impressive work has been done that uses HOAS for mechanising the metatheory
- of SML~\cite{LeeCraryHarper07}. We are, however, not aware how multiple
- binders of SML are represented in this work. Judging from the submitted
- Twelf-solution for the POPLmark challenge, HOAS cannot easily deal with
- binding constructs where the number of bound variables is not fixed. %For example
- In the second part of this challenge, @{text "Let"}s involve
- patterns that bind multiple variables at once. In such situations, HOAS
- seems to have to resort to the iterated-single-binders-approach with
- all the unwanted consequences when reasoning about the resulting terms.
-
- %Two formalisations involving general binders have been
- %performed in older
- %versions of Nominal Isabelle (one about Psi-calculi and one about algorithm W
- %\cite{BengtsonParow09,UrbanNipkow09}). Both
- %use the approach based on iterated single binders. Our experience with
- %the latter formalisation has been disappointing. The major pain arose from
- %the need to ``unbind'' variables. This can be done in one step with our
- %general binders described in this paper, but needs a cumbersome
- %iteration with single binders. The resulting formal reasoning turned out to
- %be rather unpleasant. The hope is that the extension presented in this paper
- %is a substantial improvement.
-
- The most closely related work to the one presented here is the Ott-tool
- \cite{ott-jfp} and the C$\alpha$ml language \cite{Pottier06}. Ott is a nifty
- front-end for creating \LaTeX{} documents from specifications of
- term-calculi involving general binders. For a subset of the specifications
- Ott can also generate theorem prover code using a raw representation of
- terms, and in Coq also a locally nameless representation. The developers of
- this tool have also put forward (on paper) a definition for
- $\alpha$-equivalence of terms that can be specified in Ott. This definition is
- rather different from ours, not using any nominal techniques. To our
- knowledge there is no concrete mathematical result concerning this
- notion of $\alpha$-equivalence. Also the definition for the
- notion of free variables
- is work in progress.
-
- Although we were heavily inspired by the syntax of Ott,
- its definition of $\alpha$-equi\-valence is unsuitable for our extension of
- Nominal Isabelle. First, it is far too complicated to be a basis for
- automated proofs implemented on the ML-level of Isabelle/HOL. Second, it
- covers cases of binders depending on other binders, which just do not make
- sense for our $\alpha$-equated terms. Third, it allows empty types that have no
- meaning in a HOL-based theorem prover. We also had to generalise slightly Ott's
- binding clauses. In Ott you specify binding clauses with a single body; we
- allow more than one. We have to do this, because this makes a difference
- for our notion of $\alpha$-equivalence in case of \isacommand{bind (set)} and
- \isacommand{bind (set+)}.
- %
- %Consider the examples
- %
- %\begin{center}
- %\begin{tabular}{@ {}l@ {\hspace{2mm}}l@ {}}
- %@{text "Foo\<^isub>1 xs::name fset t::trm s::trm"} &
- % \isacommand{bind (set)} @{text "xs"} \isacommand{in} @{text "t s"}\\
- %@{text "Foo\<^isub>2 xs::name fset t::trm s::trm"} &
- % \isacommand{bind (set)} @{text "xs"} \isacommand{in} @{text "t"},
- % \isacommand{bind (set)} @{text "xs"} \isacommand{in} @{text "s"}\\
- %\end{tabular}
- %\end{center}
- %
- %\noindent
- %In the first term-constructor we have a single
- %body that happens to be ``spread'' over two arguments; in the second term-constructor we have
- %two independent bodies in which the same variables are bound. As a result we
- %have
- %
- %\begin{center}
- %\begin{tabular}{r@ {\hspace{1.5mm}}c@ {\hspace{1.5mm}}l}
- %@{text "Foo\<^isub>1 {a, b} (a, b) (a, b)"} & $\not=$ &
- %@{text "Foo\<^isub>1 {a, b} (a, b) (b, a)"}\\
- %@{text "Foo\<^isub>2 {a, b} (a, b) (a, b)"} & $=$ &
- %@{text "Foo\<^isub>2 {a, b} (a, b) (b, a)"}\\
- %\end{tabular}
- %\end{center}
- %
- %\noindent
- %and therefore need the extra generality to be able to distinguish between
- %both specifications.
- Because of how we set up our definitions, we also had to impose some restrictions
- (like a single binding function for a deep binder) that are not present in Ott.
- %Our
- %expectation is that we can still cover many interesting term-calculi from
- %programming language research, for example Core-Haskell.
-
- Pottier presents in \cite{Pottier06} a language, called C$\alpha$ml, for
- representing terms with general binders inside OCaml. This language is
- implemented as a front-end that can be translated to OCaml with the help of
- a library. He presents a type-system in which the scope of general binders
- can be specified using special markers, written @{text "inner"} and
- @{text "outer"}. It seems our and his specifications can be
- inter-translated as long as ours use the binding mode
- \isacommand{bind} only.
- However, we have not proved this. Pottier gives a definition for
- $\alpha$-equivalence, which also uses a permutation operation (like ours).
- Still, this definition is rather different from ours and he only proves that
- it defines an equivalence relation. A complete
- reasoning infrastructure is well beyond the purposes of his language.
- Similar work for Haskell with similar results was reported by Cheney \cite{Cheney05a}.
-
- In a slightly different domain (programming with dependent types), the
- paper \cite{Altenkirch10} presents a calculus with a notion of
- $\alpha$-equivalence related to our binding mode \isacommand{bind (set+)}.
- The definition in \cite{Altenkirch10} is similar to the one by Pottier, except that it
- has a more operational flavour and calculates a partial (renaming) map.
- In this way, the definition can deal with vacuous binders. However, to our
- best knowledge, no concrete mathematical result concerning this
- definition of $\alpha$-equivalence has been proved.\\[-7mm]
-*}
-
-section {* Conclusion *}
-
-text {*
- We have presented an extension of Nominal Isabelle for dealing with
- general binders, that is term-constructors having multiple bound
- variables. For this extension we introduced new definitions of
- $\alpha$-equivalence and automated all necessary proofs in Isabelle/HOL.
- To specify general binders we used the specifications from Ott, but extended them
- in some places and restricted
- them in others so that they make sense in the context of $\alpha$-equated terms.
- We also introduced two binding modes (set and set+) that do not
- exist in Ott.
- We have tried out the extension with calculi such as Core-Haskell, type-schemes
- and approximately a dozen of other typical examples from programming
- language research~\cite{SewellBestiary}.
- %The code
- %will eventually become part of the next Isabelle distribution.\footnote{For the moment
- %it can be downloaded from the Mercurial repository linked at
- %\href{http://isabelle.in.tum.de/nominal/download}
- %{http://isabelle.in.tum.de/nominal/download}.}
-
- We have left out a discussion about how functions can be defined over
- $\alpha$-equated terms involving general binders. In earlier versions of Nominal
- Isabelle this turned out to be a thorny issue. We
- hope to do better this time by using the function package that has recently
- been implemented in Isabelle/HOL and also by restricting function
- definitions to equivariant functions (for them we can
- provide more automation).
-
- %There are some restrictions we imposed in this paper that we would like to lift in
- %future work. One is the exclusion of nested datatype definitions. Nested
- %datatype definitions allow one to specify, for instance, the function kinds
- %in Core-Haskell as @{text "TFun string (ty list)"} instead of the unfolded
- %version @{text "TFun string ty_list"} (see Figure~\ref{nominalcorehas}). To
- %achieve this, we need a slightly more clever implementation than we have at the moment.
-
- %A more interesting line of investigation is whether we can go beyond the
- %simple-minded form of binding functions that we adopted from Ott. At the moment, binding
- %functions can only return the empty set, a singleton atom set or unions
- %of atom sets (similarly for lists). It remains to be seen whether
- %properties like
- %%
- %\begin{center}
- %@{text "fa_ty x = bn x \<union> fa_bn x"}.
- %\end{center}
- %
- %\noindent
- %allow us to support more interesting binding functions.
- %
- %We have also not yet played with other binding modes. For example we can
- %imagine that there is need for a binding mode
- %where instead of lists, we abstract lists of distinct elements.
- %Once we feel confident about such binding modes, our implementation
- %can be easily extended to accommodate them.
- %
- \smallskip
- \noindent
- {\bf Acknowledgements:} %We are very grateful to Andrew Pitts for
- %many discussions about Nominal Isabelle.
- We thank Peter Sewell for
- making the informal notes \cite{SewellBestiary} available to us and
- also for patiently explaining some of the finer points of the Ott-tool.\\[-7mm]
- %Stephanie Weirich suggested to separate the subgrammars
- %of kinds and types in our Core-Haskell example. \\[-6mm]
-*}
-
-
-(*<*)
-end
-(*>*)
--- a/ESOP-Paper/ROOT.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-quick_and_dirty := true;
-no_document use_thys ["~~/src/HOL/Library/LaTeXsugar",
- "../Nominal/Nominal2"];
-use_thys ["Paper"];
\ No newline at end of file
--- a/ESOP-Paper/ROOTa.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-quick_and_dirty := true;
-no_document use_thys ["~~/src/HOL/Library/LaTeXsugar",
- "../Nominal/Nominal2"];
-use_thys ["Appendix"];
\ No newline at end of file
--- a/ESOP-Paper/document/llncs.cls Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1207 +0,0 @@
-% LLNCS DOCUMENT CLASS -- version 2.17 (12-Jul-2010)
-% Springer Verlag LaTeX2e support for Lecture Notes in Computer Science
-%
-%%
-%% \CharacterTable
-%% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
-%% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
-%% Digits \0\1\2\3\4\5\6\7\8\9
-%% Exclamation \! Double quote \" Hash (number) \#
-%% Dollar \$ Percent \% Ampersand \&
-%% Acute accent \' Left paren \( Right paren \)
-%% Asterisk \* Plus \+ Comma \,
-%% Minus \- Point \. Solidus \/
-%% Colon \: Semicolon \; Less than \<
-%% Equals \= Greater than \> Question mark \?
-%% Commercial at \@ Left bracket \[ Backslash \\
-%% Right bracket \] Circumflex \^ Underscore \_
-%% Grave accent \` Left brace \{ Vertical bar \|
-%% Right brace \} Tilde \~}
-%%
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-\ProvidesClass{llncs}[2010/07/12 v2.17
-^^J LaTeX document class for Lecture Notes in Computer Science]
-% Options
-\let\if@envcntreset\iffalse
-\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue}
-\DeclareOption{citeauthoryear}{\let\citeauthoryear=Y}
-\DeclareOption{oribibl}{\let\oribibl=Y}
-\let\if@custvec\iftrue
-\DeclareOption{orivec}{\let\if@custvec\iffalse}
-\let\if@envcntsame\iffalse
-\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue}
-\let\if@envcntsect\iffalse
-\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue}
-\let\if@runhead\iffalse
-\DeclareOption{runningheads}{\let\if@runhead\iftrue}
-
-\let\if@openright\iftrue
-\let\if@openbib\iffalse
-\DeclareOption{openbib}{\let\if@openbib\iftrue}
-
-% languages
-\let\switcht@@therlang\relax
-\def\ds@deutsch{\def\switcht@@therlang{\switcht@deutsch}}
-\def\ds@francais{\def\switcht@@therlang{\switcht@francais}}
-
-\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}}
-
-\ProcessOptions
-
-\LoadClass[twoside]{article}
-\RequirePackage{multicol} % needed for the list of participants, index
-\RequirePackage{aliascnt}
-
-\setlength{\textwidth}{12.2cm}
-\setlength{\textheight}{19.3cm}
-\renewcommand\@pnumwidth{2em}
-\renewcommand\@tocrmarg{3.5em}
-%
-\def\@dottedtocline#1#2#3#4#5{%
- \ifnum #1>\c@tocdepth \else
- \vskip \z@ \@plus.2\p@
- {\leftskip #2\relax \rightskip \@tocrmarg \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \parindent #2\relax\@afterindenttrue
- \interlinepenalty\@M
- \leavevmode
- \@tempdima #3\relax
- \advance\leftskip \@tempdima \null\nobreak\hskip -\leftskip
- {#4}\nobreak
- \leaders\hbox{$\m@th
- \mkern \@dotsep mu\hbox{.}\mkern \@dotsep
- mu$}\hfill
- \nobreak
- \hb@xt@\@pnumwidth{\hfil\normalfont \normalcolor #5}%
- \par}%
- \fi}
-%
-\def\switcht@albion{%
-\def\abstractname{Abstract.}
-\def\ackname{Acknowledgement.}
-\def\andname{and}
-\def\lastandname{\unskip, and}
-\def\appendixname{Appendix}
-\def\chaptername{Chapter}
-\def\claimname{Claim}
-\def\conjecturename{Conjecture}
-\def\contentsname{Table of Contents}
-\def\corollaryname{Corollary}
-\def\definitionname{Definition}
-\def\examplename{Example}
-\def\exercisename{Exercise}
-\def\figurename{Fig.}
-\def\keywordname{{\bf Keywords:}}
-\def\indexname{Index}
-\def\lemmaname{Lemma}
-\def\contriblistname{List of Contributors}
-\def\listfigurename{List of Figures}
-\def\listtablename{List of Tables}
-\def\mailname{{\it Correspondence to\/}:}
-\def\noteaddname{Note added in proof}
-\def\notename{Note}
-\def\partname{Part}
-\def\problemname{Problem}
-\def\proofname{Proof}
-\def\propertyname{Property}
-\def\propositionname{Proposition}
-\def\questionname{Question}
-\def\remarkname{Remark}
-\def\seename{see}
-\def\solutionname{Solution}
-\def\subclassname{{\it Subject Classifications\/}:}
-\def\tablename{Table}
-\def\theoremname{Theorem}}
-\switcht@albion
-% Names of theorem like environments are already defined
-% but must be translated if another language is chosen
-%
-% French section
-\def\switcht@francais{%\typeout{On parle francais.}%
- \def\abstractname{R\'esum\'e.}%
- \def\ackname{Remerciements.}%
- \def\andname{et}%
- \def\lastandname{ et}%
- \def\appendixname{Appendice}
- \def\chaptername{Chapitre}%
- \def\claimname{Pr\'etention}%
- \def\conjecturename{Hypoth\`ese}%
- \def\contentsname{Table des mati\`eres}%
- \def\corollaryname{Corollaire}%
- \def\definitionname{D\'efinition}%
- \def\examplename{Exemple}%
- \def\exercisename{Exercice}%
- \def\figurename{Fig.}%
- \def\keywordname{{\bf Mots-cl\'e:}}
- \def\indexname{Index}
- \def\lemmaname{Lemme}%
- \def\contriblistname{Liste des contributeurs}
- \def\listfigurename{Liste des figures}%
- \def\listtablename{Liste des tables}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}%
- \def\notename{Remarque}%
- \def\partname{Partie}%
- \def\problemname{Probl\`eme}%
- \def\proofname{Preuve}%
- \def\propertyname{Caract\'eristique}%
-%\def\propositionname{Proposition}%
- \def\questionname{Question}%
- \def\remarkname{Remarque}%
- \def\seename{voir}
- \def\solutionname{Solution}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tableau}%
- \def\theoremname{Th\'eor\`eme}%
-}
-%
-% German section
-\def\switcht@deutsch{%\typeout{Man spricht deutsch.}%
- \def\abstractname{Zusammenfassung.}%
- \def\ackname{Danksagung.}%
- \def\andname{und}%
- \def\lastandname{ und}%
- \def\appendixname{Anhang}%
- \def\chaptername{Kapitel}%
- \def\claimname{Behauptung}%
- \def\conjecturename{Hypothese}%
- \def\contentsname{Inhaltsverzeichnis}%
- \def\corollaryname{Korollar}%
-%\def\definitionname{Definition}%
- \def\examplename{Beispiel}%
- \def\exercisename{\"Ubung}%
- \def\figurename{Abb.}%
- \def\keywordname{{\bf Schl\"usselw\"orter:}}
- \def\indexname{Index}
-%\def\lemmaname{Lemma}%
- \def\contriblistname{Mitarbeiter}
- \def\listfigurename{Abbildungsverzeichnis}%
- \def\listtablename{Tabellenverzeichnis}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Nachtrag}%
- \def\notename{Anmerkung}%
- \def\partname{Teil}%
-%\def\problemname{Problem}%
- \def\proofname{Beweis}%
- \def\propertyname{Eigenschaft}%
-%\def\propositionname{Proposition}%
- \def\questionname{Frage}%
- \def\remarkname{Anmerkung}%
- \def\seename{siehe}
- \def\solutionname{L\"osung}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tabelle}%
-%\def\theoremname{Theorem}%
-}
-
-% Ragged bottom for the actual page
-\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil
-\global\let\@textbottom\relax}}
-
-\renewcommand\small{%
- \@setfontsize\small\@ixpt{11}%
- \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
- \abovedisplayshortskip \z@ \@plus2\p@
- \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
- \def\@listi{\leftmargin\leftmargini
- \parsep 0\p@ \@plus1\p@ \@minus\p@
- \topsep 8\p@ \@plus2\p@ \@minus4\p@
- \itemsep0\p@}%
- \belowdisplayskip \abovedisplayskip
-}
-
-\frenchspacing
-\widowpenalty=10000
-\clubpenalty=10000
-
-\setlength\oddsidemargin {63\p@}
-\setlength\evensidemargin {63\p@}
-\setlength\marginparwidth {90\p@}
-
-\setlength\headsep {16\p@}
-
-\setlength\footnotesep{7.7\p@}
-\setlength\textfloatsep{8mm\@plus 2\p@ \@minus 4\p@}
-\setlength\intextsep {8mm\@plus 2\p@ \@minus 2\p@}
-
-\setcounter{secnumdepth}{2}
-
-\newcounter {chapter}
-\renewcommand\thechapter {\@arabic\c@chapter}
-
-\newif\if@mainmatter \@mainmattertrue
-\newcommand\frontmatter{\cleardoublepage
- \@mainmatterfalse\pagenumbering{Roman}}
-\newcommand\mainmatter{\cleardoublepage
- \@mainmattertrue\pagenumbering{arabic}}
-\newcommand\backmatter{\if@openright\cleardoublepage\else\clearpage\fi
- \@mainmatterfalse}
-
-\renewcommand\part{\cleardoublepage
- \thispagestyle{empty}%
- \if@twocolumn
- \onecolumn
- \@tempswatrue
- \else
- \@tempswafalse
- \fi
- \null\vfil
- \secdef\@part\@spart}
-
-\def\@part[#1]#2{%
- \ifnum \c@secnumdepth >-2\relax
- \refstepcounter{part}%
- \addcontentsline{toc}{part}{\thepart\hspace{1em}#1}%
- \else
- \addcontentsline{toc}{part}{#1}%
- \fi
- \markboth{}{}%
- {\centering
- \interlinepenalty \@M
- \normalfont
- \ifnum \c@secnumdepth >-2\relax
- \huge\bfseries \partname~\thepart
- \par
- \vskip 20\p@
- \fi
- \Huge \bfseries #2\par}%
- \@endpart}
-\def\@spart#1{%
- {\centering
- \interlinepenalty \@M
- \normalfont
- \Huge \bfseries #1\par}%
- \@endpart}
-\def\@endpart{\vfil\newpage
- \if@twoside
- \null
- \thispagestyle{empty}%
- \newpage
- \fi
- \if@tempswa
- \twocolumn
- \fi}
-
-\newcommand\chapter{\clearpage
- \thispagestyle{empty}%
- \global\@topnum\z@
- \@afterindentfalse
- \secdef\@chapter\@schapter}
-\def\@chapter[#1]#2{\ifnum \c@secnumdepth >\m@ne
- \if@mainmatter
- \refstepcounter{chapter}%
- \typeout{\@chapapp\space\thechapter.}%
- \addcontentsline{toc}{chapter}%
- {\protect\numberline{\thechapter}#1}%
- \else
- \addcontentsline{toc}{chapter}{#1}%
- \fi
- \else
- \addcontentsline{toc}{chapter}{#1}%
- \fi
- \chaptermark{#1}%
- \addtocontents{lof}{\protect\addvspace{10\p@}}%
- \addtocontents{lot}{\protect\addvspace{10\p@}}%
- \if@twocolumn
- \@topnewpage[\@makechapterhead{#2}]%
- \else
- \@makechapterhead{#2}%
- \@afterheading
- \fi}
-\def\@makechapterhead#1{%
-% \vspace*{50\p@}%
- {\centering
- \ifnum \c@secnumdepth >\m@ne
- \if@mainmatter
- \large\bfseries \@chapapp{} \thechapter
- \par\nobreak
- \vskip 20\p@
- \fi
- \fi
- \interlinepenalty\@M
- \Large \bfseries #1\par\nobreak
- \vskip 40\p@
- }}
-\def\@schapter#1{\if@twocolumn
- \@topnewpage[\@makeschapterhead{#1}]%
- \else
- \@makeschapterhead{#1}%
- \@afterheading
- \fi}
-\def\@makeschapterhead#1{%
-% \vspace*{50\p@}%
- {\centering
- \normalfont
- \interlinepenalty\@M
- \Large \bfseries #1\par\nobreak
- \vskip 40\p@
- }}
-
-\renewcommand\section{\@startsection{section}{1}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {12\p@ \@plus 4\p@ \@minus 4\p@}%
- {\normalfont\large\bfseries\boldmath
- \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsection{\@startsection{subsection}{2}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {8\p@ \@plus 4\p@ \@minus 4\p@}%
- {\normalfont\normalsize\bfseries\boldmath
- \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {-0.5em \@plus -0.22em \@minus -0.1em}%
- {\normalfont\normalsize\bfseries\boldmath}}
-\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}%
- {-12\p@ \@plus -4\p@ \@minus -4\p@}%
- {-0.5em \@plus -0.22em \@minus -0.1em}%
- {\normalfont\normalsize\itshape}}
-\renewcommand\subparagraph[1]{\typeout{LLNCS warning: You should not use
- \string\subparagraph\space with this class}\vskip0.5cm
-You should not use \verb|\subparagraph| with this class.\vskip0.5cm}
-
-\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00}
-\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01}
-\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02}
-\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03}
-\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04}
-\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05}
-\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06}
-\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07}
-\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08}
-\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09}
-\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A}
-
-\let\footnotesize\small
-
-\if@custvec
-\def\vec#1{\mathchoice{\mbox{\boldmath$\displaystyle#1$}}
-{\mbox{\boldmath$\textstyle#1$}}
-{\mbox{\boldmath$\scriptstyle#1$}}
-{\mbox{\boldmath$\scriptscriptstyle#1$}}}
-\fi
-
-\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}}
-\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil
-\penalty50\hskip1em\null\nobreak\hfil\squareforqed
-\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi}
-
-\def\getsto{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr\gets\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-\gets\cr\to\cr}}}}}
-\def\lid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr<\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr<\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr<\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-<\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\gid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr>\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr>\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\grole{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr
->\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.8pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.3pt}<\cr}}}}}
-\def\bbbr{{\rm I\!R}} %reelle Zahlen
-\def\bbbm{{\rm I\!M}}
-\def\bbbn{{\rm I\!N}} %natuerliche Zahlen
-\def\bbbf{{\rm I\!F}}
-\def\bbbh{{\rm I\!H}}
-\def\bbbk{{\rm I\!K}}
-\def\bbbp{{\rm I\!P}}
-\def\bbbone{{\mathchoice {\rm 1\mskip-4mu l} {\rm 1\mskip-4mu l}
-{\rm 1\mskip-4.5mu l} {\rm 1\mskip-5mu l}}}
-\def\bbbc{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbq{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}}}
-\def\bbbt{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-T$}\hbox{\hbox to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbs{{\mathchoice
-{\setbox0=\hbox{$\displaystyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.5\wd0\vrule height0.45\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.4\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.55\wd0\vrule height0.45\ht0\hss}\box0}}}}
-\def\bbbz{{\mathchoice {\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\scriptstyle Z\kern-0.3em Z$}}
-{\hbox{$\mathsf\scriptscriptstyle Z\kern-0.2em Z$}}}}
-
-\let\ts\,
-
-\setlength\leftmargini {17\p@}
-\setlength\leftmargin {\leftmargini}
-\setlength\leftmarginii {\leftmargini}
-\setlength\leftmarginiii {\leftmargini}
-\setlength\leftmarginiv {\leftmargini}
-\setlength \labelsep {.5em}
-\setlength \labelwidth{\leftmargini}
-\addtolength\labelwidth{-\labelsep}
-
-\def\@listI{\leftmargin\leftmargini
- \parsep 0\p@ \@plus1\p@ \@minus\p@
- \topsep 8\p@ \@plus2\p@ \@minus4\p@
- \itemsep0\p@}
-\let\@listi\@listI
-\@listi
-\def\@listii {\leftmargin\leftmarginii
- \labelwidth\leftmarginii
- \advance\labelwidth-\labelsep
- \topsep 0\p@ \@plus2\p@ \@minus\p@}
-\def\@listiii{\leftmargin\leftmarginiii
- \labelwidth\leftmarginiii
- \advance\labelwidth-\labelsep
- \topsep 0\p@ \@plus\p@\@minus\p@
- \parsep \z@
- \partopsep \p@ \@plus\z@ \@minus\p@}
-
-\renewcommand\labelitemi{\normalfont\bfseries --}
-\renewcommand\labelitemii{$\m@th\bullet$}
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\def\tableofcontents{\chapter*{\contentsname\@mkboth{{\contentsname}}%
- {{\contentsname}}}
- \def\authcount##1{\setcounter{auco}{##1}\setcounter{@auth}{1}}
- \def\lastand{\ifnum\value{auco}=2\relax
- \unskip{} \andname\
- \else
- \unskip \lastandname\
- \fi}%
- \def\and{\stepcounter{@auth}\relax
- \ifnum\value{@auth}=\value{auco}%
- \lastand
- \else
- \unskip,
- \fi}%
- \@starttoc{toc}\if@restonecol\twocolumn\fi}
-
-\def\l@part#1#2{\addpenalty{\@secpenalty}%
- \addvspace{2em plus\p@}% % space above part line
- \begingroup
- \parindent \z@
- \rightskip \z@ plus 5em
- \hrule\vskip5pt
- \large % same size as for a contribution heading
- \bfseries\boldmath % set line in boldface
- \leavevmode % TeX command to enter horizontal mode.
- #1\par
- \vskip5pt
- \hrule
- \vskip1pt
- \nobreak % Never break after part entry
- \endgroup}
-
-\def\@dotsep{2}
-
-\let\phantomsection=\relax
-
-\def\hyperhrefextend{\ifx\hyper@anchor\@undefined\else
-{}\fi}
-
-\def\addnumcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{\protect\numberline
- {\thechapter}#3}{\thepage}\hyperhrefextend}}%
-\def\addcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{\thepage}\hyperhrefextend}}%
-\def\addcontentsmarkwop#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{0}\hyperhrefextend}}%
-
-\def\@adcmk[#1]{\ifcase #1 \or
-\def\@gtempa{\addnumcontentsmark}%
- \or \def\@gtempa{\addcontentsmark}%
- \or \def\@gtempa{\addcontentsmarkwop}%
- \fi\@gtempa{toc}{chapter}%
-}
-\def\addtocmark{%
-\phantomsection
-\@ifnextchar[{\@adcmk}{\@adcmk[3]}%
-}
-
-\def\l@chapter#1#2{\addpenalty{-\@highpenalty}
- \vskip 1.0em plus 1pt \@tempdima 1.5em \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- {\large\bfseries\boldmath#1}\ifx0#2\hfil\null
- \else
- \nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}%
- \fi\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@title#1#2{\addpenalty{-\@highpenalty}
- \addvspace{8pt plus 1pt}
- \@tempdima \z@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- #1\nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@author#1#2{\addpenalty{\@highpenalty}
- \@tempdima=15\p@ %\z@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima %\hskip -\leftskip
- \textit{#1}\par
- \penalty\@highpenalty \endgroup}
-
-\setcounter{tocdepth}{0}
-\newdimen\tocchpnum
-\newdimen\tocsecnum
-\newdimen\tocsectotal
-\newdimen\tocsubsecnum
-\newdimen\tocsubsectotal
-\newdimen\tocsubsubsecnum
-\newdimen\tocsubsubsectotal
-\newdimen\tocparanum
-\newdimen\tocparatotal
-\newdimen\tocsubparanum
-\tocchpnum=\z@ % no chapter numbers
-\tocsecnum=15\p@ % section 88. plus 2.222pt
-\tocsubsecnum=23\p@ % subsection 88.8 plus 2.222pt
-\tocsubsubsecnum=27\p@ % subsubsection 88.8.8 plus 1.444pt
-\tocparanum=35\p@ % paragraph 88.8.8.8 plus 1.666pt
-\tocsubparanum=43\p@ % subparagraph 88.8.8.8.8 plus 1.888pt
-\def\calctocindent{%
-\tocsectotal=\tocchpnum
-\advance\tocsectotal by\tocsecnum
-\tocsubsectotal=\tocsectotal
-\advance\tocsubsectotal by\tocsubsecnum
-\tocsubsubsectotal=\tocsubsectotal
-\advance\tocsubsubsectotal by\tocsubsubsecnum
-\tocparatotal=\tocsubsubsectotal
-\advance\tocparatotal by\tocparanum}
-\calctocindent
-
-\def\l@section{\@dottedtocline{1}{\tocchpnum}{\tocsecnum}}
-\def\l@subsection{\@dottedtocline{2}{\tocsectotal}{\tocsubsecnum}}
-\def\l@subsubsection{\@dottedtocline{3}{\tocsubsectotal}{\tocsubsubsecnum}}
-\def\l@paragraph{\@dottedtocline{4}{\tocsubsubsectotal}{\tocparanum}}
-\def\l@subparagraph{\@dottedtocline{5}{\tocparatotal}{\tocsubparanum}}
-
-\def\listoffigures{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}}
- \@starttoc{lof}\if@restonecol\twocolumn\fi}
-\def\l@figure{\@dottedtocline{1}{0em}{1.5em}}
-
-\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}}
- \@starttoc{lot}\if@restonecol\twocolumn\fi}
-\let\l@table\l@figure
-
-\renewcommand\listoffigures{%
- \section*{\listfigurename
- \@mkboth{\listfigurename}{\listfigurename}}%
- \@starttoc{lof}%
- }
-
-\renewcommand\listoftables{%
- \section*{\listtablename
- \@mkboth{\listtablename}{\listtablename}}%
- \@starttoc{lot}%
- }
-
-\ifx\oribibl\undefined
-\ifx\citeauthoryear\undefined
-\renewenvironment{thebibliography}[1]
- {\section*{\refname}
- \def\@biblabel##1{##1.}
- \small
- \list{\@biblabel{\@arabic\c@enumiv}}%
- {\settowidth\labelwidth{\@biblabel{#1}}%
- \leftmargin\labelwidth
- \advance\leftmargin\labelsep
- \if@openbib
- \advance\leftmargin\bibindent
- \itemindent -\bibindent
- \listparindent \itemindent
- \parsep \z@
- \fi
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{\@arabic\c@enumiv}}%
- \if@openbib
- \renewcommand\newblock{\par}%
- \else
- \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
- \fi
- \sloppy\clubpenalty4000\widowpenalty4000%
- \sfcode`\.=\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
-\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw
- {\let\protect\noexpand\immediate
- \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
-\newcount\@tempcntc
-\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi
- \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do
- {\@ifundefined
- {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries
- ?}\@warning
- {Citation `\@citeb' on page \thepage \space undefined}}%
- {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}%
- \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne
- \@citea\def\@citea{,}\hbox{\csname b@\@citeb\endcsname}%
- \else
- \advance\@tempcntb\@ne
- \ifnum\@tempcntb=\@tempcntc
- \else\advance\@tempcntb\m@ne\@citeo
- \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}}
-\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else
- \@citea\def\@citea{,\,\hskip\z@skip}%
- \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else
- {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else
- \def\@citea{--}\fi
- \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi}
-\else
-\renewenvironment{thebibliography}[1]
- {\section*{\refname}
- \small
- \list{}%
- {\settowidth\labelwidth{}%
- \leftmargin\parindent
- \itemindent=-\parindent
- \labelsep=\z@
- \if@openbib
- \advance\leftmargin\bibindent
- \itemindent -\bibindent
- \listparindent \itemindent
- \parsep \z@
- \fi
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{}}%
- \if@openbib
- \renewcommand\newblock{\par}%
- \else
- \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
- \fi
- \sloppy\clubpenalty4000\widowpenalty4000%
- \sfcode`\.=\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
- \def\@cite#1{#1}%
- \def\@lbibitem[#1]#2{\item[]\if@filesw
- {\def\protect##1{\string ##1\space}\immediate
- \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
- \fi
-\else
-\@cons\@openbib@code{\noexpand\small}
-\fi
-
-\def\idxquad{\hskip 10\p@}% space that divides entry from number
-
-\def\@idxitem{\par\hangindent 10\p@}
-
-\def\subitem{\par\setbox0=\hbox{--\enspace}% second order
- \noindent\hangindent\wd0\box0}% index entry
-
-\def\subsubitem{\par\setbox0=\hbox{--\,--\enspace}% third
- \noindent\hangindent\wd0\box0}% order index entry
-
-\def\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax}
-
-\renewenvironment{theindex}
- {\@mkboth{\indexname}{\indexname}%
- \thispagestyle{empty}\parindent\z@
- \parskip\z@ \@plus .3\p@\relax
- \let\item\par
- \def\,{\relax\ifmmode\mskip\thinmuskip
- \else\hskip0.2em\ignorespaces\fi}%
- \normalfont\small
- \begin{multicols}{2}[\@makeschapterhead{\indexname}]%
- }
- {\end{multicols}}
-
-\renewcommand\footnoterule{%
- \kern-3\p@
- \hrule\@width 2truecm
- \kern2.6\p@}
- \newdimen\fnindent
- \fnindent1em
-\long\def\@makefntext#1{%
- \parindent \fnindent%
- \leftskip \fnindent%
- \noindent
- \llap{\hb@xt@1em{\hss\@makefnmark\ }}\ignorespaces#1}
-
-\long\def\@makecaption#1#2{%
- \small
- \vskip\abovecaptionskip
- \sbox\@tempboxa{{\bfseries #1.} #2}%
- \ifdim \wd\@tempboxa >\hsize
- {\bfseries #1.} #2\par
- \else
- \global \@minipagefalse
- \hb@xt@\hsize{\hfil\box\@tempboxa\hfil}%
- \fi
- \vskip\belowcaptionskip}
-
-\def\fps@figure{htbp}
-\def\fnum@figure{\figurename\thinspace\thefigure}
-\def \@floatboxreset {%
- \reset@font
- \small
- \@setnobreak
- \@setminipage
-}
-\def\fps@table{htbp}
-\def\fnum@table{\tablename~\thetable}
-\renewenvironment{table}
- {\setlength\abovecaptionskip{0\p@}%
- \setlength\belowcaptionskip{10\p@}%
- \@float{table}}
- {\end@float}
-\renewenvironment{table*}
- {\setlength\abovecaptionskip{0\p@}%
- \setlength\belowcaptionskip{10\p@}%
- \@dblfloat{table}}
- {\end@dblfloat}
-
-\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname
- ext@#1\endcsname}{#1}{\protect\numberline{\csname
- the#1\endcsname}{\ignorespaces #2}}\begingroup
- \@parboxrestore
- \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
- \endgroup}
-
-% LaTeX does not provide a command to enter the authors institute
-% addresses. The \institute command is defined here.
-
-\newcounter{@inst}
-\newcounter{@auth}
-\newcounter{auco}
-\newdimen\instindent
-\newbox\authrun
-\newtoks\authorrunning
-\newtoks\tocauthor
-\newbox\titrun
-\newtoks\titlerunning
-\newtoks\toctitle
-
-\def\clearheadinfo{\gdef\@author{No Author Given}%
- \gdef\@title{No Title Given}%
- \gdef\@subtitle{}%
- \gdef\@institute{No Institute Given}%
- \gdef\@thanks{}%
- \global\titlerunning={}\global\authorrunning={}%
- \global\toctitle={}\global\tocauthor={}}
-
-\def\institute#1{\gdef\@institute{#1}}
-
-\def\institutename{\par
- \begingroup
- \parskip=\z@
- \parindent=\z@
- \setcounter{@inst}{1}%
- \def\and{\par\stepcounter{@inst}%
- \noindent$^{\the@inst}$\enspace\ignorespaces}%
- \setbox0=\vbox{\def\thanks##1{}\@institute}%
- \ifnum\c@@inst=1\relax
- \gdef\fnnstart{0}%
- \else
- \xdef\fnnstart{\c@@inst}%
- \setcounter{@inst}{1}%
- \noindent$^{\the@inst}$\enspace
- \fi
- \ignorespaces
- \@institute\par
- \endgroup}
-
-\def\@fnsymbol#1{\ensuremath{\ifcase#1\or\star\or{\star\star}\or
- {\star\star\star}\or \dagger\or \ddagger\or
- \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger
- \or \ddagger\ddagger \else\@ctrerr\fi}}
-
-\def\inst#1{\unskip$^{#1}$}
-\def\fnmsep{\unskip$^,$}
-\def\email#1{{\tt#1}}
-\AtBeginDocument{\@ifundefined{url}{\def\url#1{#1}}{}%
-\@ifpackageloaded{babel}{%
-\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}%
-\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}%
-\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}%
-}{\switcht@@therlang}%
-\providecommand{\keywords}[1]{\par\addvspace\baselineskip
-\noindent\keywordname\enspace\ignorespaces#1}%
-}
-\def\homedir{\~{ }}
-
-\def\subtitle#1{\gdef\@subtitle{#1}}
-\clearheadinfo
-%
-%%% to avoid hyperref warnings
-\providecommand*{\toclevel@author}{999}
-%%% to make title-entry parent of section-entries
-\providecommand*{\toclevel@title}{0}
-%
-\renewcommand\maketitle{\newpage
-\phantomsection
- \refstepcounter{chapter}%
- \stepcounter{section}%
- \setcounter{section}{0}%
- \setcounter{subsection}{0}%
- \setcounter{figure}{0}
- \setcounter{table}{0}
- \setcounter{equation}{0}
- \setcounter{footnote}{0}%
- \begingroup
- \parindent=\z@
- \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
- \if@twocolumn
- \ifnum \col@number=\@ne
- \@maketitle
- \else
- \twocolumn[\@maketitle]%
- \fi
- \else
- \newpage
- \global\@topnum\z@ % Prevents figures from going at top of page.
- \@maketitle
- \fi
- \thispagestyle{empty}\@thanks
-%
- \def\\{\unskip\ \ignorespaces}\def\inst##1{\unskip{}}%
- \def\thanks##1{\unskip{}}\def\fnmsep{\unskip}%
- \instindent=\hsize
- \advance\instindent by-\headlineindent
- \if!\the\toctitle!\addcontentsline{toc}{title}{\@title}\else
- \addcontentsline{toc}{title}{\the\toctitle}\fi
- \if@runhead
- \if!\the\titlerunning!\else
- \edef\@title{\the\titlerunning}%
- \fi
- \global\setbox\titrun=\hbox{\small\rm\unboldmath\ignorespaces\@title}%
- \ifdim\wd\titrun>\instindent
- \typeout{Title too long for running head. Please supply}%
- \typeout{a shorter form with \string\titlerunning\space prior to
- \string\maketitle}%
- \global\setbox\titrun=\hbox{\small\rm
- Title Suppressed Due to Excessive Length}%
- \fi
- \xdef\@title{\copy\titrun}%
- \fi
-%
- \if!\the\tocauthor!\relax
- {\def\and{\noexpand\protect\noexpand\and}%
- \protected@xdef\toc@uthor{\@author}}%
- \else
- \def\\{\noexpand\protect\noexpand\newline}%
- \protected@xdef\scratch{\the\tocauthor}%
- \protected@xdef\toc@uthor{\scratch}%
- \fi
- \addtocontents{toc}{\noexpand\protect\noexpand\authcount{\the\c@auco}}%
- \addcontentsline{toc}{author}{\toc@uthor}%
- \if@runhead
- \if!\the\authorrunning!
- \value{@inst}=\value{@auth}%
- \setcounter{@auth}{1}%
- \else
- \edef\@author{\the\authorrunning}%
- \fi
- \global\setbox\authrun=\hbox{\small\unboldmath\@author\unskip}%
- \ifdim\wd\authrun>\instindent
- \typeout{Names of authors too long for running head. Please supply}%
- \typeout{a shorter form with \string\authorrunning\space prior to
- \string\maketitle}%
- \global\setbox\authrun=\hbox{\small\rm
- Authors Suppressed Due to Excessive Length}%
- \fi
- \xdef\@author{\copy\authrun}%
- \markboth{\@author}{\@title}%
- \fi
- \endgroup
- \setcounter{footnote}{\fnnstart}%
- \clearheadinfo}
-%
-\def\@maketitle{\newpage
- \markboth{}{}%
- \def\lastand{\ifnum\value{@inst}=2\relax
- \unskip{} \andname\
- \else
- \unskip \lastandname\
- \fi}%
- \def\and{\stepcounter{@auth}\relax
- \ifnum\value{@auth}=\value{@inst}%
- \lastand
- \else
- \unskip,
- \fi}%
- \begin{center}%
- \let\newline\\
- {\Large \bfseries\boldmath
- \pretolerance=10000
- \@title \par}\vskip .8cm
-\if!\@subtitle!\else {\large \bfseries\boldmath
- \vskip -.65cm
- \pretolerance=10000
- \@subtitle \par}\vskip .8cm\fi
- \setbox0=\vbox{\setcounter{@auth}{1}\def\and{\stepcounter{@auth}}%
- \def\thanks##1{}\@author}%
- \global\value{@inst}=\value{@auth}%
- \global\value{auco}=\value{@auth}%
- \setcounter{@auth}{1}%
-{\lineskip .5em
-\noindent\ignorespaces
-\@author\vskip.35cm}
- {\small\institutename}
- \end{center}%
- }
-
-% definition of the "\spnewtheorem" command.
-%
-% Usage:
-%
-% \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font}
-% or \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font}
-% or \spnewtheorem*{env_nam}{caption}{cap_font}{body_font}
-%
-% New is "cap_font" and "body_font". It stands for
-% fontdefinition of the caption and the text itself.
-%
-% "\spnewtheorem*" gives a theorem without number.
-%
-% A defined spnewthoerem environment is used as described
-% by Lamport.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\def\@thmcountersep{}
-\def\@thmcounterend{.}
-
-\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}}
-
-% definition of \spnewtheorem with number
-
-\def\@spnthm#1#2{%
- \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}}
-\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}}
-
-\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}\@addtoreset{#1}{#3}%
- \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand
- \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}%
- \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spothm#1[#2]#3#4#5{%
- \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}%
- {\expandafter\@ifdefinable\csname #1\endcsname
- {\newaliascnt{#1}{#2}%
- \expandafter\xdef\csname #1name\endcsname{#3}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}}
-
-\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\refstepcounter{#1}%
-\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}}
-
-\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}%
- \ignorespaces}
-
-\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname
- the#1\endcsname}{#5}{#3}{#4}\ignorespaces}
-
-\def\@spbegintheorem#1#2#3#4{\trivlist
- \item[\hskip\labelsep{#3#1\ #2\@thmcounterend}]#4}
-
-\def\@spopargbegintheorem#1#2#3#4#5{\trivlist
- \item[\hskip\labelsep{#4#1\ #2}]{#4(#3)\@thmcounterend\ }#5}
-
-% definition of \spnewtheorem* without number
-
-\def\@sthm#1#2{\@Ynthm{#1}{#2}}
-
-\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}}
-
-\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces}
-
-\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1}
- {#4}{#2}{#3}\ignorespaces}
-
-\def\@Begintheorem#1#2#3{#3\trivlist
- \item[\hskip\labelsep{#2#1\@thmcounterend}]}
-
-\def\@Opargbegintheorem#1#2#3#4{#4\trivlist
- \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }}
-
-\if@envcntsect
- \def\@thmcountersep{.}
- \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape}
-\else
- \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape}
- \if@envcntreset
- \@addtoreset{theorem}{section}
- \else
- \@addtoreset{theorem}{chapter}
- \fi
-\fi
-
-%definition of divers theorem environments
-\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily}
-\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily}
-\if@envcntsame % alle Umgebungen wie Theorem.
- \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}}
-\else % alle Umgebungen mit eigenem Zaehler
- \if@envcntsect % mit section numeriert
- \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}}
- \else % nicht mit section numeriert
- \if@envcntreset
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{section}}
- \else
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{chapter}}%
- \fi
- \fi
-\fi
-\spn@wtheorem{case}{Case}{\itshape}{\rmfamily}
-\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily}
-\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape}
-\spn@wtheorem{definition}{Definition}{\bfseries}{\itshape}
-\spn@wtheorem{example}{Example}{\itshape}{\rmfamily}
-\spn@wtheorem{exercise}{Exercise}{\itshape}{\rmfamily}
-\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape}
-\spn@wtheorem{note}{Note}{\itshape}{\rmfamily}
-\spn@wtheorem{problem}{Problem}{\itshape}{\rmfamily}
-\spn@wtheorem{property}{Property}{\itshape}{\rmfamily}
-\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape}
-\spn@wtheorem{question}{Question}{\itshape}{\rmfamily}
-\spn@wtheorem{solution}{Solution}{\itshape}{\rmfamily}
-\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily}
-
-\def\@takefromreset#1#2{%
- \def\@tempa{#1}%
- \let\@tempd\@elt
- \def\@elt##1{%
- \def\@tempb{##1}%
- \ifx\@tempa\@tempb\else
- \@addtoreset{##1}{#2}%
- \fi}%
- \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
- \expandafter\def\csname cl@#2\endcsname{}%
- \@tempc
- \let\@elt\@tempd}
-
-\def\theopargself{\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5}
- \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
- \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }}
- }
-
-\renewenvironment{abstract}{%
- \list{}{\advance\topsep by0.35cm\relax\small
- \leftmargin=1cm
- \labelwidth=\z@
- \listparindent=\z@
- \itemindent\listparindent
- \rightmargin\leftmargin}\item[\hskip\labelsep
- \bfseries\abstractname]}
- {\endlist}
-
-\newdimen\headlineindent % dimension for space between
-\headlineindent=1.166cm % number and text of headings.
-
-\def\ps@headings{\let\@mkboth\@gobbletwo
- \let\@oddfoot\@empty\let\@evenfoot\@empty
- \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
- \leftmark\hfil}
- \def\@oddhead{\normalfont\small\hfil\rightmark\hspace{\headlineindent}%
- \llap{\thepage}}
- \def\chaptermark##1{}%
- \def\sectionmark##1{}%
- \def\subsectionmark##1{}}
-
-\def\ps@titlepage{\let\@mkboth\@gobbletwo
- \let\@oddfoot\@empty\let\@evenfoot\@empty
- \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
- \hfil}
- \def\@oddhead{\normalfont\small\hfil\hspace{\headlineindent}%
- \llap{\thepage}}
- \def\chaptermark##1{}%
- \def\sectionmark##1{}%
- \def\subsectionmark##1{}}
-
-\if@runhead\ps@headings\else
-\ps@empty\fi
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\endinput
-%end of file llncs.cls
--- a/ESOP-Paper/document/proof.sty Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,278 +0,0 @@
-% proof.sty (Proof Figure Macros)
-%
-% version 3.0 (for both LaTeX 2.09 and LaTeX 2e)
-% Mar 6, 1997
-% Copyright (C) 1990 -- 1997, Makoto Tatsuta (tatsuta@kusm.kyoto-u.ac.jp)
-%
-% This program is free software; you can redistribute it or modify
-% it under the terms of the GNU General Public License as published by
-% the Free Software Foundation; either versions 1, or (at your option)
-% any later version.
-%
-% This program 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.
-%
-% Usage:
-% In \documentstyle, specify an optional style `proof', say,
-% \documentstyle[proof]{article}.
-%
-% The following macros are available:
-%
-% In all the following macros, all the arguments such as
-% <Lowers> and <Uppers> are processed in math mode.
-%
-% \infer<Lower><Uppers>
-% draws an inference.
-%
-% Use & in <Uppers> to delimit upper formulae.
-% <Uppers> consists more than 0 formulae.
-%
-% \infer returns \hbox{ ... } or \vbox{ ... } and
-% sets \@LeftOffset and \@RightOffset globally.
-%
-% \infer[<Label>]<Lower><Uppers>
-% draws an inference labeled with <Label>.
-%
-% \infer*<Lower><Uppers>
-% draws a many step deduction.
-%
-% \infer*[<Label>]<Lower><Uppers>
-% draws a many step deduction labeled with <Label>.
-%
-% \infer=<Lower><Uppers>
-% draws a double-ruled deduction.
-%
-% \infer=[<Label>]<Lower><Uppers>
-% draws a double-ruled deduction labeled with <Label>.
-%
-% \deduce<Lower><Uppers>
-% draws an inference without a rule.
-%
-% \deduce[<Proof>]<Lower><Uppers>
-% draws a many step deduction with a proof name.
-%
-% Example:
-% If you want to write
-% B C
-% -----
-% A D
-% ----------
-% E
-% use
-% \infer{E}{
-% A
-% &
-% \infer{D}{B & C}
-% }
-%
-
-% Style Parameters
-
-\newdimen\inferLineSkip \inferLineSkip=2pt
-\newdimen\inferLabelSkip \inferLabelSkip=5pt
-\def\inferTabSkip{\quad}
-
-% Variables
-
-\newdimen\@LeftOffset % global
-\newdimen\@RightOffset % global
-\newdimen\@SavedLeftOffset % safe from users
-
-\newdimen\UpperWidth
-\newdimen\LowerWidth
-\newdimen\LowerHeight
-\newdimen\UpperLeftOffset
-\newdimen\UpperRightOffset
-\newdimen\UpperCenter
-\newdimen\LowerCenter
-\newdimen\UpperAdjust
-\newdimen\RuleAdjust
-\newdimen\LowerAdjust
-\newdimen\RuleWidth
-\newdimen\HLabelAdjust
-\newdimen\VLabelAdjust
-\newdimen\WidthAdjust
-
-\newbox\@UpperPart
-\newbox\@LowerPart
-\newbox\@LabelPart
-\newbox\ResultBox
-
-% Flags
-
-\newif\if@inferRule % whether \@infer draws a rule.
-\newif\if@DoubleRule % whether \@infer draws doulbe rules.
-\newif\if@ReturnLeftOffset % whether \@infer returns \@LeftOffset.
-\newif\if@MathSaved % whether inner math mode where \infer or
- % \deduce appears.
-
-% Special Fonts
-
-\def\DeduceSym{\vtop{\baselineskip4\p@ \lineskiplimit\z@
- \vbox{\hbox{.}\hbox{.}\hbox{.}}\hbox{.}}}
-
-% Math Save Macros
-%
-% \@SaveMath is called in the very begining of toplevel macros
-% which are \infer and \deduce.
-% \@RestoreMath is called in the very last before toplevel macros end.
-% Remark \infer and \deduce ends calling \@infer.
-
-\def\@SaveMath{\@MathSavedfalse \ifmmode \ifinner
- \relax $\relax \@MathSavedtrue \fi\fi }
-
-\def\@RestoreMath{\if@MathSaved \relax $\relax\fi }
-
-% Macros
-
-% Renaming @ifnextchar and @ifnch of LaTeX2e to @IFnextchar and @IFnch.
-
-\def\@IFnextchar#1#2#3{%
- \let\reserved@e=#1\def\reserved@a{#2}\def\reserved@b{#3}\futurelet
- \reserved@c\@IFnch}
-\def\@IFnch{\ifx \reserved@c \@sptoken \let\reserved@d\@xifnch
- \else \ifx \reserved@c \reserved@e\let\reserved@d\reserved@a\else
- \let\reserved@d\reserved@b\fi
- \fi \reserved@d}
-
-\def\@ifEmpty#1#2#3{\def\@tempa{\@empty}\def\@tempb{#1}\relax
- \ifx \@tempa \@tempb #2\else #3\fi }
-
-\def\infer{\@SaveMath \@IFnextchar *{\@inferSteps}{\relax
- \@IFnextchar ={\@inferDoubleRule}{\@inferOneStep}}}
-
-\def\@inferOneStep{\@inferRuletrue \@DoubleRulefalse
- \@IFnextchar [{\@infer}{\@infer[\@empty]}}
-
-\def\@inferDoubleRule={\@inferRuletrue \@DoubleRuletrue
- \@IFnextchar [{\@infer}{\@infer[\@empty]}}
-
-\def\@inferSteps*{\@IFnextchar [{\@@inferSteps}{\@@inferSteps[\@empty]}}
-
-\def\@@inferSteps[#1]{\@deduce{#1}[\DeduceSym]}
-
-\def\deduce{\@SaveMath \@IFnextchar [{\@deduce{\@empty}}
- {\@inferRulefalse \@infer[\@empty]}}
-
-% \@deduce<Proof Label>[<Proof>]<Lower><Uppers>
-
-\def\@deduce#1[#2]#3#4{\@inferRulefalse
- \@infer[\@empty]{#3}{\@SaveMath \@infer[{#1}]{#2}{#4}}}
-
-% \@infer[<Label>]<Lower><Uppers>
-% If \@inferRuletrue, it draws a rule and <Label> is right to
-% a rule. In this case, if \@DoubleRuletrue, it draws
-% double rules.
-%
-% Otherwise, draws no rule and <Label> is right to <Lower>.
-
-\def\@infer[#1]#2#3{\relax
-% Get parameters
- \if@ReturnLeftOffset \else \@SavedLeftOffset=\@LeftOffset \fi
- \setbox\@LabelPart=\hbox{$#1$}\relax
- \setbox\@LowerPart=\hbox{$#2$}\relax
-%
- \global\@LeftOffset=0pt
- \setbox\@UpperPart=\vbox{\tabskip=0pt \halign{\relax
- \global\@RightOffset=0pt \@ReturnLeftOffsettrue $##$&&
- \inferTabSkip
- \global\@RightOffset=0pt \@ReturnLeftOffsetfalse $##$\cr
- #3\cr}}\relax
-% Here is a little trick.
-% \@ReturnLeftOffsettrue(false) influences on \infer or
-% \deduce placed in ## locally
-% because of \@SaveMath and \@RestoreMath.
- \UpperLeftOffset=\@LeftOffset
- \UpperRightOffset=\@RightOffset
-% Calculate Adjustments
- \LowerWidth=\wd\@LowerPart
- \LowerHeight=\ht\@LowerPart
- \LowerCenter=0.5\LowerWidth
-%
- \UpperWidth=\wd\@UpperPart \advance\UpperWidth by -\UpperLeftOffset
- \advance\UpperWidth by -\UpperRightOffset
- \UpperCenter=\UpperLeftOffset
- \advance\UpperCenter by 0.5\UpperWidth
-%
- \ifdim \UpperWidth > \LowerWidth
- % \UpperCenter > \LowerCenter
- \UpperAdjust=0pt
- \RuleAdjust=\UpperLeftOffset
- \LowerAdjust=\UpperCenter \advance\LowerAdjust by -\LowerCenter
- \RuleWidth=\UpperWidth
- \global\@LeftOffset=\LowerAdjust
-%
- \else % \UpperWidth <= \LowerWidth
- \ifdim \UpperCenter > \LowerCenter
-%
- \UpperAdjust=0pt
- \RuleAdjust=\UpperCenter \advance\RuleAdjust by -\LowerCenter
- \LowerAdjust=\RuleAdjust
- \RuleWidth=\LowerWidth
- \global\@LeftOffset=\LowerAdjust
-%
- \else % \UpperWidth <= \LowerWidth
- % \UpperCenter <= \LowerCenter
-%
- \UpperAdjust=\LowerCenter \advance\UpperAdjust by -\UpperCenter
- \RuleAdjust=0pt
- \LowerAdjust=0pt
- \RuleWidth=\LowerWidth
- \global\@LeftOffset=0pt
-%
- \fi\fi
-% Make a box
- \if@inferRule
-%
- \setbox\ResultBox=\vbox{
- \moveright \UpperAdjust \box\@UpperPart
- \nointerlineskip \kern\inferLineSkip
- \if@DoubleRule
- \moveright \RuleAdjust \vbox{\hrule width\RuleWidth
- \kern 1pt\hrule width\RuleWidth}\relax
- \else
- \moveright \RuleAdjust \vbox{\hrule width\RuleWidth}\relax
- \fi
- \nointerlineskip \kern\inferLineSkip
- \moveright \LowerAdjust \box\@LowerPart }\relax
-%
- \@ifEmpty{#1}{}{\relax
-%
- \HLabelAdjust=\wd\ResultBox \advance\HLabelAdjust by -\RuleAdjust
- \advance\HLabelAdjust by -\RuleWidth
- \WidthAdjust=\HLabelAdjust
- \advance\WidthAdjust by -\inferLabelSkip
- \advance\WidthAdjust by -\wd\@LabelPart
- \ifdim \WidthAdjust < 0pt \WidthAdjust=0pt \fi
-%
- \VLabelAdjust=\dp\@LabelPart
- \advance\VLabelAdjust by -\ht\@LabelPart
- \VLabelAdjust=0.5\VLabelAdjust \advance\VLabelAdjust by \LowerHeight
- \advance\VLabelAdjust by \inferLineSkip
-%
- \setbox\ResultBox=\hbox{\box\ResultBox
- \kern -\HLabelAdjust \kern\inferLabelSkip
- \raise\VLabelAdjust \box\@LabelPart \kern\WidthAdjust}\relax
-%
- }\relax % end @ifEmpty
-%
- \else % \@inferRulefalse
-%
- \setbox\ResultBox=\vbox{
- \moveright \UpperAdjust \box\@UpperPart
- \nointerlineskip \kern\inferLineSkip
- \moveright \LowerAdjust \hbox{\unhbox\@LowerPart
- \@ifEmpty{#1}{}{\relax
- \kern\inferLabelSkip \unhbox\@LabelPart}}}\relax
- \fi
-%
- \global\@RightOffset=\wd\ResultBox
- \global\advance\@RightOffset by -\@LeftOffset
- \global\advance\@RightOffset by -\LowerWidth
- \if@ReturnLeftOffset \else \global\@LeftOffset=\@SavedLeftOffset \fi
-%
- \box\ResultBox
- \@RestoreMath
-}
--- a/ESOP-Paper/document/root.bib Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,319 +0,0 @@
-
-@Unpublished{KaliszykUrban11,
- author = {C.~Kaliszyk and C.~Urban},
- title = {{Q}uotients {R}evisited for {I}sabelle/{HOL}},
- note = {To appear in the Proc.~of the 26th ACM Symposium On Applied Computing},
- year = {2011}
-}
-
-@InProceedings{cheney05a,
- author = {J.~Cheney},
- title = {{S}crap your {N}ameplate ({F}unctional {P}earl)},
- booktitle = {Proc.~of the 10th ICFP Conference},
- pages = {180--191},
- year = {2005}
-}
-
-@Inproceedings{Altenkirch10,
- author = {T.~Altenkirch and N.~A.~Danielsson and A.~L\"oh and N.~Oury},
- title = {{PiSigma}: {D}ependent {T}ypes {W}ithout the {S}ugar},
- booktitle = "Proc.~of the 10th FLOPS Conference",
- year = 2010,
- series = "LNCS",
- pages = "40--55",
- volume = 6009
-}
-
-
-@InProceedings{ UrbanTasson05,
- author = "C. Urban and C. Tasson",
- title = "{N}ominal {T}echniques in {I}sabelle/{HOL}",
- booktitle = "Proc.~of the 20th CADE Conference",
- year = 2005,
- series = "LNCS",
- pages = "38--53",
- volume = 3632
-}
-
-@InProceedings{ UrbanBerghofer06,
- author = "C. Urban and S. Berghofer",
- title = "{A} {R}ecursion {C}ombinator for {N}ominal {D}atatypes {I}mplemented in {I}sabelle/{HOL}",
- booktitle = "Proc.~of the 3rd IJCAR Conference",
- year = 2006,
- series = "LNAI",
- volume = 4130,
- pages = "498--512"
-}
-
-@InProceedings{LeeCraryHarper07,
- author = {D.~K.~Lee and K.~Crary and R.~Harper},
- title = {{T}owards a {M}echanized {M}etatheory of {Standard ML}},
- booktitle = {Proc.~of the 34th POPL Symposium},
- year = 2007,
- pages = {173--184}
-}
-
-@Unpublished{chargueraud09,
- author = "A.~Chargu{\'e}raud",
- title = "{T}he {L}ocally {N}ameless {R}epresentation",
- Note = "To appear in J.~of Automated Reasoning."
-}
-
-@article{NaraschewskiNipkow99,
- author={W.~Naraschewski and T.~Nipkow},
- title={{T}ype {I}nference {V}erified: {A}lgorithm {W} in {Isabelle/HOL}},
- journal={J.~of Automated Reasoning},
- year=1999,
- volume=23,
- pages={299--318}}
-
-@InProceedings{Berghofer99,
- author = {S.~Berghofer and M.~Wenzel},
- title = {{I}nductive {D}atatypes in {HOL} - {L}essons {L}earned in
- {F}ormal-{L}ogic {E}ngineering},
- booktitle = {Proc.~of the 12th TPHOLs conference},
- pages = {19--36},
- year = 1999,
- volume = 1690,
- series = {LNCS}
-}
-
-@InProceedings{CoreHaskell,
- author = {M.~Sulzmann and M.~Chakravarty and S.~Peyton Jones and K.~Donnelly},
- title = {{S}ystem {F} with {T}ype {E}quality {C}oercions},
- booktitle = {Proc.~of the TLDI Workshop},
- pages = {53-66},
- year = {2007}
-}
-
-@inproceedings{cheney05,
- author = {J.~Cheney},
- title = {{T}oward a {G}eneral {T}heory of {N}ames: {B}inding and {S}cope},
- booktitle = {Proc.~of the 3rd MERLIN workshop},
- year = {2005},
- pages = {33-40}
-}
-
-@Unpublished{Pitts04,
- author = {A.~Pitts},
- title = {{N}otes on the {R}estriction {M}onad for {N}ominal {S}ets and {C}pos},
- note = {Unpublished notes for an invited talk given at CTCS},
- year = {2004}
-}
-
-@incollection{UrbanNipkow09,
- author = {C.~Urban and T.~Nipkow},
- title = {{N}ominal {V}erification of {A}lgorithm {W}},
- booktitle={From Semantics to Computer Science. Essays in Honour of Gilles Kahn},
- editor={G.~Huet and J.-J.~L{\'e}vy and G.~Plotkin},
- publisher={Cambridge University Press},
- pages={363--382},
- year=2009
-}
-
-@InProceedings{Homeier05,
- author = {P.~Homeier},
- title = {{A} {D}esign {S}tructure for {H}igher {O}rder {Q}uotients},
- booktitle = {Proc.~of the 18th TPHOLs Conference},
- pages = {130--146},
- year = {2005},
- volume = {3603},
- series = {LNCS}
-}
-
-@article{ott-jfp,
- author = {P.~Sewell and
- F.~Z.~Nardelli and
- S.~Owens and
- G.~Peskine and
- T.~Ridge and
- S.~Sarkar and
- R.~Strni\v{s}a},
- title = {{Ott}: {E}ffective {T}ool {S}upport for the {W}orking {S}emanticist},
- journal = {J.~of Functional Programming},
- year = {2010},
- volume = {20},
- number = {1},
- pages = {70--122}
-}
-
-@INPROCEEDINGS{Pottier06,
- author = {F.~Pottier},
- title = {{A}n {O}verview of {C$\alpha$ml}},
- year = {2006},
- booktitle = {ACM Workshop on ML},
- pages = {27--52},
- volume = {148},
- number = {2},
- series = {ENTCS}
-}
-
-@inproceedings{HuffmanUrban10,
- author = {B.~Huffman and C.~Urban},
- title = {{P}roof {P}earl: {A} {N}ew {F}oundation for {N}ominal {I}sabelle},
- booktitle = {Proc.~of the 1st ITP Conference},
- pages = {35--50},
- volume = {6172},
- series = {LNCS},
- year = {2010}
-}
-
-@PhdThesis{Leroy92,
- author = {X.~Leroy},
- title = {{P}olymorphic {T}yping of an {A}lgorithmic {L}anguage},
- school = {University Paris 7},
- year = {1992},
- note = {INRIA Research Report, No~1778}
-}
-
-@Unpublished{SewellBestiary,
- author = {P.~Sewell},
- title = {{A} {B}inding {B}estiary},
- note = {Unpublished notes.}
-}
-
-@InProceedings{challenge05,
- author = {B.~E.~Aydemir and A.~Bohannon and M.~Fairbairn and
- J.~N.~Foster and B.~C.~Pierce and P.~Sewell and
- D.~Vytiniotis and G.~Washburn and S.~Weirich and
- S.~Zdancewic},
- title = {{M}echanized {M}etatheory for the {M}asses: {T}he \mbox{Popl}{M}ark
- {C}hallenge},
- booktitle = {Proc.~of the 18th TPHOLs Conference},
- pages = {50--65},
- year = {2005},
- volume = {3603},
- series = {LNCS}
-}
-
-@article{MckinnaPollack99,
- author = {J.~McKinna and R.~Pollack},
- title = {{S}ome {T}ype {T}heory and {L}ambda {C}alculus {F}ormalised},
- journal = {J.~of Automated Reasoning},
- volume = 23,
- number = {1-4},
- year = 1999
-}
-
-@article{SatoPollack10,
- author = {M.~Sato and R.~Pollack},
- title = {{E}xternal and {I}nternal {S}yntax of the {L}ambda-{C}alculus},
- journal = {J.~of Symbolic Computation},
- volume = 45,
- pages = {598--616},
- year = 2010
-}
-
-@article{GabbayPitts02,
- author = {M.~J.~Gabbay and A.~M.~Pitts},
- title = {A New Approach to Abstract Syntax with Variable
- Binding},
- journal = {Formal Aspects of Computing},
- volume = {13},
- year = 2002,
- pages = {341--363}
-}
-
-@article{Pitts03,
- author = {A.~M.~Pitts},
- title = {{N}ominal {L}ogic, {A} {F}irst {O}rder {T}heory of {N}ames and
- {B}inding},
- journal = {Information and Computation},
- year = {2003},
- volume = {183},
- pages = {165--193}
-}
-
-@InProceedings{BengtsonParrow07,
- author = {J.~Bengtson and J.~Parrow},
- title = {Formalising the pi-{C}alculus using {N}ominal {L}ogic},
- booktitle = {Proc.~of the 10th FOSSACS Conference},
- year = 2007,
- pages = {63--77},
- series = {LNCS},
- volume = {4423}
-}
-
-@inproceedings{BengtsonParow09,
- author = {J.~Bengtson and J.~Parrow},
- title = {{P}si-{C}alculi in {I}sabelle},
- booktitle = {Proc of the 22nd TPHOLs Conference},
- year = 2009,
- pages = {99--114},
- series = {LNCS},
- volume = {5674}
-}
-
-@inproceedings{TobinHochstadtFelleisen08,
- author = {S.~Tobin-Hochstadt and M.~Felleisen},
- booktitle = {Proc.~of the 35rd POPL Symposium},
- title = {{T}he {D}esign and {I}mplementation of {T}yped {S}cheme},
- year = {2008},
- pages = {395--406}
-}
-
-@InProceedings{UrbanCheneyBerghofer08,
- author = "C.~Urban and J.~Cheney and S.~Berghofer",
- title = "{M}echanizing the {M}etatheory of {LF}",
- pages = "45--56",
- year = 2008,
- booktitle = "Proc.~of the 23rd LICS Symposium"
-}
-
-@InProceedings{UrbanZhu08,
- title = "{R}evisiting {C}ut-{E}limination: {O}ne {D}ifficult {P}roof is {R}eally a {P}roof",
- author = "C.~Urban and B.~Zhu",
- booktitle = "Proc.~of the 9th RTA Conference",
- year = "2008",
- pages = "409--424",
- series = "LNCS",
- volume = 5117
-}
-
-@Article{UrbanPittsGabbay04,
- title = "{N}ominal {U}nification",
- author = "C.~Urban and A.M.~Pitts and M.J.~Gabbay",
- journal = "Theoretical Computer Science",
- pages = "473--497",
- volume = "323",
- number = "1-3",
- year = "2004"
-}
-
-@Article{Church40,
- author = {A.~Church},
- title = {{A} {F}ormulation of the {S}imple {T}heory of {T}ypes},
- journal = {Journal of Symbolic Logic},
- year = {1940},
- volume = {5},
- number = {2},
- pages = {56--68}
-}
-
-
-@Manual{PittsHOL4,
- title = {{S}yntax and {S}emantics},
- author = {A.~M.~Pitts},
- note = {Part of the documentation for the HOL4 system.}
-}
-
-
-@book{PaulsonBenzmueller,
- year={2009},
- author={Benzm{\"u}ller, Christoph and Paulson, Lawrence C.},
- title={Quantified Multimodal Logics in Simple Type Theory},
- note={{http://arxiv.org/abs/0905.2435}},
- series={{SEKI Report SR--2009--02 (ISSN 1437-4447)}},
- publisher={{SEKI Publications}}
-}
-
-@Article{Cheney06,
- author = {J.~Cheney},
- title = {{C}ompleteness and {H}erbrand theorems for {N}ominal {L}ogic},
- journal = {Journal of Symbolic Logic},
- year = {2006},
- volume = {71},
- number = {1},
- pages = {299--320}
-}
-
--- a/ESOP-Paper/document/root.tex Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,112 +0,0 @@
-\documentclass{llncs}
-\usepackage{times}
-\usepackage{isabelle}
-\usepackage{isabellesym}
-\usepackage{amsmath}
-\usepackage{amssymb}
-%%\usepackage{amsthm}
-\usepackage{tikz}
-\usepackage{pgf}
-\usepackage{pdfsetup}
-\usepackage{ot1patch}
-\usepackage{times}
-\usepackage{boxedminipage}
-\usepackage{proof}
-\usepackage{setspace}
-
-\allowdisplaybreaks
-\urlstyle{rm}
-\isabellestyle{it}
-\renewcommand{\isastyleminor}{\it}%
-\renewcommand{\isastyle}{\normalsize\it}%
-
-\DeclareRobustCommand{\flqq}{\mbox{\guillemotleft}}
-\DeclareRobustCommand{\frqq}{\mbox{\guillemotright}}
-\renewcommand{\isacharunderscore}{\mbox{$\_\!\_$}}
-\renewcommand{\isasymbullet}{{\raisebox{-0.4mm}{\Large$\boldsymbol{\hspace{-0.5mm}\cdot\hspace{-0.5mm}}$}}}
-\def\dn{\,\stackrel{\mbox{\scriptsize def}}{=}\,}
-\renewcommand{\isasymequiv}{$\dn$}
-%%\renewcommand{\isasymiota}{}
-\renewcommand{\isasymxi}{$..$}
-\renewcommand{\isasymemptyset}{$\varnothing$}
-\newcommand{\isasymnotapprox}{$\not\approx$}
-\newcommand{\isasymLET}{$\mathtt{let}$}
-\newcommand{\isasymAND}{$\mathtt{and}$}
-\newcommand{\isasymIN}{$\mathtt{in}$}
-\newcommand{\isasymEND}{$\mathtt{end}$}
-\newcommand{\isasymBIND}{$\mathtt{bind}$}
-\newcommand{\isasymANIL}{$\mathtt{anil}$}
-\newcommand{\isasymACONS}{$\mathtt{acons}$}
-\newcommand{\isasymCASE}{$\mathtt{case}$}
-\newcommand{\isasymOF}{$\mathtt{of}$}
-\newcommand{\isasymAL}{\makebox[0mm][l]{$^\alpha$}}
-\newcommand{\isasymPRIME}{\makebox[0mm][l]{$'$}}
-\newcommand{\isasymFRESH}{\#}
-\newcommand{\LET}{\;\mathtt{let}\;}
-\newcommand{\IN}{\;\mathtt{in}\;}
-\newcommand{\END}{\;\mathtt{end}\;}
-\newcommand{\AND}{\;\mathtt{and}\;}
-\newcommand{\fv}{\mathit{fv}}
-
-\newcommand{\numbered}[1]{\refstepcounter{equation}{\rm(\arabic{equation})}\label{#1}}
-%----------------- theorem definitions ----------
-%%\theoremstyle{plain}
-%%\spnewtheorem{thm}[section]{Theorem}
-%%\newtheorem{property}[thm]{Property}
-%%\newtheorem{lemma}[thm]{Lemma}
-%%\spnewtheorem{defn}[theorem]{Definition}
-%%\spnewtheorem{exmple}[theorem]{Example}
-\spnewtheorem{myproperty}{Property}{\bfseries}{\rmfamily}
-%-------------------- environment definitions -----------------
-\newenvironment{proof-of}[1]{{\em Proof of #1:}}{}
-
-%\addtolength{\textwidth}{2mm}
-\addtolength{\parskip}{-0.33mm}
-\begin{document}
-
-\title{General Bindings and Alpha-Equivalence\\ in Nominal Isabelle}
-\author{Christian Urban and Cezary Kaliszyk}
-\institute{TU Munich, Germany}
-%%%{\{urbanc, kaliszyk\}@in.tum.de}
-\maketitle
-
-\begin{abstract}
-Nominal Isabelle is a definitional extension of the Isabelle/HOL theorem
-prover. It provides a proving infrastructure for reasoning about
-programming language calculi involving named bound variables (as
-opposed to de-Bruijn indices). In this paper we present an extension of
-Nominal Isabelle for dealing with general bindings, that means
-term-constructors where multiple variables are bound at once. Such general
-bindings are ubiquitous in programming language research and only very
-poorly supported with single binders, such as lambda-abstractions. Our
-extension includes new definitions of $\alpha$-equivalence and establishes
-automatically the reasoning infrastructure for $\alpha$-equated terms. We
-also prove strong induction principles that have the usual variable
-convention already built in.
-\end{abstract}
-
-%\category{F.4.1}{subcategory}{third-level}
-
-%\terms
-%formal reasoning, programming language calculi
-
-%\keywords
-%nominal logic work, variable convention
-
-
-\input{session}
-
-\begin{spacing}{0.9}
- \bibliographystyle{plain}
- \bibliography{root}
-\end{spacing}
-
-%\pagebreak
-%\input{Appendix}
-
-\end{document}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
--- a/Fun-Paper/Paper.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,21 +0,0 @@
-(*<*)
-theory Paper
-imports "~~/src/HOL/Library/LaTeXsugar"
-begin
-
-declare [[show_question_marks = false]]
-
-section {* Introduction *}
-
-text {*
-mention Russo paper which concludes that technology is not
-ready beyond core-calculi.
-
-
-
-*}
-
-
-(*<*)
-end
-(*>*)
\ No newline at end of file
--- a/Fun-Paper/ROOT.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,3 +0,0 @@
-no_document use_thy "~~/src/HOL/Library/LaTeXsugar";
-
-use_thy "Paper"
\ No newline at end of file
--- a/Fun-Paper/document/root.tex Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,71 +0,0 @@
-\documentclass[preprint]{sigplanconf}
-\usepackage{isabelle}
-\usepackage{isabellesym}
-\usepackage{amsmath}
-\usepackage{amssymb}
-\usepackage{pdfsetup}
-\usepackage{ot1patch}
-\usepackage{times}
-\usepackage{stmaryrd}
-
-\urlstyle{rm}
-\isabellestyle{it}
-\renewcommand{\isastyleminor}{\it}%
-\renewcommand{\isastyle}{\normalsize\it}%
-
-
-\def\dn{\,\stackrel{\mbox{\scriptsize def}}{=}\,}
-\renewcommand{\isasymequiv}{$\dn$}
-\renewcommand{\isasymemptyset}{$\varnothing$}
-\renewcommand{\isacharunderscore}{\mbox{$\_\!\_$}}
-
-\newcommand{\isasymcalL}{\ensuremath{\cal{L}}}
-\newcommand{\isasymbigplus}{\ensuremath{\bigplus}}
-
-\newcommand{\bigplus}{\mbox{\Large\bf$+$}}
-\begin{document}
-
-%\titlebanner{Nominal Functions} % These are ignored unless
-%\preprintfooter{Nominal Functions} % 'preprint' option specified.
-
-\title{Nominal Functions}
-\authorinfo{?}
- {?}
- {?}
-
-\maketitle
-
-\begin{abstract}
-Nominal Isabelle provides an infrastructure for formal reasoning about terms
-involving named bound variables. This infrastructure allows bound variables to
-be analysed and manipulated directly. In this way it can mimic informal
-pen-and-pencil reasoning. However, this ability makes defining functions the
-most difficult aspect of the ``nominal approach''. In this paper we present a
-new way of defining recursive functions over terms with bound variables. For
-example, we are able to define without difficulties functions for CPS
-translations or the evaluation of lambda-terms, which were previously not
-definable in Nominal Isabelle. We also generalise the
-freshness-condition-for-binders to general binders recently introduced in
-Nominal Isabelle.
-\end{abstract}
-
-\category{CR-number}{subcategory}{third-level}
-
-\terms
-term1, term2
-
-\keywords
-keyword1, keyword2
-
-
-\input{session}
-
-%\bibliographystyle{plain}
-%\bibliography{root}
-
-\end{document}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
--- a/Fun-Paper/document/sigplanconf.cls Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1251 +0,0 @@
-%-----------------------------------------------------------------------------
-%
-% LaTeX Class/Style File
-%
-% Name: sigplanconf.cls
-% Purpose: A LaTeX 2e class file for SIGPLAN conference proceedings.
-% This class file supercedes acm_proc_article-sp,
-% sig-alternate, and sigplan-proc.
-%
-% Author: Paul C. Anagnostopoulos
-% Windfall Software
-% 978 371-2316
-% sigplan-style [atsign] acm.org
-%
-% Created: 12 September 2004
-%
-% Revisions: See end of file.
-%
-%-----------------------------------------------------------------------------
-
-
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-\ProvidesClass{sigplanconf}[2010/05/24 v2.4 ACM SIGPLAN Proceedings]
-
-% The following few pages contain LaTeX programming extensions adapted
-% from the ZzTeX macro package.
-
-% Token Hackery
-% ----- -------
-
-
-\def \@expandaftertwice {\expandafter\expandafter\expandafter}
-\def \@expandafterthrice {\expandafter\expandafter\expandafter\expandafter
- \expandafter\expandafter\expandafter}
-
-% This macro discards the next token.
-
-\def \@discardtok #1{}% token
-
-% This macro removes the `pt' following a dimension.
-
-{\catcode `\p = 12 \catcode `\t = 12
-
-\gdef \@remover #1pt{#1}
-
-} % \catcode
-
-% This macro extracts the contents of a macro and returns it as plain text.
-% Usage: \expandafter\@defof \meaning\macro\@mark
-
-\def \@defof #1:->#2\@mark{#2}
-
-% Control Sequence Names
-% ------- -------- -----
-
-
-\def \@name #1{% {\tokens}
- \csname \expandafter\@discardtok \string#1\endcsname}
-
-\def \@withname #1#2{% {\command}{\tokens}
- \expandafter#1\csname \expandafter\@discardtok \string#2\endcsname}
-
-% Flags (Booleans)
-% ----- ----------
-
-% The boolean literals \@true and \@false are appropriate for use with
-% the \if command, which tests the codes of the next two characters.
-
-\def \@true {TT}
-\def \@false {FL}
-
-\def \@setflag #1=#2{\edef #1{#2}}% \flag = boolean
-
-% IF and Predicates
-% -- --- ----------
-
-% A "predicate" is a macro that returns \@true or \@false as its value.
-% Such values are suitable for use with the \if conditional. For example:
-%
-% \if \@oddp{\x} <then-clause> \else <else-clause> \fi
-
-% A predicate can be used with \@setflag as follows:
-%
-% \@setflag \flag = {<predicate>}
-
-% Here are the predicates for TeX's repertoire of conditional
-% commands. These might be more appropriately interspersed with
-% other definitions in this module, but what the heck.
-% Some additional "obvious" predicates are defined.
-
-\def \@eqlp #1#2{\ifnum #1 = #2\@true \else \@false \fi}
-\def \@neqlp #1#2{\ifnum #1 = #2\@false \else \@true \fi}
-\def \@lssp #1#2{\ifnum #1 < #2\@true \else \@false \fi}
-\def \@gtrp #1#2{\ifnum #1 > #2\@true \else \@false \fi}
-\def \@zerop #1{\ifnum #1 = 0\@true \else \@false \fi}
-\def \@onep #1{\ifnum #1 = 1\@true \else \@false \fi}
-\def \@posp #1{\ifnum #1 > 0\@true \else \@false \fi}
-\def \@negp #1{\ifnum #1 < 0\@true \else \@false \fi}
-\def \@oddp #1{\ifodd #1\@true \else \@false \fi}
-\def \@evenp #1{\ifodd #1\@false \else \@true \fi}
-\def \@rangep #1#2#3{\if \@orp{\@lssp{#1}{#2}}{\@gtrp{#1}{#3}}\@false \else
- \@true \fi}
-\def \@tensp #1{\@rangep{#1}{10}{19}}
-
-\def \@dimeqlp #1#2{\ifdim #1 = #2\@true \else \@false \fi}
-\def \@dimneqlp #1#2{\ifdim #1 = #2\@false \else \@true \fi}
-\def \@dimlssp #1#2{\ifdim #1 < #2\@true \else \@false \fi}
-\def \@dimgtrp #1#2{\ifdim #1 > #2\@true \else \@false \fi}
-\def \@dimzerop #1{\ifdim #1 = 0pt\@true \else \@false \fi}
-\def \@dimposp #1{\ifdim #1 > 0pt\@true \else \@false \fi}
-\def \@dimnegp #1{\ifdim #1 < 0pt\@true \else \@false \fi}
-
-\def \@vmodep {\ifvmode \@true \else \@false \fi}
-\def \@hmodep {\ifhmode \@true \else \@false \fi}
-\def \@mathmodep {\ifmmode \@true \else \@false \fi}
-\def \@textmodep {\ifmmode \@false \else \@true \fi}
-\def \@innermodep {\ifinner \@true \else \@false \fi}
-
-\long\def \@codeeqlp #1#2{\if #1#2\@true \else \@false \fi}
-
-\long\def \@cateqlp #1#2{\ifcat #1#2\@true \else \@false \fi}
-
-\long\def \@tokeqlp #1#2{\ifx #1#2\@true \else \@false \fi}
-\long\def \@xtokeqlp #1#2{\expandafter\ifx #1#2\@true \else \@false \fi}
-
-\long\def \@definedp #1{%
- \expandafter\ifx \csname \expandafter\@discardtok \string#1\endcsname
- \relax \@false \else \@true \fi}
-
-\long\def \@undefinedp #1{%
- \expandafter\ifx \csname \expandafter\@discardtok \string#1\endcsname
- \relax \@true \else \@false \fi}
-
-\def \@emptydefp #1{\ifx #1\@empty \@true \else \@false \fi}% {\name}
-
-\let \@emptylistp = \@emptydefp
-
-\long\def \@emptyargp #1{% {#n}
- \@empargp #1\@empargq\@mark}
-\long\def \@empargp #1#2\@mark{%
- \ifx #1\@empargq \@true \else \@false \fi}
-\def \@empargq {\@empargq}
-
-\def \@emptytoksp #1{% {\tokenreg}
- \expandafter\@emptoksp \the#1\@mark}
-
-\long\def \@emptoksp #1\@mark{\@emptyargp{#1}}
-
-\def \@voidboxp #1{\ifvoid #1\@true \else \@false \fi}
-\def \@hboxp #1{\ifhbox #1\@true \else \@false \fi}
-\def \@vboxp #1{\ifvbox #1\@true \else \@false \fi}
-
-\def \@eofp #1{\ifeof #1\@true \else \@false \fi}
-
-
-% Flags can also be used as predicates, as in:
-%
-% \if \flaga <then-clause> \else <else-clause> \fi
-
-
-% Now here we have predicates for the common logical operators.
-
-\def \@notp #1{\if #1\@false \else \@true \fi}
-
-\def \@andp #1#2{\if #1%
- \if #2\@true \else \@false \fi
- \else
- \@false
- \fi}
-
-\def \@orp #1#2{\if #1%
- \@true
- \else
- \if #2\@true \else \@false \fi
- \fi}
-
-\def \@xorp #1#2{\if #1%
- \if #2\@false \else \@true \fi
- \else
- \if #2\@true \else \@false \fi
- \fi}
-
-% Arithmetic
-% ----------
-
-\def \@increment #1{\advance #1 by 1\relax}% {\count}
-
-\def \@decrement #1{\advance #1 by -1\relax}% {\count}
-
-% Options
-% -------
-
-
-\@setflag \@authoryear = \@false
-\@setflag \@blockstyle = \@false
-\@setflag \@copyrightwanted = \@true
-\@setflag \@explicitsize = \@false
-\@setflag \@mathtime = \@false
-\@setflag \@natbib = \@true
-\@setflag \@ninepoint = \@true
-\newcount{\@numheaddepth} \@numheaddepth = 3
-\@setflag \@onecolumn = \@false
-\@setflag \@preprint = \@false
-\@setflag \@reprint = \@false
-\@setflag \@tenpoint = \@false
-\@setflag \@times = \@false
-
-% Note that all the dangerous article class options are trapped.
-
-\DeclareOption{9pt}{\@setflag \@ninepoint = \@true
- \@setflag \@explicitsize = \@true}
-
-\DeclareOption{10pt}{\PassOptionsToClass{10pt}{article}%
- \@setflag \@ninepoint = \@false
- \@setflag \@tenpoint = \@true
- \@setflag \@explicitsize = \@true}
-
-\DeclareOption{11pt}{\PassOptionsToClass{11pt}{article}%
- \@setflag \@ninepoint = \@false
- \@setflag \@explicitsize = \@true}
-
-\DeclareOption{12pt}{\@unsupportedoption{12pt}}
-
-\DeclareOption{a4paper}{\@unsupportedoption{a4paper}}
-
-\DeclareOption{a5paper}{\@unsupportedoption{a5paper}}
-
-\DeclareOption{authoryear}{\@setflag \@authoryear = \@true}
-
-\DeclareOption{b5paper}{\@unsupportedoption{b5paper}}
-
-\DeclareOption{blockstyle}{\@setflag \@blockstyle = \@true}
-
-\DeclareOption{cm}{\@setflag \@times = \@false}
-
-\DeclareOption{computermodern}{\@setflag \@times = \@false}
-
-\DeclareOption{executivepaper}{\@unsupportedoption{executivepaper}}
-
-\DeclareOption{indentedstyle}{\@setflag \@blockstyle = \@false}
-
-\DeclareOption{landscape}{\@unsupportedoption{landscape}}
-
-\DeclareOption{legalpaper}{\@unsupportedoption{legalpaper}}
-
-\DeclareOption{letterpaper}{\@unsupportedoption{letterpaper}}
-
-\DeclareOption{mathtime}{\@setflag \@mathtime = \@true}
-
-\DeclareOption{natbib}{\@setflag \@natbib = \@true}
-
-\DeclareOption{nonatbib}{\@setflag \@natbib = \@false}
-
-\DeclareOption{nocopyrightspace}{\@setflag \@copyrightwanted = \@false}
-
-\DeclareOption{notitlepage}{\@unsupportedoption{notitlepage}}
-
-\DeclareOption{numberedpars}{\@numheaddepth = 4}
-
-\DeclareOption{numbers}{\@setflag \@authoryear = \@false}
-
-%%%\DeclareOption{onecolumn}{\@setflag \@onecolumn = \@true}
-
-\DeclareOption{preprint}{\@setflag \@preprint = \@true}
-
-\DeclareOption{reprint}{\@setflag \@reprint = \@true}
-
-\DeclareOption{times}{\@setflag \@times = \@true}
-
-\DeclareOption{titlepage}{\@unsupportedoption{titlepage}}
-
-\DeclareOption{twocolumn}{\@setflag \@onecolumn = \@false}
-
-\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}}
-
-\ExecuteOptions{9pt,indentedstyle,times}
-\@setflag \@explicitsize = \@false
-\ProcessOptions
-
-\if \@onecolumn
- \if \@notp{\@explicitsize}%
- \@setflag \@ninepoint = \@false
- \PassOptionsToClass{11pt}{article}%
- \fi
- \PassOptionsToClass{twoside,onecolumn}{article}
-\else
- \PassOptionsToClass{twoside,twocolumn}{article}
-\fi
-\LoadClass{article}
-
-\def \@unsupportedoption #1{%
- \ClassError{proc}{The standard '#1' option is not supported.}}
-
-% This can be used with the 'reprint' option to get the final folios.
-
-\def \setpagenumber #1{%
- \setcounter{page}{#1}}
-
-\AtEndDocument{\label{sigplanconf@finalpage}}
-
-% Utilities
-% ---------
-
-
-\newcommand{\setvspace}[2]{%
- #1 = #2
- \advance #1 by -1\parskip}
-
-% Document Parameters
-% -------- ----------
-
-
-% Page:
-
-\setlength{\hoffset}{-1in}
-\setlength{\voffset}{-1in}
-
-\setlength{\topmargin}{1in}
-\setlength{\headheight}{0pt}
-\setlength{\headsep}{0pt}
-
-\if \@onecolumn
- \setlength{\evensidemargin}{.75in}
- \setlength{\oddsidemargin}{.75in}
-\else
- \setlength{\evensidemargin}{.75in}
- \setlength{\oddsidemargin}{.75in}
-\fi
-
-% Text area:
-
-\newdimen{\standardtextwidth}
-\setlength{\standardtextwidth}{42pc}
-
-\if \@onecolumn
- \setlength{\textwidth}{40.5pc}
-\else
- \setlength{\textwidth}{\standardtextwidth}
-\fi
-
-\setlength{\topskip}{8pt}
-\setlength{\columnsep}{2pc}
-\setlength{\textheight}{54.5pc}
-
-% Running foot:
-
-\setlength{\footskip}{30pt}
-
-% Paragraphs:
-
-\if \@blockstyle
- \setlength{\parskip}{5pt plus .1pt minus .5pt}
- \setlength{\parindent}{0pt}
-\else
- \setlength{\parskip}{0pt}
- \setlength{\parindent}{12pt}
-\fi
-
-\setlength{\lineskip}{.5pt}
-\setlength{\lineskiplimit}{\lineskip}
-
-\frenchspacing
-\pretolerance = 400
-\tolerance = \pretolerance
-\setlength{\emergencystretch}{5pt}
-\clubpenalty = 10000
-\widowpenalty = 10000
-\setlength{\hfuzz}{.5pt}
-
-% Standard vertical spaces:
-
-\newskip{\standardvspace}
-\setvspace{\standardvspace}{5pt plus 1pt minus .5pt}
-
-% Margin paragraphs:
-
-\setlength{\marginparwidth}{36pt}
-\setlength{\marginparsep}{2pt}
-\setlength{\marginparpush}{8pt}
-
-
-\setlength{\skip\footins}{8pt plus 3pt minus 1pt}
-\setlength{\footnotesep}{9pt}
-
-\renewcommand{\footnoterule}{%
- \hrule width .5\columnwidth height .33pt depth 0pt}
-
-\renewcommand{\@makefntext}[1]{%
- \noindent \@makefnmark \hspace{1pt}#1}
-
-% Floats:
-
-\setcounter{topnumber}{4}
-\setcounter{bottomnumber}{1}
-\setcounter{totalnumber}{4}
-
-\renewcommand{\fps@figure}{tp}
-\renewcommand{\fps@table}{tp}
-\renewcommand{\topfraction}{0.90}
-\renewcommand{\bottomfraction}{0.30}
-\renewcommand{\textfraction}{0.10}
-\renewcommand{\floatpagefraction}{0.75}
-
-\setcounter{dbltopnumber}{4}
-
-\renewcommand{\dbltopfraction}{\topfraction}
-\renewcommand{\dblfloatpagefraction}{\floatpagefraction}
-
-\setlength{\floatsep}{18pt plus 4pt minus 2pt}
-\setlength{\textfloatsep}{18pt plus 4pt minus 3pt}
-\setlength{\intextsep}{10pt plus 4pt minus 3pt}
-
-\setlength{\dblfloatsep}{18pt plus 4pt minus 2pt}
-\setlength{\dbltextfloatsep}{20pt plus 4pt minus 3pt}
-
-% Miscellaneous:
-
-\errorcontextlines = 5
-
-% Fonts
-% -----
-
-
-\if \@times
- \renewcommand{\rmdefault}{ptm}%
- \if \@mathtime
- \usepackage[mtbold,noTS1]{mathtime}%
- \else
-%%% \usepackage{mathptm}%
- \fi
-\else
- \relax
-\fi
-
-\if \@ninepoint
-
-\renewcommand{\normalsize}{%
- \@setfontsize{\normalsize}{9pt}{10pt}%
- \setlength{\abovedisplayskip}{5pt plus 1pt minus .5pt}%
- \setlength{\belowdisplayskip}{\abovedisplayskip}%
- \setlength{\abovedisplayshortskip}{3pt plus 1pt minus 2pt}%
- \setlength{\belowdisplayshortskip}{\abovedisplayshortskip}}
-
-\renewcommand{\tiny}{\@setfontsize{\tiny}{5pt}{6pt}}
-
-\renewcommand{\scriptsize}{\@setfontsize{\scriptsize}{7pt}{8pt}}
-
-\renewcommand{\small}{%
- \@setfontsize{\small}{8pt}{9pt}%
- \setlength{\abovedisplayskip}{4pt plus 1pt minus 1pt}%
- \setlength{\belowdisplayskip}{\abovedisplayskip}%
- \setlength{\abovedisplayshortskip}{2pt plus 1pt}%
- \setlength{\belowdisplayshortskip}{\abovedisplayshortskip}}
-
-\renewcommand{\footnotesize}{%
- \@setfontsize{\footnotesize}{8pt}{9pt}%
- \setlength{\abovedisplayskip}{4pt plus 1pt minus .5pt}%
- \setlength{\belowdisplayskip}{\abovedisplayskip}%
- \setlength{\abovedisplayshortskip}{2pt plus 1pt}%
- \setlength{\belowdisplayshortskip}{\abovedisplayshortskip}}
-
-\renewcommand{\large}{\@setfontsize{\large}{11pt}{13pt}}
-
-\renewcommand{\Large}{\@setfontsize{\Large}{14pt}{18pt}}
-
-\renewcommand{\LARGE}{\@setfontsize{\LARGE}{18pt}{20pt}}
-
-\renewcommand{\huge}{\@setfontsize{\huge}{20pt}{25pt}}
-
-\renewcommand{\Huge}{\@setfontsize{\Huge}{25pt}{30pt}}
-
-\else\if \@tenpoint
-
-\relax
-
-\else
-
-\relax
-
-\fi\fi
-
-% Abstract
-% --------
-
-
-\renewenvironment{abstract}{%
- \section*{Abstract}%
- \normalsize}{%
- }
-
-% Bibliography
-% ------------
-
-
-\renewenvironment{thebibliography}[1]
- {\section*{\refname
- \@mkboth{\MakeUppercase\refname}{\MakeUppercase\refname}}%
- \list{\@biblabel{\@arabic\c@enumiv}}%
- {\settowidth\labelwidth{\@biblabel{#1}}%
- \leftmargin\labelwidth
- \advance\leftmargin\labelsep
- \@openbib@code
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{\@arabic\c@enumiv}}%
- \bibfont
- \clubpenalty4000
- \@clubpenalty \clubpenalty
- \widowpenalty4000%
- \sfcode`\.\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
-
-\if \@natbib
-
-\if \@authoryear
- \typeout{Using natbib package with 'authoryear' citation style.}
- \usepackage[authoryear,sort,square]{natbib}
- \bibpunct{[}{]}{;}{a}{}{,} % Change citation separator to semicolon,
- % eliminate comma between author and year.
- \let \cite = \citep
-\else
- \typeout{Using natbib package with 'numbers' citation style.}
- \usepackage[numbers,sort&compress,square]{natbib}
-\fi
-\setlength{\bibsep}{3pt plus .5pt minus .25pt}
-
-\fi
-
-\def \bibfont {\small}
-
-% Categories
-% ----------
-
-
-\@setflag \@firstcategory = \@true
-
-\newcommand{\category}[3]{%
- \if \@firstcategory
- \paragraph*{Categories and Subject Descriptors}%
- \@setflag \@firstcategory = \@false
- \else
- \unskip ;\hspace{.75em}%
- \fi
- \@ifnextchar [{\@category{#1}{#2}{#3}}{\@category{#1}{#2}{#3}[]}}
-
-\def \@category #1#2#3[#4]{%
- {\let \and = \relax
- #1 [\textit{#2}]%
- \if \@emptyargp{#4}%
- \if \@notp{\@emptyargp{#3}}: #3\fi
- \else
- :\space
- \if \@notp{\@emptyargp{#3}}#3---\fi
- \textrm{#4}%
- \fi}}
-
-% Copyright Notice
-% --------- ------
-
-
-\def \ftype@copyrightbox {8}
-\def \@toappear {}
-\def \@permission {}
-\def \@reprintprice {}
-
-\def \@copyrightspace {%
- \@float{copyrightbox}[b]%
- \vbox to 1in{%
- \vfill
- \parbox[b]{20pc}{%
- \scriptsize
- \if \@preprint
- [Copyright notice will appear here
- once 'preprint' option is removed.]\par
- \else
- \@toappear
- \fi
- \if \@reprint
- \noindent Reprinted from \@conferencename,
- \@proceedings,
- \@conferenceinfo,
- pp.~\number\thepage--\pageref{sigplanconf@finalpage}.\par
- \fi}}%
- \end@float}
-
-\long\def \toappear #1{%
- \def \@toappear {#1}}
-
-\toappear{%
- \noindent \@permission \par
- \vspace{2pt}
- \noindent \textsl{\@conferencename}\quad \@conferenceinfo \par
- \noindent Copyright \copyright\ \@copyrightyear\ ACM \@copyrightdata
- \dots \@reprintprice\par}
-
-\newcommand{\permission}[1]{%
- \gdef \@permission {#1}}
-
-\permission{%
- Permission to make digital or hard copies of all or
- part of this work for personal or classroom use is granted without
- fee provided that copies are not made or distributed for profit or
- commercial advantage and that copies bear this notice and the full
- citation on the first page. To copy otherwise, to republish, to
- post on servers or to redistribute to lists, requires prior specific
- permission and/or a fee.}
-
-% Here we have some alternate permission statements and copyright lines:
-
-\newcommand{\ACMCanadapermission}{%
- \permission{%
- Copyright \@copyrightyear\ Association for Computing Machinery.
- ACM acknowledges that
- this contribution was authored or co-authored by an affiliate of the
- National Research Council of Canada (NRC).
- As such, the Crown in Right of
- Canada retains an equal interest in the copyright, however granting
- nonexclusive, royalty-free right to publish or reproduce this article,
- or to allow others to do so, provided that clear attribution
- is also given to the authors and the NRC.}}
-
-\newcommand{\ACMUSpermission}{%
- \permission{%
- Copyright \@copyrightyear\ Association for
- Computing Machinery. ACM acknowledges that
- this contribution was authored or co-authored
- by a contractor or affiliate
- of the U.S. Government. As such, the Government retains a nonexclusive,
- royalty-free right to publish or reproduce this article,
- or to allow others to do so, for Government purposes only.}}
-
-\newcommand{\authorpermission}{%
- \permission{%
- Copyright is held by the author/owner(s).}
- \toappear{%
- \noindent \@permission \par
- \vspace{2pt}
- \noindent \textsl{\@conferencename}\quad \@conferenceinfo \par
- ACM \@copyrightdata.}}
-
-\newcommand{\Sunpermission}{%
- \permission{%
- Copyright is held by Sun Microsystems, Inc.}%
- \toappear{%
- \noindent \@permission \par
- \vspace{2pt}
- \noindent \textsl{\@conferencename}\quad \@conferenceinfo \par
- ACM \@copyrightdata.}}
-
-\newcommand{\USpublicpermission}{%
- \permission{%
- This paper is authored by an employee(s) of the United States
- Government and is in the public domain.}%
- \toappear{%
- \noindent \@permission \par
- \vspace{2pt}
- \noindent \textsl{\@conferencename}\quad \@conferenceinfo \par
- ACM \@copyrightdata.}}
-
-\newcommand{\reprintprice}[1]{%
- \gdef \@reprintprice {#1}}
-
-\reprintprice{\$10.00}
-
-% Enunciations
-% ------------
-
-
-\def \@begintheorem #1#2{% {name}{number}
- \trivlist
- \item[\hskip \labelsep \textsc{#1 #2.}]%
- \itshape\selectfont
- \ignorespaces}
-
-\def \@opargbegintheorem #1#2#3{% {name}{number}{title}
- \trivlist
- \item[%
- \hskip\labelsep \textsc{#1\ #2}%
- \if \@notp{\@emptyargp{#3}}\nut (#3).\fi]%
- \itshape\selectfont
- \ignorespaces}
-
-% Figures
-% -------
-
-
-\@setflag \@caprule = \@true
-
-\long\def \@makecaption #1#2{%
- \addvspace{4pt}
- \if \@caprule
- \hrule width \hsize height .33pt
- \vspace{4pt}
- \fi
- \setbox \@tempboxa = \hbox{\@setfigurenumber{#1.}\nut #2}%
- \if \@dimgtrp{\wd\@tempboxa}{\hsize}%
- \noindent \@setfigurenumber{#1.}\nut #2\par
- \else
- \centerline{\box\@tempboxa}%
- \fi}
-
-\newcommand{\nocaptionrule}{%
- \@setflag \@caprule = \@false}
-
-\def \@setfigurenumber #1{%
- {\rmfamily \bfseries \selectfont #1}}
-
-% Hierarchy
-% ---------
-
-
-\setcounter{secnumdepth}{\@numheaddepth}
-
-\newskip{\@sectionaboveskip}
-\setvspace{\@sectionaboveskip}{10pt plus 3pt minus 2pt}
-
-\newskip{\@sectionbelowskip}
-\if \@blockstyle
- \setlength{\@sectionbelowskip}{0.1pt}%
-\else
- \setlength{\@sectionbelowskip}{4pt}%
-\fi
-
-\renewcommand{\section}{%
- \@startsection
- {section}%
- {1}%
- {0pt}%
- {-\@sectionaboveskip}%
- {\@sectionbelowskip}%
- {\large \bfseries \raggedright}}
-
-\newskip{\@subsectionaboveskip}
-\setvspace{\@subsectionaboveskip}{8pt plus 2pt minus 2pt}
-
-\newskip{\@subsectionbelowskip}
-\if \@blockstyle
- \setlength{\@subsectionbelowskip}{0.1pt}%
-\else
- \setlength{\@subsectionbelowskip}{4pt}%
-\fi
-
-\renewcommand{\subsection}{%
- \@startsection%
- {subsection}%
- {2}%
- {0pt}%
- {-\@subsectionaboveskip}%
- {\@subsectionbelowskip}%
- {\normalsize \bfseries \raggedright}}
-
-\renewcommand{\subsubsection}{%
- \@startsection%
- {subsubsection}%
- {3}%
- {0pt}%
- {-\@subsectionaboveskip}
- {\@subsectionbelowskip}%
- {\normalsize \bfseries \raggedright}}
-
-\newskip{\@paragraphaboveskip}
-\setvspace{\@paragraphaboveskip}{6pt plus 2pt minus 2pt}
-
-\renewcommand{\paragraph}{%
- \@startsection%
- {paragraph}%
- {4}%
- {0pt}%
- {\@paragraphaboveskip}
- {-1em}%
- {\normalsize \bfseries \if \@times \itshape \fi}}
-
-\renewcommand{\subparagraph}{%
- \@startsection%
- {subparagraph}%
- {4}%
- {0pt}%
- {\@paragraphaboveskip}
- {-1em}%
- {\normalsize \itshape}}
-
-% Standard headings:
-
-\newcommand{\acks}{\section*{Acknowledgments}}
-
-\newcommand{\keywords}{\paragraph*{Keywords}}
-
-\newcommand{\terms}{\paragraph*{General Terms}}
-
-% Identification
-% --------------
-
-
-\def \@conferencename {}
-\def \@conferenceinfo {}
-\def \@copyrightyear {}
-\def \@copyrightdata {[to be supplied]}
-\def \@proceedings {[Unknown Proceedings]}
-
-
-\newcommand{\conferenceinfo}[2]{%
- \gdef \@conferencename {#1}%
- \gdef \@conferenceinfo {#2}}
-
-\newcommand{\copyrightyear}[1]{%
- \gdef \@copyrightyear {#1}}
-
-\let \CopyrightYear = \copyrightyear
-
-\newcommand{\copyrightdata}[1]{%
- \gdef \@copyrightdata {#1}}
-
-\let \crdata = \copyrightdata
-
-\newcommand{\proceedings}[1]{%
- \gdef \@proceedings {#1}}
-
-% Lists
-% -----
-
-
-\setlength{\leftmargini}{13pt}
-\setlength\leftmarginii{13pt}
-\setlength\leftmarginiii{13pt}
-\setlength\leftmarginiv{13pt}
-\setlength{\labelsep}{3.5pt}
-
-\setlength{\topsep}{\standardvspace}
-\if \@blockstyle
- \setlength{\itemsep}{1pt}
- \setlength{\parsep}{3pt}
-\else
- \setlength{\itemsep}{1pt}
- \setlength{\parsep}{3pt}
-\fi
-
-\renewcommand{\labelitemi}{{\small \centeroncapheight{\textbullet}}}
-\renewcommand{\labelitemii}{\centeroncapheight{\rule{2.5pt}{2.5pt}}}
-\renewcommand{\labelitemiii}{$-$}
-\renewcommand{\labelitemiv}{{\Large \textperiodcentered}}
-
-\renewcommand{\@listi}{%
- \leftmargin = \leftmargini
- \listparindent = 0pt}
-%%% \itemsep = 1pt
-%%% \parsep = 3pt}
-%%% \listparindent = \parindent}
-
-\let \@listI = \@listi
-
-\renewcommand{\@listii}{%
- \leftmargin = \leftmarginii
- \topsep = 1pt
- \labelwidth = \leftmarginii
- \advance \labelwidth by -\labelsep
- \listparindent = \parindent}
-
-\renewcommand{\@listiii}{%
- \leftmargin = \leftmarginiii
- \labelwidth = \leftmarginiii
- \advance \labelwidth by -\labelsep
- \listparindent = \parindent}
-
-\renewcommand{\@listiv}{%
- \leftmargin = \leftmarginiv
- \labelwidth = \leftmarginiv
- \advance \labelwidth by -\labelsep
- \listparindent = \parindent}
-
-% Mathematics
-% -----------
-
-
-\def \theequation {\arabic{equation}}
-
-% Miscellaneous
-% -------------
-
-
-\newcommand{\balancecolumns}{%
- \vfill\eject
- \global\@colht = \textheight
- \global\ht\@cclv = \textheight}
-
-\newcommand{\nut}{\hspace{.5em}}
-
-\newcommand{\softraggedright}{%
- \let \\ = \@centercr
- \leftskip = 0pt
- \rightskip = 0pt plus 10pt}
-
-% Program Code
-% ------- ----
-
-
-\newcommand{\mono}[1]{%
- {\@tempdima = \fontdimen2\font
- \texttt{\spaceskip = 1.1\@tempdima #1}}}
-
-% Running Heads and Feet
-% ------- ----- --- ----
-
-
-\def \@preprintfooter {}
-
-\newcommand{\preprintfooter}[1]{%
- \gdef \@preprintfooter {#1}}
-
-\if \@preprint
-
-\def \ps@plain {%
- \let \@mkboth = \@gobbletwo
- \let \@evenhead = \@empty
- \def \@evenfoot {\scriptsize \textit{\@preprintfooter}\hfil \thepage \hfil
- \textit{\@formatyear}}%
- \let \@oddhead = \@empty
- \let \@oddfoot = \@evenfoot}
-
-\else\if \@reprint
-
-\def \ps@plain {%
- \let \@mkboth = \@gobbletwo
- \let \@evenhead = \@empty
- \def \@evenfoot {\scriptsize \hfil \thepage \hfil}%
- \let \@oddhead = \@empty
- \let \@oddfoot = \@evenfoot}
-
-\else
-
-\let \ps@plain = \ps@empty
-\let \ps@headings = \ps@empty
-\let \ps@myheadings = \ps@empty
-
-\fi\fi
-
-\def \@formatyear {%
- \number\year/\number\month/\number\day}
-
-% Special Characters
-% ------- ----------
-
-
-\DeclareRobustCommand{\euro}{%
- \protect{\rlap{=}}{\sf \kern .1em C}}
-
-% Title Page
-% ----- ----
-
-
-\@setflag \@addauthorsdone = \@false
-
-\def \@titletext {\@latex@error{No title was provided}{}}
-\def \@subtitletext {}
-
-\newcount{\@authorcount}
-
-\newcount{\@titlenotecount}
-\newtoks{\@titlenotetext}
-
-\def \@titlebanner {}
-
-\renewcommand{\title}[1]{%
- \gdef \@titletext {#1}}
-
-\newcommand{\subtitle}[1]{%
- \gdef \@subtitletext {#1}}
-
-\newcommand{\authorinfo}[3]{% {names}{affiliation}{email/URL}
- \global\@increment \@authorcount
- \@withname\gdef {\@authorname\romannumeral\@authorcount}{#1}%
- \@withname\gdef {\@authoraffil\romannumeral\@authorcount}{#2}%
- \@withname\gdef {\@authoremail\romannumeral\@authorcount}{#3}}
-
-\renewcommand{\author}[1]{%
- \@latex@error{The \string\author\space command is obsolete;
- use \string\authorinfo}{}}
-
-\newcommand{\titlebanner}[1]{%
- \gdef \@titlebanner {#1}}
-
-\renewcommand{\maketitle}{%
- \pagestyle{plain}%
- \if \@onecolumn
- {\hsize = \standardtextwidth
- \@maketitle}%
- \else
- \twocolumn[\@maketitle]%
- \fi
- \@placetitlenotes
- \if \@copyrightwanted \@copyrightspace \fi}
-
-\def \@maketitle {%
- \begin{center}
- \@settitlebanner
- \let \thanks = \titlenote
- {\leftskip = 0pt plus 0.25\linewidth
- \rightskip = 0pt plus 0.25 \linewidth
- \parfillskip = 0pt
- \spaceskip = .7em
- \noindent \LARGE \bfseries \@titletext \par}
- \vskip 6pt
- \noindent \Large \@subtitletext \par
- \vskip 12pt
- \ifcase \@authorcount
- \@latex@error{No authors were specified for this paper}{}\or
- \@titleauthors{i}{}{}\or
- \@titleauthors{i}{ii}{}\or
- \@titleauthors{i}{ii}{iii}\or
- \@titleauthors{i}{ii}{iii}\@titleauthors{iv}{}{}\or
- \@titleauthors{i}{ii}{iii}\@titleauthors{iv}{v}{}\or
- \@titleauthors{i}{ii}{iii}\@titleauthors{iv}{v}{vi}\or
- \@titleauthors{i}{ii}{iii}\@titleauthors{iv}{v}{vi}%
- \@titleauthors{vii}{}{}\or
- \@titleauthors{i}{ii}{iii}\@titleauthors{iv}{v}{vi}%
- \@titleauthors{vii}{viii}{}\or
- \@titleauthors{i}{ii}{iii}\@titleauthors{iv}{v}{vi}%
- \@titleauthors{vii}{viii}{ix}\or
- \@titleauthors{i}{ii}{iii}\@titleauthors{iv}{v}{vi}%
- \@titleauthors{vii}{viii}{ix}\@titleauthors{x}{}{}\or
- \@titleauthors{i}{ii}{iii}\@titleauthors{iv}{v}{vi}%
- \@titleauthors{vii}{viii}{ix}\@titleauthors{x}{xi}{}\or
- \@titleauthors{i}{ii}{iii}\@titleauthors{iv}{v}{vi}%
- \@titleauthors{vii}{viii}{ix}\@titleauthors{x}{xi}{xii}%
- \else
- \@latex@error{Cannot handle more than 12 authors}{}%
- \fi
- \vspace{1.75pc}
- \end{center}}
-
-\def \@settitlebanner {%
- \if \@andp{\@preprint}{\@notp{\@emptydefp{\@titlebanner}}}%
- \vbox to 0pt{%
- \vskip -32pt
- \noindent \textbf{\@titlebanner}\par
- \vss}%
- \nointerlineskip
- \fi}
-
-\def \@titleauthors #1#2#3{%
- \if \@andp{\@emptyargp{#2}}{\@emptyargp{#3}}%
- \noindent \@setauthor{40pc}{#1}{\@false}\par
- \else\if \@emptyargp{#3}%
- \noindent \@setauthor{17pc}{#1}{\@false}\hspace{3pc}%
- \@setauthor{17pc}{#2}{\@false}\par
- \else
- \noindent \@setauthor{12.5pc}{#1}{\@false}\hspace{2pc}%
- \@setauthor{12.5pc}{#2}{\@false}\hspace{2pc}%
- \@setauthor{12.5pc}{#3}{\@true}\par
- \relax
- \fi\fi
- \vspace{20pt}}
-
-\def \@setauthor #1#2#3{% {width}{text}{unused}
- \vtop{%
- \def \and {%
- \hspace{16pt}}
- \hsize = #1
- \normalfont
- \centering
- \large \@name{\@authorname#2}\par
- \vspace{5pt}
- \normalsize \@name{\@authoraffil#2}\par
- \vspace{2pt}
- \textsf{\@name{\@authoremail#2}}\par}}
-
-\def \@maybetitlenote #1{%
- \if \@andp{#1}{\@gtrp{\@authorcount}{3}}%
- \titlenote{See page~\pageref{@addauthors} for additional authors.}%
- \fi}
-
-\newtoks{\@fnmark}
-
-\newcommand{\titlenote}[1]{%
- \global\@increment \@titlenotecount
- \ifcase \@titlenotecount \relax \or
- \@fnmark = {\ast}\or
- \@fnmark = {\dagger}\or
- \@fnmark = {\ddagger}\or
- \@fnmark = {\S}\or
- \@fnmark = {\P}\or
- \@fnmark = {\ast\ast}%
- \fi
- \,$^{\the\@fnmark}$%
- \edef \reserved@a {\noexpand\@appendtotext{%
- \noexpand\@titlefootnote{\the\@fnmark}}}%
- \reserved@a{#1}}
-
-\def \@appendtotext #1#2{%
- \global\@titlenotetext = \expandafter{\the\@titlenotetext #1{#2}}}
-
-\newcount{\@authori}
-
-\iffalse
-\def \additionalauthors {%
- \if \@gtrp{\@authorcount}{3}%
- \section{Additional Authors}%
- \label{@addauthors}%
- \noindent
- \@authori = 4
- {\let \\ = ,%
- \loop
- \textbf{\@name{\@authorname\romannumeral\@authori}},
- \@name{\@authoraffil\romannumeral\@authori},
- email: \@name{\@authoremail\romannumeral\@authori}.%
- \@increment \@authori
- \if \@notp{\@gtrp{\@authori}{\@authorcount}} \repeat}%
- \par
- \fi
- \global\@setflag \@addauthorsdone = \@true}
-\fi
-
-\let \addauthorsection = \additionalauthors
-
-\def \@placetitlenotes {
- \the\@titlenotetext}
-
-% Utilities
-% ---------
-
-
-\newcommand{\centeroncapheight}[1]{%
- {\setbox\@tempboxa = \hbox{#1}%
- \@measurecapheight{\@tempdima}% % Calculate ht(CAP) - ht(text)
- \advance \@tempdima by -\ht\@tempboxa % ------------------
- \divide \@tempdima by 2 % 2
- \raise \@tempdima \box\@tempboxa}}
-
-\newbox{\@measbox}
-
-\def \@measurecapheight #1{% {\dimen}
- \setbox\@measbox = \hbox{ABCDEFGHIJKLMNOPQRSTUVWXYZ}%
- #1 = \ht\@measbox}
-
-\long\def \@titlefootnote #1#2{%
- \insert\footins{%
- \reset@font\footnotesize
- \interlinepenalty\interfootnotelinepenalty
- \splittopskip\footnotesep
- \splitmaxdepth \dp\strutbox \floatingpenalty \@MM
- \hsize\columnwidth \@parboxrestore
-%%% \protected@edef\@currentlabel{%
-%%% \csname p@footnote\endcsname\@thefnmark}%
- \color@begingroup
- \def \@makefnmark {$^{#1}$}%
- \@makefntext{%
- \rule\z@\footnotesep\ignorespaces#2\@finalstrut\strutbox}%
- \color@endgroup}}
-
-% LaTeX Modifications
-% ----- -------------
-
-\def \@seccntformat #1{%
- \@name{\the#1}%
- \@expandaftertwice\@seccntformata \csname the#1\endcsname.\@mark
- \quad}
-
-\def \@seccntformata #1.#2\@mark{%
- \if \@emptyargp{#2}.\fi}
-
-% Revision History
-% -------- -------
-
-
-% Date Person Ver. Change
-% ---- ------ ---- ------
-
-% 2004.09.12 PCA 0.1--5 Preliminary development.
-
-% 2004.11.18 PCA 0.5 Start beta testing.
-
-% 2004.11.19 PCA 0.6 Obsolete \author and replace with
-% \authorinfo.
-% Add 'nocopyrightspace' option.
-% Compress article opener spacing.
-% Add 'mathtime' option.
-% Increase text height by 6 points.
-
-% 2004.11.28 PCA 0.7 Add 'cm/computermodern' options.
-% Change default to Times text.
-
-% 2004.12.14 PCA 0.8 Remove use of mathptm.sty; it cannot
-% coexist with latexsym or amssymb.
-
-% 2005.01.20 PCA 0.9 Rename class file to sigplanconf.cls.
-
-% 2005.03.05 PCA 0.91 Change default copyright data.
-
-% 2005.03.06 PCA 0.92 Add at-signs to some macro names.
-
-% 2005.03.07 PCA 0.93 The 'onecolumn' option defaults to '11pt',
-% and it uses the full type width.
-
-% 2005.03.15 PCA 0.94 Add at-signs to more macro names.
-% Allow margin paragraphs during review.
-
-% 2005.03.22 PCA 0.95 Implement \euro.
-% Remove proof and newdef environments.
-
-% 2005.05.06 PCA 1.0 Eliminate 'onecolumn' option.
-% Change footer to small italic and eliminate
-% left portion if no \preprintfooter.
-% Eliminate copyright notice if preprint.
-% Clean up and shrink copyright box.
-
-% 2005.05.30 PCA 1.1 Add alternate permission statements.
-
-% 2005.06.29 PCA 1.1 Publish final first edition of guide.
-
-% 2005.07.14 PCA 1.2 Add \subparagraph.
-% Use block paragraphs in lists, and adjust
-% spacing between items and paragraphs.
-
-% 2006.06.22 PCA 1.3 Add 'reprint' option and associated
-% commands.
-
-% 2006.08.24 PCA 1.4 Fix bug in \maketitle case command.
-
-% 2007.03.13 PCA 1.5 The title banner only displays with the
-% 'preprint' option.
-
-% 2007.06.06 PCA 1.6 Use \bibfont in \thebibliography.
-% Add 'natbib' option to load and configure
-% the natbib package.
-
-% 2007.11.20 PCA 1.7 Balance line lengths in centered article
-% title (thanks to Norman Ramsey).
-
-% 2009.01.26 PCA 1.8 Change natbib \bibpunct values.
-
-% 2009.03.24 PCA 1.9 Change natbib to use the 'numbers' option.
-% Change templates to use 'natbib' option.
-
-% 2009.09.01 PCA 2.0 Add \reprintprice command (suggested by
-% Stephen Chong).
-
-% 2009.09.08 PCA 2.1 Make 'natbib' the default; add 'nonatbib'.
-% SB Add 'authoryear' and 'numbers' (default) to
-% control citation style when using natbib.
-% Add \bibpunct to change punctuation for
-% 'authoryear' style.
-
-% 2009.09.21 PCA 2.2 Add \softraggedright to the thebibliography
-% environment. Also add to template so it will
-% happen with natbib.
-
-% 2009.09.30 PCA 2.3 Remove \softraggedright from thebibliography.
-% Just include in the template.
-
-% 2010.05.24 PCA 2.4 Obfuscate author's email address.
--- a/IsaMakefile Sat Dec 17 16:58:11 2011 +0000
+++ b/IsaMakefile Sat Dec 17 17:03:01 2011 +0000
@@ -4,7 +4,7 @@
default: tests
images:
-all: tests esop pearl pearl-jv qpaper slides
+all: tests
## global settings
@@ -23,156 +23,6 @@
$(LOG)/HOL-Nominal2.gz: Nominal/ROOT.ML Nominal/*.thy
@cd Nominal; $(USEDIR) -b -d "" HOL Nominal2
-## ESOP Paper
-
-esop: $(LOG)/HOL-ESOP-Paper.gz
-
-$(LOG)/HOL-ESOP-Paper.gz: ESOP-Paper/ROOT.ML ESOP-Paper/document/root.* ESOP-Paper/*.thy
- @$(USEDIR) -f ROOT.ML -D generated Nominal2 ESOP-Paper
- $(ISABELLE_TOOL) document -o pdf ESOP-Paper/generated
- @cp ESOP-Paper/document.pdf esop-paper.pdf
-
-## LMCS Paper
-
-lmcs: $(LOG)/HOL-LMCS-Paper.gz
-
-$(LOG)/HOL-LMCS-Paper.gz: LMCS-Paper/ROOT.ML LMCS-Paper/document/root.* LMCS-Paper/*.thy
- @$(USEDIR) -f ROOT.ML -D generated Nominal2 LMCS-Paper
- $(ISABELLE_TOOL) document -o pdf LMCS-Paper/generated
- @cp LMCS-Paper/document.pdf lmcs-paper.pdf
-
-## Pearl Paper ITP
-
-pearl: $(LOG)/HOL-Pearl.gz
-
-$(LOG)/HOL-Pearl.gz: Pearl/ROOT.ML Pearl/document/root.* Pearl/*.thy
- @$(USEDIR) -D generated HOL Pearl
- $(ISABELLE_TOOL) document -o pdf Pearl/generated
- @cp Pearl/document.pdf pearl.pdf
-
-## Pearl Journal Paper
-
-$(LOG)/HOL-Pearl-jv.gz: Pearl-jv/ROOT.ML Nominal/*.thy
- @cd Pearl-jv; $(USEDIR) -b -f ROOT.ML HOL HOL-Pearl-jv
-
-pearl-jv: $(LOG)/HOL-Pearl-jv.gz Pearl-jv/ROOT2.ML Pearl-jv/document/root.* Pearl-jv/*.thy
- @$(USEDIR) -f ROOT2.ML -D generated HOL-Pearl-jv Pearl-jv
- @$(ISABELLE_TOOL) document -o pdf Pearl-jv/generated
- @cp Pearl-jv/document.pdf pearl-jv.pdf
-
-## Quotient Paper
-
-qpaper: $(LOG)/HOL-QPaper.gz
-
-$(LOG)/HOL-QPaper.gz: Quotient-Paper/ROOT.ML Quotient-Paper/document/root.* Quotient-Paper/*.thy
- @$(USEDIR) -D generated HOL Quotient-Paper
- $(ISABELLE_TOOL) document -o pdf Quotient-Paper/generated
- @cp Quotient-Paper/document.pdf qpaper.pdf
-
-## Nominal Functions paper
-
-fnpaper: $(LOG)/HOL-FnPaper.gz
-
-$(LOG)/HOL-FnPaper.gz: Fun-Paper/ROOT.ML Fun-Paper/document/root.* Fun-Paper/*.thy
- @$(USEDIR) -D generated HOL Fun-Paper
- $(ISABELLE_TOOL) document -o pdf Fun-Paper/generated
- @cp Fun-Paper/document.pdf fnpaper.pdf
-
-## Slides
-
-session1: Slides/ROOT1.ML \
- Slides/document/root* \
- Slides/Slides1.thy
- @$(USEDIR) -D generated1 -f ROOT1.ML HOL-Nominal Slides
-
-slides1: session1
- rm -f Slides/generated1/*.aux # otherwise latex will fall over
- cd Slides/generated1 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cd Slides/generated1 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cp Slides/generated1/root.beamer.pdf Slides/slides1.pdf
-
-session2: Slides/ROOT2.ML \
- Slides/document/root* \
- Slides/Slides2.thy
- @$(USEDIR) -D generated2 -f ROOT2.ML HOL-Nominal Slides
-
-slides2: session2
- rm -f Slides/generated2/*.aux # otherwise latex will fall over
- cd Slides/generated2 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cd Slides/generated2 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cp Slides/generated2/root.beamer.pdf Slides/slides2.pdf
-
-session3: Slides/ROOT3.ML \
- Slides/document/root* \
- Slides/Slides3.thy
- @$(USEDIR) -D generated3 -f ROOT3.ML HOL-Nominal Slides
-
-slides3: session3
- rm -f Slides/generated3/*.aux # otherwise latex will fall over
- cd Slides/generated3 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cd Slides/generated3 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cp Slides/generated3/root.beamer.pdf Slides/slides3.pdf
-
-session4: Slides/ROOT4.ML \
- Slides/document/root* \
- Slides/Slides4.thy
- @$(USEDIR) -D generated4 -f ROOT4.ML HOL-Nominal Slides
-
-slides4: session4
- rm -f Slides/generated4/*.aux # otherwise latex will fall over
- cd Slides/generated4 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cd Slides/generated4 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cp Slides/generated4/root.beamer.pdf Slides/slides4.pdf
-
-session5: Slides/ROOT5.ML \
- Slides/document/root* \
- Slides/Slides5.thy
- @$(USEDIR) -D generated5 -f ROOT5.ML HOL-Nominal Slides
-
-slides5: session5
- rm -f Slides/generated5/*.aux # otherwise latex will fall over
- cd Slides/generated5 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cd Slides/generated5 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cp Slides/generated5/root.beamer.pdf Slides/slides5.pdf
-
-session6: Slides/ROOT6.ML \
- Slides/document/root* \
- Slides/Slides6.thy
- @$(USEDIR) -D generated6 -f ROOT6.ML HOL-Nominal Slides
-
-slides6: session6
- rm -f Slides/generated6/*.aux # otherwise latex will fall over
- cd Slides/generated6 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cd Slides/generated6 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cp Slides/generated6/root.beamer.pdf Slides/slides6.pdf
-
-session7: Slides/ROOT7.ML \
- Slides/document/root* \
- Slides/Slides6.thy
- @$(USEDIR) -D generated7 -f ROOT7.ML HOL Slides
-
-slides7: session7
- rm -f Slides/generated7/*.aux # otherwise latex will fall over
- cd Slides/generated7 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cd Slides/generated7 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cp Slides/generated7/root.beamer.pdf Slides/slides7.pdf
-
-session8: Slides/ROOT8.ML \
- Slides/document/root* \
- Slides/Slides6.thy
- @$(USEDIR) -D generated8 -f ROOT8.ML HOL-Nominal Slides
-
-slides8: session8
- rm -f Slides/generated8/*.aux # otherwise latex will fall over
- cd Slides/generated8 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cd Slides/generated8 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
- cp Slides/generated8/root.beamer.pdf Slides/slides8.pdf
-
-slides: slides1 slides2 slides3 slides4 slides5 slides6 slides7 slides8
-
-
-
-
## clean
clean:
--- a/LMCS-Paper/Appendix.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,135 +0,0 @@
-(*<*)
-theory Appendix
-imports "../Nominal/Nominal2" "~~/src/HOL/Library/LaTeXsugar"
-begin
-
-consts
- fv :: "'a \<Rightarrow> 'b"
- abs_set :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
- alpha_bn :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
- abs_set2 :: "'a \<Rightarrow> perm \<Rightarrow> 'b \<Rightarrow> 'c"
- Abs_dist :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
- Abs_print :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
-
-definition
- "equal \<equiv> (op =)"
-
-notation (latex output)
- swap ("'(_ _')" [1000, 1000] 1000) and
- fresh ("_ # _" [51, 51] 50) and
- fresh_star ("_ #\<^sup>* _" [51, 51] 50) and
- supp ("supp _" [78] 73) and
- uminus ("-_" [78] 73) and
- If ("if _ then _ else _" 10) and
- alpha_set ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{set}}$}}>\<^bsup>_, _, _\<^esup> _") and
- alpha_lst ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{list}}$}}>\<^bsup>_, _, _\<^esup> _") and
- alpha_res ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{res}}$}}>\<^bsup>_, _, _\<^esup> _") and
- abs_set ("_ \<approx>\<^raw:{$\,_{\textit{abs\_set}}$}> _") and
- abs_set2 ("_ \<approx>\<^raw:\raisebox{-1pt}{\makebox[0mm][l]{$\,_{\textit{list}}$}}>\<^bsup>_\<^esup> _") and
- fv ("fa'(_')" [100] 100) and
- equal ("=") and
- alpha_abs_set ("_ \<approx>\<^raw:{$\,_{\textit{abs\_set}}$}> _") and
- Abs_set ("[_]\<^bsub>set\<^esub>._" [20, 101] 999) and
- Abs_lst ("[_]\<^bsub>list\<^esub>._") and
- Abs_dist ("[_]\<^bsub>#list\<^esub>._") and
- Abs_res ("[_]\<^bsub>res\<^esub>._") and
- Abs_print ("_\<^bsub>set\<^esub>._") and
- Cons ("_::_" [78,77] 73) and
- supp_set ("aux _" [1000] 10) and
- alpha_bn ("_ \<approx>bn _")
-
-consts alpha_trm ::'a
-consts fa_trm :: 'a
-consts alpha_trm2 ::'a
-consts fa_trm2 :: 'a
-consts ast :: 'a
-consts ast' :: 'a
-notation (latex output)
- alpha_trm ("\<approx>\<^bsub>trm\<^esub>") and
- fa_trm ("fa\<^bsub>trm\<^esub>") and
- alpha_trm2 ("'(\<approx>\<^bsub>assn\<^esub>, \<approx>\<^bsub>trm\<^esub>')") and
- fa_trm2 ("'(fa\<^bsub>assn\<^esub>, fa\<^bsub>trm\<^esub>')") and
- ast ("'(as, t')") and
- ast' ("'(as', t\<PRIME> ')")
-
-(*>*)
-
-text {*
-\appendix
-\section*{Appendix}
-
- Details for one case in Theorem \ref{suppabs}, which the reader might like to ignore.
- By definition of the abstraction type @{text "abs_set"}
- we have
- %
- \begin{equation}\label{abseqiff}
- @{thm (lhs) Abs_eq_iff(1)[where bs="as" and cs="bs", no_vars]} \;\;\text{if and only if}\;\;
- @{thm (rhs) Abs_eq_iff(1)[where bs="as" and cs="bs", no_vars]}
- \end{equation}
-
- \noindent
- and also
-
- \begin{equation}\label{absperm}
- @{thm permute_Abs(1)[no_vars]}%
- \end{equation}
-
- \noindent
- The second fact derives from the definition of permutations acting on pairs
- and $\alpha$-equivalence being equivariant. With these two facts at our disposal, we can show
- the following lemma about swapping two atoms in an abstraction.
-
- \begin{lemma}
- @{thm[mode=IfThen] Abs_swap1(1)[where bs="as", no_vars]}
- \end{lemma}
-
- \begin{proof}
- This lemma is straightforward using \eqref{abseqiff} and observing that
- the assumptions give us @{term "(a \<rightleftharpoons> b) \<bullet> (supp x - as) = (supp x - as)"}.
- Moreover @{text supp} and set difference are equivariant (see \cite{HuffmanUrban10}).
- \end{proof}
-
- \noindent
- Assuming that @{text "x"} has finite support, this lemma together
- with \eqref{absperm} allows us to show
-
- \begin{equation}\label{halfone}
- @{thm Abs_supports(1)[no_vars]}
- \end{equation}
-
- \noindent
- which gives us ``one half'' of
- Theorem~\ref{suppabs} (the notion of supports is defined in \cite{HuffmanUrban10}).
- The ``other half'' is a bit more involved. To establish
- it, we use a trick from \cite{Pitts04} and first define an auxiliary
- function @{text aux}, taking an abstraction as argument:
- @{thm supp_set.simps[THEN eq_reflection, no_vars]}.
-
- We can show that
- @{text "aux"} is equivariant (since @{term "p \<bullet> (supp x - as) = (supp (p \<bullet> x)) - (p \<bullet> as)"})
- and therefore has empty support.
- This in turn means
-
- \begin{center}
- @{text "supp (aux ([as]\<^bsub>set\<^esub>. x)) \<subseteq> supp ([as]\<^bsub>set\<^esub> x)"}
- \end{center}
-
- \noindent
- Assuming @{term "supp x - as"} is a finite set,
- we further obtain
-
- \begin{equation}\label{halftwo}
- @{thm (concl) Abs_supp_subset1(1)[no_vars]}
- \end{equation}
-
- \noindent
- since for finite sets of atoms, @{text "bs"}, we have
- @{thm (concl) supp_finite_atom_set[where S="bs", no_vars]}.
- Finally, taking \eqref{halfone} and \eqref{halftwo} together establishes
- Theorem~\ref{suppabs}.
-
-*}
-
-(*<*)
-end
-(*>*)
--- a/LMCS-Paper/Paper.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2672 +0,0 @@
-(*<*)
-theory Paper
-imports "../Nominal/Nominal2"
- "~~/src/HOL/Library/LaTeXsugar"
-begin
-
-consts
- fv :: "'a \<Rightarrow> 'b"
- abs_set :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
- alpha_bn :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
- abs_set2 :: "'a \<Rightarrow> perm \<Rightarrow> 'b \<Rightarrow> 'c"
- equ2 :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
- Abs_dist :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
- Abs_print :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
-
-definition
- "equal \<equiv> (op =)"
-
-fun alpha_set_ex where
- "alpha_set_ex (bs, x) R f (cs, y) = (\<exists>pi. alpha_set (bs, x) R f pi (cs, y))"
-
-fun alpha_res_ex where
- "alpha_res_ex (bs, x) R f pi (cs, y) = (\<exists>pi. alpha_res (bs, x) R f pi (cs, y))"
-
-fun alpha_lst_ex where
- "alpha_lst_ex (bs, x) R f pi (cs, y) = (\<exists>pi. alpha_lst (bs, x) R f pi (cs, y))"
-
-
-
-notation (latex output)
- swap ("'(_ _')" [1000, 1000] 1000) and
- fresh ("_ # _" [51, 51] 50) and
- fresh_star ("_ #\<^sup>* _" [51, 51] 50) and
- supp ("supp _" [78] 73) and
- uminus ("-_" [78] 73) and
- If ("if _ then _ else _" 10) and
- alpha_set_ex ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{set}}$}}>\<^bsup>_, _\<^esup> _") and
- alpha_lst_ex ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{list}}$}}>\<^bsup>_, _\<^esup> _") and
- alpha_res_ex ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{set+}}$}}>\<^bsup>_, _\<^esup> _") and
- abs_set ("_ \<approx>\<^raw:{$\,_{\textit{abs\_set}}$}> _") and
- abs_set2 ("_ \<approx>\<^raw:\raisebox{-1pt}{\makebox[0mm][l]{$\,_{\textit{list}}$}}>\<^bsup>_\<^esup> _") and
- fv ("fa'(_')" [100] 100) and
- equal ("=") and
- alpha_abs_set ("_ \<approx>\<^raw:{$\,_{\textit{abs\_set}}$}> _") and
- alpha_abs_lst ("_ \<approx>\<^raw:{$\,_{\textit{abs\_list}}$}> _") and
- alpha_abs_res ("_ \<approx>\<^raw:{$\,_{\textit{abs\_set+}}$}> _") and
- Abs_set ("[_]\<^bsub>set\<^esub>._" [20, 101] 999) and
- Abs_lst ("[_]\<^bsub>list\<^esub>._" [20, 101] 999) and
- Abs_dist ("[_]\<^bsub>#list\<^esub>._" [20, 101] 999) and
- Abs_res ("[_]\<^bsub>set+\<^esub>._") and
- Abs_print ("_\<^bsub>set\<^esub>._") and
- Cons ("_::_" [78,77] 73) and
- supp_set ("aux _" [1000] 10) and
- alpha_bn ("_ \<approx>bn _")
-
-consts alpha_trm ::'a
-consts fa_trm :: 'a
-consts fa_trm_al :: 'a
-consts alpha_trm2 ::'a
-consts fa_trm2 :: 'a
-consts fa_trm2_al :: 'a
-consts supp2 :: 'a
-consts ast :: 'a
-consts ast' :: 'a
-consts bn_al :: "'b \<Rightarrow> 'a"
-notation (latex output)
- alpha_trm ("\<approx>\<^bsub>trm\<^esub>") and
- fa_trm ("fa\<^bsub>trm\<^esub>") and
- fa_trm_al ("fa\<AL>\<^bsub>trm\<^esub>") and
- alpha_trm2 ("'(\<approx>\<^bsub>assn\<^esub>, \<approx>\<^bsub>trm\<^esub>')") and
- fa_trm2 ("'(fa\<^bsub>assn\<^esub>, fa\<^bsub>trm\<^esub>')") and
- fa_trm2_al ("'(fa\<AL>\<^bsub>assn\<^esub>, fa\<AL>\<^bsub>trm\<^esub>')") and
- ast ("'(as, t')") and
- ast' ("'(as', t\<PRIME> ')") and
- equ2 ("'(=, =')") and
- supp2 ("'(supp, supp')") and
- bn_al ("bn\<^sup>\<alpha> _" [100] 100)
-(*>*)
-
-
-section {* Introduction *}
-
-text {*
- So far, Nominal Isabelle provided a mechanism for constructing alpha-equated
- terms, for example lambda-terms
-
- \[
- @{text "t ::= x | t t | \<lambda>x. t"}
- \]\smallskip
-
- \noindent
- where free and bound variables have names. For such alpha-equated terms,
- Nominal Isabelle derives automatically a reasoning infrastructure that has
- been used successfully in formalisations of an equivalence checking
- algorithm for LF \cite{UrbanCheneyBerghofer08}, Typed
- Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for concurrency
- \cite{BengtsonParow09} and a strong normalisation result for cut-elimination
- in classical logic \cite{UrbanZhu08}. It has also been used by Pollack for
- formalisations in the locally-nameless approach to binding
- \cite{SatoPollack10}.
-
- However, Nominal Isabelle has fared less well in a formalisation of the
- algorithm W \cite{UrbanNipkow09}, where types and type-schemes are,
- respectively, of the form
-
- \begin{equation}\label{tysch}
- \begin{array}{l}
- @{text "T ::= x | T \<rightarrow> T"}\hspace{15mm}
- @{text "S ::= \<forall>{x\<^isub>1,\<dots>, x\<^isub>n}. T"}
- \end{array}
- \end{equation}\smallskip
-
- \noindent
- and the @{text "\<forall>"}-quantification binds a finite (possibly empty) set of
- type-variables. While it is possible to implement this kind of more general
- binders by iterating single binders, this leads to a rather clumsy
- formalisation of W. The need of iterating single binders is also one reason
- why Nominal Isabelle and similar theorem provers that only provide
- mechanisms for binding single variables have not fared extremely well with
- the more advanced tasks in the POPLmark challenge \cite{challenge05},
- because also there one would like to bind multiple variables at once.
-
- Binding multiple variables has interesting properties that cannot be captured
- easily by iterating single binders. For example in the case of type-schemes we do not
- want to make a distinction about the order of the bound variables. Therefore
- we would like to regard in \eqref{ex1} below the first pair of type-schemes as alpha-equivalent,
- but assuming that @{text x}, @{text y} and @{text z} are distinct variables,
- the second pair should \emph{not} be alpha-equivalent:
-
- \begin{equation}\label{ex1}
- @{text "\<forall>{x, y}. x \<rightarrow> y \<approx>\<^isub>\<alpha> \<forall>{y, x}. y \<rightarrow> x"}\hspace{10mm}
- @{text "\<forall>{x, y}. x \<rightarrow> y \<notapprox>\<^isub>\<alpha> \<forall>{z}. z \<rightarrow> z"}
- \end{equation}\smallskip
-
- \noindent
- Moreover, we like to regard type-schemes as alpha-equivalent, if they differ
- only on \emph{vacuous} binders, such as
-
- \begin{equation}\label{ex3}
- @{text "\<forall>{x}. x \<rightarrow> y \<approx>\<^isub>\<alpha> \<forall>{x, z}. x \<rightarrow> y"}
- \end{equation}\smallskip
-
- \noindent
- where @{text z} does not occur freely in the type. In this paper we will
- give a general binding mechanism and associated notion of alpha-equivalence
- that can be used to faithfully represent this kind of binding in Nominal
- Isabelle. The difficulty of finding the right notion for alpha-equivalence
- can be appreciated in this case by considering that the definition given for
- type-schemes by Leroy in \cite[Page 18--19]{Leroy92} is incorrect (it omits a side-condition).
-
- However, the notion of alpha-equivalence that is preserved by vacuous
- binders is not always wanted. For example in terms like
-
- \begin{equation}\label{one}
- @{text "\<LET> x = 3 \<AND> y = 2 \<IN> x - y \<END>"}
- \end{equation}\smallskip
-
- \noindent
- we might not care in which order the assignments @{text "x = 3"} and
- \mbox{@{text "y = 2"}} are given, but it would be often unusual (particularly
- in strict languages) to regard \eqref{one} as alpha-equivalent with
-
- \[
- @{text "\<LET> x = 3 \<AND> y = 2 \<AND> z = foo \<IN> x - y \<END>"}
- \]\smallskip
-
- \noindent
- Therefore we will also provide a separate binding mechanism for cases in
- which the order of binders does not matter, but the `cardinality' of the
- binders has to agree.
-
- However, we found that this is still not sufficient for dealing with
- language constructs frequently occurring in programming language
- research. For example in @{text "\<LET>"}s containing patterns like
-
- \begin{equation}\label{two}
- @{text "\<LET> (x, y) = (3, 2) \<IN> x - y \<END>"}
- \end{equation}\smallskip
-
- \noindent
- we want to bind all variables from the pattern inside the body of the
- $\mathtt{let}$, but we also care about the order of these variables, since
- we do not want to regard \eqref{two} as alpha-equivalent with
-
- \[
- @{text "\<LET> (y, x) = (3, 2) \<IN> x - y \<END>"}
- \]\smallskip
-
- \noindent
- As a result, we provide three general binding mechanisms each of which binds
- multiple variables at once, and let the user chose which one is intended
- when formalising a term-calculus.
-
- By providing these general binding mechanisms, however, we have to work
- around a problem that has been pointed out by Pottier \cite{Pottier06} and
- Cheney \cite{Cheney05}: in @{text "\<LET>"}-constructs of the form
-
- \[
- @{text "\<LET> x\<^isub>1 = t\<^isub>1 \<AND> \<dots> \<AND> x\<^isub>n = t\<^isub>n \<IN> s \<END>"}
- \]\smallskip
-
- \noindent
- we care about the information that there are as many bound variables @{text
- "x\<^isub>i"} as there are @{text "t\<^isub>i"}. We lose this information if
- we represent the @{text "\<LET>"}-constructor by something like
-
- \[
- @{text "\<LET> (\<lambda>x\<^isub>1\<dots>x\<^isub>n . s) [t\<^isub>1,\<dots>,t\<^isub>n]"}
- \]\smallskip
-
- \noindent
- where the notation @{text "\<lambda>_ . _"} indicates that the list of @{text
- "x\<^isub>i"} becomes bound in @{text s}. In this representation the term
- \mbox{@{text "\<LET> (\<lambda>x . s) [t\<^isub>1, t\<^isub>2]"}} is a perfectly
- legal instance, but the lengths of the two lists do not agree. To exclude
- such terms, additional predicates about well-formed terms are needed in
- order to ensure that the two lists are of equal length. This can result in
- very messy reasoning (see for example~\cite{BengtsonParow09}). To avoid
- this, we will allow type specifications for @{text "\<LET>"}s as follows
-
- \[
- \mbox{\begin{tabular}{r@ {\hspace{2mm}}r@ {\hspace{2mm}}ll}
- @{text trm} & @{text "::="} & @{text "\<dots>"} \\
- & @{text "|"} & @{text "\<LET> as::assn s::trm"}\hspace{2mm}
- \isacommand{binds} @{text "bn(as)"} \isacommand{in} @{text "s"}\\[1mm]
- @{text assn} & @{text "::="} & @{text "\<ANIL>"}\\
- & @{text "|"} & @{text "\<ACONS> name trm assn"}
- \end{tabular}}
- \]\smallskip
-
- \noindent
- where @{text assn} is an auxiliary type representing a list of assignments
- and @{text bn} an auxiliary function identifying the variables to be bound
- by the @{text "\<LET>"}. This function can be defined by recursion over @{text
- assn} as follows
-
- \[
- @{text "bn(\<ANIL>) ="}~@{term "{}"} \hspace{10mm}
- @{text "bn(\<ACONS> x t as) = {x} \<union> bn(as)"}
- \]\smallskip
-
- \noindent
- The scope of the binding is indicated by labels given to the types, for
- example @{text "s::trm"}, and a binding clause, in this case
- \isacommand{binds} @{text "bn(as)"} \isacommand{in} @{text "s"}. This binding
- clause states that all the names the function @{text "bn(as)"} returns
- should be bound in @{text s}. This style of specifying terms and bindings
- is heavily inspired by the syntax of the Ott-tool \cite{ott-jfp}. Our work
- extends Ott in several aspects: one is that we support three binding
- modes---Ott has only one, namely the one where the order of binders matters.
- Another is that our reasoning infrastructure, like strong induction principles
- and the notion of free variables, is derived from first principles within
- the Isabelle/HOL theorem prover.
-
- However, we will not be able to cope with all specifications that are
- allowed by Ott. One reason is that Ott lets the user specify `empty' types
- like \mbox{@{text "t ::= t t | \<lambda>x. t"}} where no clause for variables is
- given. Arguably, such specifications make some sense in the context of Coq's
- type theory (which Ott supports), but not at all in a HOL-based environment
- where every datatype must have a non-empty set-theoretic model
- \cite{Berghofer99}. Another reason is that we establish the reasoning
- infrastructure for alpha-\emph{equated} terms. In contrast, Ott produces a
- reasoning infrastructure in Isabelle/HOL for \emph{non}-alpha-equated, or
- `raw', terms. While our alpha-equated terms and the `raw' terms produced by
- Ott use names for bound variables, there is a key difference: working with
- alpha-equated terms means, for example, that the two type-schemes
-
- \[
- @{text "\<forall>{x}. x \<rightarrow> y = \<forall>{x, z}. x \<rightarrow> y"}
- \]\smallskip
-
- \noindent
- are not just alpha-equal, but actually \emph{equal}! As a result, we can
- only support specifications that make sense on the level of alpha-equated
- terms (offending specifications, which for example bind a variable according
- to a variable bound somewhere else, are not excluded by Ott, but we have
- to).
-
- Our insistence on reasoning with alpha-equated terms comes from the
- wealth of experience we gained with the older version of Nominal Isabelle:
- for non-trivial properties, reasoning with alpha-equated terms is much
- easier than reasoning with `raw' terms. The fundamental reason for this is
- that the HOL-logic underlying Nominal Isabelle allows us to replace
- `equals-by-equals'. In contrast, replacing
- `alpha-equals-by-alpha-equals' in a representation based on `raw' terms
- requires a lot of extra reasoning work.
-
- Although in informal settings a reasoning infrastructure for alpha-equated
- terms is nearly always taken for granted, establishing it automatically in
- Isabelle/HOL is a rather non-trivial task. For every
- specification we will need to construct type(s) containing as elements the
- alpha-equated terms. To do so, we use the standard HOL-technique of defining
- a new type by identifying a non-empty subset of an existing type. The
- construction we perform in Isabelle/HOL can be illustrated by the following picture:
-
- \begin{equation}\label{picture}
- \mbox{\begin{tikzpicture}[scale=1.1]
- %\draw[step=2mm] (-4,-1) grid (4,1);
-
- \draw[very thick] (0.7,0.4) circle (4.25mm);
- \draw[rounded corners=1mm, very thick] ( 0.0,-0.8) rectangle ( 1.8, 0.9);
- \draw[rounded corners=1mm, very thick] (-1.95,0.85) rectangle (-2.85,-0.05);
-
- \draw (-2.0, 0.845) -- (0.7,0.845);
- \draw (-2.0,-0.045) -- (0.7,-0.045);
-
- \draw ( 0.7, 0.5) node {\footnotesize\begin{tabular}{@ {}c@ {}}$\alpha$-\\[-1mm]classes\end{tabular}};
- \draw (-2.4, 0.5) node {\footnotesize\begin{tabular}{@ {}c@ {}}$\alpha$-eq.\\[-1mm]terms\end{tabular}};
- \draw (1.8, 0.48) node[right=-0.1mm]
- {\small\begin{tabular}{@ {}l@ {}}existing\\[-1mm] type\\ (sets of raw terms)\end{tabular}};
- \draw (0.9, -0.35) node {\footnotesize\begin{tabular}{@ {}l@ {}}non-empty\\[-1mm]subset\end{tabular}};
- \draw (-3.25, 0.55) node {\small\begin{tabular}{@ {}l@ {}}new\\[-1mm]type\end{tabular}};
-
- \draw[<->, very thick] (-1.8, 0.3) -- (-0.1,0.3);
- \draw (-0.95, 0.3) node[above=0mm] {\footnotesize{}isomorphism};
-
- \end{tikzpicture}}
- \end{equation}\smallskip
-
- \noindent
- We take as the starting point a definition of raw terms (defined as a
- datatype in Isabelle/HOL); then identify the alpha-equivalence classes in
- the type of sets of raw terms according to our alpha-equivalence relation,
- and finally define the new type as these alpha-equivalence classes (the
- non-emptiness requirement is always satisfied whenever the raw terms are
- definable as datatype in Isabelle/HOL and our relation for alpha-equivalence
- is an equivalence relation).
-
- The fact that we obtain an isomorphism between the new type and the
- non-empty subset shows that the new type is a faithful representation of
- alpha-equated terms. That is not the case for example for terms using the
- locally nameless representation of binders \cite{McKinnaPollack99}: in this
- representation there are `junk' terms that need to be excluded by
- reasoning about a well-formedness predicate.
-
- The problem with introducing a new type in Isabelle/HOL is that in order to
- be useful, a reasoning infrastructure needs to be `lifted' from the
- underlying subset to the new type. This is usually a tricky and arduous
- task. To ease it, we re-implemented in Isabelle/HOL \cite{KaliszykUrban11}
- the quotient package described by Homeier \cite{Homeier05} for the HOL4
- system. This package allows us to lift definitions and theorems involving
- raw terms to definitions and theorems involving alpha-equated terms. For
- example if we define the free-variable function over raw lambda-terms
- as follows
-
- \[
- \mbox{\begin{tabular}{l@ {\hspace{1mm}}r@ {\hspace{1mm}}l}
- @{text "fv(x)"} & @{text "\<equiv>"} & @{text "{x}"}\\
- @{text "fv(t\<^isub>1 t\<^isub>2)"} & @{text "\<equiv>"} & @{text "fv(t\<^isub>1) \<union> fv(t\<^isub>2)"}\\
- @{text "fv(\<lambda>x.t)"} & @{text "\<equiv>"} & @{text "fv(t) - {x}"}
- \end{tabular}}
- \]\smallskip
-
- \noindent
- then with the help of the quotient package we can obtain a function @{text "fv\<^sup>\<alpha>"}
- operating on quotients, that is alpha-equivalence classes of lambda-terms. This
- lifted function is characterised by the equations
-
- \[
- \mbox{\begin{tabular}{l@ {\hspace{1mm}}r@ {\hspace{1mm}}l}
- @{text "fv\<^sup>\<alpha>(x)"} & @{text "="} & @{text "{x}"}\\
- @{text "fv\<^sup>\<alpha>(t\<^isub>1 t\<^isub>2)"} & @{text "="} & @{text "fv\<^sup>\<alpha>(t\<^isub>1) \<union> fv\<^sup>\<alpha>(t\<^isub>2)"}\\
- @{text "fv\<^sup>\<alpha>(\<lambda>x.t)"} & @{text "="} & @{text "fv\<^sup>\<alpha>(t) - {x}"}
- \end{tabular}}
- \]\smallskip
-
- \noindent
- (Note that this means also the term-constructors for variables, applications
- and lambda are lifted to the quotient level.) This construction, of course,
- only works if alpha-equivalence is indeed an equivalence relation, and the
- `raw' definitions and theorems are respectful w.r.t.~alpha-equivalence.
- For example, we will not be able to lift a bound-variable function. Although
- this function can be defined for raw terms, it does not respect
- alpha-equivalence and therefore cannot be lifted.
- To sum up, every lifting
- of theorems to the quotient level needs proofs of some respectfulness
- properties (see \cite{Homeier05}). In the paper we show that we are able to
- automate these proofs and as a result can automatically establish a reasoning
- infrastructure for alpha-equated terms.\smallskip
-
- The examples we have in mind where our reasoning infrastructure will be
- helpful include the term language of Core-Haskell (see
- Figure~\ref{corehas}). This term language involves patterns that have lists
- of type-, coercion- and term-variables, all of which are bound in @{text
- "\<CASE>"}-expressions. In these patterns we do not know in advance how many
- variables need to be bound. Another example is the algorithm W,
- which includes multiple binders in type-schemes.\medskip
-
- \noindent
- {\bf Contributions:} We provide three new definitions for when terms
- involving general binders are alpha-equivalent. These definitions are
- inspired by earlier work of Pitts \cite{Pitts04}. By means of automatic
- proofs, we establish a reasoning infrastructure for alpha-equated terms,
- including properties about support, freshness and equality conditions for
- alpha-equated terms. We are also able to automatically derive strong
- induction principles that have the variable convention already built in.
- For this we simplify the earlier automated proofs by using the proving tools
- from the function package~\cite{Krauss09} of Isabelle/HOL. The method
- behind our specification of general binders is taken from the Ott-tool, but
- we introduce crucial restrictions, and also extensions, so that our
- specifications make sense for reasoning about alpha-equated terms. The main
- improvement over Ott is that we introduce three binding modes (only one is
- present in Ott), provide formalised definitions for alpha-equivalence and
- for free variables of our terms, and also derive a reasoning infrastructure
- for our specifications from `first principles' inside a theorem prover.
-
-
- \begin{figure}[t]
- \begin{boxedminipage}{\linewidth}
- \begin{center}
- \begin{tabular}{@ {\hspace{8mm}}r@ {\hspace{2mm}}r@ {\hspace{2mm}}l}
- \multicolumn{3}{@ {}l}{Type Kinds}\\
- @{text "\<kappa>"} & @{text "::="} & @{text "\<star> | \<kappa>\<^isub>1 \<rightarrow> \<kappa>\<^isub>2"}\smallskip\\
- \multicolumn{3}{@ {}l}{Coercion Kinds}\\
- @{text "\<iota>"} & @{text "::="} & @{text "\<sigma>\<^isub>1 \<sim> \<sigma>\<^isub>2"}\smallskip\\
- \multicolumn{3}{@ {}l}{Types}\\
- @{text "\<sigma>"} & @{text "::="} & @{text "a | T | \<sigma>\<^isub>1 \<sigma>\<^isub>2 | S\<^isub>n"}$\;\overline{@{text "\<sigma>"}}$@{text "\<^sup>n"}
- @{text "| \<forall>a:\<kappa>. \<sigma> | \<iota> \<Rightarrow> \<sigma>"}\smallskip\\
- \multicolumn{3}{@ {}l}{Coercion Types}\\
- @{text "\<gamma>"} & @{text "::="} & @{text "c | C | \<gamma>\<^isub>1 \<gamma>\<^isub>2 | S\<^isub>n"}$\;\overline{@{text "\<gamma>"}}$@{text "\<^sup>n"}
- @{text "| \<forall>c:\<iota>. \<gamma> | \<iota> \<Rightarrow> \<gamma> | refl \<sigma> | sym \<gamma> | \<gamma>\<^isub>1 \<circ> \<gamma>\<^isub>2"}\\
- & @{text "|"} & @{text "\<gamma> @ \<sigma> | left \<gamma> | right \<gamma> | \<gamma>\<^isub>1 \<sim> \<gamma>\<^isub>2 | rightc \<gamma> | leftc \<gamma> | \<gamma>\<^isub>1 \<triangleright> \<gamma>\<^isub>2"}\smallskip\\
- \multicolumn{3}{@ {}l}{Terms}\\
- @{text "e"} & @{text "::="} & @{text "x | K | \<Lambda>a:\<kappa>. e | \<Lambda>c:\<iota>. e | e \<sigma> | e \<gamma> | \<lambda>x:\<sigma>. e | e\<^isub>1 e\<^isub>2"}\\
- & @{text "|"} & @{text "\<LET> x:\<sigma> = e\<^isub>1 \<IN> e\<^isub>2 | \<CASE> e\<^isub>1 \<OF>"}$\;\overline{@{text "p \<rightarrow> e\<^isub>2"}}$ @{text "| e \<triangleright> \<gamma>"}\smallskip\\
- \multicolumn{3}{@ {}l}{Patterns}\\
- @{text "p"} & @{text "::="} & @{text "K"}$\;\overline{@{text "a:\<kappa>"}}\;\overline{@{text "c:\<iota>"}}\;\overline{@{text "x:\<sigma>"}}$\smallskip\\
- \multicolumn{3}{@ {}l}{Constants}\\
- & @{text C} & coercion constants\\
- & @{text T} & value type constructors\\
- & @{text "S\<^isub>n"} & n-ary type functions (which need to be fully applied)\\
- & @{text K} & data constructors\smallskip\\
- \multicolumn{3}{@ {}l}{Variables}\\
- & @{text a} & type variables\\
- & @{text c} & coercion variables\\
- & @{text x} & term variables\\
- \end{tabular}
- \end{center}
- \end{boxedminipage}
- \caption{The System @{text "F\<^isub>C"}
- \cite{CoreHaskell}, also often referred to as \emph{Core-Haskell}. In this
- version of @{text "F\<^isub>C"} we made a modification by separating the
- grammars for type kinds and coercion kinds, as well as for types and coercion
- types. For this paper the interesting term-constructor is @{text "\<CASE>"},
- which binds multiple type-, coercion- and term-variables (the overlines stand for lists).\label{corehas}}
- \end{figure}
-*}
-
-section {* A Short Review of the Nominal Logic Work *}
-
-text {*
- At its core, Nominal Isabelle is an adaption of the nominal logic work by
- Pitts \cite{Pitts03}. This adaptation for Isabelle/HOL is described in
- \cite{HuffmanUrban10} (including proofs). We shall briefly review this work
- to aid the description of what follows.
-
- Two central notions in the nominal logic work are sorted atoms and
- sort-respecting permutations of atoms. We will use the letters @{text "a, b,
- c, \<dots>"} to stand for atoms and @{text "\<pi>, \<pi>\<^isub>1, \<dots>"} to stand for permutations,
- which in Nominal Isabelle have type @{typ perm}. The purpose of atoms is to
- represent variables, be they bound or free. The sorts of atoms can be used
- to represent different kinds of variables, such as the term-, coercion- and
- type-variables in Core-Haskell. It is assumed that there is an infinite
- supply of atoms for each sort. In the interest of brevity, we shall restrict
- ourselves in what follows to only one sort of atoms.
-
- Permutations are bijective functions from atoms to atoms that are
- the identity everywhere except on a finite number of atoms. There is a
- two-place permutation operation written
- @{text "_ \<bullet> _ :: perm \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
- where the generic type @{text "\<beta>"} is the type of the object
- over which the permutation
- acts. In Nominal Isabelle, the identity permutation is written as @{term "0::perm"},
- the composition of two permutations @{term "\<pi>\<^isub>1"} and @{term "\<pi>\<^isub>2"} as \mbox{@{term "\<pi>\<^isub>1 + \<pi>\<^isub>2"}},
- and the inverse permutation of @{term "\<pi>"} as @{text "- \<pi>"}. The permutation
- operation is defined over Isabelle/HOL's type-hierarchy \cite{HuffmanUrban10};
- for example permutations acting on atoms, products, lists, permutations, sets,
- functions and booleans are given by:
-
- \begin{equation}\label{permute}
- \mbox{\begin{tabular}{@ {}c@ {\hspace{10mm}}c@ {}}
- \begin{tabular}{@ {}l@ {}}
- @{text "\<pi> \<bullet> a \<equiv> \<pi> a"}\\
- @{thm permute_prod.simps[where p="\<pi>", no_vars, THEN eq_reflection]}\\[2mm]
- @{thm permute_list.simps(1)[where p="\<pi>", no_vars, THEN eq_reflection]}\\
- @{thm permute_list.simps(2)[where p="\<pi>", no_vars, THEN eq_reflection]}\\
- \end{tabular} &
- \begin{tabular}{@ {}l@ {}}
- @{thm permute_perm_def[where p="\<pi>" and q="\<pi>'", no_vars, THEN eq_reflection]}\\
- @{thm permute_set_eq[where p="\<pi>", no_vars, THEN eq_reflection]}\\
- @{text "\<pi> \<bullet> f \<equiv> \<lambda>x. \<pi> \<bullet> (f (- \<pi> \<bullet> x))"}\\
- @{thm permute_bool_def[where p="\<pi>", no_vars, THEN eq_reflection]}
- \end{tabular}
- \end{tabular}}
- \end{equation}\smallskip
-
- \noindent
- Concrete permutations in Nominal Isabelle are built up from swappings,
- written as \mbox{@{text "(a b)"}}, which are permutations that behave
- as follows:
-
- \[
- @{text "(a b) = \<lambda>c. if a = c then b else if b = c then a else c"}
- \]\smallskip
-
- The most original aspect of the nominal logic work of Pitts is a general
- definition for the notion of the `set of free variables of an object @{text
- "x"}'. This notion, written @{term "supp x"}, is general in the sense that
- it applies not only to lambda-terms (alpha-equated or not), but also to lists,
- products, sets and even functions. Its definition depends only on the
- permutation operation and on the notion of equality defined for the type of
- @{text x}, namely:
-
- \begin{equation}\label{suppdef}
- @{thm supp_def[no_vars, THEN eq_reflection]}
- \end{equation}\smallskip
-
- \noindent
- There is also the derived notion for when an atom @{text a} is \emph{fresh}
- for an @{text x}, defined as
-
- \[
- @{thm fresh_def[no_vars]}
- \]\smallskip
-
- \noindent
- We use for sets of atoms the abbreviation
- @{thm (lhs) fresh_star_def[no_vars]}, defined as
- @{thm (rhs) fresh_star_def[no_vars]}.
- A striking consequence of these definitions is that we can prove
- without knowing anything about the structure of @{term x} that
- swapping two fresh atoms, say @{text a} and @{text b}, leaves
- @{text x} unchanged, namely
-
- \begin{prop}\label{swapfreshfresh}
- If @{thm (prem 1) swap_fresh_fresh[no_vars]} and @{thm (prem 2) swap_fresh_fresh[no_vars]}
- then @{thm (concl) swap_fresh_fresh[no_vars]}.
- \end{prop}
-
- While often the support of an object can be relatively easily
- described, for example for atoms, products, lists, function applications,
- booleans and permutations as follows
-
- \begin{equation}\label{supps}\mbox{
- \begin{tabular}{c@ {\hspace{10mm}}c}
- \begin{tabular}{rcl}
- @{term "supp a"} & $=$ & @{term "{a}"}\\
- @{term "supp (x, y)"} & $=$ & @{term "supp x \<union> supp y"}\\
- @{term "supp []"} & $=$ & @{term "{}"}\\
- @{term "supp (x#xs)"} & $=$ & @{term "supp x \<union> supp xs"}\\
- \end{tabular}
- &
- \begin{tabular}{rcl}
- @{text "supp (f x)"} & @{text "\<subseteq>"} & @{term "supp f \<union> supp x"}\\
- @{term "supp b"} & $=$ & @{term "{}"}\\
- @{term "supp \<pi>"} & $=$ & @{term "{a. \<pi> \<bullet> a \<noteq> a}"}
- \end{tabular}
- \end{tabular}}
- \end{equation}\smallskip
-
- \noindent
- in some cases it can be difficult to characterise the support precisely, and
- only an approximation can be established (as for function applications
- above). Reasoning about such approximations can be simplified with the
- notion \emph{supports}, defined as follows:
-
- \begin{defi}
- A set @{text S} \emph{supports} @{text x}, if for all atoms @{text a} and @{text b}
- not in @{text S} we have @{term "(a \<rightleftharpoons> b) \<bullet> x = x"}.
- \end{defi}
-
- \noindent
- The main point of @{text supports} is that we can establish the following
- two properties.
-
- \begin{prop}\label{supportsprop}
- Given a set @{text "as"} of atoms.\\
- {\it (i)} If @{thm (prem 1) supp_is_subset[where S="as", no_vars]}
- and @{thm (prem 2) supp_is_subset[where S="as", no_vars]} then
- @{thm (concl) supp_is_subset[where S="as", no_vars]}.\\
- {\it (ii)} @{thm supp_supports[no_vars]}.
- \end{prop}
-
- Another important notion in the nominal logic work is \emph{equivariance}.
- For a function @{text f}, say of type @{text "\<alpha> \<Rightarrow> \<beta>"}, to be equivariant
- it is required that every permutation leaves @{text f} unchanged, that is
-
- \begin{equation}\label{equivariancedef}
- @{term "\<forall>\<pi>. \<pi> \<bullet> f = f"}
- \end{equation}\smallskip
-
- \noindent or equivalently that a permutation applied to the application
- @{text "f x"} can be moved to the argument @{text x}. That means for equivariant
- functions @{text f}, we have for all permutations @{text "\<pi>"}:
-
- \begin{equation}\label{equivariance}
- @{text "\<pi> \<bullet> f = f"} \;\;\;\;\textit{if and only if}\;\;\;\;
- @{text "\<forall>x. \<pi> \<bullet> (f x) = f (\<pi> \<bullet> x)"}
- \end{equation}\smallskip
-
- \noindent
- From property \eqref{equivariancedef} and the definition of @{text supp}, we
- can easily deduce that equivariant functions have empty support. There is
- also a similar notion for equivariant relations, say @{text R}, namely the property
- that
-
- \[
- @{text "x R y"} \;\;\textit{implies}\;\; @{text "(\<pi> \<bullet> x) R (\<pi> \<bullet> y)"}
- \]\smallskip
-
- Using freshness, the nominal logic work provides us with general means for renaming
- binders.
-
- \noindent
- While in the older version of Nominal Isabelle, we used extensively
- Property~\ref{swapfreshfresh} to rename single binders, this property
- proved too unwieldy for dealing with multiple binders. For such binders the
- following generalisations turned out to be easier to use.
-
- \begin{prop}\label{supppermeq}
- @{thm[mode=IfThen] supp_perm_eq[where p="\<pi>", no_vars]}
- \end{prop}
-
- \begin{prop}\label{avoiding}
- For a finite set @{text as} and a finitely supported @{text x} with
- @{term "as \<sharp>* x"} and also a finitely supported @{text c}, there
- exists a permutation @{text "\<pi>"} such that @{term "(\<pi> \<bullet> as) \<sharp>* c"} and
- @{term "supp x \<sharp>* \<pi>"}.
- \end{prop}
-
- \noindent
- The idea behind the second property is that given a finite set @{text as}
- of binders (being bound, or fresh, in @{text x} is ensured by the
- assumption @{term "as \<sharp>* x"}), then there exists a permutation @{text "\<pi>"} such that
- the renamed binders @{term "\<pi> \<bullet> as"} avoid @{text c} (which can be arbitrarily chosen
- as long as it is finitely supported) and also @{text "\<pi>"} does not affect anything
- in the support of @{text x} (that is @{term "supp x \<sharp>* \<pi>"}). The last
- fact and Property~\ref{supppermeq} allow us to `rename' just the binders
- @{text as} in @{text x}, because @{term "\<pi> \<bullet> x = x"}.
-
- Note that @{term "supp x \<sharp>* \<pi>"}
- is equivalent with @{term "supp \<pi> \<sharp>* x"}, which means we could also formulate
- Propositions \ref{supppermeq} and \ref{avoiding} in the other `direction'; however the
- reasoning infrastructure of Nominal Isabelle is set up so that it provides more
- automation for the formulation given above.
-
- Most properties given in this section are described in detail in \cite{HuffmanUrban10}
- and all are formalised in Isabelle/HOL. In the next sections we will make
- use of these properties in order to define alpha-equivalence in
- the presence of multiple binders.
-*}
-
-
-section {* Abstractions\label{sec:binders} *}
-
-text {*
- In Nominal Isabelle, the user is expected to write down a specification of a
- term-calculus and then a reasoning infrastructure is automatically derived
- from this specification (remember that Nominal Isabelle is a definitional
- extension of Isabelle/HOL, which does not introduce any new axioms).
-
- In order to keep our work with deriving the reasoning infrastructure
- manageable, we will wherever possible state definitions and perform proofs
- on the `user-level' of Isabelle/HOL, as opposed to writing custom ML-code that
- generates them anew for each specification.
- To that end, we will consider
- first pairs @{text "(as, x)"} of type @{text "(atom set) \<times> \<beta>"}. These pairs
- are intended to represent the abstraction, or binding, of the set of atoms @{text
- "as"} in the body @{text "x"}.
-
- The first question we have to answer is when two pairs @{text "(as, x)"} and
- @{text "(bs, y)"} are alpha-equivalent? (For the moment we are interested in
- the notion of alpha-equivalence that is \emph{not} preserved by adding
- vacuous binders.) To answer this question, we identify four conditions: {\it (i)}
- given a free-atom function @{text "fa"} of type \mbox{@{text "\<beta> \<Rightarrow> atom
- set"}}, then @{text x} and @{text y} need to have the same set of free
- atoms; moreover there must be a permutation @{text \<pi>} such that {\it
- (ii)} @{text \<pi>} leaves the free atoms of @{text x} and @{text y} unchanged, but
- {\it (iii)} `moves' their bound names so that we obtain modulo a relation,
- say \mbox{@{text "_ R _"}}, two equivalent terms. We also require that {\it (iv)}
- @{text \<pi>} makes the sets of abstracted atoms @{text as} and @{text bs} equal. The
- requirements {\it (i)} to {\it (iv)} can be stated formally as:
-
- \begin{defi}[Alpha-Equivalence for Set-Bindings]\label{alphaset}\mbox{}\\
- \begin{tabular}{@ {\hspace{10mm}}l@ {\hspace{5mm}}rl}
- @{term "alpha_set_ex (as, x) R fa (bs, y)"}\hspace{2mm}@{text "\<equiv>"} &
- \multicolumn{2}{@ {}l}{if there exists a @{text "\<pi>"} such that:}\\
- & \mbox{\it (i)} & @{term "fa(x) - as = fa(y) - bs"}\\
- & \mbox{\it (ii)} & @{term "(fa(x) - as) \<sharp>* \<pi>"}\\
- & \mbox{\it (iii)} & @{text "(\<pi> \<bullet> x) R y"} \\
- & \mbox{\it (iv)} & @{term "(\<pi> \<bullet> as) = bs"} \\
- \end{tabular}
- \end{defi}
-
- \noindent
- Note that the relation is
- dependent on a free-atom function @{text "fa"} and a relation @{text
- "R"}. The reason for this extra generality is that we will use
- $\approx_{\,\textit{set}}^{\textit{R}, \textit{fa}}$ for both raw terms and
- alpha-equated terms. In
- the latter case, @{text R} will be replaced by equality @{text "="} and we
- will prove that @{text "fa"} is equal to @{text "supp"}.
-
- Definition \ref{alphaset} does not make any distinction between the
- order of abstracted atoms. If we want this, then we can define alpha-equivalence
- for pairs of the form \mbox{@{text "(as, x)"}} with type @{text "(atom list) \<times> \<beta>"}
- as follows
-
- \begin{defi}[Alpha-Equivalence for List-Bindings]\label{alphalist}\mbox{}\\
- \begin{tabular}{@ {\hspace{10mm}}l@ {\hspace{5mm}}rl}
- @{term "alpha_lst_ex (as, x) R fa (bs, y)"}\hspace{2mm}@{text "\<equiv>"} &
- \multicolumn{2}{@ {}l}{if there exists a @{text "\<pi>"} such that:}\\
- & \mbox{\it (i)} & @{term "fa(x) - (set as) = fa(y) - (set bs)"}\\
- & \mbox{\it (ii)} & @{term "(fa(x) - set as) \<sharp>* \<pi>"}\\
- & \mbox{\it (iii)} & @{text "(\<pi> \<bullet> x) R y"}\\
- & \mbox{\it (iv)} & @{term "(\<pi> \<bullet> as) = bs"}\\
- \end{tabular}
- \end{defi}
-
- \noindent
- where @{term set} is the function that coerces a list of atoms into a set of atoms.
- Now the last clause ensures that the order of the binders matters (since @{text as}
- and @{text bs} are lists of atoms).
-
- If we do not want to make any difference between the order of binders \emph{and}
- also allow vacuous binders, that means according to Pitts~\cite{Pitts04}
- \emph{restrict} atoms, then we keep sets of binders, but drop
- condition {\it (iv)} in Definition~\ref{alphaset}:
-
- \begin{defi}[Alpha-Equivalence for Set+-Bindings]\label{alphares}\mbox{}\\
- \begin{tabular}{@ {\hspace{10mm}}l@ {\hspace{5mm}}rl}
- @{term "alpha_res_ex (as, x) R fa (bs, y)"}\hspace{2mm}@{text "\<equiv>"} &
- \multicolumn{2}{@ {}l}{if there exists a @{text "\<pi>"} such that:}\\
- & \mbox{\it (i)} & @{term "fa(x) - as = fa(y) - bs"}\\
- & \mbox{\it (ii)} & @{term "(fa(x) - as) \<sharp>* \<pi>"}\\
- & \mbox{\it (iii)} & @{text "(\<pi> \<bullet> x) R y"}\\
- \end{tabular}
- \end{defi}
-
-
- It might be useful to consider first some examples how these definitions
- of alpha-equivalence pan out in practice. For this consider the case of
- abstracting a set of atoms over types (as in type-schemes). We set
- @{text R} to be the usual equality @{text "="} and for @{text "fa(T)"} we
- define
-
- \[
- @{text "fa(x) \<equiv> {x}"} \hspace{10mm} @{text "fa(T\<^isub>1 \<rightarrow> T\<^isub>2) \<equiv> fa(T\<^isub>1) \<union> fa(T\<^isub>2)"}
- \]\smallskip
-
- \noindent
- Now recall the examples shown in \eqref{ex1} and
- \eqref{ex3}. It can be easily checked that @{text "({x, y}, x \<rightarrow> y)"} and
- @{text "({y, x}, y \<rightarrow> x)"} are alpha-equivalent according to
- $\approx_{\,\textit{set}}$ and $\approx_{\,\textit{set+}}$ by taking @{text "\<pi>"} to
- be the swapping @{term "(x \<rightleftharpoons> y)"}. In case of @{text "x \<noteq> y"}, then @{text
- "([x, y], x \<rightarrow> y)"} $\not\approx_{\,\textit{list}}$ @{text "([y, x], x \<rightarrow> y)"}
- since there is no permutation that makes the lists @{text "[x, y]"} and
- @{text "[y, x]"} equal, and also leaves the type \mbox{@{text "x \<rightarrow> y"}}
- unchanged. Another example is @{text "({x}, x)"} $\approx_{\,\textit{set+}}$
- @{text "({x, y}, x)"} which holds by taking @{text "\<pi>"} to be the identity
- permutation. However, if @{text "x \<noteq> y"}, then @{text "({x}, x)"}
- $\not\approx_{\,\textit{set}}$ @{text "({x, y}, x)"} since there is no
- permutation that makes the sets @{text "{x}"} and @{text "{x, y}"} equal
- (similarly for $\approx_{\,\textit{list}}$). It can also relatively easily be
- shown that all three notions of alpha-equivalence coincide, if we only
- abstract a single atom. In this case they also agree with the alpha-equivalence
- used in older versions of Nominal Isabelle \cite{Urban08}.\footnote{We omit a
- proof of this fact since the details are hairy and not really important for the
- purpose of this paper.}
-
- In the rest of this section we are going to show that the alpha-equivalences
- really lead to abstractions where some atoms are bound (or more precisely
- removed from the support). For this we will consider three abstraction
- types that are quotients of the relations
-
- \begin{equation}
- \begin{array}{r}
- @{term "alpha_set_ex (as, x) equal supp (bs, y)"}\smallskip\\
- @{term "alpha_res_ex (as, x) equal supp (bs, y)"}\smallskip\\
- @{term "alpha_lst_ex (as, x) equal supp (bs, y)"}\\
- \end{array}
- \end{equation}\smallskip
-
- \noindent
- Note that in these relation we replaced the free-atom function @{text "fa"}
- with @{term "supp"} and the relation @{text R} with equality. We can show
- the following two properties:
-
- \begin{lem}\label{alphaeq}
- The relations $\approx_{\,\textit{set}}^{=, \textit{supp}}$,
- $\approx_{\,\textit{set+}}^{=, \textit{supp}}$
- and $\approx_{\,\textit{list}}^{=, \textit{supp}}$ are
- equivalence relations and equivariant.
- \end{lem}
-
- \begin{proof}
- Reflexivity is by taking @{text "\<pi>"} to be @{text "0"}. For symmetry we have
- a permutation @{text "\<pi>"} and for the proof obligation take @{term "-
- \<pi>"}. In case of transitivity, we have two permutations @{text "\<pi>\<^isub>1"}
- and @{text "\<pi>\<^isub>2"}, and for the proof obligation use @{text
- "\<pi>\<^isub>1 + \<pi>\<^isub>2"}. Equivariance means @{term "alpha_set_ex (\<pi> \<bullet> as,
- \<pi> \<bullet> x) equal supp (\<pi> \<bullet> bs, \<pi> \<bullet> y)"} holds provided \mbox{@{term
- "alpha_set_ex (as, x) equal supp(bs, y)"}} holds. From the assumption we
- have a permutation @{text "\<pi>'"} and for the proof obligation use @{text "\<pi> \<bullet>
- \<pi>'"}. To show equivariance, we need to `pull out' the permutations,
- which is possible since all operators, namely as @{text "#\<^sup>*, -, =, \<bullet>,
- set"} and @{text "supp"}, are equivariant (see
- \cite{HuffmanUrban10}). Finally, we apply the permutation operation on
- booleans.
- \end{proof}
-
- \noindent
- Recall the picture shown in \eqref{picture} about new types in HOL.
- The lemma above allows us to use our quotient package for introducing
- new types @{text "\<beta> abs\<^bsub>set\<^esub>"}, @{text "\<beta> abs\<^bsub>set+\<^esub>"} and @{text "\<beta> abs\<^bsub>list\<^esub>"}
- representing alpha-equivalence classes of pairs of type
- @{text "(atom set) \<times> \<beta>"} (in the first two cases) and of type @{text "(atom list) \<times> \<beta>"}
- (in the third case).
- The elements in these types will be, respectively, written as
-
- \[
- @{term "Abs_set as x"} \hspace{10mm}
- @{term "Abs_res as x"} \hspace{10mm}
- @{term "Abs_lst as x"}
- \]\smallskip
-
- \noindent
- indicating that a set (or list) of atoms @{text as} is abstracted in @{text x}. We will
- call the types \emph{abstraction types} and their elements
- \emph{abstractions}. The important property we need to derive is the support of
- abstractions, namely:
-
- \begin{thm}[Support of Abstractions]\label{suppabs}
- Assuming @{text x} has finite support, then
-
- \[
- \begin{array}{l@ {\;=\;}l}
- @{thm (lhs) supp_Abs(1)[no_vars]} & @{thm (rhs) supp_Abs(1)[no_vars]}\\
- @{thm (lhs) supp_Abs(2)[no_vars]} & @{thm (rhs) supp_Abs(2)[no_vars]}\\
- @{thm (lhs) supp_Abs(3)[where bs="as", no_vars]} &
- @{thm (rhs) supp_Abs(3)[where bs="as", no_vars]}\\
- \end{array}
- \]\smallskip
- \end{thm}
-
- \noindent
- In effect, this theorem states that the atoms @{text "as"} are bound in the
- abstraction. As stated earlier, this can be seen as a litmus test that our
- Definitions \ref{alphaset}, \ref{alphalist} and \ref{alphares} capture the
- idea of alpha-equivalence relations. Below we will give the proof for the
- first equation of Theorem \ref{suppabs}. The others follow by similar
- arguments. By definition of the abstraction type @{text
- "abs\<^bsub>set\<^esub>"} we have
-
- \begin{equation}\label{abseqiff}
- @{thm (lhs) Abs_eq_iff(1)[where bs="as" and cs="bs", no_vars]} \;\;\;\text{if and only if}\;\;\;
- @{term "alpha_set_ex (as, x) equal supp (bs, y)"}
- \end{equation}\smallskip
-
- \noindent
- and also
-
- \begin{equation}\label{absperm}
- @{thm permute_Abs(1)[where p="\<pi>", no_vars]}
- \end{equation}\smallskip
-
- \noindent
- The second fact derives from the definition of permutations acting on pairs
- \eqref{permute} and alpha-equivalence being equivariant
- (see Lemma~\ref{alphaeq}). With these two facts at our disposal, we can show
- the following lemma about swapping two atoms in an abstraction.
-
- \begin{lem}
- If @{thm (prem 1) Abs_swap1(1)[where bs="as", no_vars]} and
- @{thm (prem 2) Abs_swap1(1)[where bs="as", no_vars]} then
- @{thm (concl) Abs_swap1(1)[where bs="as", no_vars]}
- \end{lem}
-
- \begin{proof}
- If @{term "a = b"} the lemma is immediate, since @{term "(a \<rightleftharpoons> b)"} is then
- the identity permutation.
- Also in the other case the lemma is straightforward using \eqref{abseqiff}
- and observing that the assumptions give us @{term "(a \<rightleftharpoons> b) \<bullet> (supp x - as) =
- (supp x - as)"}. We therefore can use the swapping @{term "(a \<rightleftharpoons> b)"} as
- the permutation for the proof obligation.
- \end{proof}
-
- \noindent
- Assuming that @{text "x"} has finite support, this lemma together
- with \eqref{absperm} allows us to show
-
- \begin{equation}\label{halfone}
- @{thm Abs_supports(1)[no_vars]}
- \end{equation}\smallskip
-
- \noindent
- which by Property~\ref{supportsprop} gives us `one half' of
- Theorem~\ref{suppabs}. The `other half' is a bit more involved. To establish
- it, we use a trick from \cite{Pitts04} and first define an auxiliary
- function @{text aux}, taking an abstraction as argument
-
- \[
- @{thm supp_set.simps[THEN eq_reflection, no_vars]}
- \]\smallskip
-
- \noindent
- Using the second equation in \eqref{equivariance}, we can show that
- @{text "aux"} is equivariant (since @{term "\<pi> \<bullet> (supp x - as) = (supp (\<pi> \<bullet> x)) - (\<pi> \<bullet> as)"})
- and therefore has empty support.
- This in turn means
-
- \[
- @{term "supp (supp_set (Abs_set as x)) \<subseteq> supp (Abs_set as x)"}
- \]\smallskip
-
- \noindent
- using the fact about the support of function applications in \eqref{supps}. Assuming
- @{term "supp x - as"} is a finite set, we further obtain
-
- \begin{equation}\label{halftwo}
- @{thm (concl) Abs_supp_subset1(1)[no_vars]}
- \end{equation}\smallskip
-
- \noindent
- This is because for every finite set of atoms, say @{text "bs"}, we have
- @{thm (concl) supp_finite_atom_set[where S="bs", no_vars]}.
- Finally, taking \eqref{halfone} and \eqref{halftwo} together establishes
- the first equation of Theorem~\ref{suppabs}. The others are similar.
-
- Recall the definition of support given in \eqref{suppdef}, and note the difference between
- the support of a raw pair and an abstraction
-
- \[
- @{term "supp (as, x) = supp as \<union> supp x"}\hspace{15mm}
- @{term "supp (Abs_set as x) = supp x - as"}
- \]\smallskip
-
- \noindent
- While the permutation operations behave in both cases the same (a permutation
- is just moved to the arguments), the notion of equality is different for pairs and
- abstractions. Therefore we have different supports. In case of abstractions,
- we have established in Theorem~\ref{suppabs} that bound atoms are removed from
- the support of the abstractions' bodies.
-
- The method of first considering abstractions of the form @{term "Abs_set as
- x"} etc is motivated by the fact that we can conveniently establish at the
- Isabelle/HOL level properties about them. It would be extremely laborious
- to write custom ML-code that derives automatically such properties for every
- term-constructor that binds some atoms. Also the generality of the
- definitions for alpha-equivalence will help us in the next sections.
-*}
-
-section {* Specifying General Bindings\label{sec:spec} *}
-
-text {*
- Our choice of syntax for specifications is influenced by the existing
- datatype package of Isabelle/HOL \cite{Berghofer99}
- and by the syntax of the
- Ott-tool \cite{ott-jfp}. For us a specification of a term-calculus is a
- collection of (possibly mutual recursive) type declarations, say @{text
- "ty\<AL>\<^isub>1, \<dots>, ty\<AL>\<^isub>n"}, and an associated collection of
- binding functions, say @{text "bn\<AL>\<^isub>1, \<dots>, bn\<AL>\<^isub>m"}. The
- syntax in Nominal Isabelle for such specifications is schematically as follows:
-
- \begin{equation}\label{scheme}
- \mbox{\begin{tabular}{@ {}p{2.5cm}l}
- type \mbox{declaration part} &
- $\begin{cases}
- \mbox{\begin{tabular}{l}
- \isacommand{nominal\_datatype} @{text "ty\<AL>\<^isub>1 = \<dots>"}\\
- \isacommand{and} @{text "ty\<AL>\<^isub>2 = \<dots>"}\\
- \raisebox{2mm}{$\ldots$}\\[-2mm]
- \isacommand{and} @{text "ty\<AL>\<^isub>n = \<dots>"}\\
- \end{tabular}}
- \end{cases}$\\[2mm]
- binding \mbox{function part} &
- $\begin{cases}
- \mbox{\begin{tabular}{l}
- \isacommand{binder} @{text "bn\<AL>\<^isub>1"} \isacommand{and} \ldots \isacommand{and} @{text "bn\<AL>\<^isub>m"}\\
- \isacommand{where}\\
- \raisebox{2mm}{$\ldots$}\\[-2mm]
- \end{tabular}}
- \end{cases}$\\
- \end{tabular}}
- \end{equation}\smallskip
-
- \noindent
- Every type declaration @{text ty}$^\alpha_{1..n}$ consists of a collection
- of term-constructors, each of which comes with a list of labelled types that
- stand for the types of the arguments of the term-constructor. For example a
- term-constructor @{text "C\<^sup>\<alpha>"} might be specified with
-
- \[
- @{text "C\<^sup>\<alpha> label\<^isub>1::ty"}\mbox{$'_1$} @{text "\<dots> label\<^isub>l::ty"}\mbox{$'_l\;\;\;\;\;$}
- @{text "binding_clauses"}
- \]\smallskip
-
- \noindent
- whereby some of the @{text ty}$'_{1..l}$ (or their components) can be
- contained in the collection of @{text ty}$^\alpha_{1..n}$ declared in
- \eqref{scheme}. In this case we will call the corresponding argument a
- \emph{recursive argument} of @{text "C\<^sup>\<alpha>"}. The types of such
- recursive arguments need to satisfy a `positivity' restriction, which
- ensures that the type has a set-theoretic semantics (see
- \cite{Berghofer99}). The labels annotated on the types are optional. Their
- purpose is to be used in the (possibly empty) list of \emph{binding
- clauses}, which indicate the binders and their scope in a term-constructor.
- They come in three \emph{modes}:
-
-
- \[\mbox{
- \begin{tabular}{@ {}l@ {}}
- \isacommand{binds} {\it binders} \isacommand{in} {\it bodies}\\
- \isacommand{binds (set)} {\it binders} \isacommand{in} {\it bodies}\\
- \isacommand{binds (set+)} {\it binders} \isacommand{in} {\it bodies}
- \end{tabular}}
- \]\smallskip
-
- \noindent
- The first mode is for binding lists of atoms (the order of bound atoms
- matters); the second is for sets of binders (the order does not matter, but
- the cardinality does) and the last is for sets of binders (with vacuous
- binders preserving alpha-equivalence). As indicated, the labels in the
- `\isacommand{in}-part' of a binding clause will be called \emph{bodies};
- the `\isacommand{binds}-part' will be called \emph{binders}. In contrast to
- Ott, we allow multiple labels in binders and bodies. For example we allow
- binding clauses of the form:
-
- \[\mbox{
- \begin{tabular}{@ {}ll@ {}}
- @{text "Foo\<^isub>1 x::name y::name t::trm s::trm"} &
- \isacommand{binds} @{text "x y"} \isacommand{in} @{text "t s"}\\
- @{text "Foo\<^isub>2 x::name y::name t::trm s::trm"} &
- \isacommand{binds} @{text "x y"} \isacommand{in} @{text "t"},
- \isacommand{binds} @{text "x y"} \isacommand{in} @{text "s"}\\
- \end{tabular}}
- \]\smallskip
-
- \noindent
- Similarly for the other binding modes. Interestingly, in case of
- \isacommand{binds (set)} and \isacommand{binds (set+)} the binding clauses
- above will make a difference to the semantics of the specifications (the
- corresponding alpha-equivalence will differ). We will show this later with
- an example.
-
-
- There are also some restrictions we need to impose on our binding clauses in
- comparison to Ott. The main idea behind these restrictions is
- that we obtain a notion of alpha-equivalence where it is ensured
- that within a given scope an atom occurrence cannot be both bound and free
- at the same time. The first restriction is that a body can only occur in
- \emph{one} binding clause of a term constructor. So for example
-
- \[\mbox{
- @{text "Foo x::name y::name t::trm"}\hspace{3mm}
- \isacommand{binds} @{text "x"} \isacommand{in} @{text "t"},
- \isacommand{binds} @{text "y"} \isacommand{in} @{text "t"}}
- \]\smallskip
-
- \noindent
- is not allowed. This ensures that the bound atoms of a body cannot be free
- at the same time by specifying an alternative binder for the same body.
-
- For binders we distinguish between \emph{shallow} and \emph{deep} binders.
- Shallow binders are just labels. The restriction we need to impose on them
- is that in case of \isacommand{binds (set)} and \isacommand{binds (set+)} the
- labels must either refer to atom types or to sets of atom types; in case of
- \isacommand{binds} the labels must refer to atom types or to lists of atom
- types. Two examples for the use of shallow binders are the specification of
- lambda-terms, where a single name is bound, and type-schemes, where a finite
- set of names is bound:
-
- \[\mbox{
- \begin{tabular}{@ {}c@ {\hspace{8mm}}c@ {}}
- \begin{tabular}{@ {}l}
- \isacommand{nominal\_datatype} @{text lam} $=$\\
- \hspace{2mm}\phantom{$\mid$}~@{text "Var name"}\\
- \hspace{2mm}$\mid$~@{text "App lam lam"}\\
- \hspace{2mm}$\mid$~@{text "Lam x::name t::lam"}\hspace{3mm}%
- \isacommand{binds} @{text x} \isacommand{in} @{text t}\\
- \\
- \end{tabular} &
- \begin{tabular}{@ {}l@ {}}
- \isacommand{nominal\_datatype}~@{text ty} $=$\\
- \hspace{2mm}\phantom{$\mid$}~@{text "TVar name"}\\
- \hspace{2mm}$\mid$~@{text "TFun ty ty"}\\
- \isacommand{and}~@{text "tsc ="}\\
- \hspace{2mm}\phantom{$\mid$}~@{text "TAll xs::(name fset) T::ty"}\hspace{3mm}%
- \isacommand{binds (set+)} @{text xs} \isacommand{in} @{text T}\\
- \end{tabular}
- \end{tabular}}
- \]\smallskip
-
-
- \noindent
- In these specifications @{text "name"} refers to a (concrete) atom type, and @{text
- "fset"} to the type of finite sets. Note that for @{text Lam} it does not
- matter which binding mode we use. The reason is that we bind only a single
- @{text name}, in which case all three binding modes coincide. However, having
- \isacommand{binds (set)} or just \isacommand{binds}
- in the second case makes a difference to the semantics of the specification
- (which we will define in the next section).
-
- A \emph{deep} binder uses an auxiliary binding function that `picks' out
- the atoms in one argument of the term-constructor, which can be bound in
- other arguments and also in the same argument (we will call such binders
- \emph{recursive}, see below). The binding functions are
- expected to return either a set of atoms (for \isacommand{binds (set)} and
- \isacommand{binds (set+)}) or a list of atoms (for \isacommand{binds}). They need
- to be defined by recursion over the corresponding type; the equations
- must be given in the binding function part of the scheme shown in
- \eqref{scheme}. For example a term-calculus containing @{text "Let"}s with
- tuple patterns may be specified as:
-
- \begin{equation}\label{letpat}
- \mbox{%
- \begin{tabular}{l}
- \isacommand{nominal\_datatype} @{text trm} $=$\\
- \hspace{5mm}\phantom{$\mid$}~@{term "Var name"}\\
- \hspace{5mm}$\mid$~@{term "App trm trm"}\\
- \hspace{5mm}$\mid$~@{text "Lam x::name t::trm"}
- \;\;\isacommand{binds} @{text x} \isacommand{in} @{text t}\\
- \hspace{5mm}$\mid$~@{text "Let_pat p::pat trm t::trm"}
- \;\;\isacommand{binds} @{text "bn(p)"} \isacommand{in} @{text t}\\
- \isacommand{and} @{text pat} $=$\\
- \hspace{5mm}\phantom{$\mid$}~@{text "PVar name"}\\
- \hspace{5mm}$\mid$~@{text "PTup pat pat"}\\
- \isacommand{binder}~@{text "bn::pat \<Rightarrow> atom list"}\\
- \isacommand{where}~@{text "bn(PVar x) = [atom x]"}\\
- \hspace{5mm}$\mid$~@{text "bn(PTup p\<^isub>1 p\<^isub>2) = bn(p\<^isub>1) @ bn(p\<^isub>2)"}\smallskip\\
- \end{tabular}}
- \end{equation}\smallskip
-
- \noindent
- In this specification the function @{text "bn"} determines which atoms of
- the pattern @{text p} (fifth line) are bound in the argument @{text "t"}. Note that in the
- second-last @{text bn}-clause the function @{text "atom"} coerces a name
- into the generic atom type of Nominal Isabelle \cite{HuffmanUrban10}. This
- allows us to treat binders of different atom type uniformly.
-
- For deep binders we allow binding clauses such as
-
- \[\mbox{
- \begin{tabular}{ll}
- @{text "Bar p::pat t::trm"} &
- \isacommand{binds} @{text "bn(p)"} \isacommand{in} @{text "p t"} \\
- \end{tabular}}
- \]\smallskip
-
-
- \noindent
- where the argument of the deep binder also occurs in the body. We call such
- binders \emph{recursive}. To see the purpose of such recursive binders,
- compare `plain' @{text "Let"}s and @{text "Let_rec"}s in the following
- specification:
-
- \begin{equation}\label{letrecs}
- \mbox{%
- \begin{tabular}{@ {}l@ {}l}
- \isacommand{nominal\_datatype}~@{text "trm ="}~\ldots\\
- \hspace{5mm}$\mid$~@{text "Let as::assn t::trm"}
- & \hspace{-19mm}\isacommand{binds} @{text "bn(as)"} \isacommand{in} @{text t}\\
- \hspace{5mm}$\mid$~@{text "Let_rec as::assn t::trm"}
- & \hspace{-19mm}\isacommand{binds} @{text "bn(as)"} \isacommand{in} @{text "as t"}\\
- \isacommand{and} @{text "assn"} $=$\\
- \hspace{5mm}\phantom{$\mid$}~@{text "ANil"}\\
- \hspace{5mm}$\mid$~@{text "ACons name trm assn"}\\
- \isacommand{binder} @{text "bn::assn \<Rightarrow> atom list"}\\
- \isacommand{where}~@{text "bn(ANil) = []"}\\
- \hspace{5mm}$\mid$~@{text "bn(ACons a t as) = [atom a] @ bn(as)"}\\
- \end{tabular}}
- \end{equation}\smallskip
-
- \noindent
- The difference is that with @{text Let} we only want to bind the atoms @{text
- "bn(as)"} in the term @{text t}, but with @{text "Let_rec"} we also want to bind the atoms
- inside the assignment. This difference has consequences for the associated
- notions of free-atoms and alpha-equivalence.
-
- To make sure that atoms bound by deep binders cannot be free at the
- same time, we cannot have more than one binding function for a deep binder.
- Consequently we exclude specifications such as
-
- \[\mbox{
- \begin{tabular}{@ {}l@ {\hspace{2mm}}l@ {}}
- @{text "Baz\<^isub>1 p::pat t::trm"} &
- \isacommand{binds} @{text "bn\<^isub>1(p) bn\<^isub>2(p)"} \isacommand{in} @{text t}\\
- @{text "Baz\<^isub>2 p::pat t\<^isub>1::trm t\<^isub>2::trm"} &
- \isacommand{binds} @{text "bn\<^isub>1(p)"} \isacommand{in} @{text "t\<^isub>1"},
- \isacommand{binds} @{text "bn\<^isub>2(p)"} \isacommand{in} @{text "t\<^isub>2"}\\
- \end{tabular}}
- \]\smallskip
-
- \noindent
- Otherwise it is possible that @{text "bn\<^isub>1"} and @{text "bn\<^isub>2"} pick
- out different atoms to become bound, respectively be free, in @{text "p"}.
- (Since the Ott-tool does not derive a reasoning infrastructure for
- alpha-equated terms with deep binders, it can permit such specifications.)
-
- We also need to restrict the form of the binding functions in order to
- ensure the @{text "bn"}-functions can be defined for alpha-equated
- terms. The main restriction is that we cannot return an atom in a binding
- function that is also bound in the corresponding term-constructor.
- Consider again the specification for @{text "trm"} and a contrived
- version for assignments @{text "assn"}:
-
- \begin{equation}\label{bnexp}
- \mbox{%
- \begin{tabular}{@ {}l@ {}}
- \isacommand{nominal\_datatype}~@{text "trm ="}~\ldots\\
- \isacommand{and} @{text "assn"} $=$\\
- \hspace{5mm}\phantom{$\mid$}~@{text "ANil'"}\\
- \hspace{5mm}$\mid$~@{text "ACons' x::name y::name t::trm assn"}
- \;\;\isacommand{binds} @{text "y"} \isacommand{in} @{text t}\\
- \isacommand{binder} @{text "bn::assn \<Rightarrow> atom list"}\\
- \isacommand{where}~@{text "bn(ANil') = []"}\\
- \hspace{5mm}$\mid$~@{text "bn(ACons' x y t as) = [atom x] @ bn(as)"}\\
- \end{tabular}}
- \end{equation}\smallskip
-
- \noindent
- In this example the term constructor @{text "ACons'"} has four arguments with
- a binding clause involving two of them. This constructor is also used in the definition
- of the binding function. The restriction we have to impose is that the
- binding function can only return free atoms, that is the ones that are \emph{not}
- mentioned in a binding clause. Therefore @{text "y"} cannot be used in the
- binding function @{text "bn"} (since it is bound in @{text "ACons'"} by the
- binding clause), but @{text x} can (since it is a free atom). This
- restriction is sufficient for lifting the binding function to alpha-equated
- terms. If we would permit @{text "bn"} to return @{text "y"},
- then it would not be respectful and therefore cannot be lifted to
- alpha-equated lambda-terms.
-
- In the version of Nominal Isabelle described here, we also adopted the
- restriction from the Ott-tool that binding functions can only return: the
- empty set or empty list (as in case @{text ANil'}), a singleton set or
- singleton list containing an atom (case @{text PVar} in \eqref{letpat}), or
- unions of atom sets or appended atom lists (case @{text ACons'}). This
- restriction will simplify some automatic definitions and proofs later on.
-
- In order to simplify our definitions of free atoms and alpha-equivalence,
- we shall assume specifications
- of term-calculi are implicitly \emph{completed}. By this we mean that
- for every argument of a term-constructor that is \emph{not}
- already part of a binding clause given by the user, we add implicitly a special \emph{empty} binding
- clause, written \isacommand{binds}~@{term "{}"}~\isacommand{in}~@{text "labels"}. In case
- of the lambda-terms, the completion produces
-
- \[\mbox{
- \begin{tabular}{@ {}l@ {\hspace{-1mm}}}
- \isacommand{nominal\_datatype} @{text lam} =\\
- \hspace{5mm}\phantom{$\mid$}~@{text "Var x::name"}
- \;\;\isacommand{binds}~@{term "{}"}~\isacommand{in}~@{text "x"}\\
- \hspace{5mm}$\mid$~@{text "App t\<^isub>1::lam t\<^isub>2::lam"}
- \;\;\isacommand{binds}~@{term "{}"}~\isacommand{in}~@{text "t\<^isub>1 t\<^isub>2"}\\
- \hspace{5mm}$\mid$~@{text "Lam x::name t::lam"}
- \;\;\isacommand{binds}~@{text x} \isacommand{in} @{text t}\\
- \end{tabular}}
- \]\smallskip
-
- \noindent
- The point of completion is that we can make definitions over the binding
- clauses and be sure to have captured all arguments of a term constructor.
-*}
-
-section {* Alpha-Equivalence and Free Atoms\label{sec:alpha} *}
-
-text {*
- Having dealt with all syntax matters, the problem now is how we can turn
- specifications into actual type definitions in Isabelle/HOL and then
- establish a reasoning infrastructure for them. As Pottier and Cheney pointed
- out \cite{Cheney05,Pottier06}, just re-arranging the arguments of
- term-constructors so that binders and their bodies are next to each other
- will result in inadequate representations in cases like \mbox{@{text "Let
- x\<^isub>1 = t\<^isub>1\<dots>x\<^isub>n = t\<^isub>n in s"}}. Therefore we will
- first extract `raw' datatype definitions from the specification and then
- define explicitly an alpha-equivalence relation over them. We subsequently
- construct the quotient of the datatypes according to our alpha-equivalence.
-
-
- The `raw' datatype definition can be obtained by stripping off the
- binding clauses and the labels from the types given by the user. We also have to invent
- new names for the types @{text "ty\<^sup>\<alpha>"} and the term-constructors @{text "C\<^sup>\<alpha>"}.
- In our implementation we just use the affix ``@{text "_raw"}''.
- But for the purpose of this paper, we use the superscript @{text "_\<^sup>\<alpha>"} to indicate
- that a notion is given for alpha-equivalence classes and leave it out
- for the corresponding notion given on the raw level. So for example
- we have @{text "ty\<^sup>\<alpha> / ty"} and @{text "C\<^sup>\<alpha> / C"}
- where @{term ty} is the type used in the quotient construction for
- @{text "ty\<^sup>\<alpha>"} and @{text "C"} is the term-constructor of the raw type @{text "ty"},
- respectively @{text "C\<^sup>\<alpha>"} is the corresponding term-constructor of @{text "ty\<^sup>\<alpha>"}.
-
- The resulting datatype definition is legal in Isabelle/HOL provided the datatypes are
- non-empty and the types in the constructors only occur in positive
- position (see \cite{Berghofer99} for an in-depth description of the datatype package
- in Isabelle/HOL).
- We subsequently define each of the user-specified binding
- functions @{term "bn"}$_{1..m}$ by recursion over the corresponding
- raw datatype. We also define permutation operations by
- recursion so that for each term constructor @{text "C"} we have that
-
- \begin{equation}\label{ceqvt}
- @{text "\<pi> \<bullet> (C z\<^isub>1 \<dots> z\<^isub>n) = C (\<pi> \<bullet> z\<^isub>1) \<dots> (\<pi> \<bullet> z\<^isub>n)"}
- \end{equation}\smallskip
-
- \noindent
- We will need this operation later when we define the notion of alpha-equivalence.
-
- The first non-trivial step we have to perform is the generation of
- \emph{free-atom functions} from the specifications.\footnote{Admittedly, the
- details of our definitions will be somewhat involved. However they are still
- conceptually simple in comparison with the `positional' approach taken in
- Ott \cite[Pages 88--95]{ott-jfp}, which uses the notions of \emph{occurrences} and
- \emph{partial equivalence relations} over sets of occurrences.} For the
- \emph{raw} types @{text "ty"}$_{1..n}$ we define the free-atom functions
-
- \begin{equation}\label{fvars}
- \mbox{@{text "fa_ty"}$_{1..n}$}
- \end{equation}\smallskip
-
- \noindent
- by recursion.
- We define these functions together with auxiliary free-atom functions for
- the binding functions. Given raw binding functions @{text "bn"}$_{1..m}$
- we define
-
- \[
- @{text "fa_bn"}\mbox{$_{1..m}$}.
- \]\smallskip
-
- \noindent
- The reason for this setup is that in a deep binder not all atoms have to be
- bound, as we saw in \eqref{letrecs} with the example of `plain' @{text Let}s. We need
- therefore functions that calculate those free atoms in deep binders.
-
- While the idea behind these free-atom functions is simple (they just
- collect all atoms that are not bound), because of our rather complicated
- binding mechanisms their definitions are somewhat involved. Given
- a raw term-constructor @{text "C"} of type @{text ty} and some associated
- binding clauses @{text "bc\<^isub>1\<dots>bc\<^isub>k"}, the result of @{text
- "fa_ty (C z\<^isub>1 \<dots> z\<^isub>n)"} will be the union @{text
- "fa(bc\<^isub>1) \<union> \<dots> \<union> fa(bc\<^isub>k)"} where we will define below what @{text "fa"} for a binding
- clause means. We only show the details for the mode \isacommand{binds (set)} (the other modes are similar).
- Suppose the binding clause @{text bc\<^isub>i} is of the form
-
- \[
- \mbox{\isacommand{binds (set)} @{text "b\<^isub>1\<dots>b\<^isub>p"} \isacommand{in} @{text "d\<^isub>1\<dots>d\<^isub>q"}}
- \]\smallskip
-
- \noindent
- in which the body-labels @{text "d"}$_{1..q}$ refer to types @{text
- ty}$_{1..q}$, and the binders @{text b}$_{1..p}$ either refer to labels of
- atom types (in case of shallow binders) or to binding functions taking a
- single label as argument (in case of deep binders). Assuming @{text "D"}
- stands for the set of free atoms of the bodies, @{text B} for the set of
- binding atoms in the binders and @{text "B'"} for the set of free atoms in
- non-recursive deep binders, then the free atoms of the binding clause @{text
- bc\<^isub>i} are
-
- \begin{equation}\label{fadef}
- \mbox{@{text "fa(bc\<^isub>i) \<equiv> (D - B) \<union> B'"}}.
- \end{equation}\smallskip
-
- \noindent
- The set @{text D} is formally defined as
-
- \[
- @{text "D \<equiv> fa_ty\<^isub>1 d\<^isub>1 \<union> ... \<union> fa_ty\<^isub>q d\<^isub>q"}
- \]\smallskip
-
- \noindent
- where in case @{text "d\<^isub>i"} refers to one of the raw types @{text "ty"}$_{1..n}$ from the
- specification, the function @{text "fa_ty\<^isub>i"} is the corresponding free-atom function
- we are defining by recursion; otherwise we set \mbox{@{text "fa_ty\<^isub>i \<equiv> supp"}}. The reason
- for the latter is that @{text "ty"}$_i$ is not a type that is part of the specification, and
- we assume @{text supp} is the generic function that characterises the free variables of
- a type (in fact in the next section we will show that the free-variable functions we
- define here, are equal to the support once lifted to alpha-equivalence classes).
-
- In order to formally define the set @{text B} we use the following auxiliary @{text "bn"}-functions
- for atom types to which shallow binders may refer\\[-4mm]
-
- \begin{equation}\label{bnaux}\mbox{
- \begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
- @{text "bn\<^bsub>atom\<^esup> a"} & @{text "\<equiv>"} & @{text "{atom a}"}\\
- @{text "bn\<^bsub>atom_set\<^esup> as"} & @{text "\<equiv>"} & @{text "atoms as"}\\
- @{text "bn\<^bsub>atom_list\<^esub> as"} & @{text "\<equiv>"} & @{text "atoms (set as)"}
- \end{tabular}}
- \end{equation}\smallskip
-
- \noindent
- Like the function @{text atom}, the function @{text "atoms"} coerces
- a set of atoms to a set of the generic atom type.
- It is defined as @{text "atoms as \<equiv> {atom a | a \<in> as}"}.
- The set @{text B} in \eqref{fadef} is then formally defined as
-
- \begin{equation}\label{bdef}
- @{text "B \<equiv> bn_ty\<^isub>1 b\<^isub>1 \<union> ... \<union> bn_ty\<^isub>p b\<^isub>p"}
- \end{equation}\smallskip
-
- \noindent
- where we use the auxiliary binding functions from \eqref{bnaux} for shallow
- binders (that means when @{text "ty"}$_i$ is of type @{text "atom"}, @{text "atom set"} or
- @{text "atom list"}).
-
- The set @{text "B'"} in \eqref{fadef} collects all free atoms in
- non-recursive deep binders. Let us assume these binders in the binding
- clause @{text "bc\<^isub>i"} are
-
- \[
- \mbox{@{text "bn\<^isub>1 l\<^isub>1, \<dots>, bn\<^isub>r l\<^isub>r"}}
- \]\smallskip
-
- \noindent
- with @{text "l"}$_{1..r}$ $\subseteq$ @{text "b"}$_{1..p}$ and
- none of the @{text "l"}$_{1..r}$ being among the bodies
- @{text "d"}$_{1..q}$. The set @{text "B'"} is defined as
-
- \begin{equation}\label{bprimedef}
- @{text "B' \<equiv> fa_bn\<^isub>1 l\<^isub>1 \<union> ... \<union> fa_bn\<^isub>r l\<^isub>r"}
- \end{equation}\smallskip
-
- \noindent
- This completes all clauses for the free-atom functions @{text "fa_ty"}$_{1..n}$.
-
- Note that for non-recursive deep binders, we have to add in \eqref{fadef}
- the set of atoms that are left unbound by the binding functions @{text
- "bn"}$_{1..m}$. We used for
- the definition of this set the functions @{text "fa_bn"}$_{1..m}$. The
- definition for those functions needs to be extracted from the clauses the
- user provided for @{text "bn"}$_{1..m}$ Assume the user specified a @{text
- bn}-clause of the form
-
- \[
- @{text "bn (C z\<^isub>1 \<dots> z\<^isub>s) = rhs"}
- \]\smallskip
-
- \noindent
- where the @{text "z"}$_{1..s}$ are of types @{text "ty"}$_{1..s}$. For
- each of the arguments we calculate the free atoms as follows:
-
- \[\mbox{
- \begin{tabular}{c@ {\hspace{2mm}}p{0.9\textwidth}}
- $\bullet$ & @{term "fa_ty\<^isub>i z\<^isub>i"} provided @{text "z\<^isub>i"} does not occur in @{text "rhs"}\\
- & (that means nothing is bound in @{text "z\<^isub>i"} by the binding function),\smallskip\\
- $\bullet$ & @{term "fa_bn\<^isub>i z\<^isub>i"} provided @{text "z\<^isub>i"} occurs in @{text "rhs"}
- with the recursive call @{text "bn\<^isub>i z\<^isub>i"}\\
- & (that means whatever is `left over' from the @{text "bn"}-function is free)\smallskip\\
- $\bullet$ & @{term "{}"} provided @{text "z\<^isub>i"} occurs in @{text "rhs"},
- but without a recursive call\\
- & (that means @{text "z\<^isub>i"} is supposed to become bound by the binding function)\\
- \end{tabular}}
- \]\smallskip
-
- \noindent
- For defining @{text "fa_bn (C z\<^isub>1 \<dots> z\<^isub>n)"} we just union up all these sets.
-
- To see how these definitions work in practice, let us reconsider the
- term-constructors @{text "Let"} and @{text "Let_rec"} shown in
- \eqref{letrecs} together with the term-constructors for assignments @{text
- "ANil"} and @{text "ACons"}. Since there is a binding function defined for
- assignments, we have three free-atom functions, namely @{text
- "fa\<^bsub>trm\<^esub>"}, @{text "fa\<^bsub>assn\<^esub>"} and @{text
- "fa\<^bsub>bn\<^esub>"} as follows:
-
- \[\mbox{
- \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
- @{text "fa\<^bsub>trm\<^esub> (Let as t)"} & @{text "\<equiv>"} & @{text "(fa\<^bsub>trm\<^esub> t - set (bn as)) \<union> fa\<^bsub>bn\<^esub> as"}\\
- @{text "fa\<^bsub>trm\<^esub> (Let_rec as t)"} & @{text "\<equiv>"} & @{text "(fa\<^bsub>assn\<^esub> as \<union> fa\<^bsub>trm\<^esub> t) - set (bn as)"}\smallskip\\
-
- @{text "fa\<^bsub>assn\<^esub> (ANil)"} & @{text "\<equiv>"} & @{term "{}"}\\
- @{text "fa\<^bsub>assn\<^esub> (ACons a t as)"} & @{text "\<equiv>"} & @{text "(supp a) \<union> (fa\<^bsub>trm\<^esub> t) \<union> (fa\<^bsub>assn\<^esub> as)"}\smallskip\\
-
- @{text "fa\<^bsub>bn\<^esub> (ANil)"} & @{text "\<equiv>"} & @{term "{}"}\\
- @{text "fa\<^bsub>bn\<^esub> (ACons a t as)"} & @{text "\<equiv>"} & @{text "(fa\<^bsub>trm\<^esub> t) \<union> (fa\<^bsub>bn\<^esub> as)"}
- \end{tabular}}
- \]\smallskip
-
-
- \noindent
- Recall that @{text ANil} and @{text "ACons"} have no binding clause in the
- specification. The corresponding free-atom function @{text
- "fa\<^bsub>assn\<^esub>"} therefore returns all free atoms of an assignment
- (in case of @{text "ACons"}, they are given in terms of @{text supp}, @{text
- "fa\<^bsub>trm\<^esub>"} and @{text "fa\<^bsub>assn\<^esub>"}). The binding
- only takes place in @{text Let} and @{text "Let_rec"}. In case of @{text
- "Let"}, the binding clause specifies that all atoms given by @{text "set (bn
- as)"} have to be bound in @{text t}. Therefore we have to subtract @{text
- "set (bn as)"} from @{text "fa\<^bsub>trm\<^esub> t"}. However, we also need
- to add all atoms that are free in @{text "as"}. This is in contrast with
- @{text "Let_rec"} where we have a recursive binder to bind all occurrences
- of the atoms in @{text "set (bn as)"} also inside @{text "as"}. Therefore we
- have to subtract @{text "set (bn as)"} from both @{text
- "fa\<^bsub>trm\<^esub> t"} and @{text "fa\<^bsub>assn\<^esub> as"}. Like the
- function @{text "bn"}, the function @{text "fa\<^bsub>bn\<^esub>"} traverses
- the list of assignments, but instead returns the free atoms, which means in
- this example the free atoms in the argument @{text "t"}.
-
-
- An interesting point in this example is that a `naked' assignment (@{text
- "ANil"} or @{text "ACons"}) does not bind any atoms, even if the binding
- function is specified over assignments. Only in the context of a @{text Let}
- or @{text "Let_rec"}, where the binding clauses are given, will some atoms
- actually become bound. This is a phenomenon that has also been pointed out
- in \cite{ott-jfp}. For us this observation is crucial, because we would not
- be able to lift the @{text "bn"}-functions to alpha-equated terms if they
- act on atoms that are bound. In that case, these functions would \emph{not}
- respect alpha-equivalence.
-
- Having the free-atom functions at our disposal, we can next define the
- alpha-equivalence relations for the raw types @{text
- "ty"}$_{1..n}$. We write them as
-
- \[
- \mbox{@{text "\<approx>ty"}$_{1..n}$}.
- \]\smallskip
-
- \noindent
- Like with the free-atom functions, we also need to
- define auxiliary alpha-equivalence relations
-
- \[
- \mbox{@{text "\<approx>bn\<^isub>"}$_{1..m}$}
- \]\smallskip
-
- \noindent
- for the binding functions @{text "bn"}$_{1..m}$,
- To simplify our definitions we will use the following abbreviations for
- \emph{compound equivalence relations} and \emph{compound free-atom functions} acting on tuples.
-
- \[\mbox{
- \begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
- @{text "(x\<^isub>1,\<dots>, x\<^isub>n) (R\<^isub>1,\<dots>, R\<^isub>n) (y\<^isub>1,\<dots>, y\<^isub>n)"} & @{text "\<equiv>"} &
- @{text "x\<^isub>1 R\<^isub>1 y\<^isub>1 \<and> \<dots> \<and> x\<^isub>n R\<^isub>n y\<^isub>n"}\\
- @{text "(fa\<^isub>1,\<dots>, fa\<^isub>n) (x\<^isub>1,\<dots>, x\<^isub>n)"} & @{text "\<equiv>"} & @{text "fa\<^isub>1 x\<^isub>1 \<union> \<dots> \<union> fa\<^isub>n x\<^isub>n"}\\
- \end{tabular}}
- \]\smallskip
-
-
- The alpha-equivalence relations are defined as inductive predicates
- having a single clause for each term-constructor. Assuming a
- term-constructor @{text C} is of type @{text ty} and has the binding clauses
- @{term "bc"}$_{1..k}$, then the alpha-equivalence clause has the form
-
- \[
- \mbox{\infer{@{text "C z\<^isub>1 \<dots> z\<^isub>n \<approx>ty C z\<PRIME>\<^isub>1 \<dots> z\<PRIME>\<^isub>n"}}
- {@{text "prems(bc\<^isub>1) \<dots> prems(bc\<^isub>k)"}}}
- \]\smallskip
-
- \noindent
- The task below is to specify what the premises corresponding to a binding
- clause are. To understand better what the general pattern is, let us first
- treat the special instance where @{text "bc\<^isub>i"} is the empty binding clause
- of the form
-
- \[
- \mbox{\isacommand{binds (set)} @{term "{}"} \isacommand{in} @{text "d\<^isub>1\<dots>d\<^isub>q"}.}
- \]\smallskip
-
- \noindent
- In this binding clause no atom is bound and we only have to `alpha-relate'
- the bodies. For this we build first the tuples @{text "D \<equiv> (d\<^isub>1,\<dots>,
- d\<^isub>q)"} and @{text "D' \<equiv> (d\<PRIME>\<^isub>1,\<dots>, d\<PRIME>\<^isub>q)"}
- whereby the labels @{text "d"}$_{1..q}$ refer to some of the arguments @{text
- "z"}$_{1..n}$ and respectively @{text "d\<PRIME>"}$_{1..q}$ to some of @{text
- "z\<PRIME>"}$_{1..n}$ of the term-constructor. In order to relate two such
- tuples we define the compound alpha-equivalence relation @{text "R"} as
- follows
-
- \begin{equation}\label{rempty}
- \mbox{@{text "R \<equiv> (R\<^isub>1,\<dots>, R\<^isub>q)"}}
- \end{equation}\smallskip
-
- \noindent
- with @{text "R\<^isub>i"} being @{text "\<approx>ty\<^isub>i"} if the corresponding
- labels @{text "d\<^isub>i"} and @{text "d\<PRIME>\<^isub>i"} refer to a
- recursive argument of @{text C} and have type @{text "ty\<^isub>i"}; otherwise
- we take @{text "R\<^isub>i"} to be the equality @{text "="}. Again the
- latter is because @{text "ty\<^isub>i"} is not part of the specified types
- and alpha-equivalence of any previously defined type is supposed to coincide
- with equality. This lets us now define the premise for an empty binding
- clause succinctly as @{text "prems(bc\<^isub>i) \<equiv> D R D'"}, which can be
- unfolded to the series of premises
-
- \[
- @{text "d\<^isub>1 R\<^isub>1 d\<PRIME>\<^isub>1 \<dots> d\<^isub>q R\<^isub>q d\<PRIME>\<^isub>q"}.
- \]\smallskip
-
- \noindent
- We will use the unfolded version in the examples below.
-
- Now suppose the binding clause @{text "bc\<^isub>i"} is of the general form
-
- \begin{equation}\label{nonempty}
- \mbox{\isacommand{binds (set)} @{text "b\<^isub>1\<dots>b\<^isub>p"} \isacommand{in} @{text "d\<^isub>1\<dots>d\<^isub>q"}.}
- \end{equation}\smallskip
-
- \noindent
- In this case we define a premise @{text P} using the relation
- $\approx_{\,\textit{set}}^{\textit{R}, \textit{fa}}$ given in Section~\ref{sec:binders} (similarly
- $\approx_{\,\textit{set+}}^{\textit{R}, \textit{fa}}$ and
- $\approx_{\,\textit{list}}^{\textit{R}, \textit{fa}}$ for the other
- binding modes). As above, we first build the tuples @{text "D"} and
- @{text "D'"} for the bodies @{text "d"}$_{1..q}$, and the corresponding
- compound alpha-relation @{text "R"} (shown in \eqref{rempty}).
- For $\approx_{\,\textit{set}}^{\textit{R}, \textit{fa}}$ we also need
- a compound free-atom function for the bodies defined as
-
- \[
- \mbox{@{text "fa \<equiv> (fa_ty\<^isub>1,\<dots>, fa_ty\<^isub>q)"}}
- \]\smallskip
-
- \noindent
- with the assumption that the @{text "d"}$_{1..q}$ refer to arguments of types @{text "ty"}$_{1..q}$.
- The last ingredient we need are the sets of atoms bound in the bodies.
- For this we take
-
- \[
- @{text "B \<equiv> bn_ty\<^isub>1 b\<^isub>1 \<union> \<dots> \<union> bn_ty\<^isub>p b\<^isub>p"}\;.\\
- \]\smallskip
-
- \noindent
- Similarly for @{text "B'"} using the labels @{text "b\<PRIME>"}$_{1..p}$. This
- lets us formally define the premise @{text P} for a non-empty binding clause as:
-
- \[
- \mbox{@{term "P \<equiv> alpha_set_ex (B, D) R fa (B', D')"}}\;.
- \]\smallskip
-
- \noindent
- This premise accounts for alpha-equivalence of the bodies of the binding
- clause. However, in case the binders have non-recursive deep binders, this
- premise is not enough: we also have to `propagate' alpha-equivalence
- inside the structure of these binders. An example is @{text "Let"} where we
- have to make sure the right-hand sides of assignments are
- alpha-equivalent. For this we use relations @{text "\<approx>bn"}$_{1..m}$ (which we
- will define shortly). Let us assume the non-recursive deep binders
- in @{text "bc\<^isub>i"} are
-
- \[
- @{text "bn\<^isub>1 l\<^isub>1, \<dots>, bn\<^isub>r l\<^isub>r"}.
- \]\smallskip
-
- \noindent
- The tuple @{text L} consists then of all these binders @{text "(l\<^isub>1,\<dots>,l\<^isub>r)"}
- (similarly @{text "L'"}) and the compound equivalence relation @{text "R'"}
- is @{text "(\<approx>bn\<^isub>1,\<dots>,\<approx>bn\<^isub>r)"}. All premises for @{text "bc\<^isub>i"} are then given by
-
- \[
- @{text "prems(bc\<^isub>i) \<equiv> P \<and> L R' L'"}
- \]\smallskip
-
- \noindent
- The auxiliary alpha-equivalence relations @{text "\<approx>bn"}$_{1..m}$
- in @{text "R'"} are defined as follows: assuming a @{text bn}-clause is of the form
-
- \[
- @{text "bn (C z\<^isub>1 \<dots> z\<^isub>s) = rhs"}
- \]\smallskip
-
- \noindent
- where the @{text "z"}$_{1..s}$ are of types @{text "ty"}$_{1..s}$,
- then the corresponding alpha-equivalence clause for @{text "\<approx>bn"} has the form
-
- \[
- \mbox{\infer{@{text "C z\<^isub>1 \<dots> z\<^isub>s \<approx>bn C z\<PRIME>\<^isub>1 \<dots> z\<PRIME>\<^isub>s"}}
- {@{text "z\<^isub>1 R\<^isub>1 z\<PRIME>\<^isub>1 \<dots> z\<^isub>s R\<^isub>s z\<PRIME>\<^isub>s"}}}
- \]\smallskip
-
- \noindent
- In this clause the relations @{text "R"}$_{1..s}$ are given by
-
- \[\mbox{
- \begin{tabular}{c@ {\hspace{2mm}}p{0.9\textwidth}}
- $\bullet$ & @{text "z\<^isub>i \<approx>ty z\<PRIME>\<^isub>i"} provided @{text "z\<^isub>i"} does not occur in @{text rhs} and
- is a recursive argument of @{text C},\smallskip\\
- $\bullet$ & @{text "z\<^isub>i = z\<PRIME>\<^isub>i"} provided @{text "z\<^isub>i"} does not occur in @{text rhs}
- and is a non-recursive argument of @{text C},\smallskip\\
- $\bullet$ & @{text "z\<^isub>i \<approx>bn\<^isub>i z\<PRIME>\<^isub>i"} provided @{text "z\<^isub>i"} occurs in @{text rhs}
- with the recursive call @{text "bn\<^isub>i x\<^isub>i"} and\smallskip\\
- $\bullet$ & @{text True} provided @{text "z\<^isub>i"} occurs in @{text rhs} but without a
- recursive call.
- \end{tabular}}
- \]\smallskip
-
- \noindent
- This completes the definition of alpha-equivalence. As a sanity check, we can show
- that the premises of empty binding clauses are a special case of the clauses for
- non-empty ones (we just have to unfold the definition of
- $\approx_{\,\textit{set}}^{\textit{R}, \textit{fa}}$ and take @{text "0"}
- for the existentially quantified permutation).
-
- Again let us take a look at a concrete example for these definitions. For
- the specification shown in \eqref{letrecs}
- we have three relations $\approx_{\textit{trm}}$, $\approx_{\textit{assn}}$ and
- $\approx_{\textit{bn}}$ with the following rules:
-
- \begin{equation}\label{rawalpha}\mbox{
- \begin{tabular}{@ {}c @ {}}
- \infer{@{text "Let as t \<approx>\<^bsub>trm\<^esub> Let as' t'"}}
- {@{term "alpha_lst_ex (bn as, t) alpha_trm fa_trm (bn as', t')"} &
- \hspace{5mm}@{text "as \<approx>\<^bsub>bn\<^esub> as'"}}\\
- \\
- \makebox[0mm]{\infer{@{text "Let_rec as t \<approx>\<^bsub>trm\<^esub> Let_rec as' t'"}}
- {@{term "alpha_lst_ex (bn as, ast) alpha_trm2 fa_trm2 (bn as', ast')"}}}\\
- \\
-
- \begin{tabular}{@ {}c @ {}}
- \infer{@{text "ANil \<approx>\<^bsub>assn\<^esub> ANil"}}{}\hspace{9mm}
- \infer{@{text "ACons a t as \<approx>\<^bsub>assn\<^esub> ACons a' t' as"}}
- {@{text "a = a'"} & \hspace{5mm}@{text "t \<approx>\<^bsub>trm\<^esub> t'"} & \hspace{5mm}@{text "as \<approx>\<^bsub>assn\<^esub> as'"}}
- \end{tabular}\\
- \\
-
- \begin{tabular}{@ {}c @ {}}
- \infer{@{text "ANil \<approx>\<^bsub>bn\<^esub> ANil"}}{}\hspace{9mm}
- \infer{@{text "ACons a t as \<approx>\<^bsub>bn\<^esub> ACons a' t' as"}}
- {@{text "t \<approx>\<^bsub>trm\<^esub> t'"} & \hspace{5mm}@{text "as \<approx>\<^bsub>bn\<^esub> as'"}}
- \end{tabular}
- \end{tabular}}
- \end{equation}\smallskip
-
- \noindent
- Notice the difference between $\approx_{\textit{assn}}$ and
- $\approx_{\textit{bn}}$: the latter only `tracks' alpha-equivalence of
- the components in an assignment that are \emph{not} bound. This is needed in the
- clause for @{text "Let"} (which has
- a non-recursive binder).
- The underlying reason is that the terms inside an assignment are not meant
- to be `under' the binder. Such a premise is \emph{not} needed in @{text "Let_rec"},
- because there all components of an assignment are `under' the binder.
- Note also that in case of more than one body (that is in the @{text "Let_rec"}-case above)
- we need to parametrise the relation $\approx_{\textit{list}}$ with a compound
- equivalence relation and a compound free-atom function. This is because the
- corresponding binding clause specifies a binder with two bodies, namely
- @{text "as"} and @{text "t"}.
-*}
-
-section {* Establishing the Reasoning Infrastructure *}
-
-text {*
- Having made all necessary definitions for raw terms, we can start with
- establishing the reasoning infrastructure for the alpha-equated types @{text
- "ty\<AL>"}$_{1..n}$, that is the types the user originally specified. We
- give in this section and the next the proofs we need for establishing this
- infrastructure. One point of our work is that we have completely
- automated these proofs in Isabelle/HOL.
-
- First we establish that the free-variable functions, the binding functions and the
- alpha-equi\-va\-lences are equivariant.
-
- \begin{lem}\mbox{}\\
- @{text "(i)"} The functions @{text "fa_ty"}$_{1..n}$, @{text "fa_bn"}$_{1..m}$ and
- @{text "bn"}$_{1..m}$ are equivariant.\\
- @{text "(ii)"} The relations @{text "\<approx>ty"}$_{1..n}$ and
- @{text "\<approx>bn"}$_{1..m}$ are equivariant.
- \end{lem}
-
- \begin{proof}
- The function package of Isabelle/HOL allows us to prove the first part by
- mutual induction over the definitions of the functions.\footnote{We have
- that the free-atom functions are terminating. From this the function
- package derives an induction principle~\cite{Krauss09}.} The second is by a
- straightforward induction over the rules of @{text "\<approx>ty"}$_{1..n}$ and
- @{text "\<approx>bn"}$_{1..m}$ using the first part.
- \end{proof}
-
- \noindent
- Next we establish that the alpha-equivalence relations defined in the
- previous section are indeed equivalence relations.
-
- \begin{lem}\label{equiv}
- The relations @{text "\<approx>ty"}$_{1..n}$ and @{text "\<approx>bn"}$_{1..m}$ are
- equivalence relations.
- \end{lem}
-
- \begin{proof}
- The proofs are by induction. The non-trivial
- cases involve premises built up by $\approx_{\textit{set}}$,
- $\approx_{\textit{set+}}$ and $\approx_{\textit{list}}$. They
- can be dealt with as in Lemma~\ref{alphaeq}. However, the transitivity
- case needs in addition the fact that the relations are equivariant.
- \end{proof}
-
- \noindent
- We can feed the last lemma into our quotient package and obtain new types
- @{text "ty"}$^\alpha_{1..n}$ representing alpha-equated terms of types
- @{text "ty"}$_{1..n}$. We also obtain definitions for the term-constructors
- @{text "C"}$^\alpha_{1..k}$ from the raw term-constructors @{text
- "C"}$_{1..k}$, and similar definitions for the free-atom functions @{text
- "fa_ty"}$^\alpha_{1..n}$ and @{text "fa_bn"}$^\alpha_{1..m}$ as well as the
- binding functions @{text "bn"}$^\alpha_{1..m}$. However, these definitions
- are not really useful to the user, since they are given in terms of the
- isomorphisms we obtained by creating new types in Isabelle/HOL (recall the
- picture shown in the Introduction).
-
- The first useful property for the user is the fact that distinct
- term-constructors are not equal, that is the property
-
- \begin{equation}\label{distinctalpha}
- \mbox{@{text "C"}$^\alpha$~@{text "x\<^isub>1 \<dots> x\<^isub>r"}~@{text "\<noteq>"}~%
- @{text "D"}$^\alpha$~@{text "y\<^isub>1 \<dots> y\<^isub>s"}}
- \end{equation}\smallskip
-
- \noindent
- whenever @{text "C"}$^\alpha$~@{text "\<noteq>"}~@{text "D"}$^\alpha$.
- In order to derive this property, we use the definition of alpha-equivalence
- and establish that
-
- \begin{equation}\label{distinctraw}
- \mbox{@{text "C x\<^isub>1 \<dots> x\<^isub>r"}\;$\not\approx$@{text ty}\;@{text "D y\<^isub>1 \<dots> y\<^isub>s"}}
- \end{equation}\smallskip
-
- \noindent
- holds for the corresponding raw term-constructors.
- In order to deduce \eqref{distinctalpha} from \eqref{distinctraw}, our quotient
- package needs to know that the raw term-constructors @{text "C"} and @{text "D"}
- are \emph{respectful} w.r.t.~the alpha-equivalence relations (see \cite{Homeier05}).
- Given, for example, @{text "C"} is of type @{text "ty"} with argument types
- @{text "ty"}$_{1..r}$, respectfulness amounts to showing that
-
- \[\mbox{
- @{text "C x\<^isub>1 \<dots> x\<^isub>r \<approx>ty C x\<PRIME>\<^isub>1 \<dots> x\<PRIME>\<^isub>r"}
- }\]\smallskip
-
- \noindent
- holds under the assumptions \mbox{@{text
- "x\<^isub>i \<approx>ty\<^isub>i x\<PRIME>\<^isub>i"}} whenever @{text "x\<^isub>i"}
- and @{text "x\<PRIME>\<^isub>i"} are recursive arguments of @{text C}, and
- @{text "x\<^isub>i = x\<PRIME>\<^isub>i"} whenever they are non-recursive arguments
- (similarly for @{text "D"}). For this we have to show
- by induction over the definitions of alpha-equivalences the following
- auxiliary implications
-
- \begin{equation}\label{fnresp}\mbox{
- \begin{tabular}{lll}
- @{text "x \<approx>ty\<^isub>i x'"} & implies & @{text "fa_ty\<^isub>i x = fa_ty\<^isub>i x'"}\\
- @{text "x \<approx>ty\<^isub>l x'"} & implies & @{text "fa_bn\<^isub>j x = fa_bn\<^isub>j x'"}\\
- @{text "x \<approx>ty\<^isub>l x'"} & implies & @{text "bn\<^isub>j x = bn\<^isub>j x'"}\\
- @{text "x \<approx>ty\<^isub>l x'"} & implies & @{text "x \<approx>bn\<^isub>j x'"}\\
- \end{tabular}
- }\end{equation}\smallskip
-
- \noindent
- whereby @{text "ty\<^isub>l"} is the type over which @{text "bn\<^isub>j"}
- is defined. Whereas the first, second and last implication are true by
- how we stated our definitions, the third \emph{only} holds because of our
- restriction imposed on the form of the binding functions---namely \emph{not}
- to return any bound atoms. In Ott, in contrast, the user may define @{text
- "bn"}$_{1..m}$ so that they return bound atoms and in this case the third
- implication is \emph{not} true. A result is that in general the lifting of the
- corresponding binding functions in Ott to alpha-equated terms is impossible.
- Having established respectfulness for the raw term-constructors, the
- quotient package is able to automatically deduce \eqref{distinctalpha} from
- \eqref{distinctraw}.
-
- Next we can lift the permutation operations defined in \eqref{ceqvt}. In
- order to make this lifting to go through, we have to show that the
- permutation operations are respectful. This amounts to showing that the
- alpha-equivalence relations are equivariant, which
- we already established in Lemma~\ref{equiv}. As a result we can add the
- equations
-
- \begin{equation}\label{calphaeqvt}
- @{text "\<pi> \<bullet> (C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>r) = C\<^sup>\<alpha> (\<pi> \<bullet> x\<^isub>1) \<dots> (\<pi> \<bullet> x\<^isub>r)"}
- \end{equation}\smallskip
-
- \noindent
- to our infrastructure. In a similar fashion we can lift the defining equations
- of the free-atom functions @{text "fa_ty\<AL>"}$_{1..n}$ and
- @{text "fa_bn\<AL>"}$_{1..m}$ as well as of the binding functions @{text
- "bn\<AL>"}$_{1..m}$ and size functions @{text "size_ty\<AL>"}$_{1..n}$.
- The latter are defined automatically for the raw types @{text "ty"}$_{1..n}$
- by the datatype package of Isabelle/HOL.
-
- We also need to lift the properties that characterise when two raw terms of the form
-
- \[
- \mbox{@{text "C x\<^isub>1 \<dots> x\<^isub>r \<approx>ty C x\<PRIME>\<^isub>1 \<dots> x\<PRIME>\<^isub>r"}}
- \]\smallskip
-
- \noindent
- are alpha-equivalent. This gives us conditions when the corresponding
- alpha-equated terms are \emph{equal}, namely
-
- \[
- @{text "C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>r = C\<^sup>\<alpha> x\<PRIME>\<^isub>1 \<dots> x\<PRIME>\<^isub>r"}.
- \]\smallskip
-
- \noindent
- We call these conditions as \emph{quasi-injectivity}. They correspond to the
- premises in our alpha-equiva\-lence relations, except that the
- relations @{text "\<approx>ty"}$_{1..n}$ are all replaced by equality (and similarly
- the free-atom and binding functions are replaced by their lifted
- counterparts). Recall the alpha-equivalence rules for @{text "Let"} and
- @{text "Let_rec"} shown in \eqref{rawalpha}. For @{text "Let\<^sup>\<alpha>"} and
- @{text "Let_rec\<^sup>\<alpha>"} we have
-
- \begin{equation}\label{alphalift}\mbox{
- \begin{tabular}{@ {}c @ {}}
- \infer{@{text "Let\<^sup>\<alpha> as t = Let\<^sup>\<alpha> as' t'"}}
- {@{term "alpha_lst_ex (bn_al as, t) equal fa_trm_al (bn as', t')"} &
- \hspace{5mm}@{text "as \<approx>\<AL>\<^bsub>bn\<^esub> as'"}}\\
- \\
- \makebox[0mm]{\infer{@{text "Let_rec\<^sup>\<alpha> as t = Let_rec\<^sup>\<alpha> as' t'"}}
- {@{term "alpha_lst_ex (bn_al as, ast) equ2 fa_trm2_al (bn_al as', ast')"}}}\\
- \end{tabular}}
- \end{equation}\smallskip
-
- We can also add to our infrastructure cases lemmas and a (mutual)
- induction principle for the types @{text "ty\<AL>"}$_{1..n}$. The cases
- lemmas allow the user to deduce a property @{text "P"} by exhaustively
- analysing how an element of a type, say @{text "ty\<AL>"}$_i$, can be
- constructed (that means one case for each of the term-constructors in @{text
- "ty\<AL>"}$_i\,$). The lifted cases lemma for the type @{text
- "ty\<AL>"}$_i\,$ looks as follows
-
- \begin{equation}\label{cases}
- \infer{P}
- {\begin{array}{l}
- @{text "\<forall>x\<^isub>1\<dots>x\<^isub>k. y = C\<AL>\<^isub>1 x\<^isub>1 \<dots> x\<^isub>k \<Rightarrow> P"}\\
- \hspace{5mm}\vdots\\
- @{text "\<forall>x\<^isub>1\<dots>x\<^isub>l. y = C\<AL>\<^isub>m x\<^isub>1 \<dots> x\<^isub>l \<Rightarrow> P"}\\
- \end{array}}
- \end{equation}\smallskip
-
- \noindent
- where @{text "y"} is a variable of type @{text "ty\<AL>"}$_i$ and @{text "P"} is the
- property that is established by the case analysis. Similarly, we have a (mutual)
- induction principle for the types @{text "ty\<AL>"}$_{1..n}$, which is of the
- form
-
- \begin{equation}\label{induct}
- \infer{@{text "P\<^isub>1 y\<^isub>1 \<and> \<dots> \<and> P\<^isub>n y\<^isub>n "}}
- {\begin{array}{l}
- @{text "\<forall>x\<^isub>1\<dots>x\<^isub>k. P\<^isub>i x\<^isub>i \<and> \<dots> \<and> P\<^isub>j x\<^isub>j \<Rightarrow> P (C\<AL>\<^isub>1 x\<^isub>1 \<dots> x\<^isub>k)"}\\
- \hspace{5mm}\vdots\\
- @{text "\<forall>x\<^isub>1\<dots>x\<^isub>l. P\<^isub>r x\<^isub>r \<and> \<dots> \<and> P\<^isub>s x\<^isub>s \<Rightarrow> P (C\<AL>\<^isub>m x\<^isub>1 \<dots> x\<^isub>l)"}\\
- \end{array}}
- \end{equation}\smallskip
-
- \noindent
- whereby the @{text P}$_{1..n}$ are the properties established by the
- induction, and the @{text y}$_{1..n}$ are of type @{text
- "ty\<AL>"}$_{1..n}$. Note that for the term constructors @{text
- "C"}$^\alpha_1$ the induction principle has a hypothesis of the form
-
- \[
- \mbox{@{text "\<forall>x\<^isub>1\<dots>x\<^isub>k. P\<^isub>i x\<^isub>i \<and> \<dots> \<and> P\<^isub>j x\<^isub>j \<Rightarrow> P (C\<AL>\<^sub>1 x\<^isub>1 \<dots> x\<^isub>k)"}}
- \]\smallskip
-
- \noindent
- in which the @{text "x"}$_{i..j}$ @{text "\<subseteq>"} @{text "x"}$_{1..k}$ are the
- recursive arguments of this term constructor (similarly for the other
- term-constructors).
-
- Recall the lambda-calculus with @{text "Let"}-patterns shown in
- \eqref{letpat}. The cases lemmas and the induction principle shown in
- \eqref{cases} and \eqref{induct} boil down in that example to the following three inference
- rules:
-
- \begin{equation}\label{inductex}\mbox{
- \begin{tabular}{c}
- \multicolumn{1}{@ {\hspace{-5mm}}l}{cases lemmas:}\smallskip\\
- \infer{@{text "P\<^bsub>trm\<^esub>"}}
- {\begin{array}{@ {}l@ {}}
- @{text "\<forall>x. y = Var\<^sup>\<alpha> x \<Rightarrow> P\<^bsub>trm\<^esub>"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2. y = App\<^sup>\<alpha> x\<^isub>1 x\<^isub>2 \<Rightarrow> P\<^bsub>trm\<^esub>"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2. y = Lam\<^sup>\<alpha> x\<^isub>1 x\<^isub>2 \<Rightarrow> P\<^bsub>trm\<^esub>"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2 x\<^isub>3. y = Let_pat\<^sup>\<alpha> x\<^isub>1 x\<^isub>2 x\<^isub>3 \<Rightarrow> P\<^bsub>trm\<^esub>"}
- \end{array}}\hspace{10mm}
-
- \infer{@{text "P\<^bsub>pat\<^esub>"}}
- {\begin{array}{@ {}l@ {}}
- @{text "\<forall>x. y = PVar\<^sup>\<alpha> x \<Rightarrow> P\<^bsub>pat\<^esub>"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2. y = PTup\<^sup>\<alpha> x\<^isub>1 x\<^isub>2 \<Rightarrow> P\<^bsub>pat\<^esub>"}
- \end{array}}\medskip\\
-
- \multicolumn{1}{@ {\hspace{-5mm}}l}{induction principle:}\smallskip\\
-
- \infer{@{text "P\<^bsub>trm\<^esub> y\<^isub>1 \<and> P\<^bsub>pat\<^esub> y\<^isub>2"}}
- {\begin{array}{@ {}l@ {}}
- @{text "\<forall>x. P\<^bsub>trm\<^esub> (Var\<^sup>\<alpha> x)"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2. P\<^bsub>trm\<^esub> x\<^isub>1 \<and> P\<^bsub>trm\<^esub> x\<^isub>2 \<Rightarrow> P\<^bsub>trm\<^esub> (App\<^sup>\<alpha> x\<^isub>1 x\<^isub>2)"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2. P\<^bsub>trm\<^esub> x\<^isub>2 \<Rightarrow> P\<^bsub>trm\<^esub> (Lam\<^sup>\<alpha> x\<^isub>1 x\<^isub>2)"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2 x\<^isub>3. P\<^bsub>pat\<^esub> x\<^isub>1 \<and> P\<^bsub>trm\<^esub> x\<^isub>2 \<and> P\<^bsub>trm\<^esub> x\<^isub>3 \<Rightarrow> P\<^bsub>trm\<^esub> (Let_pat\<^sup>\<alpha> x\<^isub>1 x\<^isub>2 x\<^isub>3)"}\\
- @{text "\<forall>x. P\<^bsub>pat\<^esub> (PVar\<^sup>\<alpha> x)"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2. P\<^bsub>pat\<^esub> x\<^isub>1 \<and> P\<^bsub>pat\<^esub> x\<^isub>2 \<Rightarrow> P\<^bsub>pat\<^esub> (PTup\<^sup>\<alpha> x\<^isub>1 x\<^isub>2)"}
- \end{array}}
- \end{tabular}}
- \end{equation}\smallskip
-
- By working now completely on the alpha-equated level, we
- can first show using \eqref{calphaeqvt} and Property~\ref{swapfreshfresh} that the support of each term
- constructor is included in the support of its arguments,
- namely
-
- \[
- @{text "(supp x\<^isub>1 \<union> \<dots> \<union> supp x\<^isub>r) supports (C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>r)"}
- \]\smallskip
-
- \noindent
- This allows us to prove using the induction principle for @{text "ty\<AL>"}$_{1..n}$
- that every element of type @{text "ty\<AL>"}$_{1..n}$ is finitely supported
- (using Proposition~\ref{supportsprop}{\it (i)}).
- Similarly, we can establish by induction that the free-atom functions and binding
- functions are equivariant, namely
-
- \[\mbox{
- \begin{tabular}{rcl}
- @{text "\<pi> \<bullet> (fa_ty\<AL>\<^isub>i x)"} & $=$ & @{text "fa_ty\<AL>\<^isub>i (\<pi> \<bullet> x)"}\\
- @{text "\<pi> \<bullet> (fa_bn\<AL>\<^isub>j x)"} & $=$ & @{text "fa_bn\<AL>\<^isub>j (\<pi> \<bullet> x)"}\\
- @{text "\<pi> \<bullet> (bn\<AL>\<^isub>j x)"} & $=$ & @{text "bn\<AL>\<^isub>j (\<pi> \<bullet> x)"}\\
- \end{tabular}}
- \]\smallskip
-
-
- \noindent
- Lastly, we can show that the support of elements in @{text
- "ty\<AL>"}$_{1..n}$ is the same as the free-atom functions @{text
- "fa_ty\<AL>"}$_{1..n}$. This fact is important in the nominal setting where
- the general theory is formulated in terms of support and freshness, but also
- provides evidence that our notions of free-atoms and alpha-equivalence
- `match up' correctly.
-
- \begin{thm}\label{suppfa}
- For @{text "x"}$_{1..n}$ with type @{text "ty\<AL>"}$_{1..n}$, we have
- @{text "supp x\<^isub>i = fa_ty\<AL>\<^isub>i x\<^isub>i"}.
- \end{thm}
-
- \begin{proof}
- The proof is by induction on @{text "x"}$_{1..n}$. In each case
- we unfold the definition of @{text "supp"}, move the swapping inside the
- term-constructors and then use the quasi-injectivity lemmas in order to complete the
- proof. For the abstraction cases we use then the facts derived in Theorem~\ref{suppabs},
- for which we have to know that every body of an abstraction is finitely supported.
- This, we have proved earlier.
- \end{proof}
-
- \noindent
- Consequently, we can replace the free-atom functions by @{text "supp"} in
- our quasi-injection lemmas. In the examples shown in \eqref{alphalift}, for instance,
- we obtain for @{text "Let\<^sup>\<alpha>"} and @{text "Let_rec\<^sup>\<alpha>"}
-
- \[\mbox{
- \begin{tabular}{@ {}c @ {}}
- \infer{@{text "Let\<^sup>\<alpha> as t = Let\<^sup>\<alpha> as' t'"}}
- {@{term "alpha_lst_ex (bn_al as, t) equal supp (bn_al as', t')"} &
- \hspace{5mm}@{text "as \<approx>\<AL>\<^bsub>bn\<^esub> as'"}}\\
- \\
- \makebox[0mm]{\infer{@{text "Let_rec\<^sup>\<alpha> as t = Let_rec\<^sup>\<alpha> as' t'"}}
- {@{term "alpha_lst_ex (bn_al as, ast) equ2 supp2 (bn_al as', ast')"}}}\\
- \end{tabular}}
- \]\smallskip
-
- \noindent
- Taking into account that the compound equivalence relation @{term
- "equ2"} and the compound free-atom function @{term "supp2"} are by
- definition equal to @{term "equal"} and @{term "supp"}, respectively, the
- above rules simplify further to
-
- \[\mbox{
- \begin{tabular}{@ {}c @ {}}
- \infer{@{text "Let\<^sup>\<alpha> as t = Let\<^sup>\<alpha> as' t'"}}
- {@{term "Abs_lst (bn_al as) t = Abs_lst (bn_al as') t'"} &
- \hspace{5mm}@{text "as \<approx>\<AL>\<^bsub>bn\<^esub> as'"}}\\
- \\
- \makebox[0mm]{\infer{@{text "Let_rec\<^sup>\<alpha> as t = Let_rec\<^sup>\<alpha> as' t'"}}
- {@{term "Abs_lst (bn_al as) ast = Abs_lst (bn_al as') ast'"}}}\\
- \end{tabular}}
- \]\smallskip
-
- \noindent
- which means we can characterise equality between term-constructors (on the
- alpha-equated level) in terms of equality between the abstractions defined
- in Section~\ref{sec:binders}. From this we can deduce the support for @{text
- "Let\<^sup>\<alpha>"} and @{text "Let_rec\<^sup>\<alpha>"}, namely
-
-
- \[\mbox{
- \begin{tabular}{l@ {\hspace{2mm}}l@ {\hspace{2mm}}l}
- @{text "supp (Let\<^sup>\<alpha> as t)"} & @{text "="} & @{text "(supp t - set (bn\<^sup>\<alpha> as)) \<union> fa\<AL>\<^bsub>bn\<^esub> as"}\\
- @{text "supp (Let_rec\<^sup>\<alpha> as t)"} & @{text "="} & @{text "(supp t \<union> supp as) - set (bn\<^sup>\<alpha> as)"}\\
- \end{tabular}}
- \]\smallskip
-
- \noindent
- using the support of abstractions derived in Theorem~\ref{suppabs}.
-
- To sum up this section, we have established a reasoning infrastructure for the
- types @{text "ty\<AL>"}$_{1..n}$ by first lifting definitions from the
- `raw' level to the quotient level and then by proving facts about
- these lifted definitions. All necessary proofs are generated automatically
- by custom ML-code.
-*}
-
-
-section {* Strong Induction Principles *}
-
-text {*
- In the previous section we derived induction principles for alpha-equated
- terms (see \eqref{induct} for the general form and \eqref{inductex} for an
- example). This was done by lifting the corresponding inductions principles
- for `raw' terms. We already employed these induction principles for
- deriving several facts about alpha-equated terms, including the property that
- the free-atom functions and the notion of support coincide. Still, we
- call these induction principles \emph{weak}, because for a term-constructor,
- say \mbox{@{text "C\<^sup>\<alpha> x\<^isub>1\<dots>x\<^isub>r"}}, the induction
- hypothesis requires us to establish (under some assumptions) a property
- @{text "P (C\<^sup>\<alpha> x\<^isub>1\<dots>x\<^isub>r)"} for \emph{all} @{text
- "x"}$_{1..r}$. The problem with this is that in the presence of binders we cannot make
- any assumptions about the atoms that are bound---for example assuming the variable convention.
- One obvious way around this
- problem is to rename bound atoms. Unfortunately, this leads to very clunky proofs
- and makes formalisations grievous experiences (especially in the context of
- multiple bound atoms).
-
- For the older versions of Nominal Isabelle we described in \cite{Urban08} a
- method for automatically strengthening weak induction principles. These
- stronger induction principles allow the user to make additional assumptions
- about bound atoms. The advantage of these assumptions is that they make in
- most cases any renaming of bound atoms unnecessary. To explain how the
- strengthening works, we use as running example the lambda-calculus with
- @{text "Let"}-patterns shown in \eqref{letpat}. Its weak induction principle
- is given in \eqref{inductex}. The stronger induction principle is as
- follows:
-
- \begin{equation}\label{stronginduct}
- \mbox{
- \begin{tabular}{@ {}c@ {}}
- \infer{@{text "P\<^bsub>trm\<^esub> c y\<^isub>1 \<and> P\<^bsub>pat\<^esub> c y\<^isub>2"}}
- {\begin{array}{l}
- @{text "\<forall>x c. P\<^bsub>trm\<^esub> c (Var\<^sup>\<alpha> x)"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2 c. (\<forall>d. P\<^bsub>trm\<^esub> d x\<^isub>1) \<and> (\<forall>d. P\<^bsub>trm\<^esub> d x\<^isub>2) \<Rightarrow> P\<^bsub>trm\<^esub> c (App\<^sup>\<alpha> x\<^isub>1 x\<^isub>2)"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2 c. atom x\<^isub>1 # c \<and> (\<forall>d. P\<^bsub>trm\<^esub> d x\<^isub>2) \<Rightarrow> P\<^bsub>trm\<^esub> c (Lam\<^sup>\<alpha> x\<^isub>1 x\<^isub>2)"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2 x\<^isub>3 c. (set (bn\<^sup>\<alpha> x\<^isub>1)) #\<^sup>* c \<and>"}\\
- \hspace{10mm}@{text "(\<forall>d. P\<^bsub>pat\<^esub> d x\<^isub>1) \<and> (\<forall>d. P\<^bsub>trm\<^esub> d x\<^isub>2) \<and> (\<forall>d. P\<^bsub>trm\<^esub> d x\<^isub>3) \<Rightarrow> P\<^bsub>trm\<^esub> c (Let_pat\<^sup>\<alpha> x\<^isub>1 x\<^isub>2 x\<^isub>3)"}\\
- @{text "\<forall>x c. P\<^bsub>pat\<^esub> c (PVar\<^sup>\<alpha> x)"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2 c. (\<forall>d. P\<^bsub>pat\<^esub> d x\<^isub>1) \<and> (\<forall>d. P\<^bsub>pat\<^esub> d x\<^isub>2) \<Rightarrow> P\<^bsub>pat\<^esub> c (PTup\<^sup>\<alpha> x\<^isub>1 x\<^isub>2)"}
- \end{array}}
- \end{tabular}}
- \end{equation}\smallskip
-
-
- \noindent
- Notice that instead of establishing two properties of the form @{text "
- P\<^bsub>trm\<^esub> y\<^isub>1 \<and> P\<^bsub>pat\<^esub> y\<^isub>2"}, as the
- weak one does, the stronger induction principle establishes the properties
- of the form @{text " P\<^bsub>trm\<^esub> c y\<^isub>1 \<and>
- P\<^bsub>pat\<^esub> c y\<^isub>2"} in which the additional parameter @{text
- c} is assumed to be of finite support. The purpose of @{text "c"} is to
- `control' which freshness assumptions the binders should satisfy in the
- @{text "Lam\<^sup>\<alpha>"} and @{text "Let_pat\<^sup>\<alpha>"} cases: for @{text
- "Lam\<^sup>\<alpha>"} we can assume the bound atom @{text "x\<^isub>1"} is fresh
- for @{text "c"} (third line); for @{text "Let_pat\<^sup>\<alpha>"} we can assume
- all bound atoms from an assignment are fresh for @{text "c"} (fourth
- line). In order to see how an instantiation for @{text "c"} in the
- conclusion `controls' the premises, one has to take into account that
- Isabelle/HOL is a typed logic. That means if @{text "c"} is instantiated
- with, for example, a pair, then this type-constraint will be propagated to
- the premises. The main point is that if @{text "c"} is instantiated
- appropriately, then the user can mimic the usual `pencil-and-paper'
- reasoning employing the variable convention about bound and free variables
- being distinct \cite{Urban08}.
-
- In what follows we will show that the weak induction principle in
- \eqref{inductex} implies the strong one \eqref{stronginduct}. This fact was established for
- single binders in \cite{Urban08} by some quite involved, nevertheless
- automated, induction proof. In this paper we simplify the proof by
- leveraging the automated proving tools from the function package of
- Isabelle/HOL \cite{Krauss09}. The reasoning principle behind these tools
- is well-founded induction. To use them in our setting, we have to discharge
- two proof obligations: one is that we have well-founded measures (one for
- each type @{text "ty"}$^\alpha_{1..n}$) that decrease in every induction
- step and the other is that we have covered all cases in the induction
- principle. Once these two proof obligations are discharged, the reasoning
- infrastructure of the function package will automatically derive the
- stronger induction principle. This way of establishing the stronger induction
- principle is considerably simpler than the earlier work presented in \cite{Urban08}.
-
- As measures we can use the size functions @{text "size_ty"}$^\alpha_{1..n}$,
- which we lifted in the previous section and which are all well-founded. It
- is straightforward to establish that the sizes decrease in every
- induction step. What is left to show is that we covered all cases.
- To do so, we have to derive stronger cases lemmas, which look in our
- running example are as follows:
-
- \[\mbox{
- \begin{tabular}{@ {}c@ {\hspace{4mm}}c@ {}}
- \infer{@{text "P\<^bsub>trm\<^esub>"}}
- {\begin{array}{@ {}l@ {}}
- @{text "\<forall>x. y = Var\<^sup>\<alpha> x \<Rightarrow> P\<^bsub>trm\<^esub>"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2. y = App\<^sup>\<alpha> x\<^isub>1 x\<^isub>2 \<Rightarrow> P\<^bsub>trm\<^esub>"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2. atom x\<^isub>1 # c \<and> y = Lam\<^sup>\<alpha> x\<^isub>1 x\<^isub>2 \<Rightarrow> P\<^bsub>trm\<^esub>"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2 x\<^isub>3. set (bn\<^sup>\<alpha> x\<^isub>1) #\<^sup>* c \<and> y = Let_pat\<^sup>\<alpha> x\<^isub>1 x\<^isub>2 x\<^isub>3 \<Rightarrow> P\<^bsub>trm\<^esub>"}
- \end{array}} &
-
- \infer{@{text "P\<^bsub>pat\<^esub>"}}
- {\begin{array}{@ {}l@ {}}
- @{text "\<forall>x. y = PVar\<^sup>\<alpha> x \<Rightarrow> P\<^bsub>pat\<^esub>"}\\
- @{text "\<forall>x\<^isub>1 x\<^isub>2. y = PTup\<^sup>\<alpha> x\<^isub>1 x\<^isub>2 \<Rightarrow> P\<^bsub>pat\<^esub>"}
- \end{array}}
- \end{tabular}}
- \]\smallskip
-
- \noindent
- They are stronger in the sense that they allow us to assume in the @{text
- "Lam\<^sup>\<alpha>"} and @{text "Let_pat\<^sup>\<alpha>"} cases that the bound atoms
- avoid, or being fresh for, a context @{text "c"} (which is assumed to be finitely supported).
-
- These stronger cases lemmas can be derived from the `weak' cases lemmas
- given in \eqref{inductex}. This is trivial in case of patterns (the one on
- the right-hand side) since the weak and strong cases lemma coincide (there
- is no binding in patterns). Interesting are only the cases for @{text
- "Lam\<^sup>\<alpha>"} and @{text "Let_pat\<^sup>\<alpha>"}, where we have some binders and
- therefore have an additional assumption about avoiding @{text "c"}. Let us
- first establish the case for @{text "Lam\<^sup>\<alpha>"}. By the weak cases lemma
- \eqref{inductex} we can assume that
-
- \begin{equation}\label{assm}
- @{text "y = Lam\<^sup>\<alpha> x\<^isub>1 x\<^isub>2"}
- \end{equation}\smallskip
-
- \noindent
- holds, and need to establish @{text "P\<^bsub>trm\<^esub>"}. The stronger cases lemma has the
- corresponding implication
-
- \begin{equation}\label{imp}
- @{text "\<forall>x\<^isub>1 x\<^isub>2. atom x\<^isub>1 # c \<and> y = Lam\<^sup>\<alpha> x\<^isub>1 x\<^isub>2 \<Rightarrow> P\<^bsub>trm\<^esub>"}
- \end{equation}\smallskip
-
- \noindent
- which we must use in order to infer @{text "P\<^bsub>trm\<^esub>"}. Clearly, we cannot
- use this implication directly, because we have no information whether or not @{text
- "x\<^isub>1"} is fresh for @{text "c"}. However, we can use Properties
- \ref{supppermeq} and \ref{avoiding} to rename @{text "x\<^isub>1"}. We know
- by Theorem~\ref{suppfa} that @{text "{atom x\<^isub>1} #\<^sup>* Lam\<^sup>\<alpha>
- x\<^isub>1 x\<^isub>2"} (since its support is @{text "supp x\<^isub>2 -
- {atom x\<^isub>1}"}). Property \ref{avoiding} provides us then with a
- permutation @{text "\<pi>"}, such that @{text "{atom (\<pi> \<bullet> x\<^isub>1)} #\<^sup>*
- c"} and \mbox{@{text "supp (Lam\<^sup>\<alpha> x\<^isub>1 x\<^isub>2) #\<^sup>* \<pi>"}} hold.
- By using Property \ref{supppermeq}, we can infer from the latter that
-
- \[
- @{text "Lam\<^sup>\<alpha> (\<pi> \<bullet> x\<^isub>1) (\<pi> \<bullet> x\<^isub>2) = Lam\<^sup>\<alpha> x\<^isub>1 x\<^isub>2"}
- \]\smallskip
-
- \noindent
- holds. We can use this equation in the assumption \eqref{assm}, and hence
- use the implication \eqref{imp} with the renamed @{text "\<pi> \<bullet> x\<^isub>1"}
- and @{text "\<pi> \<bullet> x\<^isub>2"} for concluding this case.
-
- The @{text "Let_pat\<^sup>\<alpha>"}-case involving a deep binder is slightly more complicated.
- We have the assumption
-
- \begin{equation}\label{assmtwo}
- @{text "y = Let_pat\<^sup>\<alpha> x\<^isub>1 x\<^isub>2 x\<^isub>3"}
- \end{equation}\smallskip
-
- \noindent
- and the implication from the stronger cases lemma
-
- \begin{equation}\label{impletpat}
- @{text "\<forall>x\<^isub>1 x\<^isub>2 x\<^isub>3. set (bn\<^sup>\<alpha> x\<^isub>1) #\<^sup>* c \<and> y = Let_pat\<^sup>\<alpha> x\<^isub>1 x\<^isub>2 x\<^isub>3 \<Rightarrow> P\<^bsub>trm\<^esub>"}
- \end{equation}\smallskip
-
- \noindent
- The reason that this case is more complicated is that we cannot directly apply Property
- \ref{avoiding} for obtaining a renaming permutation. Property \ref{avoiding} requires
- that the binders are fresh for the term in which we want to perform the renaming. But
- this is not true in terms such as (using an informal notation)
-
- \[
- @{text "Let (x, y) := (x, y) in (x, y)"}
- \]\smallskip
-
- \noindent
- where @{text x} and @{text y} are bound in the term, but are also free
- in the right-hand side of the assignment. We can, however, obtain such a renaming permutation, say
- @{text "\<pi>"}, for the abstraction @{term "Abs_lst (bn_al x\<^isub>1)
- x\<^isub>3"}. As a result we have \mbox{@{term "set (bn_al (\<pi> \<bullet> x\<^isub>1))
- \<sharp>* c"}} and @{term "Abs_lst (bn_al (\<pi> \<bullet> x\<^isub>1)) (\<pi> \<bullet> x\<^isub>3) =
- Abs_lst (bn_al x\<^isub>1) x\<^isub>3"} (remember @{text "set"} and @{text
- "bn\<^sup>\<alpha>"} are equivariant). Now the quasi-injective property for @{text
- "Let_pat\<^sup>\<alpha>"} states that
-
- \[
- \infer{@{text "Let_pat\<^sup>\<alpha> p t\<^isub>1 t\<^isub>2 = Let_pat\<^sup>\<alpha> p\<PRIME> t\<PRIME>\<^isub>1 t\<PRIME>\<^isub>2"}}
- {@{text "[bn\<^sup>\<alpha> p]\<^bsub>list\<^esub>. t\<^isub>2 = [bn\<^sup>\<alpha> p']\<^bsub>list\<^esub>. t\<PRIME>\<^isub>2"}\;\; &
- @{text "p \<approx>\<AL>\<^bsub>bn\<^esub> p\<PRIME>"}\;\; & @{text "t\<^isub>1 = t\<PRIME>\<^isub>1"}}
- \]\smallskip
-
- \noindent
- Since all atoms in a pattern are bound by @{text "Let_pat\<^sup>\<alpha>"}, we can infer
- that @{text "(\<pi> \<bullet> x\<^isub>1) \<approx>\<AL>\<^bsub>bn\<^esub> x\<^isub>1"} holds for every @{text "\<pi>"}. Therefore we have that
-
- \[
- @{text "Let_pat\<^sup>\<alpha> (\<pi> \<bullet> x\<^isub>1) x\<^isub>2 (\<pi> \<bullet> x\<^isub>3) = Let_pat\<^sup>\<alpha> x\<^isub>1 x\<^isub>2 x\<^isub>3"}
- \]\smallskip
-
- \noindent
- Taking the left-hand side in the assumption shown in \eqref{assmtwo}, we can use
- the implication \eqref{impletpat} from the stronger cases lemma to infer @{text "P\<^bsub>trm\<^esub>"}, as needed.
-
- The remaining difficulty is when a deep binder contains some atoms that are
- bound and some that are free. An example is @{text "Let\<^sup>\<alpha>"} in
- \eqref{letrecs}. In such cases @{text "(\<pi> \<bullet> x\<^isub>1)
- \<approx>\<AL>\<^bsub>bn\<^esub> x\<^isub>1"} does not hold in general. The idea however is
- that @{text "\<pi>"} only renames atoms that become bound. In this way @{text "\<pi>"}
- does not affect @{text "\<approx>\<AL>\<^bsub>bn\<^esub>"} (which only tracks alpha-equivalence of terms that are not
- under the binder). However, the problem is that the
- permutation operation @{text "\<pi> \<bullet> x\<^isub>1"} applies to all atoms in @{text "x\<^isub>1"}. To avoid this
- we introduce an auxiliary permutation operations, written @{text "_
- \<bullet>\<^bsub>bn\<^esub> _"}, for deep binders that only permutes bound atoms (or
- more precisely the atoms specified by the @{text "bn"}-functions) and leaves
- the other atoms unchanged. Like the functions @{text "fa_bn"}$_{1..m}$, we
- can define these permutation operations over raw terms analysing how the functions @{text
- "bn"}$_{1..m}$ are defined. Assuming the user specified a clause
-
- \[
- @{text "bn (C x\<^isub>1 \<dots> x\<^isub>r) = rhs"}
- \]\smallskip
-
- \noindent
- we define @{text "\<pi> \<bullet>\<^bsub>bn\<^esub> (C x\<^isub>1 \<dots> x\<^isub>r) \<equiv> C y\<^isub>1 \<dots> y\<^isub>r"} with @{text "y\<^isub>i"} determined as follows:
-
- \[\mbox{
- \begin{tabular}{c@ {\hspace{2mm}}p{0.9\textwidth}}
- $\bullet$ & @{text "y\<^isub>i \<equiv> x\<^isub>i"} provided @{text "x\<^isub>i"} does not occur in @{text "rhs"}\\
- $\bullet$ & @{text "y\<^isub>i \<equiv> \<pi> \<bullet>\<^bsub>bn\<^esub> x\<^isub>i"} provided @{text "bn x\<^isub>i"} is in @{text "rhs"}\\
- $\bullet$ & @{text "y\<^isub>i \<equiv> \<pi> \<bullet> x\<^isub>i"} otherwise
- \end{tabular}}
- \]\smallskip
-
- \noindent
- Using again the quotient package we can lift the auxiliary permutation operations
- @{text "_ \<bullet>\<^bsub>bn\<^esub> _"}
- to alpha-equated terms. Moreover we can prove the following two properties:
-
- \begin{lem}\label{permutebn}
- Given a binding function @{text "bn\<^sup>\<alpha>"} and auxiliary equivalence @{text "\<approx>\<AL>\<^bsub>bn\<^esub>"}
- then for all @{text "\<pi>"}\smallskip\\
- {\it (i)} @{text "\<pi> \<bullet> (bn\<^sup>\<alpha> x) = bn\<^sup>\<alpha> (\<pi> \<bullet>\<AL>\<^bsub>bn\<^esub> x)"} and\\
- {\it (ii)} @{text "(\<pi> \<bullet>\<AL>\<^bsub>bn\<^esub> x) \<approx>\<AL>\<^bsub>bn\<^esub> x"}.
- \end{lem}
-
- \begin{proof}
- By induction on @{text x}. The properties follow by unfolding of the
- definitions.
- \end{proof}
-
- \noindent
- The first property states that a permutation applied to a binding function
- is equivalent to first permuting the binders and then calculating the bound
- atoms. The second states that @{text "_ \<bullet>\<AL>\<^bsub>bn\<^esub> _"} preserves
- @{text "\<approx>\<AL>\<^bsub>bn\<^esub>"}. The main point of the auxiliary
- permutation functions is that they allow us to rename just the bound atoms in a
- term, without changing anything else.
-
- Having the auxiliary permutation function in place, we can now solve all remaining cases.
- For the @{text "Let\<^sup>\<alpha>"} term-constructor, for example, we can by Property \ref{avoiding}
- obtain a @{text "\<pi>"} such that
-
- \[
- @{text "(\<pi> \<bullet> (set (bn\<^sup>\<alpha> x\<^isub>1)) #\<^sup>* c"} \hspace{10mm}
- @{text "\<pi> \<bullet> [bn\<^sup>\<alpha> x\<^isub>1]\<^bsub>list\<^esub>. x\<^isub>2 = [bn\<^sup>\<alpha> x\<^isub>1]\<^bsub>list\<^esub>. x\<^isub>2"}
- \]\smallskip
-
- \noindent
- hold. Using the first part of Lemma \ref{permutebn}, we can simplify this
- to @{text "set (bn\<^sup>\<alpha> (\<pi> \<bullet>\<AL>\<^bsub>bn\<^esub> x\<^isub>1)) #\<^sup>* c"} and
- \mbox{@{text "[bn\<^sup>\<alpha> (\<pi> \<bullet>\<AL>\<^bsub>bn\<^esub> x\<^isub>1)]\<^bsub>list\<^esub>. (\<pi> \<bullet> x\<^isub>2) = [bn\<^sup>\<alpha> x\<^isub>1]\<^bsub>list\<^esub>. x\<^isub>2"}}. Since
- @{text "(\<pi> \<bullet>\<AL>\<^bsub>bn\<^esub> x\<^isub>1) \<approx>\<AL>\<^bsub>bn\<^esub> x\<^isub>1"} holds by the second part,
- we can infer that
-
- \[
- @{text "Let\<^sup>\<alpha> (\<pi> \<bullet>\<AL>\<^bsub>bn\<^esub> x\<^isub>1) (\<pi> \<bullet> x\<^isub>2) = Let\<^sup>\<alpha> x\<^isub>1 x\<^isub>2"}
- \]\smallskip
-
- \noindent
- holds. This allows us to use the implication from the strong cases
- lemma, and we are done.
-
- Consequently, we can discharge all proof-obligations about having covered all
- cases. This completes the proof establishing that the weak induction principles imply
- the strong induction principles. These strong induction principles have already proved
- being very useful in practice, particularly for proving properties about
- capture-avoiding substitution \cite{Urban08}.
-*}
-
-
-section {* Related Work\label{related} *}
-
-text {*
- To our knowledge the earliest usage of general binders in a theorem prover
- is described by Nara\-schew\-ski and Nipkow \cite{NaraschewskiNipkow99} with a
- formalisation of the algorithm W. This formalisation implements binding in
- type-schemes using a de-Bruijn indices representation. Since type-schemes in
- W contain only a single place where variables are bound, different indices
- do not refer to different binders (as in the usual de-Bruijn
- representation), but to different bound variables. A similar idea has been
- recently explored for general binders by Chargu\'eraud \cite{chargueraud09}
- in the locally nameless approach to
- binding. There, de-Bruijn indices consist of two
- numbers, one referring to the place where a variable is bound, and the other
- to which variable is bound. The reasoning infrastructure for both
- representations of bindings comes for free in theorem provers like
- Isabelle/HOL and Coq, since the corresponding term-calculi can be implemented
- as `normal' datatypes. However, in both approaches it seems difficult to
- achieve our fine-grained control over the `semantics' of bindings
- (i.e.~whether the order of binders should matter, or vacuous binders should
- be taken into account). To do so, one would require additional predicates
- that filter out unwanted terms. Our guess is that such predicates result in
- rather intricate formal reasoning. We are not aware of any formalisation of
- a non-trivial language that uses Chargu\'eraud's idea.
-
- Another technique for representing binding is higher-order abstract syntax
- (HOAS), which for example is implemented in the Twelf system \cite{pfenningsystem}.
- This representation technique supports very elegantly many aspects of
- \emph{single} binding, and impressive work by Lee et al~\cite{LeeCraryHarper07}
- has been done that uses HOAS for mechanising the metatheory of SML. We
- are, however, not aware how multiple binders of SML are represented in this
- work. Judging from the submitted Twelf-solution for the POPLmark challenge,
- HOAS cannot easily deal with binding constructs where the number of bound
- variables is not fixed. For example, in the second part of this challenge,
- @{text "Let"}s involve patterns that bind multiple variables at once. In
- such situations, HOAS seems to have to resort to the
- iterated-single-binders-approach with all the unwanted consequences when
- reasoning about the resulting terms.
-
-
- Two formalisations involving general binders have been
- performed in older
- versions of Nominal Isabelle (one about Psi-calculi and one about algorithm W
- \cite{BengtsonParow09,UrbanNipkow09}). Both
- use the approach based on iterated single binders. Our experience with
- the latter formalisation has been disappointing. The major pain arose from
- the need to `unbind' variables. This can be done in one step with our
- general binders described in this paper, but needs a cumbersome
- iteration with single binders. The resulting formal reasoning turned out to
- be rather unpleasant.
-
- The most closely related work to the one presented here is the Ott-tool by
- Sewell et al \cite{ott-jfp} and the C$\alpha$ml language by Pottier
- \cite{Pottier06}. Ott is a nifty front-end for creating \LaTeX{} documents
- from specifications of term-calculi involving general binders. For a subset
- of the specifications Ott can also generate theorem prover code using a `raw'
- representation of terms, and in Coq also a locally nameless
- representation. The developers of this tool have also put forward (on paper)
- a definition for alpha-equivalence and free variables for terms that can be
- specified in Ott. This definition is rather different from ours, not using
- any nominal techniques. To our knowledge there is no concrete mathematical
- result concerning this notion of alpha-equivalence and free variables. We
- have proved that our definitions lead to alpha-equated terms, whose support
- is as expected (that means bound atoms are removed from the support). We
- also showed that our specifications lift from `raw' types to types of
- alpha-equivalence classes. For this we have established (automatically) that every
- term-constructor and function defined for `raw' types
- is respectful w.r.t.~alpha-equivalence.
-
- Although we were heavily inspired by the syntax of Ott, its definition of
- alpha-equi\-valence is unsuitable for our extension of Nominal
- Isabelle. First, it is far too complicated to be a basis for automated
- proofs implemented on the ML-level of Isabelle/HOL. Second, it covers cases
- of binders depending on other binders, which just do not make sense for our
- alpha-equated terms. Third, it allows empty types that have no meaning in a
- HOL-based theorem prover. We also had to generalise slightly Ott's binding
- clauses. In Ott one specifies binding clauses with a single body; we allow
- more than one. We have to do this, because this makes a difference for our
- notion of alpha-equivalence in case of \isacommand{binds (set)} and
- \isacommand{binds (set+)}. Consider the examples
-
- \[\mbox{
- \begin{tabular}{@ {}l@ {\hspace{2mm}}l@ {}}
- @{text "Foo\<^isub>1 xs::name fset t::trm s::trm"} &
- \isacommand{binds (set)} @{text "xs"} \isacommand{in} @{text "t s"}\\
- @{text "Foo\<^isub>2 xs::name fset t::trm s::trm"} &
- \isacommand{binds (set)} @{text "xs"} \isacommand{in} @{text "t"},
- \isacommand{binds (set)} @{text "xs"} \isacommand{in} @{text "s"}\\
- \end{tabular}}
- \]\smallskip
-
- \noindent
- In the first term-constructor we have a single body that happens to be
- `spread' over two arguments; in the second term-constructor we have two
- independent bodies in which the same variables are bound. As a result we
- have\footnote{Assuming @{term "a \<noteq> b"}, there is no permutation that can
- make @{text "(a, b)"} equal with both @{text "(a, b)"} and @{text "(b, a)"}, but
- there are two permutations so that we can make @{text "(a, b)"} and @{text
- "(a, b)"} equal with one permutation, and @{text "(a, b)"} and @{text "(b,
- a)"} with the other.}
-
-
- \[\mbox{
- \begin{tabular}{r@ {\hspace{1.5mm}}c@ {\hspace{1.5mm}}l}
- @{text "Foo\<^isub>1 {a, b} (a, b) (a, b)"} & $\not=$ &
- @{text "Foo\<^isub>1 {a, b} (a, b) (b, a)"}
- \end{tabular}}
- \]\smallskip
-
- \noindent
- but
-
- \[\mbox{
- \begin{tabular}{r@ {\hspace{1.5mm}}c@ {\hspace{1.5mm}}l}
- @{text "Foo\<^isub>2 {a, b} (a, b) (a, b)"} & $=$ &
- @{text "Foo\<^isub>2 {a, b} (a, b) (b, a)"}\\
- \end{tabular}}
- \]\smallskip
-
- \noindent
- and therefore need the extra generality to be able to distinguish between
- both specifications. Because of how we set up our definitions, we also had
- to impose some restrictions (like a single binding function for a deep
- binder) that are not present in Ott. Our expectation is that we can still
- cover many interesting term-calculi from programming language research, for
- example the Core-Haskell language from the Introduction. With the work
- presented in this paper we can define it formally as shown in
- Figure~\ref{nominalcorehas} and then Nominal Isabelle derives automatically
- a corresponding reasoning infrastructure.
-
- \begin{figure}[p]
- \begin{boxedminipage}{\linewidth}
- \small
- \begin{tabular}{l}
- \isacommand{atom\_decl}~@{text "var cvar tvar"}\\[1mm]
- \isacommand{nominal\_datatype}~@{text "tkind ="}~@{text "KStar"}~$|$~@{text "KFun tkind tkind"}\\
- \isacommand{and}~@{text "ckind ="}~@{text "CKSim ty ty"}\\
- \isacommand{and}~@{text "ty ="}~@{text "TVar tvar"}~$|$~@{text "T string"}~$|$~@{text "TApp ty ty"}\\
- $|$~@{text "TFun string ty_list"}~%
- $|$~@{text "TAll tv::tvar tkind ty::ty"}\hspace{3mm}\isacommand{binds}~@{text "tv"}~\isacommand{in}~@{text ty}\\
- $|$~@{text "TArr ckind ty"}\\
- \isacommand{and}~@{text "ty_lst ="}~@{text "TNil"}~$|$~@{text "TCons ty ty_lst"}\\
- \isacommand{and}~@{text "cty ="}~@{text "CVar cvar"}~%
- $|$~@{text "C string"}~$|$~@{text "CApp cty cty"}~$|$~@{text "CFun string co_lst"}\\
- $|$~@{text "CAll cv::cvar ckind cty::cty"}\hspace{3mm}\isacommand{binds}~@{text "cv"}~\isacommand{in}~@{text cty}\\
- $|$~@{text "CArr ckind cty"}~$|$~@{text "CRefl ty"}~$|$~@{text "CSym cty"}~$|$~@{text "CCirc cty cty"}\\
- $|$~@{text "CAt cty ty"}~$|$~@{text "CLeft cty"}~$|$~@{text "CRight cty"}~$|$~@{text "CSim cty cty"}\\
- $|$~@{text "CRightc cty"}~$|$~@{text "CLeftc cty"}~$|$~@{text "Coerce cty cty"}\\
- \isacommand{and}~@{text "co_lst ="}~@{text "CNil"}~$|$~@{text "CCons cty co_lst"}\\
- \isacommand{and}~@{text "trm ="}~@{text "Var var"}~$|$~@{text "K string"}\\
- $|$~@{text "LAM_ty tv::tvar tkind t::trm"}\hspace{3mm}\isacommand{binds}~@{text "tv"}~\isacommand{in}~@{text t}\\
- $|$~@{text "LAM_cty cv::cvar ckind t::trm"}\hspace{3mm}\isacommand{binds}~@{text "cv"}~\isacommand{in}~@{text t}\\
- $|$~@{text "App_ty trm ty"}~$|$~@{text "App_cty trm cty"}~$|$~@{text "App trm trm"}\\
- $|$~@{text "Lam v::var ty t::trm"}\hspace{3mm}\isacommand{binds}~@{text "v"}~\isacommand{in}~@{text t}\\
- $|$~@{text "Let x::var ty trm t::trm"}\hspace{3mm}\isacommand{binds}~@{text x}~\isacommand{in}~@{text t}\\
- $|$~@{text "Case trm assoc_lst"}~$|$~@{text "Cast trm co"}\\
- \isacommand{and}~@{text "assoc_lst ="}~@{text ANil}~%
- $|$~@{text "ACons p::pat t::trm assoc_lst"}\hspace{3mm}\isacommand{binds}~@{text "bv p"}~\isacommand{in}~@{text t}\\
- \isacommand{and}~@{text "pat ="}~@{text "Kpat string tvtk_lst tvck_lst vt_lst"}\\
- \isacommand{and}~@{text "vt_lst ="}~@{text VTNil}~$|$~@{text "VTCons var ty vt_lst"}\\
- \isacommand{and}~@{text "tvtk_lst ="}~@{text TVTKNil}~$|$~@{text "TVTKCons tvar tkind tvtk_lst"}\\
- \isacommand{and}~@{text "tvck_lst ="}~@{text TVCKNil}~$|$ @{text "TVCKCons cvar ckind tvck_lst"}\\
- \isacommand{binder}\\
- @{text "bv :: pat \<Rightarrow> atom list"}~\isacommand{and}\\
- @{text "bv\<^isub>1 :: vt_lst \<Rightarrow> atom list"}~\isacommand{and}\\
- @{text "bv\<^isub>2 :: tvtk_lst \<Rightarrow> atom list"}~\isacommand{and}\\
- @{text "bv\<^isub>3 :: tvck_lst \<Rightarrow> atom list"}\\
- \isacommand{where}\\
- \phantom{$|$}~@{text "bv (K s tvts tvcs vs) = (bv\<^isub>3 tvts) @ (bv\<^isub>2 tvcs) @ (bv\<^isub>1 vs)"}\\
- $|$~@{text "bv\<^isub>1 VTNil = []"}\\
- $|$~@{text "bv\<^isub>1 (VTCons x ty tl) = (atom x)::(bv\<^isub>1 tl)"}\\
- $|$~@{text "bv\<^isub>2 TVTKNil = []"}\\
- $|$~@{text "bv\<^isub>2 (TVTKCons a ty tl) = (atom a)::(bv\<^isub>2 tl)"}\\
- $|$~@{text "bv\<^isub>3 TVCKNil = []"}\\
- $|$~@{text "bv\<^isub>3 (TVCKCons c cty tl) = (atom c)::(bv\<^isub>3 tl)"}\\
- \end{tabular}
- \end{boxedminipage}
- \caption{A definition for Core-Haskell in Nominal Isabelle. For the moment we
- do not support nested types; therefore we explicitly have to unfold the
- lists @{text "co_lst"}, @{text "assoc_lst"} and so on. Apart from that limitation, the
- definition follows closely the original shown in Figure~\ref{corehas}. The
- point of our work is that having made such a definition in Nominal Isabelle,
- one obtains automatically a reasoning infrastructure for Core-Haskell.
- \label{nominalcorehas}}
- \end{figure}
- \afterpage{\clearpage}
-
- Pottier presents a programming language, called C$\alpha$ml, for
- representing terms with general binders inside OCaml \cite{Pottier06}. This
- language is implemented as a front-end that can be translated to OCaml with
- the help of a library. He presents a type-system in which the scope of
- general binders can be specified using special markers, written @{text
- "inner"} and @{text "outer"}. It seems our and his specifications can be
- inter-translated as long as ours use the binding mode \isacommand{binds}
- only. However, we have not proved this. Pottier gives a definition for
- alpha-equivalence, which also uses a permutation operation (like ours).
- Still, this definition is rather different from ours and he only proves that
- it defines an equivalence relation. A complete reasoning infrastructure is
- well beyond the purposes of his language. Similar work for Haskell with
- similar results was reported by Cheney \cite{Cheney05a} and more recently
- by Weirich et al \cite{WeirichYorgeySheard11}.
-
- In a slightly different domain (programming with dependent types),
- Altenkirch et al \cite{Altenkirch10} present a calculus with a notion of
- alpha-equivalence related to our binding mode \isacommand{binds (set+)}.
- Their definition is similar to the one by Pottier, except that it has a more
- operational flavour and calculates a partial (renaming) map. In this way,
- the definition can deal with vacuous binders. However, to our best
- knowledge, no concrete mathematical result concerning this definition of
- alpha-equivalence has been proved.
-*}
-
-section {* Conclusion *}
-
-text {*
-
- We have presented an extension of Nominal Isabelle for dealing with general
- binders, that is where term-constructors have multiple bound atoms. For this
- extension we introduced new definitions of alpha-equivalence and automated
- all necessary proofs in Isabelle/HOL. To specify general binders we used
- the syntax from Ott, but extended it in some places and restricted
- it in others so that the definitions make sense in the context of alpha-equated
- terms. We also introduced two binding modes (set and set+) that do not exist
- in Ott. We have tried out the extension with calculi such as Core-Haskell,
- type-schemes and approximately a dozen of other typical examples from
- programming language research~\cite{SewellBestiary}. The code will
- eventually become part of the Isabelle distribution.\footnote{It
- can be downloaded already from \href{http://isabelle.in.tum.de/nominal/download}
- {http://isabelle.in.tum.de/nominal/download}.}
-
- We have left out a discussion about how functions can be defined over
- alpha-equated terms involving general binders. In earlier versions of
- Nominal Isabelle this turned out to be a thorny issue. We hope to do better
- this time by using the function package \cite{Krauss09} that has recently
- been implemented in Isabelle/HOL and also by restricting function
- definitions to equivariant functions (for them we can provide more
- automation).
-
- There are some restrictions we imposed in this paper that we would like to lift in
- future work. One is the exclusion of nested datatype definitions. Nested
- datatype definitions would allow one to specify, for instance, the function kinds
- in Core-Haskell as @{text "TFun string (ty list)"} instead of the unfolded
- version @{text "TFun string ty_list"} (see Figure~\ref{nominalcorehas}). To
- achieve this, we need more clever implementation than we have
- at the moment. However, really lifting this restriction will involve major
- work on the underlying datatype package of Isabelle/HOL.
-
- A more interesting line of investigation is whether we can go beyond the
- simple-minded form of binding functions that we adopted from Ott. At the moment, binding
- functions can only return the empty set, a singleton atom set or unions
- of atom sets (similarly for lists). It remains to be seen whether
- properties like
-
- \[
- \mbox{@{text "fa_ty x = bn x \<union> fa_bn x"}}
- \]\smallskip
-
- \noindent
- allow us to support more interesting binding functions.
-
- We have also not yet played with other binding modes. For example we can
- imagine that there is need for a binding mode where instead of usual lists,
- we abstract lists of distinct elements (the corresponding type @{text
- "dlist"} already exists in the library of Isabelle/HOL). We expect the
- presented work can be extended to accommodate such binding modes.\medskip
-
- \noindent
- {\bf Acknowledgements:} We are very grateful to Andrew Pitts for many
- discussions about Nominal Isabelle. We thank Peter Sewell for making the
- informal notes \cite{SewellBestiary} available to us and also for patiently
- explaining some of the finer points of the Ott-tool. Stephanie Weirich
- suggested to separate the subgrammars of kinds and types in our Core-Haskell
- example. Ramana Kumar and Andrei Popescu helped us with comments about
- an earlier version of this paper.
-*}
-
-
-(*<*)
-end
-(*>*)
--- a/LMCS-Paper/ROOT.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-quick_and_dirty := true;
-no_document use_thys ["~~/src/HOL/Library/LaTeXsugar",
- "../Nominal/Nominal2"];
-use_thys ["Paper"];
\ No newline at end of file
--- a/LMCS-Paper/document/llncs.cls Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1207 +0,0 @@
-% LLNCS DOCUMENT CLASS -- version 2.17 (12-Jul-2010)
-% Springer Verlag LaTeX2e support for Lecture Notes in Computer Science
-%
-%%
-%% \CharacterTable
-%% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
-%% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
-%% Digits \0\1\2\3\4\5\6\7\8\9
-%% Exclamation \! Double quote \" Hash (number) \#
-%% Dollar \$ Percent \% Ampersand \&
-%% Acute accent \' Left paren \( Right paren \)
-%% Asterisk \* Plus \+ Comma \,
-%% Minus \- Point \. Solidus \/
-%% Colon \: Semicolon \; Less than \<
-%% Equals \= Greater than \> Question mark \?
-%% Commercial at \@ Left bracket \[ Backslash \\
-%% Right bracket \] Circumflex \^ Underscore \_
-%% Grave accent \` Left brace \{ Vertical bar \|
-%% Right brace \} Tilde \~}
-%%
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-\ProvidesClass{llncs}[2010/07/12 v2.17
-^^J LaTeX document class for Lecture Notes in Computer Science]
-% Options
-\let\if@envcntreset\iffalse
-\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue}
-\DeclareOption{citeauthoryear}{\let\citeauthoryear=Y}
-\DeclareOption{oribibl}{\let\oribibl=Y}
-\let\if@custvec\iftrue
-\DeclareOption{orivec}{\let\if@custvec\iffalse}
-\let\if@envcntsame\iffalse
-\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue}
-\let\if@envcntsect\iffalse
-\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue}
-\let\if@runhead\iffalse
-\DeclareOption{runningheads}{\let\if@runhead\iftrue}
-
-\let\if@openright\iftrue
-\let\if@openbib\iffalse
-\DeclareOption{openbib}{\let\if@openbib\iftrue}
-
-% languages
-\let\switcht@@therlang\relax
-\def\ds@deutsch{\def\switcht@@therlang{\switcht@deutsch}}
-\def\ds@francais{\def\switcht@@therlang{\switcht@francais}}
-
-\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}}
-
-\ProcessOptions
-
-\LoadClass[twoside]{article}
-\RequirePackage{multicol} % needed for the list of participants, index
-\RequirePackage{aliascnt}
-
-\setlength{\textwidth}{12.2cm}
-\setlength{\textheight}{19.3cm}
-\renewcommand\@pnumwidth{2em}
-\renewcommand\@tocrmarg{3.5em}
-%
-\def\@dottedtocline#1#2#3#4#5{%
- \ifnum #1>\c@tocdepth \else
- \vskip \z@ \@plus.2\p@
- {\leftskip #2\relax \rightskip \@tocrmarg \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \parindent #2\relax\@afterindenttrue
- \interlinepenalty\@M
- \leavevmode
- \@tempdima #3\relax
- \advance\leftskip \@tempdima \null\nobreak\hskip -\leftskip
- {#4}\nobreak
- \leaders\hbox{$\m@th
- \mkern \@dotsep mu\hbox{.}\mkern \@dotsep
- mu$}\hfill
- \nobreak
- \hb@xt@\@pnumwidth{\hfil\normalfont \normalcolor #5}%
- \par}%
- \fi}
-%
-\def\switcht@albion{%
-\def\abstractname{Abstract.}
-\def\ackname{Acknowledgement.}
-\def\andname{and}
-\def\lastandname{\unskip, and}
-\def\appendixname{Appendix}
-\def\chaptername{Chapter}
-\def\claimname{Claim}
-\def\conjecturename{Conjecture}
-\def\contentsname{Table of Contents}
-\def\corollaryname{Corollary}
-\def\definitionname{Definition}
-\def\examplename{Example}
-\def\exercisename{Exercise}
-\def\figurename{Fig.}
-\def\keywordname{{\bf Keywords:}}
-\def\indexname{Index}
-\def\lemmaname{Lemma}
-\def\contriblistname{List of Contributors}
-\def\listfigurename{List of Figures}
-\def\listtablename{List of Tables}
-\def\mailname{{\it Correspondence to\/}:}
-\def\noteaddname{Note added in proof}
-\def\notename{Note}
-\def\partname{Part}
-\def\problemname{Problem}
-\def\proofname{Proof}
-\def\propertyname{Property}
-\def\propositionname{Proposition}
-\def\questionname{Question}
-\def\remarkname{Remark}
-\def\seename{see}
-\def\solutionname{Solution}
-\def\subclassname{{\it Subject Classifications\/}:}
-\def\tablename{Table}
-\def\theoremname{Theorem}}
-\switcht@albion
-% Names of theorem like environments are already defined
-% but must be translated if another language is chosen
-%
-% French section
-\def\switcht@francais{%\typeout{On parle francais.}%
- \def\abstractname{R\'esum\'e.}%
- \def\ackname{Remerciements.}%
- \def\andname{et}%
- \def\lastandname{ et}%
- \def\appendixname{Appendice}
- \def\chaptername{Chapitre}%
- \def\claimname{Pr\'etention}%
- \def\conjecturename{Hypoth\`ese}%
- \def\contentsname{Table des mati\`eres}%
- \def\corollaryname{Corollaire}%
- \def\definitionname{D\'efinition}%
- \def\examplename{Exemple}%
- \def\exercisename{Exercice}%
- \def\figurename{Fig.}%
- \def\keywordname{{\bf Mots-cl\'e:}}
- \def\indexname{Index}
- \def\lemmaname{Lemme}%
- \def\contriblistname{Liste des contributeurs}
- \def\listfigurename{Liste des figures}%
- \def\listtablename{Liste des tables}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}%
- \def\notename{Remarque}%
- \def\partname{Partie}%
- \def\problemname{Probl\`eme}%
- \def\proofname{Preuve}%
- \def\propertyname{Caract\'eristique}%
-%\def\propositionname{Proposition}%
- \def\questionname{Question}%
- \def\remarkname{Remarque}%
- \def\seename{voir}
- \def\solutionname{Solution}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tableau}%
- \def\theoremname{Th\'eor\`eme}%
-}
-%
-% German section
-\def\switcht@deutsch{%\typeout{Man spricht deutsch.}%
- \def\abstractname{Zusammenfassung.}%
- \def\ackname{Danksagung.}%
- \def\andname{und}%
- \def\lastandname{ und}%
- \def\appendixname{Anhang}%
- \def\chaptername{Kapitel}%
- \def\claimname{Behauptung}%
- \def\conjecturename{Hypothese}%
- \def\contentsname{Inhaltsverzeichnis}%
- \def\corollaryname{Korollar}%
-%\def\definitionname{Definition}%
- \def\examplename{Beispiel}%
- \def\exercisename{\"Ubung}%
- \def\figurename{Abb.}%
- \def\keywordname{{\bf Schl\"usselw\"orter:}}
- \def\indexname{Index}
-%\def\lemmaname{Lemma}%
- \def\contriblistname{Mitarbeiter}
- \def\listfigurename{Abbildungsverzeichnis}%
- \def\listtablename{Tabellenverzeichnis}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Nachtrag}%
- \def\notename{Anmerkung}%
- \def\partname{Teil}%
-%\def\problemname{Problem}%
- \def\proofname{Beweis}%
- \def\propertyname{Eigenschaft}%
-%\def\propositionname{Proposition}%
- \def\questionname{Frage}%
- \def\remarkname{Anmerkung}%
- \def\seename{siehe}
- \def\solutionname{L\"osung}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tabelle}%
-%\def\theoremname{Theorem}%
-}
-
-% Ragged bottom for the actual page
-\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil
-\global\let\@textbottom\relax}}
-
-\renewcommand\small{%
- \@setfontsize\small\@ixpt{11}%
- \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
- \abovedisplayshortskip \z@ \@plus2\p@
- \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
- \def\@listi{\leftmargin\leftmargini
- \parsep 0\p@ \@plus1\p@ \@minus\p@
- \topsep 8\p@ \@plus2\p@ \@minus4\p@
- \itemsep0\p@}%
- \belowdisplayskip \abovedisplayskip
-}
-
-\frenchspacing
-\widowpenalty=10000
-\clubpenalty=10000
-
-\setlength\oddsidemargin {63\p@}
-\setlength\evensidemargin {63\p@}
-\setlength\marginparwidth {90\p@}
-
-\setlength\headsep {16\p@}
-
-\setlength\footnotesep{7.7\p@}
-\setlength\textfloatsep{8mm\@plus 2\p@ \@minus 4\p@}
-\setlength\intextsep {8mm\@plus 2\p@ \@minus 2\p@}
-
-\setcounter{secnumdepth}{2}
-
-\newcounter {chapter}
-\renewcommand\thechapter {\@arabic\c@chapter}
-
-\newif\if@mainmatter \@mainmattertrue
-\newcommand\frontmatter{\cleardoublepage
- \@mainmatterfalse\pagenumbering{Roman}}
-\newcommand\mainmatter{\cleardoublepage
- \@mainmattertrue\pagenumbering{arabic}}
-\newcommand\backmatter{\if@openright\cleardoublepage\else\clearpage\fi
- \@mainmatterfalse}
-
-\renewcommand\part{\cleardoublepage
- \thispagestyle{empty}%
- \if@twocolumn
- \onecolumn
- \@tempswatrue
- \else
- \@tempswafalse
- \fi
- \null\vfil
- \secdef\@part\@spart}
-
-\def\@part[#1]#2{%
- \ifnum \c@secnumdepth >-2\relax
- \refstepcounter{part}%
- \addcontentsline{toc}{part}{\thepart\hspace{1em}#1}%
- \else
- \addcontentsline{toc}{part}{#1}%
- \fi
- \markboth{}{}%
- {\centering
- \interlinepenalty \@M
- \normalfont
- \ifnum \c@secnumdepth >-2\relax
- \huge\bfseries \partname~\thepart
- \par
- \vskip 20\p@
- \fi
- \Huge \bfseries #2\par}%
- \@endpart}
-\def\@spart#1{%
- {\centering
- \interlinepenalty \@M
- \normalfont
- \Huge \bfseries #1\par}%
- \@endpart}
-\def\@endpart{\vfil\newpage
- \if@twoside
- \null
- \thispagestyle{empty}%
- \newpage
- \fi
- \if@tempswa
- \twocolumn
- \fi}
-
-\newcommand\chapter{\clearpage
- \thispagestyle{empty}%
- \global\@topnum\z@
- \@afterindentfalse
- \secdef\@chapter\@schapter}
-\def\@chapter[#1]#2{\ifnum \c@secnumdepth >\m@ne
- \if@mainmatter
- \refstepcounter{chapter}%
- \typeout{\@chapapp\space\thechapter.}%
- \addcontentsline{toc}{chapter}%
- {\protect\numberline{\thechapter}#1}%
- \else
- \addcontentsline{toc}{chapter}{#1}%
- \fi
- \else
- \addcontentsline{toc}{chapter}{#1}%
- \fi
- \chaptermark{#1}%
- \addtocontents{lof}{\protect\addvspace{10\p@}}%
- \addtocontents{lot}{\protect\addvspace{10\p@}}%
- \if@twocolumn
- \@topnewpage[\@makechapterhead{#2}]%
- \else
- \@makechapterhead{#2}%
- \@afterheading
- \fi}
-\def\@makechapterhead#1{%
-% \vspace*{50\p@}%
- {\centering
- \ifnum \c@secnumdepth >\m@ne
- \if@mainmatter
- \large\bfseries \@chapapp{} \thechapter
- \par\nobreak
- \vskip 20\p@
- \fi
- \fi
- \interlinepenalty\@M
- \Large \bfseries #1\par\nobreak
- \vskip 40\p@
- }}
-\def\@schapter#1{\if@twocolumn
- \@topnewpage[\@makeschapterhead{#1}]%
- \else
- \@makeschapterhead{#1}%
- \@afterheading
- \fi}
-\def\@makeschapterhead#1{%
-% \vspace*{50\p@}%
- {\centering
- \normalfont
- \interlinepenalty\@M
- \Large \bfseries #1\par\nobreak
- \vskip 40\p@
- }}
-
-\renewcommand\section{\@startsection{section}{1}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {12\p@ \@plus 4\p@ \@minus 4\p@}%
- {\normalfont\large\bfseries\boldmath
- \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsection{\@startsection{subsection}{2}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {8\p@ \@plus 4\p@ \@minus 4\p@}%
- {\normalfont\normalsize\bfseries\boldmath
- \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {-0.5em \@plus -0.22em \@minus -0.1em}%
- {\normalfont\normalsize\bfseries\boldmath}}
-\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}%
- {-12\p@ \@plus -4\p@ \@minus -4\p@}%
- {-0.5em \@plus -0.22em \@minus -0.1em}%
- {\normalfont\normalsize\itshape}}
-\renewcommand\subparagraph[1]{\typeout{LLNCS warning: You should not use
- \string\subparagraph\space with this class}\vskip0.5cm
-You should not use \verb|\subparagraph| with this class.\vskip0.5cm}
-
-\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00}
-\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01}
-\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02}
-\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03}
-\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04}
-\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05}
-\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06}
-\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07}
-\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08}
-\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09}
-\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A}
-
-\let\footnotesize\small
-
-\if@custvec
-\def\vec#1{\mathchoice{\mbox{\boldmath$\displaystyle#1$}}
-{\mbox{\boldmath$\textstyle#1$}}
-{\mbox{\boldmath$\scriptstyle#1$}}
-{\mbox{\boldmath$\scriptscriptstyle#1$}}}
-\fi
-
-\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}}
-\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil
-\penalty50\hskip1em\null\nobreak\hfil\squareforqed
-\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi}
-
-\def\getsto{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr\gets\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-\gets\cr\to\cr}}}}}
-\def\lid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr<\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr<\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr<\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-<\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\gid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr>\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr>\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\grole{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr
->\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.8pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.3pt}<\cr}}}}}
-\def\bbbr{{\rm I\!R}} %reelle Zahlen
-\def\bbbm{{\rm I\!M}}
-\def\bbbn{{\rm I\!N}} %natuerliche Zahlen
-\def\bbbf{{\rm I\!F}}
-\def\bbbh{{\rm I\!H}}
-\def\bbbk{{\rm I\!K}}
-\def\bbbp{{\rm I\!P}}
-\def\bbbone{{\mathchoice {\rm 1\mskip-4mu l} {\rm 1\mskip-4mu l}
-{\rm 1\mskip-4.5mu l} {\rm 1\mskip-5mu l}}}
-\def\bbbc{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbq{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}}}
-\def\bbbt{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-T$}\hbox{\hbox to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbs{{\mathchoice
-{\setbox0=\hbox{$\displaystyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.5\wd0\vrule height0.45\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.4\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.55\wd0\vrule height0.45\ht0\hss}\box0}}}}
-\def\bbbz{{\mathchoice {\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\scriptstyle Z\kern-0.3em Z$}}
-{\hbox{$\mathsf\scriptscriptstyle Z\kern-0.2em Z$}}}}
-
-\let\ts\,
-
-\setlength\leftmargini {17\p@}
-\setlength\leftmargin {\leftmargini}
-\setlength\leftmarginii {\leftmargini}
-\setlength\leftmarginiii {\leftmargini}
-\setlength\leftmarginiv {\leftmargini}
-\setlength \labelsep {.5em}
-\setlength \labelwidth{\leftmargini}
-\addtolength\labelwidth{-\labelsep}
-
-\def\@listI{\leftmargin\leftmargini
- \parsep 0\p@ \@plus1\p@ \@minus\p@
- \topsep 8\p@ \@plus2\p@ \@minus4\p@
- \itemsep0\p@}
-\let\@listi\@listI
-\@listi
-\def\@listii {\leftmargin\leftmarginii
- \labelwidth\leftmarginii
- \advance\labelwidth-\labelsep
- \topsep 0\p@ \@plus2\p@ \@minus\p@}
-\def\@listiii{\leftmargin\leftmarginiii
- \labelwidth\leftmarginiii
- \advance\labelwidth-\labelsep
- \topsep 0\p@ \@plus\p@\@minus\p@
- \parsep \z@
- \partopsep \p@ \@plus\z@ \@minus\p@}
-
-\renewcommand\labelitemi{\normalfont\bfseries --}
-\renewcommand\labelitemii{$\m@th\bullet$}
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\def\tableofcontents{\chapter*{\contentsname\@mkboth{{\contentsname}}%
- {{\contentsname}}}
- \def\authcount##1{\setcounter{auco}{##1}\setcounter{@auth}{1}}
- \def\lastand{\ifnum\value{auco}=2\relax
- \unskip{} \andname\
- \else
- \unskip \lastandname\
- \fi}%
- \def\and{\stepcounter{@auth}\relax
- \ifnum\value{@auth}=\value{auco}%
- \lastand
- \else
- \unskip,
- \fi}%
- \@starttoc{toc}\if@restonecol\twocolumn\fi}
-
-\def\l@part#1#2{\addpenalty{\@secpenalty}%
- \addvspace{2em plus\p@}% % space above part line
- \begingroup
- \parindent \z@
- \rightskip \z@ plus 5em
- \hrule\vskip5pt
- \large % same size as for a contribution heading
- \bfseries\boldmath % set line in boldface
- \leavevmode % TeX command to enter horizontal mode.
- #1\par
- \vskip5pt
- \hrule
- \vskip1pt
- \nobreak % Never break after part entry
- \endgroup}
-
-\def\@dotsep{2}
-
-\let\phantomsection=\relax
-
-\def\hyperhrefextend{\ifx\hyper@anchor\@undefined\else
-{}\fi}
-
-\def\addnumcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{\protect\numberline
- {\thechapter}#3}{\thepage}\hyperhrefextend}}%
-\def\addcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{\thepage}\hyperhrefextend}}%
-\def\addcontentsmarkwop#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{0}\hyperhrefextend}}%
-
-\def\@adcmk[#1]{\ifcase #1 \or
-\def\@gtempa{\addnumcontentsmark}%
- \or \def\@gtempa{\addcontentsmark}%
- \or \def\@gtempa{\addcontentsmarkwop}%
- \fi\@gtempa{toc}{chapter}%
-}
-\def\addtocmark{%
-\phantomsection
-\@ifnextchar[{\@adcmk}{\@adcmk[3]}%
-}
-
-\def\l@chapter#1#2{\addpenalty{-\@highpenalty}
- \vskip 1.0em plus 1pt \@tempdima 1.5em \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- {\large\bfseries\boldmath#1}\ifx0#2\hfil\null
- \else
- \nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}%
- \fi\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@title#1#2{\addpenalty{-\@highpenalty}
- \addvspace{8pt plus 1pt}
- \@tempdima \z@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- #1\nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@author#1#2{\addpenalty{\@highpenalty}
- \@tempdima=15\p@ %\z@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima %\hskip -\leftskip
- \textit{#1}\par
- \penalty\@highpenalty \endgroup}
-
-\setcounter{tocdepth}{0}
-\newdimen\tocchpnum
-\newdimen\tocsecnum
-\newdimen\tocsectotal
-\newdimen\tocsubsecnum
-\newdimen\tocsubsectotal
-\newdimen\tocsubsubsecnum
-\newdimen\tocsubsubsectotal
-\newdimen\tocparanum
-\newdimen\tocparatotal
-\newdimen\tocsubparanum
-\tocchpnum=\z@ % no chapter numbers
-\tocsecnum=15\p@ % section 88. plus 2.222pt
-\tocsubsecnum=23\p@ % subsection 88.8 plus 2.222pt
-\tocsubsubsecnum=27\p@ % subsubsection 88.8.8 plus 1.444pt
-\tocparanum=35\p@ % paragraph 88.8.8.8 plus 1.666pt
-\tocsubparanum=43\p@ % subparagraph 88.8.8.8.8 plus 1.888pt
-\def\calctocindent{%
-\tocsectotal=\tocchpnum
-\advance\tocsectotal by\tocsecnum
-\tocsubsectotal=\tocsectotal
-\advance\tocsubsectotal by\tocsubsecnum
-\tocsubsubsectotal=\tocsubsectotal
-\advance\tocsubsubsectotal by\tocsubsubsecnum
-\tocparatotal=\tocsubsubsectotal
-\advance\tocparatotal by\tocparanum}
-\calctocindent
-
-\def\l@section{\@dottedtocline{1}{\tocchpnum}{\tocsecnum}}
-\def\l@subsection{\@dottedtocline{2}{\tocsectotal}{\tocsubsecnum}}
-\def\l@subsubsection{\@dottedtocline{3}{\tocsubsectotal}{\tocsubsubsecnum}}
-\def\l@paragraph{\@dottedtocline{4}{\tocsubsubsectotal}{\tocparanum}}
-\def\l@subparagraph{\@dottedtocline{5}{\tocparatotal}{\tocsubparanum}}
-
-\def\listoffigures{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}}
- \@starttoc{lof}\if@restonecol\twocolumn\fi}
-\def\l@figure{\@dottedtocline{1}{0em}{1.5em}}
-
-\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}}
- \@starttoc{lot}\if@restonecol\twocolumn\fi}
-\let\l@table\l@figure
-
-\renewcommand\listoffigures{%
- \section*{\listfigurename
- \@mkboth{\listfigurename}{\listfigurename}}%
- \@starttoc{lof}%
- }
-
-\renewcommand\listoftables{%
- \section*{\listtablename
- \@mkboth{\listtablename}{\listtablename}}%
- \@starttoc{lot}%
- }
-
-\ifx\oribibl\undefined
-\ifx\citeauthoryear\undefined
-\renewenvironment{thebibliography}[1]
- {\section*{\refname}
- \def\@biblabel##1{##1.}
- \small
- \list{\@biblabel{\@arabic\c@enumiv}}%
- {\settowidth\labelwidth{\@biblabel{#1}}%
- \leftmargin\labelwidth
- \advance\leftmargin\labelsep
- \if@openbib
- \advance\leftmargin\bibindent
- \itemindent -\bibindent
- \listparindent \itemindent
- \parsep \z@
- \fi
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{\@arabic\c@enumiv}}%
- \if@openbib
- \renewcommand\newblock{\par}%
- \else
- \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
- \fi
- \sloppy\clubpenalty4000\widowpenalty4000%
- \sfcode`\.=\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
-\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw
- {\let\protect\noexpand\immediate
- \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
-\newcount\@tempcntc
-\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi
- \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do
- {\@ifundefined
- {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries
- ?}\@warning
- {Citation `\@citeb' on page \thepage \space undefined}}%
- {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}%
- \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne
- \@citea\def\@citea{,}\hbox{\csname b@\@citeb\endcsname}%
- \else
- \advance\@tempcntb\@ne
- \ifnum\@tempcntb=\@tempcntc
- \else\advance\@tempcntb\m@ne\@citeo
- \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}}
-\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else
- \@citea\def\@citea{,\,\hskip\z@skip}%
- \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else
- {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else
- \def\@citea{--}\fi
- \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi}
-\else
-\renewenvironment{thebibliography}[1]
- {\section*{\refname}
- \small
- \list{}%
- {\settowidth\labelwidth{}%
- \leftmargin\parindent
- \itemindent=-\parindent
- \labelsep=\z@
- \if@openbib
- \advance\leftmargin\bibindent
- \itemindent -\bibindent
- \listparindent \itemindent
- \parsep \z@
- \fi
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{}}%
- \if@openbib
- \renewcommand\newblock{\par}%
- \else
- \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
- \fi
- \sloppy\clubpenalty4000\widowpenalty4000%
- \sfcode`\.=\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
- \def\@cite#1{#1}%
- \def\@lbibitem[#1]#2{\item[]\if@filesw
- {\def\protect##1{\string ##1\space}\immediate
- \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
- \fi
-\else
-\@cons\@openbib@code{\noexpand\small}
-\fi
-
-\def\idxquad{\hskip 10\p@}% space that divides entry from number
-
-\def\@idxitem{\par\hangindent 10\p@}
-
-\def\subitem{\par\setbox0=\hbox{--\enspace}% second order
- \noindent\hangindent\wd0\box0}% index entry
-
-\def\subsubitem{\par\setbox0=\hbox{--\,--\enspace}% third
- \noindent\hangindent\wd0\box0}% order index entry
-
-\def\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax}
-
-\renewenvironment{theindex}
- {\@mkboth{\indexname}{\indexname}%
- \thispagestyle{empty}\parindent\z@
- \parskip\z@ \@plus .3\p@\relax
- \let\item\par
- \def\,{\relax\ifmmode\mskip\thinmuskip
- \else\hskip0.2em\ignorespaces\fi}%
- \normalfont\small
- \begin{multicols}{2}[\@makeschapterhead{\indexname}]%
- }
- {\end{multicols}}
-
-\renewcommand\footnoterule{%
- \kern-3\p@
- \hrule\@width 2truecm
- \kern2.6\p@}
- \newdimen\fnindent
- \fnindent1em
-\long\def\@makefntext#1{%
- \parindent \fnindent%
- \leftskip \fnindent%
- \noindent
- \llap{\hb@xt@1em{\hss\@makefnmark\ }}\ignorespaces#1}
-
-\long\def\@makecaption#1#2{%
- \small
- \vskip\abovecaptionskip
- \sbox\@tempboxa{{\bfseries #1.} #2}%
- \ifdim \wd\@tempboxa >\hsize
- {\bfseries #1.} #2\par
- \else
- \global \@minipagefalse
- \hb@xt@\hsize{\hfil\box\@tempboxa\hfil}%
- \fi
- \vskip\belowcaptionskip}
-
-\def\fps@figure{htbp}
-\def\fnum@figure{\figurename\thinspace\thefigure}
-\def \@floatboxreset {%
- \reset@font
- \small
- \@setnobreak
- \@setminipage
-}
-\def\fps@table{htbp}
-\def\fnum@table{\tablename~\thetable}
-\renewenvironment{table}
- {\setlength\abovecaptionskip{0\p@}%
- \setlength\belowcaptionskip{10\p@}%
- \@float{table}}
- {\end@float}
-\renewenvironment{table*}
- {\setlength\abovecaptionskip{0\p@}%
- \setlength\belowcaptionskip{10\p@}%
- \@dblfloat{table}}
- {\end@dblfloat}
-
-\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname
- ext@#1\endcsname}{#1}{\protect\numberline{\csname
- the#1\endcsname}{\ignorespaces #2}}\begingroup
- \@parboxrestore
- \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
- \endgroup}
-
-% LaTeX does not provide a command to enter the authors institute
-% addresses. The \institute command is defined here.
-
-\newcounter{@inst}
-\newcounter{@auth}
-\newcounter{auco}
-\newdimen\instindent
-\newbox\authrun
-\newtoks\authorrunning
-\newtoks\tocauthor
-\newbox\titrun
-\newtoks\titlerunning
-\newtoks\toctitle
-
-\def\clearheadinfo{\gdef\@author{No Author Given}%
- \gdef\@title{No Title Given}%
- \gdef\@subtitle{}%
- \gdef\@institute{No Institute Given}%
- \gdef\@thanks{}%
- \global\titlerunning={}\global\authorrunning={}%
- \global\toctitle={}\global\tocauthor={}}
-
-\def\institute#1{\gdef\@institute{#1}}
-
-\def\institutename{\par
- \begingroup
- \parskip=\z@
- \parindent=\z@
- \setcounter{@inst}{1}%
- \def\and{\par\stepcounter{@inst}%
- \noindent$^{\the@inst}$\enspace\ignorespaces}%
- \setbox0=\vbox{\def\thanks##1{}\@institute}%
- \ifnum\c@@inst=1\relax
- \gdef\fnnstart{0}%
- \else
- \xdef\fnnstart{\c@@inst}%
- \setcounter{@inst}{1}%
- \noindent$^{\the@inst}$\enspace
- \fi
- \ignorespaces
- \@institute\par
- \endgroup}
-
-\def\@fnsymbol#1{\ensuremath{\ifcase#1\or\star\or{\star\star}\or
- {\star\star\star}\or \dagger\or \ddagger\or
- \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger
- \or \ddagger\ddagger \else\@ctrerr\fi}}
-
-\def\inst#1{\unskip$^{#1}$}
-\def\fnmsep{\unskip$^,$}
-\def\email#1{{\tt#1}}
-\AtBeginDocument{\@ifundefined{url}{\def\url#1{#1}}{}%
-\@ifpackageloaded{babel}{%
-\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}%
-\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}%
-\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}%
-}{\switcht@@therlang}%
-\providecommand{\keywords}[1]{\par\addvspace\baselineskip
-\noindent\keywordname\enspace\ignorespaces#1}%
-}
-\def\homedir{\~{ }}
-
-\def\subtitle#1{\gdef\@subtitle{#1}}
-\clearheadinfo
-%
-%%% to avoid hyperref warnings
-\providecommand*{\toclevel@author}{999}
-%%% to make title-entry parent of section-entries
-\providecommand*{\toclevel@title}{0}
-%
-\renewcommand\maketitle{\newpage
-\phantomsection
- \refstepcounter{chapter}%
- \stepcounter{section}%
- \setcounter{section}{0}%
- \setcounter{subsection}{0}%
- \setcounter{figure}{0}
- \setcounter{table}{0}
- \setcounter{equation}{0}
- \setcounter{footnote}{0}%
- \begingroup
- \parindent=\z@
- \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
- \if@twocolumn
- \ifnum \col@number=\@ne
- \@maketitle
- \else
- \twocolumn[\@maketitle]%
- \fi
- \else
- \newpage
- \global\@topnum\z@ % Prevents figures from going at top of page.
- \@maketitle
- \fi
- \thispagestyle{empty}\@thanks
-%
- \def\\{\unskip\ \ignorespaces}\def\inst##1{\unskip{}}%
- \def\thanks##1{\unskip{}}\def\fnmsep{\unskip}%
- \instindent=\hsize
- \advance\instindent by-\headlineindent
- \if!\the\toctitle!\addcontentsline{toc}{title}{\@title}\else
- \addcontentsline{toc}{title}{\the\toctitle}\fi
- \if@runhead
- \if!\the\titlerunning!\else
- \edef\@title{\the\titlerunning}%
- \fi
- \global\setbox\titrun=\hbox{\small\rm\unboldmath\ignorespaces\@title}%
- \ifdim\wd\titrun>\instindent
- \typeout{Title too long for running head. Please supply}%
- \typeout{a shorter form with \string\titlerunning\space prior to
- \string\maketitle}%
- \global\setbox\titrun=\hbox{\small\rm
- Title Suppressed Due to Excessive Length}%
- \fi
- \xdef\@title{\copy\titrun}%
- \fi
-%
- \if!\the\tocauthor!\relax
- {\def\and{\noexpand\protect\noexpand\and}%
- \protected@xdef\toc@uthor{\@author}}%
- \else
- \def\\{\noexpand\protect\noexpand\newline}%
- \protected@xdef\scratch{\the\tocauthor}%
- \protected@xdef\toc@uthor{\scratch}%
- \fi
- \addtocontents{toc}{\noexpand\protect\noexpand\authcount{\the\c@auco}}%
- \addcontentsline{toc}{author}{\toc@uthor}%
- \if@runhead
- \if!\the\authorrunning!
- \value{@inst}=\value{@auth}%
- \setcounter{@auth}{1}%
- \else
- \edef\@author{\the\authorrunning}%
- \fi
- \global\setbox\authrun=\hbox{\small\unboldmath\@author\unskip}%
- \ifdim\wd\authrun>\instindent
- \typeout{Names of authors too long for running head. Please supply}%
- \typeout{a shorter form with \string\authorrunning\space prior to
- \string\maketitle}%
- \global\setbox\authrun=\hbox{\small\rm
- Authors Suppressed Due to Excessive Length}%
- \fi
- \xdef\@author{\copy\authrun}%
- \markboth{\@author}{\@title}%
- \fi
- \endgroup
- \setcounter{footnote}{\fnnstart}%
- \clearheadinfo}
-%
-\def\@maketitle{\newpage
- \markboth{}{}%
- \def\lastand{\ifnum\value{@inst}=2\relax
- \unskip{} \andname\
- \else
- \unskip \lastandname\
- \fi}%
- \def\and{\stepcounter{@auth}\relax
- \ifnum\value{@auth}=\value{@inst}%
- \lastand
- \else
- \unskip,
- \fi}%
- \begin{center}%
- \let\newline\\
- {\Large \bfseries\boldmath
- \pretolerance=10000
- \@title \par}\vskip .8cm
-\if!\@subtitle!\else {\large \bfseries\boldmath
- \vskip -.65cm
- \pretolerance=10000
- \@subtitle \par}\vskip .8cm\fi
- \setbox0=\vbox{\setcounter{@auth}{1}\def\and{\stepcounter{@auth}}%
- \def\thanks##1{}\@author}%
- \global\value{@inst}=\value{@auth}%
- \global\value{auco}=\value{@auth}%
- \setcounter{@auth}{1}%
-{\lineskip .5em
-\noindent\ignorespaces
-\@author\vskip.35cm}
- {\small\institutename}
- \end{center}%
- }
-
-% definition of the "\spnewtheorem" command.
-%
-% Usage:
-%
-% \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font}
-% or \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font}
-% or \spnewtheorem*{env_nam}{caption}{cap_font}{body_font}
-%
-% New is "cap_font" and "body_font". It stands for
-% fontdefinition of the caption and the text itself.
-%
-% "\spnewtheorem*" gives a theorem without number.
-%
-% A defined spnewthoerem environment is used as described
-% by Lamport.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\def\@thmcountersep{}
-\def\@thmcounterend{.}
-
-\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}}
-
-% definition of \spnewtheorem with number
-
-\def\@spnthm#1#2{%
- \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}}
-\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}}
-
-\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}\@addtoreset{#1}{#3}%
- \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand
- \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}%
- \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spothm#1[#2]#3#4#5{%
- \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}%
- {\expandafter\@ifdefinable\csname #1\endcsname
- {\newaliascnt{#1}{#2}%
- \expandafter\xdef\csname #1name\endcsname{#3}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}}
-
-\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\refstepcounter{#1}%
-\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}}
-
-\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}%
- \ignorespaces}
-
-\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname
- the#1\endcsname}{#5}{#3}{#4}\ignorespaces}
-
-\def\@spbegintheorem#1#2#3#4{\trivlist
- \item[\hskip\labelsep{#3#1\ #2\@thmcounterend}]#4}
-
-\def\@spopargbegintheorem#1#2#3#4#5{\trivlist
- \item[\hskip\labelsep{#4#1\ #2}]{#4(#3)\@thmcounterend\ }#5}
-
-% definition of \spnewtheorem* without number
-
-\def\@sthm#1#2{\@Ynthm{#1}{#2}}
-
-\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}}
-
-\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces}
-
-\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1}
- {#4}{#2}{#3}\ignorespaces}
-
-\def\@Begintheorem#1#2#3{#3\trivlist
- \item[\hskip\labelsep{#2#1\@thmcounterend}]}
-
-\def\@Opargbegintheorem#1#2#3#4{#4\trivlist
- \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }}
-
-\if@envcntsect
- \def\@thmcountersep{.}
- \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape}
-\else
- \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape}
- \if@envcntreset
- \@addtoreset{theorem}{section}
- \else
- \@addtoreset{theorem}{chapter}
- \fi
-\fi
-
-%definition of divers theorem environments
-\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily}
-\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily}
-\if@envcntsame % alle Umgebungen wie Theorem.
- \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}}
-\else % alle Umgebungen mit eigenem Zaehler
- \if@envcntsect % mit section numeriert
- \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}}
- \else % nicht mit section numeriert
- \if@envcntreset
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{section}}
- \else
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{chapter}}%
- \fi
- \fi
-\fi
-\spn@wtheorem{case}{Case}{\itshape}{\rmfamily}
-\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily}
-\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape}
-\spn@wtheorem{definition}{Definition}{\bfseries}{\itshape}
-\spn@wtheorem{example}{Example}{\itshape}{\rmfamily}
-\spn@wtheorem{exercise}{Exercise}{\itshape}{\rmfamily}
-\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape}
-\spn@wtheorem{note}{Note}{\itshape}{\rmfamily}
-\spn@wtheorem{problem}{Problem}{\itshape}{\rmfamily}
-\spn@wtheorem{property}{Property}{\itshape}{\rmfamily}
-\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape}
-\spn@wtheorem{question}{Question}{\itshape}{\rmfamily}
-\spn@wtheorem{solution}{Solution}{\itshape}{\rmfamily}
-\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily}
-
-\def\@takefromreset#1#2{%
- \def\@tempa{#1}%
- \let\@tempd\@elt
- \def\@elt##1{%
- \def\@tempb{##1}%
- \ifx\@tempa\@tempb\else
- \@addtoreset{##1}{#2}%
- \fi}%
- \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
- \expandafter\def\csname cl@#2\endcsname{}%
- \@tempc
- \let\@elt\@tempd}
-
-\def\theopargself{\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5}
- \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
- \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }}
- }
-
-\renewenvironment{abstract}{%
- \list{}{\advance\topsep by0.35cm\relax\small
- \leftmargin=1cm
- \labelwidth=\z@
- \listparindent=\z@
- \itemindent\listparindent
- \rightmargin\leftmargin}\item[\hskip\labelsep
- \bfseries\abstractname]}
- {\endlist}
-
-\newdimen\headlineindent % dimension for space between
-\headlineindent=1.166cm % number and text of headings.
-
-\def\ps@headings{\let\@mkboth\@gobbletwo
- \let\@oddfoot\@empty\let\@evenfoot\@empty
- \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
- \leftmark\hfil}
- \def\@oddhead{\normalfont\small\hfil\rightmark\hspace{\headlineindent}%
- \llap{\thepage}}
- \def\chaptermark##1{}%
- \def\sectionmark##1{}%
- \def\subsectionmark##1{}}
-
-\def\ps@titlepage{\let\@mkboth\@gobbletwo
- \let\@oddfoot\@empty\let\@evenfoot\@empty
- \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
- \hfil}
- \def\@oddhead{\normalfont\small\hfil\hspace{\headlineindent}%
- \llap{\thepage}}
- \def\chaptermark##1{}%
- \def\sectionmark##1{}%
- \def\subsectionmark##1{}}
-
-\if@runhead\ps@headings\else
-\ps@empty\fi
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\endinput
-%end of file llncs.cls
--- a/LMCS-Paper/document/lmcs.cls Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,512 +0,0 @@
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%% Modification log
-%%
-%% 2004/03/25 v0.1 based on amsart.cls, inspired by jair.sty
-%% 2004/09/01 v0.2 based on amsart.cls
-%% 2004/10/12 v0.3 based on amsart.cls
-%% 2004/12/16 v0.4 based on amsart.cls
-%% 2005/01/24 v0.5 based on amsart.cls
-%% 2005/03/10 v0.6 based on amsart.cls
-%% 2006/07/24 v0.7 based on amsart.cls
-%% 2007/11/13 v0.8 based on amsart.cls
-%%
-%% Juergen Koslowski, Stefan Milius
-%%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\NeedsTeXFormat{LaTeX2e}
-\ProvidesClass{lmcs}
- [2009/05/25 v0.8 LMCS Layout Editor Class]
-\DeclareOption*{\PassOptionsToClass{\CurrentOption}{amsart}}
-\ProcessOptions\relax
-
-\LoadClass[11pt,reqno]{amsart}
-\usepackage{helvet,cclicenses}
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Use of this class, cf. also lmcs-smp.tex
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This class builds upon the amsart class of AMS-LaTeX and requires use
-% of LaTeX 2e.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Start of the paper
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% \documentclass{lmcs}
-%
-% without any options followed optionally by
-%
-% \usepackage{package1,package2,...}
-%
-% loading additional macro packages you may wish to use, (eg, xypic, etc.)
-% (This also is the place to define further theorem-environments, in case
-% those provided by default do not suffice, cf. below.)
-%
-% \begin{document}
-% \title[short_title]{real title}
-%
-% and a list of author information of the form
-%
-% \author[short_author1]{Author 1}
-% \address{address 1}
-% \email{author1@email1}
-% \thanks{thanks 1}
-%
-% \author[short_author2]{Author 2}
-% \address{address 2}
-% \email{author2@email2}
-% \thanks{thanks 2}
-%
-% \author[short_author3]{Author 3}
-% \address{address 3}
-% \email{author3@email3}
-% \thanks{thanks 3}
-%
-% The \email and \thanks fields are optional. The \thanks fields appear
-% in footnotes on the title page, the addresses and email information
-% is relegated to the end of the article. The optional arguments to
-% the \title and \author macros determine a running head on the odd
-% and even pages, respectively.
-%
-% Lists of keywords and phrases as well as an ACM Subject
-% classification are mandatory; these appear in footnotes on the
-% title page, preceeding any \thanks fields.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Body of the paper
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% We encourage the use of LaTeX's crossreferencing capabilities with the
-% \label and \ref commands, for sections, subsections, theorems etc., and
-% displayed equations and figures.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Theorems, Definitions etc.
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% The following theorem-like environments are available. The
-% numbering is consecutive within sections.
-%
-% thm Theorem
-% cor Corollary
-% lem Lemma
-% prop Proposition
-% asm Asumption
-%
-% defi Definition
-% rem Remark
-% rems Remarks (intended for use with itemized remarks)
-% exa Example
-% exas Examples (intended for use with itemized examples)
-% conj Conjecture
-% prob Problem
-% oprob Open Problem
-% algo Algorithm
-% obs Observation
-%
-% If you require additional environments, you can add them before
-% \begin{document} by means of
-%
-% \theoremstyle{plain}\newtheorem{env}[thm]{Environment}
-%
-% or
-%
-% \theoremstyle{definition}\newtheorem{env}[thm]{Environment}
-%
-% In the first case the font within the new environment will be italicised.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Proofs
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% Proofs start with the command \proof and should be ended by \qed,
-% which provides an end-of-proof box at the right margin:
-%
-% \proof ... \qed
-%
-% In itemized or enumerated proofs the \qed command has to occur BEFORE
-% \end{itemize} or \end{enumerate} to ensure proper placement of the box:
-%
-% \proof
-% \begin{itemize}
-% \item[(1)] ...
-% \item[(2)] ...
-% ...
-% \item[(n)] ... \qed
-% \end{itemize}
-%
-% Similarly, the box may be used within theorem environments, when no
-% explicit proof is given:
-%
-% \begin{thm} ... \qed \end{thm}
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Itemized or enumerated environments and proofs
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% By default, the first item of an itemized (or enumerated) environment
-% or proof appears inlined on the same line as the environment title.
-% This can be prevented by placing \hfill before the itemization, e.g.:
-%
-% \begin{thm}\label{T:abc}\hfill
-% \begin{itemize} ...
-%
-% \proof\hfill
-% \begin{itemize} ...
-
-
-%
-% End of the paper
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% Acknowledgements should be placed in a non-numbered section:
-%
-% \section*{Acknowledgement}
-%
-% The bibliography uses alpha.bst where references are built from the
-% authors' initials and the year of publication. The use of bibtex
-% for creating the bibliography is strongly encouraged. Then the
-% end of the paper takes the form
-%
-% \begin{thebibliography}{key}
-% \end{thebibliography}
-% \end{document}
-%
-% where ``key'' is the longest alphanumerical key expected to occur.
-%
-% Optionally, appendices can be inserted after the bibliography by
-% means of
-%
-% \end{thebibliography}
-% \appendix
-% \section{} % Appendix A
-% \section{} % Appendix B
-% % etc.
-% \end{document}
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%%% actual macros
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\count255=\the\catcode`\@ \catcode`\@=11 \edef\catc@de{\the\count255}
-
-\newif\ifsuPer \suPertrue
-\def\rsuper#1{\ifsuPer${\,}^{\MakeLowercase #1}$\fi}%
-\def\lsuper#1{\ \hskip-2 pt\ifsuPer\llap{${}^{\MakeLowercase #1}\ $\fi}}%
-
-\def\titlecomment#1{\def\@titlecomment{#1}}
-\let\@titlecomment=\@empty
-\renewcommand{\sfdefault}{phv}
-\renewcommand*\subjclass[2][1991]{%
- \def\@subjclass{#2}%
- \@ifundefined{subjclassname@#1}{%
- \ClassWarning{\@classname}{Unknown edition (#1) of ACM
- Subject Classification; using '1991'.}%
- }{%
- \@xp\let\@xp\subjclassname\csname subjclassname@1998\endcsname
- }%
-}
-\@namedef{subjclassname@1998}{1998 ACM Subject Classification}
-\newcommand{\revisionname}{Revision Note}
-\newbox\revisionbox
-\newenvironment{revision}{%
- \ifx\maketitle\relax
- \ClassWarning{\@classname}{Revision should precede
- \protect\maketitle\space in LMCS documentclasses; reported}%
- \fi
- \global\setbox\revisionbox=\vtop \bgroup
- \normalfont\Small
- \list{}{\labelwidth\z@
- \leftmargin3pc \rightmargin\leftmargin
- \listparindent\normalparindent \itemindent\z@
- \parsep\z@ \@plus\p@
- \let\fullwidthdisplay\relax
- }%
- \item[\hskip\labelsep\scshape\revisionname.]%
-}{%
- \endlist\egroup
- \ifx\@setrevision\relax \@setrevisiona \fi
-}
-\def\@setrevision{\@setrevisiona \global\let\@setrevision\relax}
-\def\@setrevisiona{%
- \ifvoid\revisionbox
- \else
- \skip@20\p@ \advance\skip@-\lastskip
- \advance\skip@-\baselineskip \vskip\skip@
- \box\revisionbox
- \prevdepth\z@ % because \revisionbox is a vtop
- \bigskip\hrule\medskip
- \fi
-}
-\def\@setsubjclass{%
- {\itshape\subjclassname:}\enspace\@subjclass\@addpunct.}
-\def\@setkeywords{%
- {\itshape \keywordsname:}\enspace \@keywords\@addpunct.}
-\def\@settitlecomment{\@titlecomment\@addpunct.}
-\def\@maketitle{%
- \normalfont\normalsize
- \let\@makefnmark\relax \let\@thefnmark\relax
- \ifx\@empty\@date\else \@footnotetext{\@setdate}\fi
- \ifx\@empty\@subjclass\else \@footnotetext{\@setsubjclass}\fi
- \ifx\@empty\@keywords\else \@footnotetext{\@setkeywords}\fi
- \ifx\@empty\@titlecomment\else \@footnotetext{\@settitlecomment}\fi
- \ifx\@empty\thankses\else \@footnotetext{%
- \def\par{\let\par\@par}\@setthanks\par}\fi
- \@mkboth{\@nx\shortauthors}{\@nx\shorttitle}%
- \global\topskip12\p@\relax % 5.5pc " " " " "
- \topskip42 pt\@settitle
- \ifx\@empty\authors \else \@setauthors \fi
- \@setaddresses
- \ifx\@empty\@dedicatory
- \else
- \baselineskip18\p@
- \vtop{\centering{\footnotesize\itshape\@dedicatory\@@par}%
- \global\dimen@i\prevdepth}\prevdepth\dimen@i
- \fi
- \endfront@text
- \bigskip\hrule\medskip
- \@setrevision
- \@setabstract
- \vskip-\bigskipamount
- \normalsize
- \if@titlepage
- \newpage
- \else
- \dimen@34\p@ \advance\dimen@-\baselineskip
- \vskip\dimen@\relax
- \fi
-}
-\def\@setaddresses{\par
- \nobreak \begingroup
-\footnotesize
- \def\author##1{\nobreak\addvspace\bigskipamount}%
- \def\\{\unskip, \ignorespaces}%
- \interlinepenalty\@M
- \def\address##1##2{\begingroup
- \par\addvspace\bigskipamount\noindent\narrower
- \@ifnotempty{##1}{(\ignorespaces##1\unskip) }%
- {\ignorespaces##2}\par\endgroup}%
- \def\curraddr##1##2{\begingroup
- \@ifnotempty{##2}{\nobreak\indent{\itshape Current address}%
- \@ifnotempty{##1}{, \ignorespaces##1\unskip}\/:\space
- ##2\par}\endgroup}%
- \def\email##1##2{\begingroup
- \@ifnotempty{##2}{\nobreak\indent{\itshape e-mail address}%
- \@ifnotempty{##1}{, \ignorespaces##1\unskip}\/:\space
- {##2}\par}\endgroup}%
- \def\urladdr##1##2{\begingroup
- \@ifnotempty{##2}{\nobreak\indent{\itshape URL}%
- \@ifnotempty{##1}{, \ignorespaces##1\unskip}\/:\space
- \ttfamily##2\par}\endgroup}%
- \addresses
- \endgroup
-}
-\copyrightinfo{}{}
-
-\newinsert\copyins
-\skip\copyins=3pc
-\count\copyins=1000 % magnification factor, 1000 = 100%
-\dimen\copyins=.5\textheight % maximum allowed per page
-
-\renewcommand{\topfraction}{0.95} % let figure take up nearly whole page
-\renewcommand{\textfraction}{0.05} % let figure take up nearly whole page
-
-%% Specify the dimensions of each page
-
-\setlength{\oddsidemargin}{.25 in} % Note \oddsidemargin = \evensidemargin
-\setlength{\evensidemargin}{.25 in}
-\setlength{\marginparwidth}{0.07 true in}
-\setlength{\topmargin}{-0.7 in}
-\addtolength{\headheight}{1.84 pt}
-\addtolength{\headsep}{0.25in}
-\addtolength{\voffset}{0.7 in}
-\setlength{\textheight}{8.5 true in} % Height of text (including footnotes & figures)
-\setlength{\textwidth}{6.0 true in} % Width of text line.
-\setlength{\parindent}{20 pt} % Width of text line.
-\widowpenalty=10000
-\clubpenalty=10000
-\@twosidetrue \@mparswitchtrue \def\ds@draft{\overfullrule 5pt}
-\raggedbottom
-
-%% Pagestyle
-
-%% Defines the pagestyle for the title page.
-%% Usage: \lmcsheading{vol}{issue}{year}{pages}{subm}{publ}{rev}{spec_iss}{title}
-
-\def\lmcsheading#1#2#3#4#5#6#7{\def\ps@firstpage{\let\@mkboth\@gobbletwo%
-\def\@oddhead{%
-\hbox{%
- \vbox to 30 pt{\scriptsize\vfill
- \hbox{\textsf{Logical Methods in Computer Science}\hfil}
- \hbox{\textsf{Vol.~? (?:?) 2???, ? pages}}
- \hbox{\textsf{www.lmcs-online.org}}
- \rlap{\vrule width\hsize depth .4 pt}}}\hfill
-\raise 4pt
-\hbox{%
- \vbox to 30 pt{\scriptsize\vfill
- \hbox{\textsf{}}
- \hbox{\textsf{}}}}\hfill
-\raise 4pt
-\hbox{%
- \vbox to 30 pt{\scriptsize\vfill
- \hbox to 94 pt{\textsf{Submitted\hfill date}}
- \hbox to 94 pt{\textsf{Published\hfill date}}
- }}}
-\def\@evenhead{}\def\@evenfoot{}}%
-\thispagestyle{firstpage}}
-
-\def\endfront@text{%
- \insert\copyins{\hsize\textwidth
- \fontsize{6}{7\p@}\normalfont\upshape
- \noindent
-\hbox{
- \vbox{\fontsize{6}{8 pt}\baselineskip=6 pt\vss
- \hbox{\hbox to 20 pt{\hfill}
- \textsf{LOGICAL METHODS}\hfil}
- \hbox{\hbox to 20 pt{\phantom{x}}
- \textsf{IN COMPUTER SCIENCE}}}}
-\hfill\textsf{DOI:10.2168/LMCS-???}
-%\hfill\textsf{\copyright\shortauthors}
-\hfill
-\hbox{
- \vbox{\fontsize{6}{8 pt}\baselineskip=6 pt\vss
- \hbox{\textsf{\,\,\copyright\quad \shortauthors}\hfil}
- \hbox{\textsf{Creative Commons}\hfil}}}
-\par\kern\z@}%
-}
-%\def\endfront@text{}
-
-\def\enddoc@text{}
-
-%% Defines the pagestyle for the rest of the pages
-%% Usage: \ShortHeadings{Minimizing Conflicts}{Minton et al}
-%% \ShortHeadings{short title}{short authors}
-
-%\def\firstpageno#1{\setcounter{page}{#1}}
-%\def\ShortHeadings#1#2{\def\ps@lmcsps{\let\@mkboth\@gobbletwo%
-%\def\@oddhead{\hfill {\small\sc #1} \hfill}%
-%\def\@oddfoot{\hfill \small\rm \thepage \hfill}%
-%\def\@evenhead{\hfill {\small\sc #2} \hfill}%
-%\def\@evenfoot{\hfill \small\rm \thepage \hfill}}%
-%\pagestyle{lmcsps}}
-
-%% MISCELLANY
-
-\def\@startsection#1#2#3#4#5#6{\bigskip%
- \if@noskipsec \leavevmode \fi
- \par \@tempskipa #4\relax
- \@afterindentfalse
- \ifdim \@tempskipa <\z@ \@tempskipa -\@tempskipa \@afterindentfalse\fi
- \if@nobreak \everypar{}\else
- \addpenalty\@secpenalty\addvspace\@tempskipa\fi
- \@ifstar{\@dblarg{\@sect{#1}{\@m}{#3}{#4}{#5}{#6}}}%
- {\@dblarg{\@sect{#1}{#2}{#3}{#4}{#5}{#6}}}%
-}
-
-\def\figurecaption#1#2{\noindent\hangindent 40pt
- \hbox to 36pt {\small\sl #1 \hfil}
- \ignorespaces {\small #2}}
-% Figurecenter prints the caption title centered.
-\def\figurecenter#1#2{\centerline{{\sl #1} #2}}
-\def\figurecenter#1#2{\centerline{{\small\sl #1} {\small #2}}}
-
-%
-% Allow ``hanging indents'' in long captions
-%
-\long\def\@makecaption#1#2{
- \vskip 10pt
- \setbox\@tempboxa\hbox{#1: #2}
- \ifdim \wd\@tempboxa >\hsize % IF longer than one line:
- \begin{list}{#1:}{
- \settowidth{\labelwidth}{#1:}
- \setlength{\leftmargin}{\labelwidth}
- \addtolength{\leftmargin}{\labelsep}
- }\item #2 \end{list}\par % Output in quote mode
- \else % ELSE center.
- \hbox to\hsize{\hfil\box\@tempboxa\hfil}
- \fi}
-
-
-% Define strut macros for skipping spaces above and below text in a
-% tabular environment.
-\def\abovestrut#1{\rule[0in]{0in}{#1}\ignorespaces}
-\def\belowstrut#1{\rule[-#1]{0in}{#1}\ignorespaces}
-
-%%% Theorem environments
-
-% the following environments switch to a slanted font:
-\theoremstyle{plain}
-
-\newtheorem{thm}{Theorem}[section]
-\newtheorem{cor}[thm]{Corollary}
-\newtheorem{lem}[thm]{Lemma}
-\newtheorem{prop}[thm]{Proposition}
-\newtheorem{asm}[thm]{Assumption}
-
-% the following environments keep the roman font:
-\theoremstyle{definition}
-
-\newtheorem{rem}[thm]{Remark}
-\newtheorem{rems}[thm]{Remarks}
-\newtheorem{exa}[thm]{Example}
-\newtheorem{exas}[thm]{Examples}
-\newtheorem{defi}[thm]{Definition}
-\newtheorem{conv}[thm]{Convention}
-\newtheorem{conj}[thm]{Conjecture}
-\newtheorem{prob}[thm]{Problem}
-\newtheorem{oprob}[thm]{Open Problem}
-\newtheorem{algo}[thm]{Algorithm}
-\newtheorem{obs}[thm]{Observation}
-\newtheorem{qu}[thm]{Question}
-\newtheorem{fact}[thm]{Fact}
-\newtheorem{pty}[thm]{Property}
-
-\numberwithin{equation}{section}
-
-% end-of-proof sign, to appear at right margin
-% Paul Taylor and Chris Thompson, Cambridge, 1986
-%
-\def\pushright#1{{% set up
- \parfillskip=0pt % so \par doesn't push #1 to left
- \widowpenalty=10000 % so we dont break the page before #1
- \displaywidowpenalty=10000 % ditto
- \finalhyphendemerits=0 % TeXbook exercise 14.32
- %
- % horizontal
- \leavevmode % \nobreak means lines not pages
- \unskip % remove previous space or glue
- \nobreak % don't break lines
- \hfil % ragged right if we spill over
- \penalty50 % discouragement to do so
- \hskip.2em % ensure some space
- \null % anchor following \hfill
- \hfill % push #1 to right
- {#1} % the end-of-proof mark (or whatever)
- %
- % vertical
- \par}} % build paragraph
-
-\def\qEd{{\lower1 pt\hbox{\vbox{\hrule\hbox{\vrule\kern4 pt
- \vbox{\kern4 pt\hbox{\hss}\kern4 pt}\kern4 pt\vrule}\hrule}}}}
-\def\qed{\pushright{\qEd}
- \penalty-700 \par\addvspace{\medskipamount}}
-
-\newenvironment{Proof}[1][\proofname]{\par
- \pushQED{\qed}%
- \normalfont \topsep6\p@\@plus6\p@\relax
- \trivlist
- \item[\hskip\labelsep
- \itshape
- #1]\ignorespaces
-}{%
- \popQED\endtrivlist\@endpefalse
-}
-\newenvironment{iteMize}[1]{\begin{enumerate}[#1]}{\end{enumerate}}
-\newenvironment{desCription}{\begin{enumerate}[\hbox to8 pt{\hfill}]}{\end{enumerate}}
-% Bibliographystyle
-
-\bibliographystyle{alpha}
-
-\endinput
--- a/LMCS-Paper/document/proof.sty Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,278 +0,0 @@
-% proof.sty (Proof Figure Macros)
-%
-% version 3.0 (for both LaTeX 2.09 and LaTeX 2e)
-% Mar 6, 1997
-% Copyright (C) 1990 -- 1997, Makoto Tatsuta (tatsuta@kusm.kyoto-u.ac.jp)
-%
-% This program is free software; you can redistribute it or modify
-% it under the terms of the GNU General Public License as published by
-% the Free Software Foundation; either versions 1, or (at your option)
-% any later version.
-%
-% This program 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.
-%
-% Usage:
-% In \documentstyle, specify an optional style `proof', say,
-% \documentstyle[proof]{article}.
-%
-% The following macros are available:
-%
-% In all the following macros, all the arguments such as
-% <Lowers> and <Uppers> are processed in math mode.
-%
-% \infer<Lower><Uppers>
-% draws an inference.
-%
-% Use & in <Uppers> to delimit upper formulae.
-% <Uppers> consists more than 0 formulae.
-%
-% \infer returns \hbox{ ... } or \vbox{ ... } and
-% sets \@LeftOffset and \@RightOffset globally.
-%
-% \infer[<Label>]<Lower><Uppers>
-% draws an inference labeled with <Label>.
-%
-% \infer*<Lower><Uppers>
-% draws a many step deduction.
-%
-% \infer*[<Label>]<Lower><Uppers>
-% draws a many step deduction labeled with <Label>.
-%
-% \infer=<Lower><Uppers>
-% draws a double-ruled deduction.
-%
-% \infer=[<Label>]<Lower><Uppers>
-% draws a double-ruled deduction labeled with <Label>.
-%
-% \deduce<Lower><Uppers>
-% draws an inference without a rule.
-%
-% \deduce[<Proof>]<Lower><Uppers>
-% draws a many step deduction with a proof name.
-%
-% Example:
-% If you want to write
-% B C
-% -----
-% A D
-% ----------
-% E
-% use
-% \infer{E}{
-% A
-% &
-% \infer{D}{B & C}
-% }
-%
-
-% Style Parameters
-
-\newdimen\inferLineSkip \inferLineSkip=2pt
-\newdimen\inferLabelSkip \inferLabelSkip=5pt
-\def\inferTabSkip{\quad}
-
-% Variables
-
-\newdimen\@LeftOffset % global
-\newdimen\@RightOffset % global
-\newdimen\@SavedLeftOffset % safe from users
-
-\newdimen\UpperWidth
-\newdimen\LowerWidth
-\newdimen\LowerHeight
-\newdimen\UpperLeftOffset
-\newdimen\UpperRightOffset
-\newdimen\UpperCenter
-\newdimen\LowerCenter
-\newdimen\UpperAdjust
-\newdimen\RuleAdjust
-\newdimen\LowerAdjust
-\newdimen\RuleWidth
-\newdimen\HLabelAdjust
-\newdimen\VLabelAdjust
-\newdimen\WidthAdjust
-
-\newbox\@UpperPart
-\newbox\@LowerPart
-\newbox\@LabelPart
-\newbox\ResultBox
-
-% Flags
-
-\newif\if@inferRule % whether \@infer draws a rule.
-\newif\if@DoubleRule % whether \@infer draws doulbe rules.
-\newif\if@ReturnLeftOffset % whether \@infer returns \@LeftOffset.
-\newif\if@MathSaved % whether inner math mode where \infer or
- % \deduce appears.
-
-% Special Fonts
-
-\def\DeduceSym{\vtop{\baselineskip4\p@ \lineskiplimit\z@
- \vbox{\hbox{.}\hbox{.}\hbox{.}}\hbox{.}}}
-
-% Math Save Macros
-%
-% \@SaveMath is called in the very begining of toplevel macros
-% which are \infer and \deduce.
-% \@RestoreMath is called in the very last before toplevel macros end.
-% Remark \infer and \deduce ends calling \@infer.
-
-\def\@SaveMath{\@MathSavedfalse \ifmmode \ifinner
- \relax $\relax \@MathSavedtrue \fi\fi }
-
-\def\@RestoreMath{\if@MathSaved \relax $\relax\fi }
-
-% Macros
-
-% Renaming @ifnextchar and @ifnch of LaTeX2e to @IFnextchar and @IFnch.
-
-\def\@IFnextchar#1#2#3{%
- \let\reserved@e=#1\def\reserved@a{#2}\def\reserved@b{#3}\futurelet
- \reserved@c\@IFnch}
-\def\@IFnch{\ifx \reserved@c \@sptoken \let\reserved@d\@xifnch
- \else \ifx \reserved@c \reserved@e\let\reserved@d\reserved@a\else
- \let\reserved@d\reserved@b\fi
- \fi \reserved@d}
-
-\def\@ifEmpty#1#2#3{\def\@tempa{\@empty}\def\@tempb{#1}\relax
- \ifx \@tempa \@tempb #2\else #3\fi }
-
-\def\infer{\@SaveMath \@IFnextchar *{\@inferSteps}{\relax
- \@IFnextchar ={\@inferDoubleRule}{\@inferOneStep}}}
-
-\def\@inferOneStep{\@inferRuletrue \@DoubleRulefalse
- \@IFnextchar [{\@infer}{\@infer[\@empty]}}
-
-\def\@inferDoubleRule={\@inferRuletrue \@DoubleRuletrue
- \@IFnextchar [{\@infer}{\@infer[\@empty]}}
-
-\def\@inferSteps*{\@IFnextchar [{\@@inferSteps}{\@@inferSteps[\@empty]}}
-
-\def\@@inferSteps[#1]{\@deduce{#1}[\DeduceSym]}
-
-\def\deduce{\@SaveMath \@IFnextchar [{\@deduce{\@empty}}
- {\@inferRulefalse \@infer[\@empty]}}
-
-% \@deduce<Proof Label>[<Proof>]<Lower><Uppers>
-
-\def\@deduce#1[#2]#3#4{\@inferRulefalse
- \@infer[\@empty]{#3}{\@SaveMath \@infer[{#1}]{#2}{#4}}}
-
-% \@infer[<Label>]<Lower><Uppers>
-% If \@inferRuletrue, it draws a rule and <Label> is right to
-% a rule. In this case, if \@DoubleRuletrue, it draws
-% double rules.
-%
-% Otherwise, draws no rule and <Label> is right to <Lower>.
-
-\def\@infer[#1]#2#3{\relax
-% Get parameters
- \if@ReturnLeftOffset \else \@SavedLeftOffset=\@LeftOffset \fi
- \setbox\@LabelPart=\hbox{$#1$}\relax
- \setbox\@LowerPart=\hbox{$#2$}\relax
-%
- \global\@LeftOffset=0pt
- \setbox\@UpperPart=\vbox{\tabskip=0pt \halign{\relax
- \global\@RightOffset=0pt \@ReturnLeftOffsettrue $##$&&
- \inferTabSkip
- \global\@RightOffset=0pt \@ReturnLeftOffsetfalse $##$\cr
- #3\cr}}\relax
-% Here is a little trick.
-% \@ReturnLeftOffsettrue(false) influences on \infer or
-% \deduce placed in ## locally
-% because of \@SaveMath and \@RestoreMath.
- \UpperLeftOffset=\@LeftOffset
- \UpperRightOffset=\@RightOffset
-% Calculate Adjustments
- \LowerWidth=\wd\@LowerPart
- \LowerHeight=\ht\@LowerPart
- \LowerCenter=0.5\LowerWidth
-%
- \UpperWidth=\wd\@UpperPart \advance\UpperWidth by -\UpperLeftOffset
- \advance\UpperWidth by -\UpperRightOffset
- \UpperCenter=\UpperLeftOffset
- \advance\UpperCenter by 0.5\UpperWidth
-%
- \ifdim \UpperWidth > \LowerWidth
- % \UpperCenter > \LowerCenter
- \UpperAdjust=0pt
- \RuleAdjust=\UpperLeftOffset
- \LowerAdjust=\UpperCenter \advance\LowerAdjust by -\LowerCenter
- \RuleWidth=\UpperWidth
- \global\@LeftOffset=\LowerAdjust
-%
- \else % \UpperWidth <= \LowerWidth
- \ifdim \UpperCenter > \LowerCenter
-%
- \UpperAdjust=0pt
- \RuleAdjust=\UpperCenter \advance\RuleAdjust by -\LowerCenter
- \LowerAdjust=\RuleAdjust
- \RuleWidth=\LowerWidth
- \global\@LeftOffset=\LowerAdjust
-%
- \else % \UpperWidth <= \LowerWidth
- % \UpperCenter <= \LowerCenter
-%
- \UpperAdjust=\LowerCenter \advance\UpperAdjust by -\UpperCenter
- \RuleAdjust=0pt
- \LowerAdjust=0pt
- \RuleWidth=\LowerWidth
- \global\@LeftOffset=0pt
-%
- \fi\fi
-% Make a box
- \if@inferRule
-%
- \setbox\ResultBox=\vbox{
- \moveright \UpperAdjust \box\@UpperPart
- \nointerlineskip \kern\inferLineSkip
- \if@DoubleRule
- \moveright \RuleAdjust \vbox{\hrule width\RuleWidth
- \kern 1pt\hrule width\RuleWidth}\relax
- \else
- \moveright \RuleAdjust \vbox{\hrule width\RuleWidth}\relax
- \fi
- \nointerlineskip \kern\inferLineSkip
- \moveright \LowerAdjust \box\@LowerPart }\relax
-%
- \@ifEmpty{#1}{}{\relax
-%
- \HLabelAdjust=\wd\ResultBox \advance\HLabelAdjust by -\RuleAdjust
- \advance\HLabelAdjust by -\RuleWidth
- \WidthAdjust=\HLabelAdjust
- \advance\WidthAdjust by -\inferLabelSkip
- \advance\WidthAdjust by -\wd\@LabelPart
- \ifdim \WidthAdjust < 0pt \WidthAdjust=0pt \fi
-%
- \VLabelAdjust=\dp\@LabelPart
- \advance\VLabelAdjust by -\ht\@LabelPart
- \VLabelAdjust=0.5\VLabelAdjust \advance\VLabelAdjust by \LowerHeight
- \advance\VLabelAdjust by \inferLineSkip
-%
- \setbox\ResultBox=\hbox{\box\ResultBox
- \kern -\HLabelAdjust \kern\inferLabelSkip
- \raise\VLabelAdjust \box\@LabelPart \kern\WidthAdjust}\relax
-%
- }\relax % end @ifEmpty
-%
- \else % \@inferRulefalse
-%
- \setbox\ResultBox=\vbox{
- \moveright \UpperAdjust \box\@UpperPart
- \nointerlineskip \kern\inferLineSkip
- \moveright \LowerAdjust \hbox{\unhbox\@LowerPart
- \@ifEmpty{#1}{}{\relax
- \kern\inferLabelSkip \unhbox\@LabelPart}}}\relax
- \fi
-%
- \global\@RightOffset=\wd\ResultBox
- \global\advance\@RightOffset by -\@LeftOffset
- \global\advance\@RightOffset by -\LowerWidth
- \if@ReturnLeftOffset \else \global\@LeftOffset=\@SavedLeftOffset \fi
-%
- \box\ResultBox
- \@RestoreMath
-}
--- a/LMCS-Paper/document/root.bib Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,370 +0,0 @@
-@inproceedings{pfenningsystem,
- author = "F.~Pfenning and C.~Sch{\"u}rmann",
- title = "{S}ystem {D}escription: {T}welf - {A} {M}eta-{L}ogical
- {F}ramework for {D}eductive {S}ystems",
- booktitle = "Proc.~of the 16th International Conference on Automated Deduction (CADE)",
- series = "LNAI",
- volume = 1632,
- pages = "202--206",
- year = 1999
-}
-
-
-
-@Article{ Urban08,
- author = "C. Urban",
- title = "{N}ominal {T}echniques in {I}sabelle/{HOL}",
- journal = "Journal of Automated Reasoning",
- volume = "40",
- number = "4",
- pages = "327--356",
- year = "2008"
-}
-
-@PhdThesis{Krauss09,
- author = {A.~Krauss},
- title = {{A}utomating {R}ecursive {D}efinitions and {T}ermination {P}roofs in
- {H}igher-Order {L}ogic},
- school = {TU Munich},
- year = {2009}
-}
-
-@InProceedings{WeirichYorgeySheard11,
- author = {S.~Weirich and B.~Yorgey and T.~Sheard},
- title = {{B}inders {U}nbound},
- booktitle = {Proc.~of the 16th International Conference on Functional Programming (ICFP)},
- year = {2011}
-}
-
-@InProceedings{UrbanKaliszyk11,
- author = {C.~Urban and C.~Kaliszyk},
- title = {{G}eneral {B}indings and {A}lpha-{E}quivalence in {N}ominal {I}sabelle},
- booktitle = {Proc.~of the 20th European Symposium on Programming (ESOP)},
- pages = {480-500},
- year = {2011},
- volume = {6602},
- series = {LNCS}
-}
-
-
-@inproceedings{KaliszykUrban11,
- author = {C.~Kaliszyk and C.~Urban},
- title = {{Q}uotients {R}evisited for {I}sabelle/{HOL}},
- booktitle = {Proc.~of the 26th ACM Symposium on Applied Computing (SAC)},
- year = {2011},
- pages = {1639--1644}
-}
-
-@InProceedings{cheney05a,
- author = {J.~Cheney},
- title = {{S}crap {Y}our {N}ameplate ({F}unctional {P}earl)},
- booktitle = {Proc.~of the 10th International Conference on Functional Programming (ICFP)},
- pages = {180--191},
- year = {2005}
-}
-
-@Inproceedings{Altenkirch10,
- author = {T.~Altenkirch and N.~A.~Danielsson and A.~L\"oh and N.~Oury},
- title = {{PiSigma}: {D}ependent {T}ypes {W}ithout the {S}ugar},
- booktitle = "Proc.~of the 10th International Symposium on Functional and Logic Programming (FLOPS)",
- year = 2010,
- series = "LNCS",
- pages = "40--55",
- volume = 6009
-}
-
-
-@InProceedings{ UrbanTasson05,
- author = "C. Urban and C. Tasson",
- title = "{N}ominal {T}echniques in {I}sabelle/{HOL}",
- booktitle = "Proc.~of the 20th Conference on Automated Deduction (CADE)",
- year = 2005,
- series = "LNCS",
- pages = "38--53",
- volume = 3632
-}
-
-@InProceedings{ UrbanBerghofer06,
- author = "C. Urban and S. Berghofer",
- title = "{A} {R}ecursion {C}ombinator for {N}ominal {D}atatypes {I}mplemented in {I}sabelle/{HOL}",
- booktitle = "Proc.~of the 3rd International Joint Conference on Automated Deduction (IJCAR)",
- year = 2006,
- series = "LNAI",
- volume = 4130,
- pages = "498--512"
-}
-
-@InProceedings{LeeCraryHarper07,
- author = {D.~K.~Lee and K.~Crary and R.~Harper},
- title = {{T}owards a {M}echanized {M}etatheory of {Standard ML}},
- booktitle = {Proc.~of the 34th Symposium on Principles of Programming Languages (POPL)},
- year = 2007,
- pages = {173--184}
-}
-
-@Unpublished{chargueraud09,
- author = "A.~Chargu{\'e}raud",
- title = "{T}he {L}ocally {N}ameless {R}epresentation",
- Note = "To appear in Journal of Automated Reasoning."
-}
-
-@article{NaraschewskiNipkow99,
- author={W.~Naraschewski and T.~Nipkow},
- title={{T}ype {I}nference {V}erified: {A}lgorithm {W} in {Isabelle/HOL}},
- journal={Journal of Automated Reasoning},
- year=1999,
- volume=23,
- pages={299--318}}
-
-@InProceedings{Berghofer99,
- author = {S.~Berghofer and M.~Wenzel},
- title = {{I}nductive {D}atatypes in {HOL} - {L}essons {L}earned in
- {F}ormal-{L}ogic {E}ngineering},
- booktitle = {Proc.~of the 12th Conference on Theorem Proving in Higher Order Logics (TPHOLs)},
- pages = {19--36},
- year = 1999,
- volume = 1690,
- series = {LNCS}
-}
-
-@InProceedings{CoreHaskell,
- author = {M.~Sulzmann and M.~Chakravarty and S.~Peyton Jones and K.~Donnelly},
- title = {{S}ystem {F} with {T}ype {E}quality {C}oercions},
- booktitle = {Proc.~of the 3rd Workshop on Types in Language Design and Implementation (TLDI)},
- pages = {53-66},
- year = {2007}
-}
-
-@inproceedings{cheney05,
- author = {J.~Cheney},
- title = {{T}owards a {G}eneral {T}heory of {N}ames: {B}inding and {S}cope},
- booktitle = {Proc.~of the 3rd ACM Workshop on Mechanized Reasoning about Languages
- with Variable Binding and Names (MERLIN)},
- year = {2005},
- pages = {33-40}
-}
-
-@Unpublished{Pitts04,
- author = {A.~Pitts},
- title = {{N}otes on the {R}estriction {M}onad for {N}ominal {S}ets and {C}pos},
- note = {Unpublished notes for an invited talk given at CTCS},
- year = {2004}
-}
-
-@incollection{UrbanNipkow09,
- author = {C.~Urban and T.~Nipkow},
- title = {{N}ominal {V}erification of {A}lgorithm {W}},
- booktitle={From Semantics to Computer Science. Essays in Honour of Gilles Kahn},
- editor={G.~Huet and J.-J.~L{\'e}vy and G.~Plotkin},
- publisher={Cambridge University Press},
- pages={363--382},
- year=2009
-}
-
-@InProceedings{Homeier05,
- author = {P.~Homeier},
- title = {{A} {D}esign {S}tructure for {H}igher {O}rder {Q}uotients},
- booktitle = {Proc.~of the 18th Conference on Theorem Proving in Higher Order Logics (TPHOLs)},
- pages = {130--146},
- year = {2005},
- volume = {3603},
- series = {LNCS}
-}
-
-@article{ott-jfp,
- author = {P.~Sewell and
- F.~Z.~Nardelli and
- S.~Owens and
- G.~Peskine and
- T.~Ridge and
- S.~Sarkar and
- R.~Strni\v{s}a},
- title = {{Ott}: {E}ffective {T}ool {S}upport for the {W}orking {S}emanticist},
- journal = {Journal of Functional Programming},
- year = {2010},
- volume = {20},
- number = {1},
- pages = {70--122}
-}
-
-@INPROCEEDINGS{Pottier06,
- author = {F.~Pottier},
- title = {{A}n {O}verview of {C$\alpha$ml}},
- year = {2006},
- booktitle = {Proc.~of the 7th ACM Workshop on ML},
- pages = {27--52},
- volume = {148},
- number = {2},
- series = {ENTCS}
-}
-
-@inproceedings{HuffmanUrban10,
- author = {B.~Huffman and C.~Urban},
- title = {{P}roof {P}earl: {A} {N}ew {F}oundation for {N}ominal {I}sabelle},
- booktitle = {Proc.~of the 1st Conference on Interactive Theorem Proving (ITP)},
- pages = {35--50},
- volume = {6172},
- series = {LNCS},
- year = {2010}
-}
-
-@PhdThesis{Leroy92,
- author = {X.~Leroy},
- title = {{P}olymorphic {T}yping of an {A}lgorithmic {L}anguage},
- school = {University Paris 7},
- year = {1992},
- note = {INRIA Research Report, No~1778}
-}
-
-@Unpublished{SewellBestiary,
- author = {P.~Sewell},
- title = {{A} {B}inding {B}estiary},
- note = {Unpublished notes.}
-}
-
-@InProceedings{challenge05,
- author = {B.~E.~Aydemir and A.~Bohannon and M.~Fairbairn and
- J.~N.~Foster and B.~C.~Pierce and P.~Sewell and
- D.~Vytiniotis and G.~Washburn and S.~Weirich and
- S.~Zdancewic},
- title = {{M}echanized {M}etatheory for the {M}asses: {T}he \mbox{Popl}{M}ark
- {C}hallenge},
- booktitle = {Proc.~of the 18th Conference on Theorem Proving in Higher Order Logics (TPHOLs)},
- pages = {50--65},
- year = {2005},
- volume = {3603},
- series = {LNCS}
-}
-
-@article{MckinnaPollack99,
- author = {J.~McKinna and R.~Pollack},
- title = {Some {L}ambda {C}alculus and {T}ype {T}heory {F}ormalized},
- journal = {Journal of Automated Reasoning},
- volume = {23},
- number = {3-4},
- pages = {373-409},
- year = {1999}
-}
-
-@article{SatoPollack10,
- author = {M.~Sato and R.~Pollack},
- title = {{E}xternal and {I}nternal {S}yntax of the {L}ambda-{C}alculus},
- journal = {Journal of Symbolic Computation},
- volume = 45,
- pages = {598--616},
- year = 2010
-}
-
-@article{GabbayPitts02,
- author = {M.~J.~Gabbay and A.~M.~Pitts},
- title = {A New Approach to Abstract Syntax with Variable
- Binding},
- journal = {Formal Aspects of Computing},
- volume = {13},
- year = 2002,
- pages = {341--363}
-}
-
-@article{Pitts03,
- author = {A.~M.~Pitts},
- title = {{N}ominal {L}ogic, {A} {F}irst {O}rder {T}heory of {N}ames and
- {B}inding},
- journal = {Information and Computation},
- year = {2003},
- volume = {183},
- pages = {165--193}
-}
-
-@InProceedings{BengtsonParrow07,
- author = {J.~Bengtson and J.~Parrow},
- title = {Formalising the pi-{C}alculus using {N}ominal {L}ogic},
- booktitle = {Proc.~of the 10th FOSSACS Conference ???},
- year = 2007,
- pages = {63--77},
- series = {LNCS},
- volume = {4423}
-}
-
-@inproceedings{BengtsonParow09,
- author = {J.~Bengtson and J.~Parrow},
- title = {{P}si-{C}alculi in {I}sabelle},
- booktitle = {Proc of the 22nd Conference on Theorem Proving in Higher Order Logics (TPHOLs)},
- year = 2009,
- pages = {99--114},
- series = {LNCS},
- volume = {5674}
-}
-
-@inproceedings{TobinHochstadtFelleisen08,
- author = {S.~Tobin-Hochstadt and M.~Felleisen},
- booktitle = {Proc.~of the 35rd Symposium on Principles of Programming Languages (POPL)},
- title = {{T}he {D}esign and {I}mplementation of {T}yped {S}cheme},
- year = {2008},
- pages = {395--406}
-}
-
-@InProceedings{UrbanCheneyBerghofer08,
- author = "C.~Urban and J.~Cheney and S.~Berghofer",
- title = "{M}echanizing the {M}etatheory of {LF}",
- pages = "45--56",
- year = 2008,
- booktitle = "Proc.~of the 23rd Symposium on Logic in Computer Science (LICS)"
-}
-
-@InProceedings{UrbanZhu08,
- title = "{R}evisiting {C}ut-{E}limination: {O}ne {D}ifficult {P}roof is {R}eally a {P}roof",
- author = "C.~Urban and B.~Zhu",
- booktitle = "Proc.~of the 9th International Conference on Rewriting Techniques and Applications (RTA)",
- year = "2008",
- pages = "409--424",
- series = "LNCS",
- volume = 5117
-}
-
-@Article{UrbanPittsGabbay04,
- title = "{N}ominal {U}nification",
- author = "C.~Urban and A.M.~Pitts and M.J.~Gabbay",
- journal = "Theoretical Computer Science",
- pages = "473--497",
- volume = "323",
- number = "1-3",
- year = "2004"
-}
-
-@Article{Church40,
- author = {A.~Church},
- title = {{A} {F}ormulation of the {S}imple {T}heory of {T}ypes},
- journal = {Journal of Symbolic Logic},
- year = {1940},
- volume = {5},
- number = {2},
- pages = {56--68}
-}
-
-
-@Manual{PittsHOL4,
- title = {{S}yntax and {S}emantics},
- author = {A.~M.~Pitts},
- note = {Part of the documentation for the HOL4 system.}
-}
-
-
-@book{PaulsonBenzmueller,
- year={2009},
- author={Benzm{\"u}ller, Christoph and Paulson, Lawrence C.},
- title={Quantified Multimodal Logics in Simple Type Theory},
- note={{http://arxiv.org/abs/0905.2435}},
- series={{SEKI Report SR--2009--02 (ISSN 1437-4447)}},
- publisher={{SEKI Publications}}
-}
-
-@Article{Cheney06,
- author = {J.~Cheney},
- title = {{C}ompleteness and {H}erbrand theorems for {N}ominal {L}ogic},
- journal = {Journal of Symbolic Logic},
- year = {2006},
- volume = {71},
- number = {1},
- pages = {299--320}
-}
-
--- a/LMCS-Paper/document/root.tex Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,112 +0,0 @@
-\documentclass{lmcs}
-%%\usepackage{times}
-\usepackage{isabelle}
-\usepackage{isabellesym}
-\usepackage{amsmath}
-\usepackage{amssymb}
-%%\usepackage{amsthm}
-\usepackage{tikz}
-\usepackage{pgf}
-\usepackage{ot1patch}
-\usepackage{times}
-\usepackage{boxedminipage}
-\usepackage{proof}
-\usepackage{setspace}
-\usepackage{afterpage}
-\usepackage{pdfsetup}
-
-\allowdisplaybreaks
-\urlstyle{rm}
-\isabellestyle{it}
-\renewcommand{\isastyleminor}{\it}%
-\renewcommand{\isastyle}{\normalsize\it}%
-
-\DeclareRobustCommand{\flqq}{\mbox{\guillemotleft}}
-\DeclareRobustCommand{\frqq}{\mbox{\guillemotright}}
-\renewcommand{\isacharunderscore}{\mbox{$\_\!\_$}}
-\renewcommand{\isasymbullet}{{\raisebox{-0.4mm}{\Large$\boldsymbol{\hspace{-0.5mm}\cdot\hspace{-0.5mm}}$}}}
-\def\dn{\,\stackrel{\mbox{\scriptsize def}}{=}\,}
-\renewcommand{\isasymequiv}{$\dn$}
-%%\renewcommand{\isasymiota}{}
-\renewcommand{\isasymxi}{$..$}
-\renewcommand{\isasymemptyset}{$\varnothing$}
-\newcommand{\isasymnotapprox}{$\not\approx$}
-\newcommand{\isasymLET}{$\mathtt{let}$}
-\newcommand{\isasymAND}{$\mathtt{and}$}
-\newcommand{\isasymIN}{$\mathtt{in}$}
-\newcommand{\isasymEND}{$\mathtt{end}$}
-\newcommand{\isasymBIND}{$\mathtt{bind}$}
-\newcommand{\isasymANIL}{$\mathtt{anil}$}
-\newcommand{\isasymACONS}{$\mathtt{acons}$}
-\newcommand{\isasymCASE}{$\mathtt{case}$}
-\newcommand{\isasymOF}{$\mathtt{of}$}
-\newcommand{\isasymAL}{\makebox[0mm][l]{$^\alpha$}}
-\newcommand{\isasymPRIME}{\makebox[0mm][l]{$'$}}
-\newcommand{\isasymFRESH}{\#}
-\newcommand{\LET}{\;\mathtt{let}\;}
-\newcommand{\IN}{\;\mathtt{in}\;}
-\newcommand{\END}{\;\mathtt{end}\;}
-\newcommand{\AND}{\;\mathtt{and}\;}
-\newcommand{\fv}{\mathit{fv}}
-
-\newcommand{\numbered}[1]{\refstepcounter{equation}{\rm(\arabic{equation})}\label{#1}}
-%----------------- theorem definitions ----------
-%%\theoremstyle{plain}
-%%\spnewtheorem{thm}[section]{Theorem}
-%%\newtheorem{property}[thm]{Property}
-%%\newtheorem{lemma}[thm]{Lemma}
-%%\spnewtheorem{defn}[theorem]{Definition}
-%%\spnewtheorem{exmple}[theorem]{Example}
-%%\spnewtheorem{myproperty}{Property}{\bfseries}{\rmfamily}
-%-------------------- environment definitions -----------------
-\newenvironment{proof-of}[1]{{\em Proof of #1:}}{}
-
-%\addtolength{\textwidth}{2mm}
-\addtolength{\parskip}{-0.33mm}
-\begin{document}
-
-\title[Genral Bindings]{General Bindings and Alpha-Equivalence in Nominal
-Isabelle$^\star$}
-\author{Christian Urban}
-\address{Technical University of Munich, Germany}
-\email{urbanc@in.tum.de}
-
-\author{Cezary Kaliszyk}
-\address{University of Tsukuba, Japan}
-\email{kaliszyk@cs.tsukuba.ac.jp}
-\thanks{$^\star$~This is a revised and expanded version of~\cite{UrbanKaliszyk11}}
-
-\keywords{Nominal Isabelle, variable convention, theorem provers, formal reasoning, lambda-calculus}
-\subjclass{F.3.1}
-
-\begin{abstract}
-Nominal Isabelle is a definitional extension of the Isabelle/HOL theorem
-prover. It provides a proving infrastructure for reasoning about
-programming language calculi involving named bound variables (as
-opposed to de-Bruijn indices). In this paper we present an extension of
-Nominal Isabelle for dealing with general bindings, that means
-term-constructors where multiple variables are bound at once. Such general
-bindings are ubiquitous in programming language research and only very
-poorly supported with single binders, such as lambda-abstractions. Our
-extension includes new definitions of alpha-equivalence and establishes
-automatically the reasoning infrastructure for alpha-equated terms. We
-also prove strong induction principles that have the usual variable
-convention already built in.
-\end{abstract}
-
-\maketitle
-\input{session}
-
-\bibliographystyle{plain}
-\bibliography{root}
-
-
-%\pagebreak
-%\input{Appendix}
-
-\end{document}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
Binary file Literature/Quotient-Isabelle.pdf has changed
Binary file Literature/cheney05icfp.pdf has changed
Binary file Literature/core-has-tech-rep.pdf has changed
Binary file Literature/core-has.pdf has changed
Binary file Literature/equivclasses.pdf has changed
Binary file Literature/fpottier-alphacaml.pdf has changed
Binary file Literature/homeier-qs.pdf has changed
Binary file Literature/newtypes.pdf has changed
--- a/Literature/ottbinding-wmm07.ps Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,12745 +0,0 @@
-%!PS-Adobe-2.0
-%%Creator: dvips(k) 5.95a Copyright 2005 Radical Eye Software
-%%Title: bind-wmm-2.dvi
-%%Pages: 1
-%%PageOrder: Ascend
-%%BoundingBox: 0 0 595 842
-%%DocumentFonts: Times-Bold Times-Roman CMSY8 CMR8 CMSY6 CMR6 CMMI9
-%%+ CMTT9 CMSS9 CMTI9 Times-Italic CMR9 CMSY9 CMBX9 CMMI6 MSBM10 CMMI8
-%%EndComments
-%DVIPSWebPage: (www.radicaleye.com)
-%DVIPSCommandLine: /nfs/repl/tetex3/tex.bin/x86_64_linux/dvips -Ppdf -j0
-%+ -G0 bind-wmm-2
-%DVIPSParameters: dpi=8000
-%DVIPSSource: TeX output 2009.01.27:1436
-%%BeginProcSet: tex.pro 0 0
-%!
-/TeXDict 300 dict def TeXDict begin/N{def}def/B{bind def}N/S{exch}N/X{S
-N}B/A{dup}B/TR{translate}N/isls false N/vsize 11 72 mul N/hsize 8.5 72
-mul N/landplus90{false}def/@rigin{isls{[0 landplus90{1 -1}{-1 1}ifelse 0
-0 0]concat}if 72 Resolution div 72 VResolution div neg scale isls{
-landplus90{VResolution 72 div vsize mul 0 exch}{Resolution -72 div hsize
-mul 0}ifelse TR}if Resolution VResolution vsize -72 div 1 add mul TR[
-matrix currentmatrix{A A round sub abs 0.00001 lt{round}if}forall round
-exch round exch]setmatrix}N/@landscape{/isls true N}B/@manualfeed{
-statusdict/manualfeed true put}B/@copies{/#copies X}B/FMat[1 0 0 -1 0 0]
-N/FBB[0 0 0 0]N/nn 0 N/IEn 0 N/ctr 0 N/df-tail{/nn 8 dict N nn begin
-/FontType 3 N/FontMatrix fntrx N/FontBBox FBB N string/base X array
-/BitMaps X/BuildChar{CharBuilder}N/Encoding IEn N end A{/foo setfont}2
-array copy cvx N load 0 nn put/ctr 0 N[}B/sf 0 N/df{/sf 1 N/fntrx FMat N
-df-tail}B/dfs{div/sf X/fntrx[sf 0 0 sf neg 0 0]N df-tail}B/E{pop nn A
-definefont setfont}B/Cw{Cd A length 5 sub get}B/Ch{Cd A length 4 sub get
-}B/Cx{128 Cd A length 3 sub get sub}B/Cy{Cd A length 2 sub get 127 sub}
-B/Cdx{Cd A length 1 sub get}B/Ci{Cd A type/stringtype ne{ctr get/ctr ctr
-1 add N}if}B/CharBuilder{save 3 1 roll S A/base get 2 index get S
-/BitMaps get S get/Cd X pop/ctr 0 N Cdx 0 Cx Cy Ch sub Cx Cw add Cy
-setcachedevice Cw Ch true[1 0 0 -1 -.1 Cx sub Cy .1 sub]{Ci}imagemask
-restore}B/D{/cc X A type/stringtype ne{]}if nn/base get cc ctr put nn
-/BitMaps get S ctr S sf 1 ne{A A length 1 sub A 2 index S get sf div put
-}if put/ctr ctr 1 add N}B/I{cc 1 add D}B/bop{userdict/bop-hook known{
-bop-hook}if/SI save N @rigin 0 0 moveto/V matrix currentmatrix A 1 get A
-mul exch 0 get A mul add .99 lt{/QV}{/RV}ifelse load def pop pop}N/eop{
-SI restore userdict/eop-hook known{eop-hook}if showpage}N/@start{
-userdict/start-hook known{start-hook}if pop/VResolution X/Resolution X
-1000 div/DVImag X/IEn 256 array N 2 string 0 1 255{IEn S A 360 add 36 4
-index cvrs cvn put}for pop 65781.76 div/vsize X 65781.76 div/hsize X}N
-/p{show}N/RMat[1 0 0 -1 0 0]N/BDot 260 string N/Rx 0 N/Ry 0 N/V{}B/RV/v{
-/Ry X/Rx X V}B statusdict begin/product where{pop false[(Display)(NeXT)
-(LaserWriter 16/600)]{A length product length le{A length product exch 0
-exch getinterval eq{pop true exit}if}{pop}ifelse}forall}{false}ifelse
-end{{gsave TR -.1 .1 TR 1 1 scale Rx Ry false RMat{BDot}imagemask
-grestore}}{{gsave TR -.1 .1 TR Rx Ry scale 1 1 false RMat{BDot}
-imagemask grestore}}ifelse B/QV{gsave newpath transform round exch round
-exch itransform moveto Rx 0 rlineto 0 Ry neg rlineto Rx neg 0 rlineto
-fill grestore}B/a{moveto}B/delta 0 N/tail{A/delta X 0 rmoveto}B/M{S p
-delta add tail}B/b{S p tail}B/c{-4 M}B/d{-3 M}B/e{-2 M}B/f{-1 M}B/g{0 M}
-B/h{1 M}B/i{2 M}B/j{3 M}B/k{4 M}B/w{0 rmoveto}B/l{p -4 w}B/m{p -3 w}B/n{
-p -2 w}B/o{p -1 w}B/q{p 1 w}B/r{p 2 w}B/s{p 3 w}B/t{p 4 w}B/x{0 S
-rmoveto}B/y{3 2 roll p a}B/bos{/SS save N}B/eos{SS restore}B end
-
-%%EndProcSet
-%%BeginProcSet: alt-rule.pro 0 0
-%!
-% Patch by TVZ
-% Makes dvips files draw rules with stroke rather than fill.
-% Makes narrow rules more predictable at low resolutions
-% after distilling to PDF.
-% May have unknown consequences for very thick rules.
-% Tested only with dvips 5.85(k).
-TeXDict begin
-/QV {
- gsave newpath /ruleY X /ruleX X
- Rx Ry gt
- { ruleX ruleY Ry 2 div sub moveto Rx 0 rlineto Ry }
- { ruleX Rx 2 div add ruleY moveto 0 Ry neg rlineto Rx }
- ifelse
- setlinewidth 0 setlinecap stroke grestore
-} bind def
-end
-
-%%EndProcSet
-%%BeginProcSet: pstricks.pro 0 0
-%!
-% PostScript prologue for pstricks.tex.
-% Version 97 patch 4, 04/05/10
-% For distribution, see pstricks.tex.
-%
-/tx@Dict 200 dict def tx@Dict begin
-/ADict 25 dict def
-/CM { matrix currentmatrix } bind def
-/SLW /setlinewidth load def
-/CLW /currentlinewidth load def
-/CP /currentpoint load def
-/ED { exch def } bind def
-/L /lineto load def
-/T /translate load def
-/TMatrix { } def
-/RAngle { 0 } def
-/Atan { /atan load stopped { pop pop 0 } if } def
-/Div { dup 0 eq { pop } { div } ifelse } def
-/NET { neg exch neg exch T } def
-/Pyth { dup mul exch dup mul add sqrt } def
-/PtoC { 2 copy cos mul 3 1 roll sin mul } def
-/PathLength@ { /z z y y1 sub x x1 sub Pyth add def /y1 y def /x1 x def }
-def
-/PathLength { flattenpath /z 0 def { /y1 ED /x1 ED /y2 y1 def /x2 x1 def
-} { /y ED /x ED PathLength@ } {} { /y y2 def /x x2 def PathLength@ }
-/pathforall load stopped { pop pop pop pop } if z } def
-/STP { .996264 dup scale } def
-/STV { SDict begin normalscale end STP } def
-%
-%%-------------- DG begin patch 15 ---------------%%
-%/DashLine { dup 0 gt { /a .5 def PathLength exch div } { pop /a 1 def
-%PathLength } ifelse /b ED /x ED /y ED /z y x add def b a .5 sub 2 mul y
-%mul sub z Div round z mul a .5 sub 2 mul y mul add b exch Div dup y mul
-%/y ED x mul /x ED x 0 gt y 0 gt and { [ y x ] 1 a sub y mul } { [ 1 0 ]
-%0 } ifelse setdash stroke } def
-/DashLine {
- dup 0 gt { /a .5 def PathLength exch div } { pop /a 1 def PathLength } ifelse
- /b ED /x1 ED /y1 ED /x ED /y ED
- /z y x add y1 add x1 add def
- /Coef b a .5 sub 2 mul y mul sub z Div round
- z mul a .5 sub 2 mul y mul add b exch Div def
- /y y Coef mul def /x x Coef mul def /y1 y1 Coef mul def /x1 x1 Coef mul def
- x1 0 gt y1 0 gt x 0 gt y 0 gt and { [ y x y1 x1 ] 1 a sub y mul}
- { [ 1 0] 0 } ifelse setdash stroke
-} def
-%%-------------- DG end patch 15 ---------------%%
-/DotLine { /b PathLength def /a ED /z ED /y CLW def /z y z add def a 0 gt
-{ /b b a div def } { a 0 eq { /b b y sub def } { a -3 eq { /b b y add
-def } if } ifelse } ifelse [ 0 b b z Div round Div dup 0 le { pop 1 } if
-] a 0 gt { 0 } { y 2 div a -2 gt { neg } if } ifelse setdash 1
-setlinecap stroke } def
-/LineFill { gsave abs CLW add /a ED a 0 dtransform round exch round exch
-2 copy idtransform exch Atan rotate idtransform pop /a ED .25 .25
-% DG/SR modification begin - Dec. 12, 1997 - Patch 2
-%itransform translate pathbbox /y2 ED a Div ceiling cvi /x2 ED /y1 ED a
-itransform pathbbox /y2 ED a Div ceiling cvi /x2 ED /y1 ED a
-% DG/SR modification end
-Div cvi /x1 ED /y2 y2 y1 sub def clip newpath 2 setlinecap systemdict
-/setstrokeadjust known { true setstrokeadjust } if x2 x1 sub 1 add { x1
-% DG/SR modification begin - Jun. 1, 1998 - Patch 3 (from Michael Vulis)
-% a mul y1 moveto 0 y2 rlineto stroke /x1 x1 1 add def } repeat grestore }
-% def
-a mul y1 moveto 0 y2 rlineto stroke /x1 x1 1 add def } repeat grestore
-pop pop } def
-% DG/SR modification end
-/BeginArrow { ADict begin /@mtrx CM def gsave 2 copy T 2 index sub neg
-exch 3 index sub exch Atan rotate newpath } def
-/EndArrow { @mtrx setmatrix CP grestore end } def
-/Arrow { CLW mul add dup 2 div /w ED mul dup /h ED mul /a ED { 0 h T 1 -1
-scale } if w neg h moveto 0 0 L w h L w neg a neg rlineto gsave fill
-grestore } def
-/Tbar { CLW mul add /z ED z -2 div CLW 2 div moveto z 0 rlineto stroke 0
-CLW moveto } def
-/Bracket { CLW mul add dup CLW sub 2 div /x ED mul CLW add /y ED /z CLW 2
-div def x neg y moveto x neg CLW 2 div L x CLW 2 div L x y L stroke 0
-CLW moveto } def
-/RoundBracket { CLW mul add dup 2 div /x ED mul /y ED /mtrx CM def 0 CLW
-2 div T x y mul 0 ne { x y scale } if 1 1 moveto .85 .5 .35 0 0 0
-curveto -.35 0 -.85 .5 -1 1 curveto mtrx setmatrix stroke 0 CLW moveto }
-def
-/SD { 0 360 arc fill } def
-/EndDot { { /z DS def } { /z 0 def } ifelse /b ED 0 z DS SD b { 0 z DS
-CLW sub SD } if 0 DS z add CLW 4 div sub moveto } def
-/Shadow { [ { /moveto load } { /lineto load } { /curveto load } {
-/closepath load } /pathforall load stopped { pop pop pop pop CP /moveto
-load } if ] cvx newpath 3 1 roll T exec } def
-/NArray { aload length 2 div dup dup cvi eq not { exch pop } if /n exch
-cvi def } def
-/NArray { /f ED counttomark 2 div dup cvi /n ED n eq not { exch pop } if
-f { ] aload /Points ED } { n 2 mul 1 add -1 roll pop } ifelse } def
-/Line { NArray n 0 eq not { n 1 eq { 0 0 /n 2 def } if ArrowA /n n 2 sub
-def n { Lineto } repeat CP 4 2 roll ArrowB L pop pop } if } def
-/Arcto { /a [ 6 -2 roll ] cvx def a r /arcto load stopped { 5 } { 4 }
-ifelse { pop } repeat a } def
-/CheckClosed { dup n 2 mul 1 sub index eq 2 index n 2 mul 1 add index eq
-and { pop pop /n n 1 sub def } if } def
-/Polygon { NArray n 2 eq { 0 0 /n 3 def } if n 3 lt { n { pop pop }
-repeat } { n 3 gt { CheckClosed } if n 2 mul -2 roll /y0 ED /x0 ED /y1
-ED /x1 ED x1 y1 /x1 x0 x1 add 2 div def /y1 y0 y1 add 2 div def x1 y1
-moveto /n n 2 sub def n { Lineto } repeat x1 y1 x0 y0 6 4 roll Lineto
-Lineto pop pop closepath } ifelse } def
-/Diamond { /mtrx CM def T rotate /h ED /w ED dup 0 eq { pop } { CLW mul
-neg /d ED /a w h Atan def /h d a sin Div h add def /w d a cos Div w add
-def } ifelse mark w 2 div h 2 div w 0 0 h neg w neg 0 0 h w 2 div h 2
-div /ArrowA { moveto } def /ArrowB { } def false Line closepath mtrx
-setmatrix } def
-% DG modification begin - Jan. 15, 1997
-%/Triangle { /mtrx CM def translate rotate /h ED 2 div /w ED dup 0 eq {
-%pop } { CLW mul /d ED /h h d w h Atan sin Div sub def /w w d h w Atan 2
-%div dup cos exch sin Div mul sub def } ifelse mark 0 d w neg d 0 h w d 0
-%d /ArrowA { moveto } def /ArrowB { } def false Line closepath mtrx
-%setmatrix } def
-/Triangle { /mtrx CM def translate rotate /h ED 2 div /w ED dup
-CLW mul /d ED /h h d w h Atan sin Div sub def /w w d h w Atan 2
-div dup cos exch sin Div mul sub def mark 0 d w neg d 0 h w d 0
-d /ArrowA { moveto } def /ArrowB { } def false Line closepath mtrx
-% DG/SR modification begin - Jun. 1, 1998 - Patch 3 (from Michael Vulis)
-% setmatrix } def
-setmatrix pop } def
-% DG/SR modification end
-/CCA { /y ED /x ED 2 copy y sub /dy1 ED x sub /dx1 ED /l1 dx1 dy1 Pyth
-def } def
-/CCA { /y ED /x ED 2 copy y sub /dy1 ED x sub /dx1 ED /l1 dx1 dy1 Pyth
-def } def
-/CC { /l0 l1 def /x1 x dx sub def /y1 y dy sub def /dx0 dx1 def /dy0 dy1
-def CCA /dx dx0 l1 c exp mul dx1 l0 c exp mul add def /dy dy0 l1 c exp
-mul dy1 l0 c exp mul add def /m dx0 dy0 Atan dx1 dy1 Atan sub 2 div cos
-abs b exp a mul dx dy Pyth Div 2 div def /x2 x l0 dx mul m mul sub def
-/y2 y l0 dy mul m mul sub def /dx l1 dx mul m mul neg def /dy l1 dy mul
-m mul neg def } def
-/IC { /c c 1 add def c 0 lt { /c 0 def } { c 3 gt { /c 3 def } if }
-ifelse /a a 2 mul 3 div 45 cos b exp div def CCA /dx 0 def /dy 0 def }
-def
-/BOC { IC CC x2 y2 x1 y1 ArrowA CP 4 2 roll x y curveto } def
-/NC { CC x1 y1 x2 y2 x y curveto } def
-/EOC { x dx sub y dy sub 4 2 roll ArrowB 2 copy curveto } def
-/BAC { IC CC x y moveto CC x1 y1 CP ArrowA } def
-/NAC { x2 y2 x y curveto CC x1 y1 } def
-/EAC { x2 y2 x y ArrowB curveto pop pop } def
-/OpenCurve { NArray n 3 lt { n { pop pop } repeat } { BOC /n n 3 sub def
- n { NC } repeat EOC } ifelse } def
-/AltCurve { { false NArray n 2 mul 2 roll [ n 2 mul 3 sub 1 roll ] aload
-/Points ED n 2 mul -2 roll } { false NArray } ifelse n 4 lt { n { pop
-pop } repeat } { BAC /n n 4 sub def n { NAC } repeat EAC } ifelse } def
-/ClosedCurve { NArray n 3 lt { n { pop pop } repeat } { n 3 gt {
-CheckClosed } if 6 copy n 2 mul 6 add 6 roll IC CC x y moveto n { NC }
-repeat closepath pop pop } ifelse } def
-/SQ { /r ED r r moveto r r neg L r neg r neg L r neg r L fill } def
-/ST { /y ED /x ED x y moveto x neg y L 0 x L fill } def
-/SP { /r ED gsave 0 r moveto 4 { 72 rotate 0 r L } repeat fill grestore }
-def
-/FontDot { DS 2 mul dup matrix scale matrix concatmatrix exch matrix
-rotate matrix concatmatrix exch findfont exch makefont setfont } def
-/Rect { x1 y1 y2 add 2 div moveto x1 y2 lineto x2 y2 lineto x2 y1 lineto
-x1 y1 lineto closepath } def
-/OvalFrame { x1 x2 eq y1 y2 eq or { pop pop x1 y1 moveto x2 y2 L } { y1
-y2 sub abs x1 x2 sub abs 2 copy gt { exch pop } { pop } ifelse 2 div
-exch { dup 3 1 roll mul exch } if 2 copy lt { pop } { exch pop } ifelse
-/b ED x1 y1 y2 add 2 div moveto x1 y2 x2 y2 b arcto x2 y2 x2 y1 b arcto
-x2 y1 x1 y1 b arcto x1 y1 x1 y2 b arcto 16 { pop } repeat closepath }
-ifelse } def
-/Frame { CLW mul /a ED 3 -1 roll 2 copy gt { exch } if a sub /y2 ED a add
-/y1 ED 2 copy gt { exch } if a sub /x2 ED a add /x1 ED 1 index 0 eq {
-pop pop Rect } { OvalFrame } ifelse } def
-/BezierNArray { /f ED counttomark 2 div dup cvi /n ED n eq not { exch pop
-} if n 1 sub neg 3 mod 3 add 3 mod { 0 0 /n n 1 add def } repeat f { ]
-aload /Points ED } { n 2 mul 1 add -1 roll pop } ifelse } def
-/OpenBezier { BezierNArray n 1 eq { pop pop } { ArrowA n 4 sub 3 idiv { 6
-2 roll 4 2 roll curveto } repeat 6 2 roll 4 2 roll ArrowB curveto }
-ifelse } def
-/ClosedBezier { BezierNArray n 1 eq { pop pop } { moveto n 1 sub 3 idiv {
-6 2 roll 4 2 roll curveto } repeat closepath } ifelse } def
-/BezierShowPoints { gsave Points aload length 2 div cvi /n ED moveto n 1
-sub { lineto } repeat CLW 2 div SLW [ 4 4 ] 0 setdash stroke grestore }
-def
-/Parab { /y0 exch def /x0 exch def /y1 exch def /x1 exch def /dx x0 x1
-sub 3 div def /dy y0 y1 sub 3 div def x0 dx sub y0 dy add x1 y1 ArrowA
-x0 dx add y0 dy add x0 2 mul x1 sub y1 ArrowB curveto /Points [ x1 y1 x0
-y0 x0 2 mul x1 sub y1 ] def } def
-/Grid { newpath /a 4 string def /b ED /c ED /n ED cvi dup 1 lt { pop 1 }
-if /s ED s div dup 0 eq { pop 1 } if /dy ED s div dup 0 eq { pop 1 } if
-/dx ED dy div round dy mul /y0 ED dx div round dx mul /x0 ED dy div
-round cvi /y2 ED dx div round cvi /x2 ED dy div round cvi /y1 ED dx div
-round cvi /x1 ED /h y2 y1 sub 0 gt { 1 } { -1 } ifelse def /w x2 x1 sub
-0 gt { 1 } { -1 } ifelse def b 0 gt { /z1 b 4 div CLW 2 div add def
-/Helvetica findfont b scalefont setfont /b b .95 mul CLW 2 div add def }
-if systemdict /setstrokeadjust known { true setstrokeadjust /t { } def }
-{ /t { transform 0.25 sub round 0.25 add exch 0.25 sub round 0.25 add
-exch itransform } bind def } ifelse gsave n 0 gt { 1 setlinecap [ 0 dy n
-div ] dy n div 2 div setdash } { 2 setlinecap } ifelse /i x1 def /f y1
-dy mul n 0 gt { dy n div 2 div h mul sub } if def /g y2 dy mul n 0 gt {
-dy n div 2 div h mul add } if def x2 x1 sub w mul 1 add dup 1000 gt {
-pop 1000 } if { i dx mul dup y0 moveto b 0 gt { gsave c i a cvs dup
-stringwidth pop /z2 ED w 0 gt {z1} {z1 z2 add neg} ifelse h 0 gt {b neg}
-{z1} ifelse rmoveto show grestore } if dup t f moveto g t L stroke /i i
-w add def } repeat grestore gsave n 0 gt
-% DG/SR modification begin - Nov. 7, 1997 - Patch 1
-%{ 1 setlinecap [ 0 dx n div ] dy n div 2 div setdash }
-{ 1 setlinecap [ 0 dx n div ] dx n div 2 div setdash }
-% DG/SR modification end
-{ 2 setlinecap } ifelse /i y1 def /f x1 dx mul
-n 0 gt { dx n div 2 div w mul sub } if def /g x2 dx mul n 0 gt { dx n
-div 2 div w mul add } if def y2 y1 sub h mul 1 add dup 1000 gt { pop
-1000 } if { newpath i dy mul dup x0 exch moveto b 0 gt { gsave c i a cvs
-dup stringwidth pop /z2 ED w 0 gt {z1 z2 add neg} {z1} ifelse h 0 gt
-{z1} {b neg} ifelse rmoveto show grestore } if dup f exch t moveto g
-exch t L stroke /i i h add def } repeat grestore } def
-/ArcArrow { /d ED /b ED /a ED gsave newpath 0 -1000 moveto clip newpath 0
-1 0 0 b grestore c mul /e ED pop pop pop r a e d PtoC y add exch x add
-exch r a PtoC y add exch x add exch b pop pop pop pop a e d CLW 8 div c
-mul neg d } def
-/Ellipse { /mtrx CM def T scale 0 0 1 5 3 roll arc mtrx setmatrix } def
-/Rot { CP CP translate 3 -1 roll neg rotate NET } def
-/RotBegin { tx@Dict /TMatrix known not { /TMatrix { } def /RAngle { 0 }
-def } if /TMatrix [ TMatrix CM ] cvx def /a ED a Rot /RAngle [ RAngle
-dup a add ] cvx def } def
-/RotEnd { /TMatrix [ TMatrix setmatrix ] cvx def /RAngle [ RAngle pop ]
-cvx def } def
-/PutCoor { gsave CP T CM STV exch exec moveto setmatrix CP grestore } def
-/PutBegin { /TMatrix [ TMatrix CM ] cvx def CP 4 2 roll T moveto } def
-/PutEnd { CP /TMatrix [ TMatrix setmatrix ] cvx def moveto } def
-/Uput { /a ED add 2 div /h ED 2 div /w ED /s a sin def /c a cos def /b s
-abs c abs 2 copy gt dup /q ED { pop } { exch pop } ifelse def /w1 c b
-div w mul def /h1 s b div h mul def q { w1 abs w sub dup c mul abs } {
-h1 abs h sub dup s mul abs } ifelse } def
-/UUput { /z ED abs /y ED /x ED q { x s div c mul abs y gt } { x c div s
-mul abs y gt } ifelse { x x mul y y mul sub z z mul add sqrt z add } { q
-{ x s div } { x c div } ifelse abs } ifelse a PtoC h1 add exch w1 add
-exch } def
-/BeginOL { dup (all) eq exch TheOL eq or { IfVisible not { Visible
-/IfVisible true def } if } { IfVisible { Invisible /IfVisible false def
-} if } ifelse } def
-/InitOL { /OLUnit [ 3000 3000 matrix defaultmatrix dtransform ] cvx def
-/Visible { CP OLUnit idtransform T moveto } def /Invisible { CP OLUnit
-neg exch neg exch idtransform T moveto } def /BOL { BeginOL } def
-/IfVisible true def } def
-end
-% END pstricks.pro
-
-%%EndProcSet
-%%BeginProcSet: pst-dots.pro 0 0
-%!PS-Adobe-2.0
-%%Title: Dot Font for PSTricks
-%%Creator: Timothy Van Zandt <tvz@Princeton.EDU>
-%%Creation Date: May 7, 1993
-%% Version 97 patch 1, 99/12/16
-%% Modified by Etienne Riga <etienne.riga@skynet.be> - Dec. 16, 1999
-%% to add /Diamond, /SolidDiamond and /BoldDiamond
-10 dict dup begin
- /FontType 3 def
- /FontMatrix [ .001 0 0 .001 0 0 ] def
- /FontBBox [ 0 0 0 0 ] def
- /Encoding 256 array def
- 0 1 255 { Encoding exch /.notdef put } for
- Encoding
- dup (b) 0 get /Bullet put
- dup (c) 0 get /Circle put
- dup (C) 0 get /BoldCircle put
- dup (u) 0 get /SolidTriangle put
- dup (t) 0 get /Triangle put
- dup (T) 0 get /BoldTriangle put
- dup (r) 0 get /SolidSquare put
- dup (s) 0 get /Square put
- dup (S) 0 get /BoldSquare put
- dup (q) 0 get /SolidPentagon put
- dup (p) 0 get /Pentagon put
- dup (P) 0 get /BoldPentagon put
-% DG/SR modification begin - Dec. 16, 1999 - From Etienne Riga
- dup (l) 0 get /SolidDiamond put
- dup (d) 0 get /Diamond put
- (D) 0 get /BoldDiamond put
-% DG/SR modification end
- /Metrics 13 dict def
- Metrics begin
- /Bullet 1000 def
- /Circle 1000 def
- /BoldCircle 1000 def
- /SolidTriangle 1344 def
- /Triangle 1344 def
- /BoldTriangle 1344 def
- /SolidSquare 886 def
- /Square 886 def
- /BoldSquare 886 def
- /SolidPentagon 1093.2 def
- /Pentagon 1093.2 def
- /BoldPentagon 1093.2 def
-% DG/SR modification begin - Dec. 16, 1999 - From Etienne Riga
- /SolidDiamond 1008 def
- /Diamond 1008 def
- /BoldDiamond 1008 def
-% DG/SR modification end
- /.notdef 0 def
- end
- /BBoxes 13 dict def
- BBoxes begin
- /Circle { -550 -550 550 550 } def
- /BoldCircle /Circle load def
- /Bullet /Circle load def
- /Triangle { -571.5 -330 571.5 660 } def
- /BoldTriangle /Triangle load def
- /SolidTriangle /Triangle load def
- /Square { -450 -450 450 450 } def
- /BoldSquare /Square load def
- /SolidSquare /Square load def
- /Pentagon { -546.6 -465 546.6 574.7 } def
- /BoldPentagon /Pentagon load def
- /SolidPentagon /Pentagon load def
-% DG/SR modification begin - Dec. 16, 1999 - From Etienne Riga
- /Diamond { -428.5 -742.5 428.5 742.5 } def
- /BoldDiamond /Diamond load def
- /SolidDiamond /Diamond load def
-% DG/SR modification end
- /.notdef { 0 0 0 0 } def
- end
- /CharProcs 20 dict def
- CharProcs begin
- /Adjust {
- 2 copy dtransform floor .5 add exch floor .5 add exch idtransform
- 3 -1 roll div 3 1 roll exch div exch scale
- } def
- /CirclePath { 0 0 500 0 360 arc closepath } def
- /Bullet { 500 500 Adjust CirclePath fill } def
- /Circle { 500 500 Adjust CirclePath .9 .9 scale CirclePath
- eofill } def
- /BoldCircle { 500 500 Adjust CirclePath .8 .8 scale CirclePath
- eofill } def
- /BoldCircle { CirclePath .8 .8 scale CirclePath eofill } def
- /TrianglePath { 0 660 moveto -571.5 -330 lineto 571.5 -330 lineto
- closepath } def
- /SolidTriangle { TrianglePath fill } def
- /Triangle { TrianglePath .85 .85 scale TrianglePath eofill } def
- /BoldTriangle { TrianglePath .7 .7 scale TrianglePath eofill } def
- /SquarePath { -450 450 moveto 450 450 lineto 450 -450 lineto
- -450 -450 lineto closepath } def
- /SolidSquare { SquarePath fill } def
- /Square { SquarePath .89 .89 scale SquarePath eofill } def
- /BoldSquare { SquarePath .78 .78 scale SquarePath eofill } def
- /PentagonPath {
- -337.8 -465 moveto
- 337.8 -465 lineto
- 546.6 177.6 lineto
- 0 574.7 lineto
- -546.6 177.6 lineto
- closepath
- } def
- /SolidPentagon { PentagonPath fill } def
- /Pentagon { PentagonPath .89 .89 scale PentagonPath eofill } def
- /BoldPentagon { PentagonPath .78 .78 scale PentagonPath eofill } def
-% DG/SR modification begin - Dec. 16, 1999 - From Etienne Riga
- /DiamondPath { 0 742.5 moveto -428.5 0 lineto 0 -742.5 lineto
- 428.5 0 lineto closepath } def
- /SolidDiamond { DiamondPath fill } def
- /Diamond { DiamondPath .85 .85 scale DiamondPath eofill } def
- /BoldDiamond { DiamondPath .7 .7 scale DiamondPath eofill } def
-% DG/SR modification end
- /.notdef { } def
- end
- /BuildGlyph {
- exch
- begin
- Metrics 1 index get exec 0
- BBoxes 3 index get exec
- setcachedevice
- CharProcs begin load exec end
- end
- } def
- /BuildChar {
- 1 index /Encoding get exch get
- 1 index /BuildGlyph get exec
- } bind def
-end
-/PSTricksDotFont exch definefont pop
-%END pst-dots.pro
-
-%%EndProcSet
-%%BeginProcSet: pst-node.pro 0 0
-%!
-% PostScript prologue for pst-node.tex.
-% Version 97 patch 1, 97/05/09.
-% For distribution, see pstricks.tex.
-%
-/tx@NodeDict 400 dict def tx@NodeDict begin
-tx@Dict begin /T /translate load def end
-/NewNode { gsave /next ED dict dup 3 1 roll def exch { dup 3 1 roll def }
-if begin tx@Dict begin STV CP T exec end /NodeMtrx CM def next end
-grestore } def
-/InitPnode { /Y ED /X ED /NodePos { NodeSep Cos mul NodeSep Sin mul } def
-} def
-/InitCnode { /r ED /Y ED /X ED /NodePos { NodeSep r add dup Cos mul exch
-Sin mul } def } def
-/GetRnodePos { Cos 0 gt { /dx r NodeSep add def } { /dx l NodeSep sub def
-} ifelse Sin 0 gt { /dy u NodeSep add def } { /dy d NodeSep sub def }
-ifelse dx Sin mul abs dy Cos mul abs gt { dy Cos mul Sin div dy } { dx
-dup Sin mul Cos Div } ifelse } def
-/InitRnode { /Y ED /X ED X sub /r ED /l X neg def Y add neg /d ED Y sub
-/u ED /NodePos { GetRnodePos } def } def
-/DiaNodePos { w h mul w Sin mul abs h Cos mul abs add Div NodeSep add dup
-Cos mul exch Sin mul } def
-/TriNodePos { Sin s lt { d NodeSep sub dup Cos mul Sin Div exch } { w h
-mul w Sin mul h Cos abs mul add Div NodeSep add dup Cos mul exch Sin mul
-} ifelse } def
-/InitTriNode { sub 2 div exch 2 div exch 2 copy T 2 copy 4 index index /d
-ED pop pop pop pop -90 mul rotate /NodeMtrx CM def /X 0 def /Y 0 def d
-sub abs neg /d ED d add /h ED 2 div h mul h d sub Div /w ED /s d w Atan
-sin def /NodePos { TriNodePos } def } def
-/OvalNodePos { /ww w NodeSep add def /hh h NodeSep add def Sin ww mul Cos
-hh mul Atan dup cos ww mul exch sin hh mul } def
-/GetCenter { begin X Y NodeMtrx transform CM itransform end } def
-/XYPos { dup sin exch cos Do /Cos ED /Sin ED /Dist ED Cos 0 gt { Dist
-Dist Sin mul Cos div } { Cos 0 lt { Dist neg Dist Sin mul Cos div neg }
-{ 0 Dist Sin mul } ifelse } ifelse Do } def
-/GetEdge { dup 0 eq { pop begin 1 0 NodeMtrx dtransform CM idtransform
-exch atan sub dup sin /Sin ED cos /Cos ED /NodeSep ED NodePos NodeMtrx
-dtransform CM idtransform end } { 1 eq {{exch}} {{}} ifelse /Do ED pop
-XYPos } ifelse } def
-/AddOffset { 1 index 0 eq { pop pop } { 2 copy 5 2 roll cos mul add 4 1
-roll sin mul sub exch } ifelse } def
-/GetEdgeA { NodeSepA AngleA NodeA NodeSepTypeA GetEdge OffsetA AngleA
-AddOffset yA add /yA1 ED xA add /xA1 ED } def
-/GetEdgeB { NodeSepB AngleB NodeB NodeSepTypeB GetEdge OffsetB AngleB
-AddOffset yB add /yB1 ED xB add /xB1 ED } def
-/GetArmA { ArmTypeA 0 eq { /xA2 ArmA AngleA cos mul xA1 add def /yA2 ArmA
-AngleA sin mul yA1 add def } { ArmTypeA 1 eq {{exch}} {{}} ifelse /Do ED
-ArmA AngleA XYPos OffsetA AngleA AddOffset yA add /yA2 ED xA add /xA2 ED
-} ifelse } def
-/GetArmB { ArmTypeB 0 eq { /xB2 ArmB AngleB cos mul xB1 add def /yB2 ArmB
-AngleB sin mul yB1 add def } { ArmTypeB 1 eq {{exch}} {{}} ifelse /Do ED
-ArmB AngleB XYPos OffsetB AngleB AddOffset yB add /yB2 ED xB add /xB2 ED
-} ifelse } def
-/InitNC { /b ED /a ED /NodeSepTypeB ED /NodeSepTypeA ED /NodeSepB ED
-/NodeSepA ED /OffsetB ED /OffsetA ED tx@NodeDict a known tx@NodeDict b
-known and dup { /NodeA a load def /NodeB b load def NodeA GetCenter /yA
-ED /xA ED NodeB GetCenter /yB ED /xB ED } if } def
-/LPutLine { 4 copy 3 -1 roll sub neg 3 1 roll sub Atan /NAngle ED 1 t sub
-mul 3 1 roll 1 t sub mul 4 1 roll t mul add /Y ED t mul add /X ED } def
-/LPutLines { mark LPutVar counttomark 2 div 1 sub /n ED t floor dup n gt
-{ pop n 1 sub /t 1 def } { dup t sub neg /t ED } ifelse cvi 2 mul { pop
-} repeat LPutLine cleartomark } def
-/BezierMidpoint { /y3 ED /x3 ED /y2 ED /x2 ED /y1 ED /x1 ED /y0 ED /x0 ED
-/t ED /cx x1 x0 sub 3 mul def /cy y1 y0 sub 3 mul def /bx x2 x1 sub 3
-mul cx sub def /by y2 y1 sub 3 mul cy sub def /ax x3 x0 sub cx sub bx
-sub def /ay y3 y0 sub cy sub by sub def ax t 3 exp mul bx t t mul mul
-add cx t mul add x0 add ay t 3 exp mul by t t mul mul add cy t mul add
-y0 add 3 ay t t mul mul mul 2 by t mul mul add cy add 3 ax t t mul mul
-mul 2 bx t mul mul add cx add atan /NAngle ED /Y ED /X ED } def
-/HPosBegin { yB yA ge { /t 1 t sub def } if /Y yB yA sub t mul yA add def
-} def
-/HPosEnd { /X Y yyA sub yyB yyA sub Div xxB xxA sub mul xxA add def
-/NAngle yyB yyA sub xxB xxA sub Atan def } def
-/HPutLine { HPosBegin /yyA ED /xxA ED /yyB ED /xxB ED HPosEnd } def
-/HPutLines { HPosBegin yB yA ge { /check { le } def } { /check { ge } def
-} ifelse /xxA xA def /yyA yA def mark xB yB LPutVar { dup Y check { exit
-} { /yyA ED /xxA ED } ifelse } loop /yyB ED /xxB ED cleartomark HPosEnd
-} def
-/VPosBegin { xB xA lt { /t 1 t sub def } if /X xB xA sub t mul xA add def
-} def
-/VPosEnd { /Y X xxA sub xxB xxA sub Div yyB yyA sub mul yyA add def
-/NAngle yyB yyA sub xxB xxA sub Atan def } def
-/VPutLine { VPosBegin /yyA ED /xxA ED /yyB ED /xxB ED VPosEnd } def
-/VPutLines { VPosBegin xB xA ge { /check { le } def } { /check { ge } def
-} ifelse /xxA xA def /yyA yA def mark xB yB LPutVar { 1 index X check {
-exit } { /yyA ED /xxA ED } ifelse } loop /yyB ED /xxB ED cleartomark
-VPosEnd } def
-/HPutCurve { gsave newpath /SaveLPutVar /LPutVar load def LPutVar 8 -2
-roll moveto curveto flattenpath /LPutVar [ {} {} {} {} pathforall ] cvx
-def grestore exec /LPutVar /SaveLPutVar load def } def
-/NCCoor { /AngleA yB yA sub xB xA sub Atan def /AngleB AngleA 180 add def
-GetEdgeA GetEdgeB /LPutVar [ xB1 yB1 xA1 yA1 ] cvx def /LPutPos {
-LPutVar LPutLine } def /HPutPos { LPutVar HPutLine } def /VPutPos {
-LPutVar VPutLine } def LPutVar } def
-/NCLine { NCCoor tx@Dict begin ArrowA CP 4 2 roll ArrowB lineto pop pop
-end } def
-/NCLines { false NArray n 0 eq { NCLine } { 2 copy yA sub exch xA sub
-Atan /AngleA ED n 2 mul dup index exch index yB sub exch xB sub Atan
-/AngleB ED GetEdgeA GetEdgeB /LPutVar [ xB1 yB1 n 2 mul 4 add 4 roll xA1
-yA1 ] cvx def mark LPutVar tx@Dict begin false Line end /LPutPos {
-LPutLines } def /HPutPos { HPutLines } def /VPutPos { VPutLines } def }
-ifelse } def
-/NCCurve { GetEdgeA GetEdgeB xA1 xB1 sub yA1 yB1 sub Pyth 2 div dup 3 -1
-roll mul /ArmA ED mul /ArmB ED /ArmTypeA 0 def /ArmTypeB 0 def GetArmA
-GetArmB xA2 yA2 xA1 yA1 tx@Dict begin ArrowA end xB2 yB2 xB1 yB1 tx@Dict
-begin ArrowB end curveto /LPutVar [ xA1 yA1 xA2 yA2 xB2 yB2 xB1 yB1 ]
-cvx def /LPutPos { t LPutVar BezierMidpoint } def /HPutPos { { HPutLines
-} HPutCurve } def /VPutPos { { VPutLines } HPutCurve } def } def
-/NCAngles { GetEdgeA GetEdgeB GetArmA GetArmB /mtrx AngleA matrix rotate
-def xA2 yA2 mtrx transform pop xB2 yB2 mtrx transform exch pop mtrx
-itransform /y0 ED /x0 ED mark ArmB 0 ne { xB1 yB1 } if xB2 yB2 x0 y0 xA2
-yA2 ArmA 0 ne { xA1 yA1 } if tx@Dict begin false Line end /LPutVar [ xB1
-yB1 xB2 yB2 x0 y0 xA2 yA2 xA1 yA1 ] cvx def /LPutPos { LPutLines } def
-/HPutPos { HPutLines } def /VPutPos { VPutLines } def } def
-/NCAngle { GetEdgeA GetEdgeB GetArmB /mtrx AngleA matrix rotate def xB2
-yB2 mtrx itransform pop xA1 yA1 mtrx itransform exch pop mtrx transform
-/y0 ED /x0 ED mark ArmB 0 ne { xB1 yB1 } if xB2 yB2 x0 y0 xA1 yA1
-tx@Dict begin false Line end /LPutVar [ xB1 yB1 xB2 yB2 x0 y0 xA1 yA1 ]
-cvx def /LPutPos { LPutLines } def /HPutPos { HPutLines } def /VPutPos {
-VPutLines } def } def
-/NCBar { GetEdgeA GetEdgeB GetArmA GetArmB /mtrx AngleA matrix rotate def
-xA2 yA2 mtrx itransform pop xB2 yB2 mtrx itransform pop sub dup 0 mtrx
-transform 3 -1 roll 0 gt { /yB2 exch yB2 add def /xB2 exch xB2 add def }
-{ /yA2 exch neg yA2 add def /xA2 exch neg xA2 add def } ifelse mark ArmB
-0 ne { xB1 yB1 } if xB2 yB2 xA2 yA2 ArmA 0 ne { xA1 yA1 } if tx@Dict
-begin false Line end /LPutVar [ xB1 yB1 xB2 yB2 xA2 yA2 xA1 yA1 ] cvx
-def /LPutPos { LPutLines } def /HPutPos { HPutLines } def /VPutPos {
-VPutLines } def } def
-/NCDiag { GetEdgeA GetEdgeB GetArmA GetArmB mark ArmB 0 ne { xB1 yB1 } if
-xB2 yB2 xA2 yA2 ArmA 0 ne { xA1 yA1 } if tx@Dict begin false Line end
-/LPutVar [ xB1 yB1 xB2 yB2 xA2 yA2 xA1 yA1 ] cvx def /LPutPos {
-LPutLines } def /HPutPos { HPutLines } def /VPutPos { VPutLines } def }
-def
-/NCDiagg { GetEdgeA GetArmA yB yA2 sub xB xA2 sub Atan 180 add /AngleB ED
-GetEdgeB mark xB1 yB1 xA2 yA2 ArmA 0 ne { xA1 yA1 } if tx@Dict begin
-false Line end /LPutVar [ xB1 yB1 xA2 yA2 xA1 yA1 ] cvx def /LPutPos {
-LPutLines } def /HPutPos { HPutLines } def /VPutPos { VPutLines } def }
-def
-/NCLoop { GetEdgeA GetEdgeB GetArmA GetArmB /mtrx AngleA matrix rotate
-def xA2 yA2 mtrx transform loopsize add /yA3 ED /xA3 ED /xB3 xB2 yB2
-mtrx transform pop def xB3 yA3 mtrx itransform /yB3 ED /xB3 ED xA3 yA3
-mtrx itransform /yA3 ED /xA3 ED mark ArmB 0 ne { xB1 yB1 } if xB2 yB2
-xB3 yB3 xA3 yA3 xA2 yA2 ArmA 0 ne { xA1 yA1 } if tx@Dict begin false
-Line end /LPutVar [ xB1 yB1 xB2 yB2 xB3 yB3 xA3 yA3 xA2 yA2 xA1 yA1 ]
-cvx def /LPutPos { LPutLines } def /HPutPos { HPutLines } def /VPutPos {
-VPutLines } def } def
-% DG/SR modification begin - May 9, 1997 - Patch 1
-%/NCCircle { 0 0 NodesepA nodeA \tx@GetEdge pop xA sub 2 div dup 2 exp r
-%r mul sub abs sqrt atan 2 mul /a ED r AngleA 90 add PtoC yA add exch xA add
-%exch 2 copy /LPutVar [ 4 2 roll r AngleA ] cvx def /LPutPos { LPutVar t 360
-%mul add dup 5 1 roll 90 sub \tx@PtoC 3 -1 roll add /Y ED add /X ED /NAngle ED
-/NCCircle { NodeSepA 0 NodeA 0 GetEdge pop 2 div dup 2 exp r
-r mul sub abs sqrt atan 2 mul /a ED r AngleA 90 add PtoC yA add exch xA add
-exch 2 copy /LPutVar [ 4 2 roll r AngleA ] cvx def /LPutPos { LPutVar t 360
-mul add dup 5 1 roll 90 sub PtoC 3 -1 roll add /Y ED add /X ED /NAngle ED
-% DG/SR modification end
-} def /HPutPos { LPutPos } def /VPutPos { LPutPos } def r AngleA 90 sub a add
-AngleA 270 add a sub tx@Dict begin /angleB ED /angleA ED /r ED /c 57.2957 r
-Div def /y ED /x ED } def
-/NCBox { /d ED /h ED /AngleB yB yA sub xB xA sub Atan def /AngleA AngleB
-180 add def GetEdgeA GetEdgeB /dx d AngleB sin mul def /dy d AngleB cos
-mul neg def /hx h AngleB sin mul neg def /hy h AngleB cos mul def
-/LPutVar [ xA1 hx add yA1 hy add xB1 hx add yB1 hy add xB1 dx add yB1 dy
-add xA1 dx add yA1 dy add ] cvx def /LPutPos { LPutLines } def /HPutPos
-{ xB yB xA yA LPutLine } def /VPutPos { HPutPos } def mark LPutVar
-tx@Dict begin false Polygon end } def
-/NCArcBox { /l ED neg /d ED /h ED /a ED /AngleA yB yA sub xB xA sub Atan
-def /AngleB AngleA 180 add def /tA AngleA a sub 90 add def /tB tA a 2
-mul add def /r xB xA sub tA cos tB cos sub Div dup 0 eq { pop 1 } if def
-/x0 xA r tA cos mul add def /y0 yA r tA sin mul add def /c 57.2958 r div
-def /AngleA AngleA a sub 180 add def /AngleB AngleB a add 180 add def
-GetEdgeA GetEdgeB /AngleA tA 180 add yA yA1 sub xA xA1 sub Pyth c mul
-sub def /AngleB tB 180 add yB yB1 sub xB xB1 sub Pyth c mul add def l 0
-eq { x0 y0 r h add AngleA AngleB arc x0 y0 r d add AngleB AngleA arcn }
-{ x0 y0 translate /tA AngleA l c mul add def /tB AngleB l c mul sub def
-0 0 r h add tA tB arc r h add AngleB PtoC r d add AngleB PtoC 2 copy 6 2
-roll l arcto 4 { pop } repeat r d add tB PtoC l arcto 4 { pop } repeat 0
-0 r d add tB tA arcn r d add AngleA PtoC r h add AngleA PtoC 2 copy 6 2
-roll l arcto 4 { pop } repeat r h add tA PtoC l arcto 4 { pop } repeat }
-ifelse closepath /LPutVar [ x0 y0 r AngleA AngleB h d ] cvx def /LPutPos
-{ LPutVar /d ED /h ED /AngleB ED /AngleA ED /r ED /y0 ED /x0 ED t 1 le {
-r h add AngleA 1 t sub mul AngleB t mul add dup 90 add /NAngle ED PtoC }
-{ t 2 lt { /NAngle AngleB 180 add def r 2 t sub h mul t 1 sub d mul add
-add AngleB PtoC } { t 3 lt { r d add AngleB 3 t sub mul AngleA 2 t sub
-mul add dup 90 sub /NAngle ED PtoC } { /NAngle AngleA 180 add def r 4 t
-sub d mul t 3 sub h mul add add AngleA PtoC } ifelse } ifelse } ifelse
-y0 add /Y ED x0 add /X ED } def /HPutPos { LPutPos } def /VPutPos {
-LPutPos } def } def
-/Tfan { /AngleA yB yA sub xB xA sub Atan def GetEdgeA w xA1 xB sub yA1 yB
-sub Pyth Pyth w Div CLW 2 div mul 2 div dup AngleA sin mul yA1 add /yA1
-ED AngleA cos mul xA1 add /xA1 ED /LPutVar [ xA1 yA1 m { xB w add yB xB
-w sub yB } { xB yB w sub xB yB w add } ifelse xA1 yA1 ] cvx def /LPutPos
-{ LPutLines } def /VPutPos@ { LPutVar flag { 8 4 roll pop pop pop pop }
-{ pop pop pop pop 4 2 roll } ifelse } def /VPutPos { VPutPos@ VPutLine }
-def /HPutPos { VPutPos@ HPutLine } def mark LPutVar tx@Dict begin
-/ArrowA { moveto } def /ArrowB { } def false Line closepath end } def
-/LPutCoor { NAngle tx@Dict begin /NAngle ED end gsave CM STV CP Y sub neg
-exch X sub neg exch moveto setmatrix CP grestore } def
-/LPut { tx@NodeDict /LPutPos known { LPutPos } { CP /Y ED /X ED /NAngle 0
-def } ifelse LPutCoor } def
-/HPutAdjust { Sin Cos mul 0 eq { 0 } { d Cos mul Sin div flag not { neg }
-if h Cos mul Sin div flag { neg } if 2 copy gt { pop } { exch pop }
-ifelse } ifelse s add flag { r add neg } { l add } ifelse X add /X ED }
-def
-/VPutAdjust { Sin Cos mul 0 eq { 0 } { l Sin mul Cos div flag { neg } if
-r Sin mul Cos div flag not { neg } if 2 copy gt { pop } { exch pop }
-ifelse } ifelse s add flag { d add } { h add neg } ifelse Y add /Y ED }
-def
-end
-% END pst-node.pro
-
-%%EndProcSet
-%%BeginProcSet: 8r.enc 0 0
-% File 8r.enc TeX Base 1 Encoding Revision 2.0 2002-10-30
-%
-% @@psencodingfile@{
-% author = "S. Rahtz, P. MacKay, Alan Jeffrey, B. Horn, K. Berry,
-% W. Schmidt, P. Lehman",
-% version = "2.0",
-% date = "30 October 2002",
-% filename = "8r.enc",
-% email = "tex-fonts@@tug.org",
-% docstring = "This is the encoding vector for Type1 and TrueType
-% fonts to be used with TeX. This file is part of the
-% PSNFSS bundle, version 9"
-% @}
-%
-% The idea is to have all the characters normally included in Type 1 fonts
-% available for typesetting. This is effectively the characters in Adobe
-% Standard encoding, ISO Latin 1, Windows ANSI including the euro symbol,
-% MacRoman, and some extra characters from Lucida.
-%
-% Character code assignments were made as follows:
-%
-% (1) the Windows ANSI characters are almost all in their Windows ANSI
-% positions, because some Windows users cannot easily reencode the
-% fonts, and it makes no difference on other systems. The only Windows
-% ANSI characters not available are those that make no sense for
-% typesetting -- rubout (127 decimal), nobreakspace (160), softhyphen
-% (173). quotesingle and grave are moved just because it's such an
-% irritation not having them in TeX positions.
-%
-% (2) Remaining characters are assigned arbitrarily to the lower part
-% of the range, avoiding 0, 10 and 13 in case we meet dumb software.
-%
-% (3) Y&Y Lucida Bright includes some extra text characters; in the
-% hopes that other PostScript fonts, perhaps created for public
-% consumption, will include them, they are included starting at 0x12.
-% These are /dotlessj /ff /ffi /ffl.
-%
-% (4) hyphen appears twice for compatibility with both ASCII and Windows.
-%
-% (5) /Euro was assigned to 128, as in Windows ANSI
-%
-% (6) Missing characters from MacRoman encoding incorporated as follows:
-%
-% PostScript MacRoman TeXBase1
-% -------------- -------------- --------------
-% /notequal 173 0x16
-% /infinity 176 0x17
-% /lessequal 178 0x18
-% /greaterequal 179 0x19
-% /partialdiff 182 0x1A
-% /summation 183 0x1B
-% /product 184 0x1C
-% /pi 185 0x1D
-% /integral 186 0x81
-% /Omega 189 0x8D
-% /radical 195 0x8E
-% /approxequal 197 0x8F
-% /Delta 198 0x9D
-% /lozenge 215 0x9E
-%
-/TeXBase1Encoding [
-% 0x00
- /.notdef /dotaccent /fi /fl
- /fraction /hungarumlaut /Lslash /lslash
- /ogonek /ring /.notdef /breve
- /minus /.notdef /Zcaron /zcaron
-% 0x10
- /caron /dotlessi /dotlessj /ff
- /ffi /ffl /notequal /infinity
- /lessequal /greaterequal /partialdiff /summation
- /product /pi /grave /quotesingle
-% 0x20
- /space /exclam /quotedbl /numbersign
- /dollar /percent /ampersand /quoteright
- /parenleft /parenright /asterisk /plus
- /comma /hyphen /period /slash
-% 0x30
- /zero /one /two /three
- /four /five /six /seven
- /eight /nine /colon /semicolon
- /less /equal /greater /question
-% 0x40
- /at /A /B /C
- /D /E /F /G
- /H /I /J /K
- /L /M /N /O
-% 0x50
- /P /Q /R /S
- /T /U /V /W
- /X /Y /Z /bracketleft
- /backslash /bracketright /asciicircum /underscore
-% 0x60
- /quoteleft /a /b /c
- /d /e /f /g
- /h /i /j /k
- /l /m /n /o
-% 0x70
- /p /q /r /s
- /t /u /v /w
- /x /y /z /braceleft
- /bar /braceright /asciitilde /.notdef
-% 0x80
- /Euro /integral /quotesinglbase /florin
- /quotedblbase /ellipsis /dagger /daggerdbl
- /circumflex /perthousand /Scaron /guilsinglleft
- /OE /Omega /radical /approxequal
-% 0x90
- /.notdef /.notdef /.notdef /quotedblleft
- /quotedblright /bullet /endash /emdash
- /tilde /trademark /scaron /guilsinglright
- /oe /Delta /lozenge /Ydieresis
-% 0xA0
- /.notdef /exclamdown /cent /sterling
- /currency /yen /brokenbar /section
- /dieresis /copyright /ordfeminine /guillemotleft
- /logicalnot /hyphen /registered /macron
-% 0xD0
- /degree /plusminus /twosuperior /threesuperior
- /acute /mu /paragraph /periodcentered
- /cedilla /onesuperior /ordmasculine /guillemotright
- /onequarter /onehalf /threequarters /questiondown
-% 0xC0
- /Agrave /Aacute /Acircumflex /Atilde
- /Adieresis /Aring /AE /Ccedilla
- /Egrave /Eacute /Ecircumflex /Edieresis
- /Igrave /Iacute /Icircumflex /Idieresis
-% 0xD0
- /Eth /Ntilde /Ograve /Oacute
- /Ocircumflex /Otilde /Odieresis /multiply
- /Oslash /Ugrave /Uacute /Ucircumflex
- /Udieresis /Yacute /Thorn /germandbls
-% 0xE0
- /agrave /aacute /acircumflex /atilde
- /adieresis /aring /ae /ccedilla
- /egrave /eacute /ecircumflex /edieresis
- /igrave /iacute /icircumflex /idieresis
-% 0xF0
- /eth /ntilde /ograve /oacute
- /ocircumflex /otilde /odieresis /divide
- /oslash /ugrave /uacute /ucircumflex
- /udieresis /yacute /thorn /ydieresis
-] def
-
-
-%%EndProcSet
-%%BeginProcSet: texps.pro 0 0
-%!
-TeXDict begin/rf{findfont dup length 1 add dict begin{1 index/FID ne 2
-index/UniqueID ne and{def}{pop pop}ifelse}forall[1 index 0 6 -1 roll
-exec 0 exch 5 -1 roll VResolution Resolution div mul neg 0 0]FontType 0
-ne{/Metrics exch def dict begin Encoding{exch dup type/integertype ne{
-pop pop 1 sub dup 0 le{pop}{[}ifelse}{FontMatrix 0 get div Metrics 0 get
-div def}ifelse}forall Metrics/Metrics currentdict end def}{{1 index type
-/nametype eq{exit}if exch pop}loop}ifelse[2 index currentdict end
-definefont 3 -1 roll makefont/setfont cvx]cvx def}def/ObliqueSlant{dup
-sin S cos div neg}B/SlantFont{4 index mul add}def/ExtendFont{3 -1 roll
-mul exch}def/ReEncodeFont{CharStrings rcheck{/Encoding false def dup[
-exch{dup CharStrings exch known not{pop/.notdef/Encoding true def}if}
-forall Encoding{]exch pop}{cleartomark}ifelse}if/Encoding exch def}def
-end
-
-%%EndProcSet
-%%BeginProcSet: special.pro 0 0
-%!
-TeXDict begin/SDict 200 dict N SDict begin/@SpecialDefaults{/hs 612 N
-/vs 792 N/ho 0 N/vo 0 N/hsc 1 N/vsc 1 N/ang 0 N/CLIP 0 N/rwiSeen false N
-/rhiSeen false N/letter{}N/note{}N/a4{}N/legal{}N}B/@scaleunit 100 N
-/@hscale{@scaleunit div/hsc X}B/@vscale{@scaleunit div/vsc X}B/@hsize{
-/hs X/CLIP 1 N}B/@vsize{/vs X/CLIP 1 N}B/@clip{/CLIP 2 N}B/@hoffset{/ho
-X}B/@voffset{/vo X}B/@angle{/ang X}B/@rwi{10 div/rwi X/rwiSeen true N}B
-/@rhi{10 div/rhi X/rhiSeen true N}B/@llx{/llx X}B/@lly{/lly X}B/@urx{
-/urx X}B/@ury{/ury X}B/magscale true def end/@MacSetUp{userdict/md known
-{userdict/md get type/dicttype eq{userdict begin md length 10 add md
-maxlength ge{/md md dup length 20 add dict copy def}if end md begin
-/letter{}N/note{}N/legal{}N/od{txpose 1 0 mtx defaultmatrix dtransform S
-atan/pa X newpath clippath mark{transform{itransform moveto}}{transform{
-itransform lineto}}{6 -2 roll transform 6 -2 roll transform 6 -2 roll
-transform{itransform 6 2 roll itransform 6 2 roll itransform 6 2 roll
-curveto}}{{closepath}}pathforall newpath counttomark array astore/gc xdf
-pop ct 39 0 put 10 fz 0 fs 2 F/|______Courier fnt invertflag{PaintBlack}
-if}N/txpose{pxs pys scale ppr aload pop por{noflips{pop S neg S TR pop 1
--1 scale}if xflip yflip and{pop S neg S TR 180 rotate 1 -1 scale ppr 3
-get ppr 1 get neg sub neg ppr 2 get ppr 0 get neg sub neg TR}if xflip
-yflip not and{pop S neg S TR pop 180 rotate ppr 3 get ppr 1 get neg sub
-neg 0 TR}if yflip xflip not and{ppr 1 get neg ppr 0 get neg TR}if}{
-noflips{TR pop pop 270 rotate 1 -1 scale}if xflip yflip and{TR pop pop
-90 rotate 1 -1 scale ppr 3 get ppr 1 get neg sub neg ppr 2 get ppr 0 get
-neg sub neg TR}if xflip yflip not and{TR pop pop 90 rotate ppr 3 get ppr
-1 get neg sub neg 0 TR}if yflip xflip not and{TR pop pop 270 rotate ppr
-2 get ppr 0 get neg sub neg 0 S TR}if}ifelse scaleby96{ppr aload pop 4
--1 roll add 2 div 3 1 roll add 2 div 2 copy TR .96 dup scale neg S neg S
-TR}if}N/cp{pop pop showpage pm restore}N end}if}if}N/normalscale{
-Resolution 72 div VResolution 72 div neg scale magscale{DVImag dup scale
-}if 0 setgray}N/psfts{S 65781.76 div N}N/startTexFig{/psf$SavedState
-save N userdict maxlength dict begin/magscale true def normalscale
-currentpoint TR/psf$ury psfts/psf$urx psfts/psf$lly psfts/psf$llx psfts
-/psf$y psfts/psf$x psfts currentpoint/psf$cy X/psf$cx X/psf$sx psf$x
-psf$urx psf$llx sub div N/psf$sy psf$y psf$ury psf$lly sub div N psf$sx
-psf$sy scale psf$cx psf$sx div psf$llx sub psf$cy psf$sy div psf$ury sub
-TR/showpage{}N/erasepage{}N/setpagedevice{pop}N/copypage{}N/p 3 def
-@MacSetUp}N/doclip{psf$llx psf$lly psf$urx psf$ury currentpoint 6 2 roll
-newpath 4 copy 4 2 roll moveto 6 -1 roll S lineto S lineto S lineto
-closepath clip newpath moveto}N/endTexFig{end psf$SavedState restore}N
-/@beginspecial{SDict begin/SpecialSave save N gsave normalscale
-currentpoint TR @SpecialDefaults count/ocount X/dcount countdictstack N}
-N/@setspecial{CLIP 1 eq{newpath 0 0 moveto hs 0 rlineto 0 vs rlineto hs
-neg 0 rlineto closepath clip}if ho vo TR hsc vsc scale ang rotate
-rwiSeen{rwi urx llx sub div rhiSeen{rhi ury lly sub div}{dup}ifelse
-scale llx neg lly neg TR}{rhiSeen{rhi ury lly sub div dup scale llx neg
-lly neg TR}if}ifelse CLIP 2 eq{newpath llx lly moveto urx lly lineto urx
-ury lineto llx ury lineto closepath clip}if/showpage{}N/erasepage{}N
-/setpagedevice{pop}N/copypage{}N newpath}N/@endspecial{count ocount sub{
-pop}repeat countdictstack dcount sub{end}repeat grestore SpecialSave
-restore end}N/@defspecial{SDict begin}N/@fedspecial{end}B/li{lineto}B
-/rl{rlineto}B/rc{rcurveto}B/np{/SaveX currentpoint/SaveY X N 1
-setlinecap newpath}N/st{stroke SaveX SaveY moveto}N/fil{fill SaveX SaveY
-moveto}N/ellipse{/endangle X/startangle X/yrad X/xrad X/savematrix
-matrix currentmatrix N TR xrad yrad scale 0 0 1 startangle endangle arc
-savematrix setmatrix}N end
-
-%%EndProcSet
-%%BeginProcSet: color.pro 0 0
-%!
-TeXDict begin/setcmykcolor where{pop}{/setcmykcolor{dup 10 eq{pop
-setrgbcolor}{1 sub 4 1 roll 3{3 index add neg dup 0 lt{pop 0}if 3 1 roll
-}repeat setrgbcolor pop}ifelse}B}ifelse/TeXcolorcmyk{setcmykcolor}def
-/TeXcolorrgb{setrgbcolor}def/TeXcolorgrey{setgray}def/TeXcolorgray{
-setgray}def/TeXcolorhsb{sethsbcolor}def/currentcmykcolor where{pop}{
-/currentcmykcolor{currentrgbcolor 10}B}ifelse/DC{exch dup userdict exch
-known{pop pop}{X}ifelse}B/GreenYellow{0.15 0 0.69 0 setcmykcolor}DC
-/Yellow{0 0 1 0 setcmykcolor}DC/Goldenrod{0 0.10 0.84 0 setcmykcolor}DC
-/Dandelion{0 0.29 0.84 0 setcmykcolor}DC/Apricot{0 0.32 0.52 0
-setcmykcolor}DC/Peach{0 0.50 0.70 0 setcmykcolor}DC/Melon{0 0.46 0.50 0
-setcmykcolor}DC/YellowOrange{0 0.42 1 0 setcmykcolor}DC/Orange{0 0.61
-0.87 0 setcmykcolor}DC/BurntOrange{0 0.51 1 0 setcmykcolor}DC
-/Bittersweet{0 0.75 1 0.24 setcmykcolor}DC/RedOrange{0 0.77 0.87 0
-setcmykcolor}DC/Mahogany{0 0.85 0.87 0.35 setcmykcolor}DC/Maroon{0 0.87
-0.68 0.32 setcmykcolor}DC/BrickRed{0 0.89 0.94 0.28 setcmykcolor}DC/Red{
-0 1 1 0 setcmykcolor}DC/OrangeRed{0 1 0.50 0 setcmykcolor}DC/RubineRed{
-0 1 0.13 0 setcmykcolor}DC/WildStrawberry{0 0.96 0.39 0 setcmykcolor}DC
-/Salmon{0 0.53 0.38 0 setcmykcolor}DC/CarnationPink{0 0.63 0 0
-setcmykcolor}DC/Magenta{0 1 0 0 setcmykcolor}DC/VioletRed{0 0.81 0 0
-setcmykcolor}DC/Rhodamine{0 0.82 0 0 setcmykcolor}DC/Mulberry{0.34 0.90
-0 0.02 setcmykcolor}DC/RedViolet{0.07 0.90 0 0.34 setcmykcolor}DC
-/Fuchsia{0.47 0.91 0 0.08 setcmykcolor}DC/Lavender{0 0.48 0 0
-setcmykcolor}DC/Thistle{0.12 0.59 0 0 setcmykcolor}DC/Orchid{0.32 0.64 0
-0 setcmykcolor}DC/DarkOrchid{0.40 0.80 0.20 0 setcmykcolor}DC/Purple{
-0.45 0.86 0 0 setcmykcolor}DC/Plum{0.50 1 0 0 setcmykcolor}DC/Violet{
-0.79 0.88 0 0 setcmykcolor}DC/RoyalPurple{0.75 0.90 0 0 setcmykcolor}DC
-/BlueViolet{0.86 0.91 0 0.04 setcmykcolor}DC/Periwinkle{0.57 0.55 0 0
-setcmykcolor}DC/CadetBlue{0.62 0.57 0.23 0 setcmykcolor}DC
-/CornflowerBlue{0.65 0.13 0 0 setcmykcolor}DC/MidnightBlue{0.98 0.13 0
-0.43 setcmykcolor}DC/NavyBlue{0.94 0.54 0 0 setcmykcolor}DC/RoyalBlue{1
-0.50 0 0 setcmykcolor}DC/Blue{1 1 0 0 setcmykcolor}DC/Cerulean{0.94 0.11
-0 0 setcmykcolor}DC/Cyan{1 0 0 0 setcmykcolor}DC/ProcessBlue{0.96 0 0 0
-setcmykcolor}DC/SkyBlue{0.62 0 0.12 0 setcmykcolor}DC/Turquoise{0.85 0
-0.20 0 setcmykcolor}DC/TealBlue{0.86 0 0.34 0.02 setcmykcolor}DC
-/Aquamarine{0.82 0 0.30 0 setcmykcolor}DC/BlueGreen{0.85 0 0.33 0
-setcmykcolor}DC/Emerald{1 0 0.50 0 setcmykcolor}DC/JungleGreen{0.99 0
-0.52 0 setcmykcolor}DC/SeaGreen{0.69 0 0.50 0 setcmykcolor}DC/Green{1 0
-1 0 setcmykcolor}DC/ForestGreen{0.91 0 0.88 0.12 setcmykcolor}DC
-/PineGreen{0.92 0 0.59 0.25 setcmykcolor}DC/LimeGreen{0.50 0 1 0
-setcmykcolor}DC/YellowGreen{0.44 0 0.74 0 setcmykcolor}DC/SpringGreen{
-0.26 0 0.76 0 setcmykcolor}DC/OliveGreen{0.64 0 0.95 0.40 setcmykcolor}
-DC/RawSienna{0 0.72 1 0.45 setcmykcolor}DC/Sepia{0 0.83 1 0.70
-setcmykcolor}DC/Brown{0 0.81 1 0.60 setcmykcolor}DC/Tan{0.14 0.42 0.56 0
-setcmykcolor}DC/Gray{0 0 0 0.50 setcmykcolor}DC/Black{0 0 0 1
-setcmykcolor}DC/White{0 0 0 0 setcmykcolor}DC end
-
-%%EndProcSet
-%%BeginFont: CMMI8
-%!PS-AdobeFont-1.1: CMMI8 1.100
-%%CreationDate: 1996 Jul 23 07:53:54
-
-% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
-
-11 dict begin
-/FontInfo 7 dict dup begin
-/version (1.100) readonly def
-/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
-/FullName (CMMI8) readonly def
-/FamilyName (Computer Modern) readonly def
-/Weight (Medium) readonly def
-/ItalicAngle -14.04 def
-/isFixedPitch false def
-end readonly def
-/FontName /CMMI8 def
-/PaintType 0 def
-/FontType 1 def
-/FontMatrix [0.001 0 0 0.001 0 0] readonly def
-/Encoding 256 array
-0 1 255 {1 index exch /.notdef put} for
-dup 161 /Gamma put
-dup 162 /Delta put
-dup 163 /Theta put
-dup 164 /Lambda put
-dup 165 /Xi put
-dup 166 /Pi put
-dup 167 /Sigma put
-dup 168 /Upsilon put
-dup 169 /Phi put
-dup 170 /Psi put
-dup 173 /Omega put
-dup 174 /alpha put
-dup 175 /beta put
-dup 176 /gamma put
-dup 177 /delta put
-dup 178 /epsilon1 put
-dup 179 /zeta put
-dup 180 /eta put
-dup 181 /theta put
-dup 182 /iota put
-dup 183 /kappa put
-dup 184 /lambda put
-dup 185 /mu put
-dup 186 /nu put
-dup 187 /xi put
-dup 188 /pi put
-dup 189 /rho put
-dup 190 /sigma put
-dup 191 /tau put
-dup 192 /upsilon put
-dup 193 /phi put
-dup 194 /chi put
-dup 195 /psi put
-dup 196 /tie put
-dup 0 /Gamma put
-dup 1 /Delta put
-dup 2 /Theta put
-dup 3 /Lambda put
-dup 4 /Xi put
-dup 5 /Pi put
-dup 6 /Sigma put
-dup 7 /Upsilon put
-dup 8 /Phi put
-dup 9 /Psi put
-dup 10 /Omega put
-dup 11 /alpha put
-dup 12 /beta put
-dup 13 /gamma put
-dup 14 /delta put
-dup 15 /epsilon1 put
-dup 16 /zeta put
-dup 17 /eta put
-dup 18 /theta put
-dup 19 /iota put
-dup 20 /kappa put
-dup 21 /lambda put
-dup 22 /mu put
-dup 23 /nu put
-dup 24 /xi put
-dup 25 /pi put
-dup 26 /rho put
-dup 27 /sigma put
-dup 28 /tau put
-dup 29 /upsilon put
-dup 30 /phi put
-dup 31 /chi put
-dup 32 /psi put
-dup 33 /omega put
-dup 34 /epsilon put
-dup 35 /theta1 put
-dup 36 /pi1 put
-dup 37 /rho1 put
-dup 38 /sigma1 put
-dup 39 /phi1 put
-dup 40 /arrowlefttophalf put
-dup 41 /arrowleftbothalf put
-dup 42 /arrowrighttophalf put
-dup 43 /arrowrightbothalf put
-dup 44 /arrowhookleft put
-dup 45 /arrowhookright put
-dup 46 /triangleright put
-dup 47 /triangleleft put
-dup 48 /zerooldstyle put
-dup 49 /oneoldstyle put
-dup 50 /twooldstyle put
-dup 51 /threeoldstyle put
-dup 52 /fouroldstyle put
-dup 53 /fiveoldstyle put
-dup 54 /sixoldstyle put
-dup 55 /sevenoldstyle put
-dup 56 /eightoldstyle put
-dup 57 /nineoldstyle put
-dup 58 /period put
-dup 59 /comma put
-dup 60 /less put
-dup 61 /slash put
-dup 62 /greater put
-dup 63 /star put
-dup 64 /partialdiff put
-dup 65 /A put
-dup 66 /B put
-dup 67 /C put
-dup 68 /D put
-dup 69 /E put
-dup 70 /F put
-dup 71 /G put
-dup 72 /H put
-dup 73 /I put
-dup 74 /J put
-dup 75 /K put
-dup 76 /L put
-dup 77 /M put
-dup 78 /N put
-dup 79 /O put
-dup 80 /P put
-dup 81 /Q put
-dup 82 /R put
-dup 83 /S put
-dup 84 /T put
-dup 85 /U put
-dup 86 /V put
-dup 87 /W put
-dup 88 /X put
-dup 89 /Y put
-dup 90 /Z put
-dup 91 /flat put
-dup 92 /natural put
-dup 93 /sharp put
-dup 94 /slurbelow put
-dup 95 /slurabove put
-dup 96 /lscript put
-dup 97 /a put
-dup 98 /b put
-dup 99 /c put
-dup 100 /d put
-dup 101 /e put
-dup 102 /f put
-dup 103 /g put
-dup 104 /h put
-dup 105 /i put
-dup 106 /j put
-dup 107 /k put
-dup 108 /l put
-dup 109 /m put
-dup 110 /n put
-dup 111 /o put
-dup 112 /p put
-dup 113 /q put
-dup 114 /r put
-dup 115 /s put
-dup 116 /t put
-dup 117 /u put
-dup 118 /v put
-dup 119 /w put
-dup 120 /x put
-dup 121 /y put
-dup 122 /z put
-dup 123 /dotlessi put
-dup 124 /dotlessj put
-dup 125 /weierstrass put
-dup 126 /vector put
-dup 127 /tie put
-dup 128 /psi put
-dup 160 /space put
-readonly def
-/FontBBox{-24 -250 1110 750}readonly def
-/UniqueID 5087383 def
-currentdict end
-currentfile eexec
-
-80347982AB3942D930E069A70D0D48311D725E830D1C76FBA12E12486E989C9874C2B527F092
-5722787027F44470D484262C360CDFDDDF3657533A57BB16F73048BFBBFCB73A650484015441
-FDC837ADD94AC8FBD2022E3EC8F115D4B4BB7B7F15388F22CC6198EFE768BD9FCEB3446EE4A8
-DC27D6CD152485384EF5F59381FFDA43F2D20C8FB08AA27AB2015B774DB10DACFDCD33E60F17
-8C461553146AB427BDD7DA12534BA078AD3D780414930E72218B3075925CE1192F11FC8530FC
-D5E3038E3A6A6DB2DCFBAE3B4653E7E02730314E02B54A1E296D2BEF8A79411D9225DAD7B4E6
-D6F9CF0688B69BA21193BF1495807E7A196CF14C95A4E02F9CD2DA8CDB2546C6DF52E5247459
-92E18D9FF87AA25E4E1800BBE4EBB357C6EF55ED6D036D3A00C1EE8073266C21D2F0AC85D656
-ABF61D7E5A4FA87DA8EC3B5329E434D0D2ADAB706B42A2E5331BE5295399D803CCAC03F631F0
-1F39A022FCDF63486B687D15EF284A77DEF7FDE4898543E7B5F7EC267756103E477F547CFB8D
-2311C4B009DEFF56085F5D419697AF1846C8B88C1BBBAE149F0F19CA3C8DAFE19CEC48FE6B38
-357246D8B5CEF80B53BAB4CBBBE8BC8318CC462A4158258F456E697F605886010493F252073E
-E4E4C9531F23485415AD6BC901BBAEA659EF701F825F4B6D5C9BED0DCE7B63810F906B33EA49
-55531CA1ED1FC120CC56DC1E277A33222CC6C3A97552607B865783FAF0895CBECCD0DF54F57B
-D162A88494768435EDA71A3969718C0C5912D3167BB0971546D733E735AA9BCD89F5AA154656
-031AB3F6FE4194EC225AD03AB8B67A9DA43AB365C4DF17C23E79F685C3B6FD5B654E4A6F2603
-6CEC3A8641B0B57F93EB95F9E238832AEAC2061F3153E0F375513C974F1C38F7B59EDCC94716
-F2862CF71C9A2BE8430985DA1CD813B1E1C1B9310564A2E9BDB8C536DB7CAEB9C37A0693A01D
-F0045080F8236ECE54DCFA7B743562771F452398100DD95B678B6F190CDEF25DF0E2B10423AE
-F1D27DFD5B022A0F895423AEBB5034EA52F46B077F2022E28ECFA6137FDD4B13935AC1976960
-0B2E0DC86B0A23594675106E26EEA6B5033525C1D79E046EE941C9AD410AD3284092F9FB725D
-366EE88A7A4CEAA6CAF2FF6A2C80749780481DB49BF09D5B7DD77CB83ED6C74FE2230076EB41
-E5E84C5F714B9CD7AE6B36396BB68C992AC812A85B87DB809BCAC7B8AEBE4FA450F3972A02C9
-8CD7E3886F6ECDB7EEC5E4750E0C92DDA7524BC07C2F9B490DDD2745C5E5C11F91238F191C16
-3E415712865A87E3C4D9041C83C0A8A053499502E18C3820F9E42B848FA1945B487E843408F6
-61F241B4B031522377CCC6037F8F186A11B2109407572BC6A17758172EE6724A3E3A7FA2064D
-3757A10084DADD5897EF81BE6BE9F61397B2CEF1E54822D37111253D9AC674EA1B3EA88A5D97
-FD4849B21A83A45A0FC5C25871FF71EE70EDF23145F56D4062540EF6ED21FDF96125B78DEDEA
-2E1CA95096C3156901717EA706132B78A1E00786B6A87597365C0DF9DC7D352B300C5184B76A
-B47271153B20D41BA460A4F4982E8A3B4329F0096EC55605C4BBBD73FDB1F62DEAFFBFAC7BA6
-B5716BDF2A84E7FA69C28821A1F90D13AB5AEE9F8FB5C3AD28B1089EAD4D3E88BE4D1C157700
-6B84D0CD9D3DCABC42D504E08318ECC2A5B59AF00655187E7D310B3DEC996BD60F91F6838033
-07F7094CC81B492A28EE42436E264321563159665C9F6BF38BF12BE1823BA5C106A6880A0744
-8A419E3D1C660CFCD6630E6E8460520F7F69F53E4C323F4F99616A1F953AA2FF01783C273235
-191DBF884B007499301177CB4D967BC717A13C64F4C005BADE5FEC790F4721A35A38C862B931
-1EB75E117344A936E509ACA1EECA895D9B26B67FFE820D4B52AAC413444943FD48045C586C76
-2AF97082EA7E891F4CD1D3248E42ACF8A1ACDADF02D6866C9562B301C363A3532D2B8D051E40
-1D648D67B8BCDEECE8AE3B24F122E104A8EA927144451BA60F083C2B94671B90340DB323AFAF
-35C11D61FA61CACD929B91B9B67FB794A72CB99BA76ECBF0618F3B479C7D1C83CF0ACF309E49
-589894C9BBE1BDC3CF3BD9A88086C41DF6E4C8ED5F7B724F3F638D65FECED0541993E87607EA
-CACCF80977DD7F955E303DF2409FB0C5328A92F5FB92C4DCC5CA232ECBDBFEA6EC03F1487EA0
-7BAC036E28AFBB8AD36C82CEAC84A4168DB3278C41D483EBCB14789C3DB2B5A6E115D0CDAE92
-C32DF9EC4C1FDAF1953E9917317444854B2645086145E06365F2891F5F58F411073E0F4F25FC
-BE255C67F14B565E5C5555F7A18FB5AFF16D3C10DF69E8D17F214180CF0B40FE2F1E7B31F119
-1FCB7B6615D4299A0CA45743F76B41AA8234FBEAA686DFAB4BFAC8490A28F259E124322950F4
-1164A4D3D0F3E1901F4180FDDB1EF63EFCBB04AC462AD7F2394F905502A7FF456656460DDCDA
-301E2D153C9D35714CBBB84E25EE5411FE12DE30A7D1975C26A01E9574A18EB9B0756D29C56F
-E9C93D0DA658AF001B8A7724A7E5C5355823F708FD1708D62DA0F4E791DC977B9C8A823DCF4B
-A8AF0A4C51A581BAF56B1A2E68106089A696E12DD1D87D881698E60DB68F8A7A473FC3EF195B
-5AECA56A7692D03F8EE1BE4F595D24186FF1A93E8CF82E9ACBD399F3B0D1EE751EB93FADCB1B
-3EC5D4D33CCC2013672C286298E426DB9C748E7525FCD3F9E936CD4267E620D6169577382FB0
-CB1B055CF1003EDC44D4F1848230B971E5D1F7A8A950D92673105FD054F431EC868F8B7F0C2D
-F6CAD1A30CC5A46885CDEDAAD4A389DE970F03708D1935988D448DC8B80250B691AC08AC7712
-6A7462E235A40C163606A14F4E153EF075A67875B5DDDB2B735048780302877B4AD2B93064C3
-85C61B8921DFB3FA912079CE90A4EBA8F2368FDF26924A04BFEA7F551EC65BB043B89FE998A6
-E6E5B30017887AD2AB34034D0F7F7D3E26532393267734F2475028410E68F1A106B43714DC5C
-1D1CC9291FC94C4AAE49668421B0F3346810A0EEBE3E2BB4944079B433AF8DD5D49EDED4B10B
-8F530CCB272E27E53154936B592379465F2F890B1C15B495706E8C121B17BCC938FF5617714C
-58CF6FDDBED98C949FF96CB7DD0BF70364B2F565D3F025854683F346CE05812C1089DEBD17C7
-50FB74BB6C087C77E2A1E9E1C6A09585B7CF4F5F409EE014E30C07194391FA554B9EF624FB52
-3C514D245F460F6FBB4D60875A3A368C2B6E63463052726794E9E0636BD7239F047BFCE33EF7
-94525ADF0B5F16DEEB12CA620DED491A6161C07503DDE543B842A540C9CFF77491EDE3520DE9
-B1CF4B575860BA8653983E256AF9601BBE3BE20443C0CC62FA456930D104F76A7B705AFB8B5A
-3AEF6C275F8EF285AC83790DB351453EF5B5D2A07D4050CD6C7401152287A25E591E25B66513
-EFA8684DED5D36AD7520DB8198A350C941614659405CFA12F73AD2E10200E9BC4E0BA8B8F787
-57369F42535B42111A9F6775E6966DA739644AF07A17B7CDE7A03E2B8A3F14AE01865776F592
-A317CF87584562F64DFE19B46317CB7FA01B58E5A560DCFF08CA43F9E243ACBB06856D0B70D3
-6F1576B1E8378BB8F236230CE5A6E3717611950C939C123A3A6875768C447A59161645E4BB6A
-EEFE87C150351D15A4D4CEA6C7AC7E060FD85EB81434EB0B655F380056077447117CD473022A
-C4CBE8422B1CE97DC4B66D280898C9C4B380727E98E0735F63CCFEB139A54854AAB27A005D94
-376E162C08B48DC6E1E0639AAF3777E129CF2FD4B4B93AE817553EDCA15C33589D4A35C0D89F
-4DB4CF4AD1400FD853E1B5DFE91F7A20ED8E369B545C9331ECB2A39E90CB425995DA9ED2F069
-15A229BA9016E0168B3E58AEE80AE760E2BDE915D17217CCD5DB144A9B5EEC057CAECB2C3C0E
-712B0E7D863964190F788DCF4C34101F560F18C3B75DA91679656BA9EE0C2DF31B7810E918BF
-E48B9190CE0B8EA35F1D8AF05B10A27C16252691AF9B5AF4CA62B7EE0449573D7B9D26465C2A
-BF40DD47415D201FB1164084ABAD77EEE21523E9B115A12F553227696EE8D19AF4F46A34B748
-DA1DE8143BA3A9D5F0AEE56D8C2070DF45C0F62E37A26C2DA65C9D4689AB67AC1B2610B2FCCD
-4CE71A411FA56325BD2B9FB8E31CD3491DCA1D4994FBA0EA5E08B30F172F9827DEA7C2C7A37B
-AD7436EFD6D51449195406A4FE36197DD9D896A58C6AB377AF2CF9442A6054E9687C1017B50A
-608FA68189E694E1B38E6D8BAE8FADCB16133833BF7E3976EC4BD721F39E00F50971F100AB95
-796071EC38E07FE1CBF5B3F0877C0758CC91EA7972C4B46DA775755CFCE62EB3F8BC10234EE5
-77459EA528203EAE467B240043891ACBDDEEAA85D00C5447DDEE9D9246DC9457ACB80C70FD4C
-1AEE90162DB366D277AE35BCBD1A2AD2B830BA352866AC93AA2C9A23E15906CE132370553E86
-4256E47890012502BC8EBFF4EC3570368577626CA74281BDEBC80DD89BA4762361B8D5E43A50
-2266234088008D0E5C8663F355185F5F60457C48D0D4A69024FE6AB25107C926B0EEBC1E4B60
-19473C8229F4B7AB2A83DCEA90C81F0014C9166327D73EF6D8CAE2671D40066A2D7E2585B4E2
-75001ADD5A3AEA5DE8D26A795AA7C17499DFF99F5602D8993E7334842201F469B70A86199987
-63A2F89B66F7E1E4A7894FE842DD9074EAF03462A4B7AABA49A6B7AE47E6BDBA6EE5D7B9A495
-989263EAB3E81F2E73DE1E5142D342C83DE3669F11DC7FC0E62315140FF9988A37CE5774B891
-CE11CC9CB9A4778352731AB4C4DA49689257FC1F8442C6468A01B3730C6E4AF32E56E0DAE8BA
-578BF15D916D09C0057A1CB6854C853A1BAC39224A75EB05288259E1DEA6D051AC146CAFD4D3
-5C39BA33531E4FBB98DA13C8660F7EA10E654D89A7BFEFF03BA73637A9AD5E5E66D8F8330FD1
-2975FBA0933FFBDB20000CE066CFC004A2AF5D5FDEC77282FE6641C032A9C089104A3A2B9CBB
-0DE87C67A2A146E32CFD0B1214FE0A3C84D116D82DFBD0CCA285807F039547D4CF63A327C436
-97337A51D923B7C6F081528B3E04D0A7D431B5836A826B81C49247B0AF87312605921DC3A9BE
-6E1D4CCA9B3A21AAEBE931249BC14EC553BBF853A4A2CDB715BF866C5C04E5D6C9FF999C36E0
-CD93FC6AB45BCD6257A63FAB956EA9999A99F001FF7EFB665810C5545055DB26627BB635583F
-8CF416257DDCA3B72353B61B4ABD6127FDFEEBFC5CF09161A0952A1DF4FBC59C3C34F088A56F
-3D0696C3E24F644E6766FF0862181240CF1CA9DC2B05A3D9B7F0F0ABD54DD46C633CC049AFF1
-4A52787C461DBB58A6F7B2874F37B069FC8174C93C185BDC3878C4396476CB1A8F6C54BA48EC
-78F9E3F54C6E746CFD4F51443536368B943FB7AEE86FE26C8D1C4CF907149AD4A9CFB4F515E0
-2B4A5F8FC24F883179B557452DE61F6A4BFB447150B311E216B8DB2FBC3E9E196B404615A337
-829DD12EDCFF850823C18705DA57D398CAF45D68E1640FE6CB3D5D41244F6B447494772688A7
-2F50FD25C52485556B23495B8455516ABCCB6DAECA6853352B19E1D3EDC2EBAF9B381A090EC1
-7E5B0B8684A4DD7934E89B245B6CB92718D2484A03CBC5E5403A164282D6BEE8FE209DC15300
-CAFF39203BE084F60497E813F145CD6888CF3079516B5CCA2EAEC2D58C7422A88F477F2406B4
-C32815B4EEE2AA7A6C5F8F9086E9D87F6D29012BDA185D95C3FA7414FE81349114BA98297FBB
-E19B95E12234F7FDD50C627CB423678A61467C7EEBDFD27D6CB7457D58A189633B9F6D8AE952
-39F81BDE87F21B8590671E3F898A76AC6C2698762016058504FC3C4B178F53D8FEB0F6C54E74
-F7112CB3D4CAC4BF49DA601409C5DFE5055E441EE34161E6E6900EFA481B5BF3FF9989EE7992
-E673A91F9F55F914876680FEDEC9B7E59DA2ABB649C6D694B3EB5957E19C6311D4C8F99F6B41
-B562B699EE9917C0B42322D0868C64B61058A2DB2DDA9EBA015C466D1A81A6658FC6EAF00656
-FE0079CCCCCCD498B66036A50894D37339B87CCE3054CB9FE42E02BEDA2C9AD3F57B16A2B6D8
-FE81CC9559AB89F718EEC1C1E3E2CA102CD4C46EC0FA214D5A019392CE61E6F7C96FDCE82E0F
-C0E17A2B6687AFDD905FE6D74A1D3CDD1EA1E17289FB9E9F6417257F26F96F66378D08733F0B
-D609ABF4E687A297C442D5C8404713C73B7E0320F3E30744A0D4C8F1A70494731C0E8175971A
-A9AED8BEE0CF0F76130F535E9709E214E7146ED8DA74439E4835E35172CBC9C3F5434051143A
-A9DC3C22FD4DB37E9C44A25A6AA7A645FFD79CFA631F2756A538043FA6703211A16280E2669C
-4CE0C04756341BC07208A9B1D03054893BE713A35FFCCC98CFB5162766E56C07A0CE9A8CE9A5
-2B9F3EC0B3D66F3A2DEEE7C9A1FC7EC5AAB6E3C5C26015E5F856985F83A237D30193A2191748
-132C81E4EC4EA4A8484B4E02379715DAEBE30E3E5FD7B0425F1BE0A447174A9C2F0E30D16E9F
-70A863023D165AD87414D7AC27A6CBE1B899023B705ACE1CD24EB3B14A8FF4E8DF153E4A4FCF
-E141CE42B8C90CD5CB5D15F8BE4ECC1549BCC81C185CE0F5E1435FE426EE1C937F122CE1B0D1
-3F65D4111CB080422D8A929EA3597AF86C31DE61F514EEC975B2D6367EDACF64014916256D3C
-4406754080BA4C982F5F736FA49B03CB75848B7C01A80AF2699FAA1B44AA2617B6064B8978B5
-F4EC53BE7FC5BAD83F30AA88A7DF4B65516883AEEFE6EAC9A6A0A895A4545DA55A0F25F34A1F
-F7AAA2C8B92AD133A12C780BE16DD13D64174D93E3FDAF081290FCED4CF1844CF32F4A170B58
-4F628834635E6E0A491CC69882E155AA7C857749FB4C8A7DE625772702644C8ACC4F32D66FAF
-4926BE9460091D1C38651B328B168EEB2A938E01ADFF8BF6D0D6EE32746C4F08A61D66AD40AE
-BC967729B3283BCD60195F4F7FD40CF06C278CD345D6837918E48A333DC099721CCBD40FC8C1
-C71E2A35F0A5C53AEB290015E75C3820F73CCAD88E9F7CDDC1644025570E2BDA8660CF1A6637
-2341D40169F6D00E208D5DDBC2ABCB7B485F4CEB8083A5598DFAA64B5740C2E7B3FE830E1E3A
-AAE4C0DCDA2AD14C5DD5D89D2C37EFFF592CF84D951E6CADACE8A2FAAE799A9284805F843288
-147E7F8EEA50AAEFC280052971ED4CDEF14CEDED125D5D92DE66D34810AA65783819C97F7561
-EBA49576EDC1DD2A8B4A2338D91C8207C49BAE526566F6E56A1A4E00D07209F3C79C706726AF
-FD231869A18B834EB7C95A32552A5E2CD1B7C99720BE9D8DCC18D7772AB7EAFAEB7F7FE46DDB
-F1A5731C3F94AD158D8F30AE9EA0F6D45E560BAF538550FB6CDEC0C191C0E54141981D2EDA95
-22B83E3617A6F57CADE6CECCA233E47EA3806C23CA4AC9A98F12AD7F414B498A96898C297445
-711ED46A990A8EF412EC97231EFA6BD36A3003E0562E4F4884CB5EB15F4A76BA7F6321D4EE88
-BF0013734CAA1A638555B2B4EDB34F762CF514A76B0667A9569FF862F18868174A989842704D
-AEA558E9C6BBAE69F5172374B183139681E1D540DE08E0CE365385B3DB3BD7872D107C6C3EFB
-72C1193F06FDE5144763533C58C9C29ED0A6C881282A9F712FDB219DD243B8930F977F3E1A65
-8400308DF7D53BF3829780BB9D93C6615B675130C8B5976B1913566F47E4112A3E4D5DFB99CF
-07DE8A5912F2FB97D4312AEA212797BB82022EB1A68E23898D4903A2CB848CABE2807B3DE9EF
-64D49C50D007B0722BE8A294D87E2A00CB078B9A66B70B5518555961C226EC45E847C8B378C2
-7E1122DE06B64ECD2342609CAE2CDF8BFE31F6BF8CB165046A21B776E453FF1E2025C293A698
-42723B75E2255042D873F7B558263E3238D6F80B41868039559F3F7077A9C8D72C12AD911192
-3E0112EAC5CE4C766B95CE5DFC26C59EDEEF9479A6A536CCE3CE644185E62702A8EE0D816199
-B08331F0C818F8699D5A3AD6FFEF80D739F7A3E7183E021665DE716EE98632AF465EE621AD75
-2A325E684234FFAF67EA765080657F9178D26B053A2217B9A210B00827291A4A96D03AAB28F2
-AD8F7A24DED288DAB40E65EFF496DAC43ECD452D31F4A117BAD0E093D1E431FA39F346F5CA4A
-97EFDBAB177CFDC2AB3EC7A0BF9BB4EE43FB05777469D044D3541AFF4B0D2B17B59C9B549F4B
-A05D4BA420C4E4B06C787099FF49B6C0BB02A5DDDDC6275E656E5929FCCBA301635DC74F89BF
-528829086676AE7B548525A6B0A4275A40B87ECF9F1558E7FD45EE43BFFD2E94091D8E6FAFD7
-FE9823961B113A72166BF341C1B6E0A7D5A2CC0A64318F73A1816F337EB5283D4F03074D32B7
-592922E75CB426475974EFB096ED5B63E492DA2D9718A8E12CF476892EDA8C75A85297B4BD07
-7AB76F0CEBBD8D4DC81AC255B161F1BEA1666FDBBF9435ABB83D139D3D40F1ECAE563FEBAFDD
-A486E514A8A073D96D94B87286413E500820ADAF94C6924250A57AC85934C8F0B31C61FF6D29
-5F483C6E3626CEF2E9923E70A8AF0B91D0E26A0EC081F764213B1C5C5924C2E47451850C961C
-08DD7B14C87480C2FBCE304874342A07CD196BD35A4BD671421C0FBDC5CB5CCC53F3AB50F0CD
-05CAD126FCECF166B3E8D93EF1A1180F2E303E7078EDFC925B9CCE37F6888A0106E34AE71B3C
-E4FECE042041280999AE3F79EA2C168B31AA5BEAE02766BF17F4BFA76E7E1195B9F851234CD3
-4FF562FB3638AC094946941FDF88CB41D0A8B6C8E600A0277212F00936B20D17DA316C2C7470
-313580AF1C889C86391B7C5C757221386BF21E5B48CE62396A07906B19195ACDFB961953F85D
-145146ADFA07BEBD256F2C92373786D85CB7086DEB9E9D6510F66604BB43F98225F8389580CB
-542B0CBC71ADB3AEB152BF5F9C3D2CE4EA2E2C3BF86E9A2F834F090988FEE04558BE6BE4B831
-9E4FF33E0AC93E0EC5D773C03B58A30FEA6F76090BA941E0088903F23870E863F1B8AD9B431B
-583521EEFCE47AEF7CC1F98E068DB46BEC579D7B477B482C07F6ED17117D45AD80A1E6E4F53D
-E0EC731F72ABC520EB376DFC6FB138DBCDB6ED42B8F8F87D4EFE406FAC6B289657CAC941D94A
-C2F6AD7987C0C80A71A936C0DEEF1BB23BAEC645547E5217767AF52D52C96F9196AFF160BF8B
-BCA46AC1ACF4781B877FC6784539B03DED56592E4223B5C875311296A20107F21DBC334F89D8
-63788E871AAAEE8B3B57395DF0E9FB0473670EC325FF95379BAC7AC23A6819E5D8F8D0A2D1E5
-222D3DF649E83C8D18AF824951CD9A955D39E21106D0C063C617B6F039371B96746263ECA4C7
-AE4EB0BBC09F95BB4D72F0E31D6D70F7E005F5872E8B033D1EF30E71EFFE49B4E5DF12F0523A
-285B2C46CBAE02E05825FF7EA18AAC8CB3A645CC1C21288D947942F452D3A3A61848D57AD74C
-57B237295880B99F7E7D86BD239759D69203C4EF4523AEC4D7B4BBC74340464C526259ED4B44
-6495ECAAB057710DFE9A739209424167167BE9E5204A88E85ACB547F8F0E59D62391851ADE34
-34A5BC4B59E44E27711ED87F58450258CB1843441509100E8048E3C314B89EBD110D2AC6F3CA
-860E9CDD288D43C408C59857A605CDC750988D483E1527187537475CA44FD747009B0E714C19
-DE98DC15A919EC47E41EFA0EA594BE2085BE206D1AAD27C0967B6E377AC5498BD2C7EA26E3E2
-CFFB19B969FD41A8CAE35A5BB3DA5EDE55A0F096B72473DF9B403DE8C408B1C8F1D1292DE934
-39666DBECA73EF0F3745FA39AA37073EE77287A58E265A16F07C70D32BD03844DD9B6F2C73A5
-0284A2B95C183F826B8851E2804EAFA7892567DB6DC04A99416B52AE96BED9118ED3C4217DC4
-85269A637AC13FFFBF9B3545D4927B02F94720C3B9DFD0DDDD1BBE28AC9B491BF3B389B14177
-E68096A67290CD1E1CB19968BF2947F95C643D98CC83235B570EBB281B804FFB1F0B52494C7D
-3679AFAE38B5172C0E5619845A42444BFDC1514652B4F906A6EB8F5F3E269F8DF2A56C8B8436
-F6800DD8F681718B43A4F933F6B7A4CB3DB49D0F7A22C4BC219360BC3CCA19226FC2273F61F8
-EEF629A9CCD01FFFA16107600F84FF6B9DB820AA9BACC6E6C9D5C2E7AD04ECFE8A7236D3A041
-E5BCBF7161F225312AA5E5B4F0400B8E43F1BB368B114EAC60511E5EE601B90355F8610EAE0F
-C0CA483063D068354C80C3BE7376E087A357BF45F1B474EC1C6D8FD7A80426B19A872B77D102
-3C58E1E02A1E1E7F5E6E4353C073374CF0F58510A77A8F186A973379279D1FFD993269EA91F7
-9AEC39F1F1DA6DEBFE095A0347112B88CA374684813147164D5715998DE7E2A7D90E1934D97E
-A533653B0C70BD643EA703EF954FC24B6B3827AF6CE53E247A61E181F0F4F56C1D8A83C6E97E
-D498A39D7B601710B2B78950233A65BC7D397617199E9B182C64CF5625B7491054BFCA5D7EC6
-C9BB1220F9748BE2D1981AFD7278BA513B453A8230CC4025E3DF3892A6B0DBDCB0F2AA4E5FC8
-02B9B624E0B7D946C6A34BEF21A7CDA336383333D8E1D7ED5CF21C67B0DF1361C74D4583913A
-8F45EDC330A97634C9F9093F50FA76E91EDD23FCB87DED6AA41DBFB6CAE5E72BD4132DBAC407
-A6F2111014FA2C71B45D5075090EE7844D6A27CB783CC417A9DC3EC7C17E616B27261EDB98F9
-0756D4F4F0E8402AA0D05C5481CF3CAA9A638D48E63EDB5BC6F442498AACC6436A6D25E40699
-81420532ABD6410179CFC51A67D12AA42431E155B3F9E98E850E70D5CBC53BF2CD10D851A8D1
-5C6A6B0E6B69A3A396063DC36A22C73CB2F315B20AE4878C3903D10A64297F191B2607DCBAFF
-FA1922D411B313120AE25BAD96736AAF1AFAE27260E7D9A1E3A9879CD37782A912F2F956CEF3
-ED044B7243C7F571D3EE1971A85EA9A4C4C8F5E9B46A9DBF667DC725ECD8BD5E2ECDE296F932
-F1C67BCAC2F54432C4DD4E0C2C7553B790AC12F3580692C827119C27DED38E51AECA9A1D565B
-8B84F9355F84B4463AC1B365852D77FFCF08F9DF4496D2B6C65913A750B86A8CDE8E5571A8A4
-AF74018DF4AD90E68F4D64C6EBABE6D9986E3DF319C7AAD0B970355CC9D6F7BA7975C31C37B2
-ED0D794F2AD68E8167F2F398016D9271B4589E6AF07EAD6D4C423503F626C7601559DE9EDF6C
-41C8CFA5D7C4A096613788B37FE3417488F5338F1EFD0ED0CF91B0F1002A357CD14E4B34D006
-0CC2453E4B0A51D8EDAE472FD0844E21EAAB90041612FB1703D88C5FCE4C90C0B997A491D8FF
-A9489FFD213D73A83D8A0499CBFC4272755B03150E7145F950FA2DC82C41C0DCA5D67102D20D
-7A984C4DAAF87FDC2FF1F182AB011215400A86ECE94A9C5713CB99DBC99599E43982B7F487AD
-A57869537ABFAE60A8D57E58ECDDB7B87796EEDA6B7634E912C22D95C571C68D860B54A3BF4A
-556ADB6CB53ADE484FED8EA857E03DFA41A4A3EEBEB18DFEA22C3FF260C0D827EC3538409315
-026BCE7238FA6756528D58196A95B074730D7C5239578E32ED03A33E984EFDB1678B573D324F
-FA81C0B4D06CB384EEAB0531CC9B414D50C67E033338679D5DE7F37BACC1CFC649C64C19A37E
-F486FB2C811C01695BAD4CD3E238510A72B209D8B9511E5B17DFD04B9C2E3F35F9F9DA0967D3
-6A5E37B8C61D360E2D84C6F835CA388794CA79C0826A379E7524AB22E282F1FF32AC84ACEDFC
-3BFEF33CA08F4B4184A5D58B8671B82BCDB185B0B43B9F7E3A6110F498A9AD35B483E80C29B5
-615642AB7B21E960E1F2553474E57A473A7186BF23E047C13059FB9F8E86CE52ACB9A41C0811
-A93A0DF5BF503021D319A4709AA58DD11DD80BBF78C90A86DFB57A22571A0A1402D69DFF8B7E
-38C7EC2EC85D656026FAB9E28904F7FEEE50042BFEF6BE32B746BFAE00CECB24B3D81D642F95
-6981B8DBE764D9EA219157875AE31B6C94C799EF283332CE591D887CBAE6922C75357AD86C5E
-C307F6758606D7B6A426BABDE0FE4A0057DC0D61A48DF51A7ABB811F9C0CC3DB2DB3D0F5E39D
-2AFA7377B55A0E605864C51B9D7AD593057D231057C3488C5A71FF4D4D3781FFDE1F9EEA6B77
-A0A0A686A7F9BF9E11F2CFBF7D62F576B38C1A52B1A501AADD48FDFB8161D3E10B344FA8BE7B
-A214E01E16EE1A7A8AAEAE48FEECF4CC31B2E1682A23A5831DDD358D51B326C3B25BDD7E19FE
-F608F4521A7853AAD6A7D4B94E5FEE730432C8C6B8314DD9FF1164E401D3DA92B659063A0C64
-44DEE4F5847ED65504C9F63CB48B6F06B48BC78267DF7346FCF7CB13E272F0D67A644E947664
-FB9AB8C09FD3DF934F7820A81B3AB23F04C9040AD1C0DC9180B015F2226D7F4CDA756195778B
-5FE40A1AFDAEF10AA4ECC1B16BD33E15BEC296B362FD22B25D57D2E1BDC7D6BB0597942A2782
-D42B57410D1D1CED882B355CC887C99AA31A52645ADE3EE123A675BA331AC2408B4EE0563F7D
-7777E47AD8BE376E6AFB13C27A9EB17813107C994926FA18D305C3FD3386B9EBB69AE03F7D84
-0809D2F154AD4920CA1976DF47371F0A48DA5DD30789D1630B7E5948E95FEE16ECE2BEFB4D47
-AE9AE63EB02492C6F10AC74ECF0620C4C303CA57A5029962202E1F5035B1AB33994AA3E15E91
-BBACE30B5D1F513EB2F6DB677D0CECC10F2D7489DDC6B118B95124746417D9E6DC4B375E124C
-2B0C1C2A21035FADD644D4D447B16F25DB357C31B964E3331ADCDE77C27277CF5CAA4E69BE25
-C6E50DC86B1CE6DDFD29554E59DF4733D82F30C07CD826E1AC4089DB1C486FCB8D76C7D12757
-1BCFBA2939E3A2437B1CB79366C3B25F3CF11AD60D3966701360D3066B80990D8B84BE030986
-7865A178C6DE3102AB3A6203FC555A08189288E25379FE3AC5729B71B86712C5D9FC923D8F84
-91ADAEEE87339025085401B84304BDAEF91EE64A680AD6A94B42395DCD168F0D8CF21765FA23
-EE195EA57500F294DCB1481CD9F98950E030C1CACB0A482C7A13124AC03A1CADAEDEFDFFF8A3
-2B906A37E070AE5F0FBC905E0C4FAAF84AF996AE18265949FB79BC4C553E5E552C2D7ED1CC6A
-3925F98C6C8B9BCE3A946F92B586A9CD9E5E4DB3F3E402D713E747AF6887E2A570D457C9D203
-E80762D2D2EAFA6677A02987C464B4FC900E5DF7367146226BD0C1119D514342A0146E34B83A
-5F36B703EF97E5D87784D16C3E7A26F1D1DF410A39C2088EEB4074C0FB3635F48EEBAE060FC9
-4F09B809519085F3D262E7A559CB31FAD1521A31513E1F8F031F61E285F89B459C2C30FB7727
-C422B89E0064C8A2C7CB93FE454F0490640E933FED7E6F39B66CABCC9E789095BCD30BA19B4A
-AE801FB075ED765F3706C09E30C55FBBD17E7C130E279A105199FB4A718924A1EF199811D9FA
-597E148C85A753A8F1E602C7D2BFF520F260A3A41F3B7D2BF483966BB786BEB7F64C7B670947
-28AA5CF3A0B609033ADF351F7243162B0FE44300B547F06B49442CD4E61454CC657957343EC0
-9E5568D932AD7D5FC83363D0F2A143BED3E20CFEC074497F7A0E91968F483F4648669A253F8A
-9A06AB87623AFCB1B94DE1396BEA2C14664D073654ED4B0804A0F8EE6040EB1652D0E6335784
-58684DBAB90FE9A9DD844F70F7268FE9CC14D5121286F5F12AC4F221163084F48DC3975705E0
-7A50B48B9DBD119EA0E95EDDFD8B0CC1ACC19F3957AC02C5B4FA10018C13DBF223C97DA61037
-B677E5BDFB530F35FBFC9804B665D5CF74E69F7C6B10BF6EB2E9C2746D1F4DA8AAAA7CE3121C
-37D517E6FDAEA9DAE9F9F211CCCBEA2C2DB6CBBEA0DFE86D67C211006EB61F8076FDE1CD1B06
-24C746D02D44BE627D357AEFB26DA11461A89F94C2B26021063E5590F30C7BFCF61E4F405980
-18875727072861E7E5760045B4B55C01BD0A2DD983262310918D9250B3469ACD72E3E89367CD
-206039A294D8FBFC903D0CE772CC731D7733EC0DC9E804AAB72564231FA7BE9C200582C6EF06
-48BB734FE8458B144097223F6D452E99DAF9FBB951456066BD1A267063FA073D671D8DE013CE
-AD831B28447AF63B2823054A6BA5A0061BF5642C36A58B343BE2B843DC5A3E9221008ADD5A7D
-C70B28535706BF6443CF280322D1DB259380911CDC00425C4DE749C6FDFBDB86FADA3A2934E4
-1FE55E1197032C012625510E962F98F987BD59D8F18B421EDB1342E60CE93EAF1EF8C2554E1E
-3BA20593A14E5EC4A53C1031C5CF0199A8CED6B7A72701FA32B4CB1B2D788C0ADF463B841CE6
-8AA61FE24FE98BD6D7001E97B29BAEA69BC0E972DEF9CB5BC8BB3F25E81F7B7AB926933EC37A
-9C21EA8D7A2EABCF5AE24401E4D4377BD57C1279BDE798F420CB2CF5D3D694BA3876865655DC
-47919A41ECEED258279176783D8AA2E15AA0159C4CBFC07D5A14B138D42315E353675E858A7B
-D57C8140297379B84BEFA7A3D52BC6F7F9989BCD73F61C38BA3A944F4021AE77D4E4C5D9F14F
-FA15467A24AF7F83025F17811030AF9134C58A9863A57995E7C3B1A297036F846957F08594C1
-803408F1993E7BB8639F238B484C6CCE750FFF37AA0D6C055235CFD7280794797D9414888913
-D7C528C4B4E27BCF5C7F72B09B862A24124B48B9F7B24E637595B101D58C9A4462CB6DE0EEFA
-712362D6A0D4B0DA1CFCD8A1C11FA494F5182C42A2B3F2C28D9003423ED6F4FAE32B9B127DC9
-0D1518495540CEEF5A9F08A3F55BF12D0D8861E86DB66F5D16A2B4C7D75F21123D29ABEBC8F7
-C29927C234FC401C1D8AE12138490F04477B214804F09CECFDB294419FD71FB0D6165E50C95A
-37096BB9B24300D186FF9E17B2EE3E17EA8BD531287809BF7F5EDC90DE7B94495C11202E9055
-11F358434A83D45305DE02139FB536F892F9E72DBEBCC0F30A6337284C126004361629198067
-034A5B73F13622BEE8947C53059634F5175BC4504926AA624475190816A1527BC8A5EB5F389C
-0E085EB89DF813212B98F5EAD661FB50E9BB7A2BD70D2165E10B8EF6F8783CB4E5CB7488AA9F
-72FA7275DD377AB270061C5713D1F3D96AB16975E55B338F19D9F83CB160CB5A214CADB888A7
-A0D278600860EDF74AEECC1487FC54975D1BE5E324069F7D4A85C9E542C13B8D07DFD7938A67
-5C8693F63AF1F8D933AD22F2AC2FD2FB9FEFE5E61D3255A7FD3D6A79235417EE24A84DD9F896
-4E5BD92F56259F5AC2CA9E3CCAD94FB3982AA01C29AC6FAB277289080B7B41374E9CE0119747
-16CF387C973D26F0B5DF7289710E2994A665CBDDE770611EB961610DB629B1961FC173859B8C
-5878997AA782EC76A45D1704CB136774BDE44ED00C5C1A2F6293F2DA484AF72E228B46759A0F
-7DA87CECE11886E00786B1AFF1EA773C0F03B608AE06D7FDE8EA262B3AD5413FE112175EBA19
-208A17F4E0F2E5E94E5BCCEF26861D4E0F13B243EEE9877A935AC186854872C03E20908B9405
-BA66B8EB5E3758504B715F6C78B8A9940A9429429C52047FE3577EE2ADFF2D58508D62DDCBC4
-D87CA273D89283528BAAFAEF4953D7752E2EA4D37791395184C03E5EE182F75FA3EAB8F3D5ED
-E01EDDA7F070BBA05E04BC90A5B36FF981EE164138067509D083082E274A2E1821A9EB3C2C7E
-5C1562D7296E0AEF01A2ED40D0D283B88AA700D646A25AE42F8771341277122B9151376CC039
-2EBF0F0EAA6820475DE0ACD9E63EE877DBD9905CCF3991A56FB44CE984DFC3DF53B01152DE6A
-4FDBD8362473D967EEA50A99CBAD3BA984FE0E72F3C3510A79CE1B4CE5AF0C063C54C872B036
-DEC6D0A99592047FC1121D857EE32F0405054713C753AFE764FAF70F6FE1DDEF89C9EA11437E
-74FCDAB35BD1F6255D6A61A59462A18014E5D3C33EB5B3F694829BF84839825D4252C8D745FC
-EF6F8959F7C5B29FA2D98B3A939CCF578DF50C5A8FD57E62720677EF515B3B395489C05873D0
-65462B3717A1EAA058372B23CD4405BD18F91E7893626F0B234C3C030B50BF64662ECF43691A
-B897300EA7B40C2F6CA5D30619E3506C90122E453C93798E008D7293368FA805A4E4FD092613
-CC9B8A3304F467B22C727EAFF6B359632986AF28F1921395D063A858311AE92F1A8CA9D5D9CF
-0F4B8B26A8D119CD4E03A520702A2814BBDC8430039ACD9144126F41ADCC59E8EC68922B9408
-232D777A7A3A0B4E7B7C9A18F1224FDDDA5814DD24A996EA75EF702F1CE7B3D2CE7B5EE06A6F
-2E6CF197DF0543D302ACDD8433E7C9BE488792869D7B96C79352DA37810268290269C1697F18
-16FD54166A5F7BAA5E7A1FD09C10802B01AA534493B2A3AAD9370327121D3FC08A95ACA51314
-8D212A472C4C7128F73020F2904AB15D929017CE21C6918818D5967ADB21815A949F67D9222B
-0DF9373B66A6A9EB5604FE9C2738D5CD7002E1DD8B14720D4F3241FDBE7DF3F323D7D0589951
-1F1D689086570FF29725B3942184810D0C975E7004D4DE2E51D2256E671D898EF50A53ADE3D4
-0778BAB0835E97EB09008854568C9607EEFAE569263F8197AE29178476CB36499AC7B06DB4CB
-AA9BB0E88859A87D2E99006D47B80E0B027479D9FEC64E4FD991ED8933013DC71CF7E781F8C0
-899A1ABE750503F1861A2C28500BE46F86CC5C6B2C53D4481456D5A92C5ED002BBAD72A96FD8
-BAC252DCFC36B8A19ACCEEF158C963BAB3A007951EF3452498F3C741FE75E909E83C5F394AC7
-D0939AA66F758F0F9596EADFC43E2D4A36FCADFD2CC4D1F2728EEDEAA27010ACDFFB125C651A
-34588E8FBB4EAD79C23AA077FDD4A0A0DA56B50D34B6A4C892F4698D0C3812CD8BEF308903C1
-ABCD83B9CC93514DB2E865D693B6544A56130A87CB1BE962CC9837D2DE51CC4BA421E1A47F24
-BEE3B76587AD5F0C2972459C7F87DA4349706698D5480A486D52DD1948FF4C620872D8C50031
-3066B74B4954AEE3E86428724F0ABC5B899249D05CB384921FD6DE0FC265B60C56F84AE7FEAB
-BC2D2D2890AFFCEF9895C78046F24DADD28BEBFEBD1F9A2CD4B10EA9D5D69E7657FD4791AD58
-018EE2F34A930E683A063A94CE1F658D1D1703591D5FB7E26AA557818BB7D69F4D6A4450621D
-71B2D0F9368742160F8F8AA1B7AE5A190B96DF452A6398969D745A1FBD8DCBA2C783DDD522D6
-F11B48EC0EB9A27DDA805EF92E9BDD537D98FAB79974541C60D62CA5030E4DC19A825E0006F6
-ACBDD6BE2897C062AD1E94B1AD62B268E1134391204CAA3DCBF07E27C8AC02C81E801FA51458
-2648E49AFADB6C4C579E47EF3B785366FFD140BA59D8D4C1E06BEAB583BA3135A8E54AAA4EF2
-E9843FFBC4C22E939A0F08F915E5B70F21E9B29F932D138BE0C3395656DA974C42F421C28A4C
-0DBC19DAB081C4FEFB21A705790EB097259EB425A48709BA1B44563DCAAF38556BAEF6F8139F
-B41697F0784A9E688A59A1DA22B97609CFB60CB0EFC95FB3785DDDB9B0EABC90DC1D31B06960
-7E4209F4EEFF16771683AF28F2C3C8765CF4E35C58454BC5B0170CD13B265E72C5F05EC54C3F
-AF70D7D93AE29A53C8A1A027E1729B630B6EDDA2EE61AC560BACD93D5DA69FBE63C200298407
-A6469479B44902679C3EF93C77599D9A6D5E1909DFE3B19ADB96F660172A61703EA6F8BE0C26
-2CF8BE408A79CF42244E52A35C23EC35A6D972C82E781ECEF2E03F6DCEE4BA7D4D45814BCB1D
-70BAD0F66084F8F808DCDC9F71BB7ECDB87F0271FE4D70317373607719F2F175D4D9EB024B2F
-4108D871D04F2C2FC1E7D433BA0AD38F1DA9BCF38CC4CE33C4894187C382C80D9BF682D85B56
-DECE79142DD17B7D91C3CD87B2B3FF1C5A805FC0F8CD25B6C2F4920E79BF68E6A7EE0B9601B6
-C5B60C427AE7ECB8454530435772646F5AA973A7B4ED9CF0C00E9AD8A66B630A0D26C460EE27
-AEF5BEF4A06F6DEA4694B0960A10BB6542F963D1E1EBCF2F8530C36FA888A7CDE8EC95315304
-2B401C8659A6932AE871C6E9042EC4C8EE74B5394EE43B83423E4FE982D82BFE0AFEC56F6D19
-7619C108DB1AFF52CA5FD81994B133CA5E7A8184B9CBDAA299206FF858764E71537DAB906DEE
-5B104490FA93DE85241C8BAAC984B79CCCF366DF1A2E0B17E4C88CD3A2D6E5B8129211689788
-288159DF03516E37B2D7295817D748A094E08538A8ABBBB402827B4444911A4D4229FEFF03D2
-B607F225BB16680F716ED63838FB6E39D2A9424E08B2CE9BA668695034D53A625123F71D6D44
-CCDED51F1BA50DC36E7DC4E9B0C4A4038C856CD1BAF018E2558684C1D1CB08F4710391FD5DC4
-262BC8C9E072B42D6CAF05B889B3A191943410683CF5A28F95DF075AAB0F5CE7D73C065CD6C2
-49FB0F8E5D850F29F9D3AC40E9525051EEA43E90BEC176BC7F32F6E8413F08A0F06F1867C06D
-8D8EE7D28F819B1B53619439043AF3A251A196D6264331273A310D7E56298C92BE33ED6F6C39
-64FCA1105A375596C55AC2BE6FA30E9A5151103B242AA0B9FE24B6BB3786F98BDE102F094D81
-01390CCCA176833F889AA6630CBE3860B861A2BB9C05A520E5C258AD389B7ADB89E7A75EB69F
-578303A57CDFAEB355682B2BF568137D020A5526F9E74ADEB8253E41BBFB6D52C32E994C0DB8
-6B83B0AC179C49E76D94904156E9CD352DE707FC7853D62FAECFC4B6B403996212F4A8BD2674
-E74BE096E753D5D89E3A71940124573C49CE393E0B63825992B8E243FFDF9195F9FA085841C9
-9806A3E65D8E23E90E8EAD4A2C05893E6427F00F389DE11449A94CF69A0C0CD2E54EBC973707
-6C3BA11C4B4ED495BD133CE21539C7B65F8E7FFA3752EDFEF0899F18829FFE0AC7A30AEB019B
-710948158EB1DC77BC5E5F38F6C57AC533F39C2C97DC05911EAE9A1F5249052BC44497F1F701
-D1552348D0F2ABA4D571045A8727ABC50112ADBF8744CD241C210F5FD7E1219BD2053ECB654F
-9249D2221278BB9C8F7C1E15D7341F7AFBC1A191A30951377254259298E872CFE7FB346A982F
-8BD26446898E73843409F8EE405F8489BFC406F82D6712469FCFB6AE99B9E1D2EAC3914F22BB
-A0C5E6E90A151E7BAB2CDD6C49667170B32664CD0EABDFB2E32F889FD397C2F1C4103696DF0C
-B4AF7AEC317D577AE6EE7F411EF863BE2AC953B83F24572C3719CB40936520AA91A16E960F5C
-DA6CEB0C90B092F14F70CB9409CC7A9B75F9752A639C7BFCE53C8FACB59DB2C0360B5DBFD4E0
-763358593E11D5D884260DFEE1F05FBD49649FA6AD192057CEDE7DA28EDC746ABBCEE5239D8A
-B00176CF7D15B2D34BCC6A85727A1F12B7AB128594D453B7C8D2644418689C57313EEA6E7934
-65F76A66D87FC7481F58377B09A8873EBD48261D62E36CE1142215B8084F4B3FCDAD88BEF2BA
-BF12C8EE4B0FE49FF2D02144F1D73E34A1CDF217AC19A910DEBA6EB746F1637AF1BFF1809298
-A0C27E9CEDB26999C220BA3CDE6A31457AAA0B09614B50E3BBE7E02764A79F83FC22AACF4B10
-65CD43437AE475B70CD3817C745B169F3817A54F6DDB55B0891044D46C95FAC683A7D39EAF0D
-C782FE96B1DAB92D364D926AA0C60803B3F9C67201B7DBFAB1DEA07A41D6D4AFFD43A0C69C4E
-9B11CCCCA33805C40CAEA7C9926A4BE534203416207A22D220A11241699DE3067D7D4100F2DA
-D1257F0400BDA70DFE849E1F528692C4C12FCCAC86E7C4AEFB800B52FDDE6031E2B4740B2CCA
-23B5E5E4C992CE50192DFEDF15D0A97DCD206827801A910BC30B0800957255B81E639492F205
-1FC6AF0E6E033A44548B9FF66D0A0AA7517EEAD72B4EA6190F84A825F4EEF71B57343E706446
-F523E8CCEB0942A5DC80B8C76A416AB4B73129AB24B8D219DF112195756EC91F5F7838564B6F
-81809E8350F9103ABC61FE1DD46E3D34ABBC34E71674B8AC43E91366F07CB36A4D0E182F4B67
-0E660C55218133627E77541228A1697DBF0CF4D01C11DC56161E2AD777A9D3659B2AD55C758D
-C18D80BEDA49304C629BC3D2E735442CC6E57239124955DC4AA41057D73F8B0B558D5CF24C42
-08A3B504D283C99C06A909D32234678FC90B558E9555114A3C79754EA5BE5B17083773FCD80D
-49A3111424F46815D12B80E9C2FCF95CA30780EBC7A79890AB9680DE6AE9B1DC18F7CDE90E89
-479891E8953EC309E572D63C994AE3C94BB56D9B94EDA02F7F863070BB5DAB56C8D797180685
-AFA62104C23E55ACC1D37643C75863979E2008141EFEC535ED5E20EA6141597B9C9A72ED1CAF
-A94E476348744EEC26E571580D7DA484DDBFF737FB722D032D081832D3D8EC8BCAE9E6B3D534
-978315C2AA3E537528B4FD0AA4CE0857F55C01FCF07AF06437F95383545671EEBF117DAEC679
-0BA1111AAAA89E6D3675E96E40E7EC48F34FA8324B492C96D5D1EBFAF7EF4468F9B749F06445
-0C03CB3CB3294CA325BA29CAB5C18E8FB2D24E63D4FCAC7876C1202EB65810B7D25885F92A3D
-B1A4B6AECEEC0711B2ECDC3B253895D220C154D5EC29FBD136275363E85A0728384D2242BFEE
-5AFCEF3DCE113EEEAC171E91B841E52A81BA74E8D3F3571D2E0B7934B0A04FCE6CCDD84502B9
-E31844653053AD77ED1A94E4A4FF04BD688174369C73158604F32BFBD91A6E77B72C175D9B18
-CEAAEB2D3EFE3501CF551172D74638D53F5E6F2C8D9A49E6A87AA945DCA05CE031A56CE97357
-33C1919CCEEDF60A2AEE6F8B8CC0FB4E0B0009F62A27CD207405E32DACDECCE6BF6FDEADCC41
-83C6C16B2D41E938EFD2AB98E9548957BABFBB7B9BD1A504198FBC48E18F52831F831BF9EB3F
-2FDAF3882C46A9CC26A0E07D4AC221CA50AD88AF3173E59400C9EB446E1924D1F8BDC44E49AB
-F146D81F23DE03F3EC46BE4B5EBE0DE485CB6130FF047265B912213729F0993F09EA4E74A632
-306286A29E93A5B9B94A303F2A141B0B5882A19CFD30C7ABE507EF76FC3EC2B5E5318FB84477
-BF4350EDA644129CC247D5F9171D62E6BD037162E42019C3AF993BB3062F9B2F68CA9069AA58
-8055B3E139D9315DAC5ED7E02D6A299D31250A9E8772AE5356D5906E454C5DECA6E5E902C36A
-E2F8D93A217928A14DCD0F0BAC28466DF38B64A0C0F4AC70228F546CFAF5F77A64E37B0A0F96
-B55876F49F93016345130AA81424C3139DFC08B9E80F4361B01762C56EDE6F5856D74B89BAEC
-D247A0CF40AB2D697CD5CBBCB2652CDB256E2E25D064DB0CF1B44844DB85ACDD251BAEE12C87
-456344FA3A11F060EC35C84D21FC2C316C4B8298A0CA3A62D6A7A93C05692A64264987EA0528
-0B3FEFA93209F38E10C719A941E7EB6285BCFF39DCAEC4E82FAB54D37653F55010A67A6DD0FA
-7DC1147D24646A7E4C132EAA0BA92C896638EE28A00BE09C816D7522BEE64FC34AB7BC8049D8
-0155BF8DCE92894A174D66306C618B9EEFFF24B5644F4FE1959961B880371776AB071061281D
-BB5641839753E41BE7170F29AB5666BE3BFB2BC6AE63BFAFC909C3B77EDC14106E9A748CFDDC
-2F8727E69C6B42570AC1DC0EFB2D561DF087E2AE65B5E8266F197A0E45BE2CE76BC6216D7B2E
-B64A9103770B4A51F6E7CE401089A95064148FD6ACC3027B792287B20120DE1448A64CCED2D0
-41A74122986B5FEAE94AE77CDDA9F94B92B8FCCEC50F262006BB4DC19543959305FD72CAF9EA
-923DC966C233D082D97E787D262C931BD3A6882367C1D3E5C494532786AD19D87681A2C9A62E
-E450088641804E457D98B9A730BFF04AE01331031B423434A5640B7B8282C95799482CEA4F0C
-411A64666EA177E150F74F209C5989C36638B507FB6A9BBBDFD70949EF1E07D2046EF5DBC54E
-B06819879ABB98E0EBF99C1AABE68110C1602D548A6419EA81B05C2D4A27450CCA5BE336C275
-647E00BFA7660C2F4A244FE8035EAE243BC75ED9E130CB5BB4BED6EEC75D62365916867ADDA1
-E73376F4595F74ABD79F1774DC62B57B8B818EF00C9E6B7CE7D0134F6C5DEDEF62DC274D241E
-D86C7FDDFDB4B45949F6864001AB46886A7D54E68B6D0600D8CF10D26E20B3C795EA1F0487AC
-B1C8CC8017A62CB1BEBE81B946BC2874447373EC0FC3C69C6E8357E30E53193C413AE82E319B
-3B2C211E9CF7DBFBBA7B838DDBD2CCC8C9A43380C7FBEED06F59A19865D9825A87B977CBE083
-8116EF2E2B1A9A6464DE0324FB76E96102CCDB6A84560EC491DD4F4DCEA84BF90D20F012042B
-702B5FA85903ED177A70C3EFAFEE6BEEAE38CF79B3C406511467F7AC5954C480FF800F08B397
-582441AABCAA55E71A70FA0DCCC1C47F697955398488F686FA1DC8454A6DC007EB9CAF641FA3
-3A7519C214F2879A1C602C44DD028D4251432BECB0DBD8DD6925E0CFB368E8545B462BFCD21A
-866F5628698B2033E5A467D5840AD7168E4AE550579AF26962306FE389728202AFA7B905515C
-AF6D1204CF9759B06117FA055EDF877C739BC2878181288B7AF8A1DAB6313F754A5F7F287E82
-31BAFFFD8222FE3D3982097050045D0C6A185E3D18EC1F8B63FAE72863123AF8E68A756A664A
-7CC69FAF050F8B6BD916CC255EF7F8240F9FA9D7FD58E7A3991073993B2D281DF5BD4DA790C1
-51ED7588FDA84147B0582E519ADB80A49268341EA3BEBCA2EDD140D08411D735000C1B55746A
-4935E532A1C5B8DF4FF899E6CA7E79AAA1EBBA234B355CB60AEF590CC654D2B24849EE9474B2
-A87C181175634B86A5AB40303F8F891FB71160465152C22A846DE3752EE3088507BE82553820
-207A79485E82ABD6CC3894739A783659FB9F40B24D94027F5E2C01CC2A35B2191F2C388B5056
-0DCA9B84F7BB209BF75FE3D96CFE3F0F1C58DED0DF7FF3CCBF42C7E35BBDD116ADE65586B4F8
-09AEB29223506897BB82841C4CD2FDF626E3EDD96755C4C33C077CFD66419E3C566D727A3A1A
-4619254089B2A6BA399B8750DAE68964C39ACD33AD491A482C942E9C2404A56628F4D2E21C76
-D5DBC19E69B4678D8A0574A80FDE515D06138043C849B64774A5F8E48788CC51CF943F39886C
-F0CFC28D58AC065EE2D45EA636E764D748928CCFB0687BFD590C8E9D204E6E16F498CEB0A2E0
-5A25170EEC24CF4331677C0C0D0C0E7B19248F8D8A6173FD7EBD4EA2BCE677CE8A5DC72F115D
-A4BC02DC15A0099396CC152A60D6E8DA760EAE1CF82AA67D0210656320BAC56A49DFD171A9E4
-B2794D2FFDDD423E82D5D6848515628870787FECB83C8C41235A83F89BE8B59DBA65D9B57CFB
-911ACAB22978F116E9548354113C5F2593D9F1B5E4CF47873F722404820A3447DF53303A1BF3
-9E7F754DC4FA55270DE98ED46C6E492FC8337FDD5319D3D8C3D9238AD448CD920CF9095C2228
-F7DBEADB9F0B3A207647113AEBADD70783DDB5307DEAD29BD2AF724D7D7625B988A9E693F072
-93D57337814181D1597FA1E7A9500601D62C544BD2E4708009C6B0EA4AADDD7A743ED95B3B41
-811FF6520CCE573904644346280361DEC20CE0B6295A21F6CF651FAFC485A14115255BAE886E
-523F988D4A5E49D0FB2F335B582FF1F1DFBCEA17D66BDB6D2A07CED38603A50CD5847ACB7107
-B15A699F6D4C3A67B63CDFB068758EC4E43BD8755A425D095ECFE9C82B0A6302B7D527FC575D
-16FD056649E7AF6B81E3676025537D4E0A42D43821869A04EF539F96B0B143F9DD1915236D06
-CA24277C0385DCF95A6EDD4F320C2B5241348E457FB8157A6A06166EDEB668641FCB78FF2BBC
-C938981B3D66433ABE825A9AED7D0BBE87977946543C001218C0203FB2276A17ACF38CAB3AC8
-2805D7BBEA37AAA7DC28E01650B0647B6409FE2C5A0AD2769D494771D7B8ADEE90445DB560DE
-4BD2D63FD778EDD4C2F18056CDFB264D947A30DED24C17E14A1E632AF27F71499DF056293905
-E96BC888E9CC5C297D2B70F749BFFF957D9E4B2D37719E45D6330E79EC5E6767DE2999322772
-863DC09E1ABFCE7225F0FD3EFF6A33D1912F38A55801221B96F76F711E9609D70640367DB803
-27D70C384409891F0B66B5DDD35298BC6F0C11051A30AFB2298E48DF9632DDC6E947AAE8530B
-D09EFB065BB9C5BD6F944F815877B7EF36A6F58652FE20E38B2D0A25EF05DFD91C3FAC944877
-29CC910B1A81E73B1523C7D7556F19561998D0932E50BF95207956EB33DA6334FA09D38304F7
-FFC9073AE1A49F7C347168875E468B2EAB092EC537DDCE7D052B490572AFBD4CDBDC2ABC61B8
-9F90A6EE45FC24B116A36614C0DF1C720EBFC1913AB9349A8916D6929BDF2F7CB2604911F083
-6CBF73FA5CB2117A00BFC2D3B2DAA201090D5D23F46EC9B40D22AF84867FECBCB022AA4DAC1D
-40FA2F1F62A632528E014424818A6D2A61CA4FC35089BD03AE8C6158B9E2BA9FE416194DAFA9
-1658E82B9BA3358EA4AA5DEEC0E02A096106BE00F2A2B9DC80FBE8F7D0D412E59934E0F1564D
-35068D3BD3248A4085EEED32FC0F253B5E67D3C945308F768C1A42C3F520AE1E66332F93FFD4
-36A5BE691A2AA22B7879BBC5057CF9543BE7703436618802B5CA722465D3D0E4C17454A4C19B
-382DAC0A42594416E560A6395D278C5473FC6F89E09C9D24CE047917F6155030FCB71DB43804
-D3A722BBF6562235BC23DDD6E15630800FD00ECCA8AB00264A3B8DABF1603E2ABDAACF3471FF
-C34D02FE8F07ECB17EC043DD2EE8716D50F5D99C1685BEA153153CF9BAF9A512D61017E37011
-AB5510741A8F6913BBE76AB0D1F962649ED752708E1C5AD5B943A68B36E3441791C23E1A373F
-6B6C22DE7705D8478F38C4930CCCBEE1223C6BA13BA32F6F09343191619A52D57FFE195A16D3
-952372F23C1112AE9A576A3EE771E956A76F3236D9605D5AA8124E3223536B15C40C69E8BBFA
-4841015175CBB7709F3CBF6A889D724716B77254F0119E131D5F37972D2AD038838482935CA3
-5EBC80981067427EB95E40678BDF7B8C31129F3907EA084CB2D20F800ADA8EA0EEA1FDE7DF15
-A9CFF6059FFFC9FCFF322A0EFC864ADA59B8EFB01C15ECB4ACEBAFE350E53B1A23397631EE15
-5F1231BDAC3BCDC7D921A7279D134FEEFE488ACB7F1C1F10810C62953CB73328B4B31C389B19
-0B8F4B3ABAFEC03A0BBFE22287D4DACF29156629F4F58B3FE032B89016BB8FF2EA9A1026E676
-3C6002EA66F7CF548C3F7BB6A3706189D6E0B4BC654C28451881B25A5F19B99F502F9E486368
-8B52F842ECD497BD73D8ACDEA31FBBB008B4E94A669419FA6995281894189641C3B479C3C30F
-7874796FAA1BB216276AB1B7EFEFE7DE3DB3489A61449D667C9C64AD722029D9FAF44BDA33D7
-D9F7F6C2EA7154B746D5520A74686ADE1D081010A7B2D4E63DF5396D5F0CEAC7AA18D267566D
-4D700A46CE982BE5586BE092D4484D2F39A5E578170AE584D85B15B99F36E8A79077FED4DB8D
-C1E17FC13DB67AAE94C03F97362E5B3DEC3630FB6F4EDC3919C6F1BFC070430F971ADA664DB0
-387F05237E5B963C714FCCFD68EB0208E84B09600F21CE35D0C3311E4B4E34ED136C4F062CF9
-CEDCFCAE5366CCB1672B7DD22EA0831E99B7891BCBF4AD5B621962A82221124A6239789B1830
-9AD854CEC530C151B58532EB9E163D1DDCBA81E7AE7B2B49C359384C8BC22E6376F5C497D4D0
-A173E7B753F170494EC62A61C6521567BAEE498CDED2A269C5D290153D0F08A079FF9349A3A2
-DBB04500DB3C2B0CE3D06055263D1C4415EBA6F7DEEBF173AE826AFE7A69AA4F4FFD129D5D30
-AFF35EC96016986D0A554462B4608DCE6888ACE76E33F503A759D4891A5D97773CE58812B38E
-B50A97C5B9327AAFBF4F8ED4B56701A101335D41F5F628C4C61F6114F1474BD8EFB7A4852FBF
-1A134FAA365BCA9D925126475C1A113B12DFC96289E4582D666709CC8B2E2B8754B1C0DD1287
-69968DF618B9308B099457877687727C728540DFAB6BD320E743BDCFFF22929B6FC156DED1FD
-0751F3DFE1F8A38B0DD1D78FEC2968290D509FEEEC9E50AA1C91D68935E2544320985D376648
-F0723E94E67E440DECFBC6F7DB8967AA05EE85FABA725C823FD5359D86B4EDB85D5729A254B8
-BC17A0E93A53F12AF21C8180C0F797C4D0D80E14457377335439717658B3271B0B6B3D5189DD
-649D913681AE237240136821CA55CFB5D106B55A1CC3C9433115B173EAA10B6BD8728505CA9E
-7A46CB305FA4554627BC906FE456759411A1B3435E7AC5BE79B8389DFD15D2AEC98AD8C1C0C1
-5D903046849797461AE07305DCFF3B5E6D926CD4EAA1B78ED59108EE0464F1F0AF2576145D09
-DF7081F40ADE29081DE688A6EB2B2CE6753FCEC8DE1B8BCE48067AA055A339F265384F46D2F0
-E0EF9164D1A2AA8248EB4C0C6A89598AA22F777E586B856F7AA3440D508D20E8EB4EF299966C
-45D6A639EA42CC54102E4792EC9589D8C82C5C52F626990D08B6DD9171E1FE4624F7F730C360
-0C1939AA0EBC835324D6234BAFEFFD5626924A8572D7B7DDB9DAA96D8A7FE053962FA9F41A9D
-A24F33258856CE4CC11CBE5EE5BAA6B7A52643A0CA5FCCA03484E9E2F01676CA0178D4F08BEF
-8996C286A57C058B836D8150A3F84046D74724714E46254088AF553BE10C4E1E58A49898481D
-75EEFCD5ED996EEE11AC96DEF9C03BD6ABFED0DCD21A9E8B469D2B268CF1A58C4CC522CD8AA6
-81E79717A5CDD279613F889A3279418894D891FDD747DE12CB594C36B6727FF272E87086918A
-C734A70B9C4AB3AEA8C8CE6CAA1677FB26F1E938977C39F13226E0688EFEC084405B99A3D33E
-0DC4D6CE62430A60DA7B51DEBE7AA21DE3D4D19FD64B9B3890AB148F683C9BFC693A63992C6C
-F59C149D5F642DEC0E55F4E1C0495C07C3005936CE21059AD64EFBD1AEBA413BC925E0306350
-1E47C8923C18DC9574FD42D320CEE07D6A529FDFB32739242511E41FB8B747F768B8D1819C66
-44264D5FBA86E777B41E089131242147CBD0111AB60ED77A61CCD78F599C882FFC7257AE8EE3
-C346CBC891E4CE0633166C8466F682545013C9CDB45E9AA52BB8E35747E28455CDC322F76986
-4DCDD0698CCF10606CCEA023636758D8B4A9C7FD22867AB9E09898CD1EF52CDC42A968F21876
-9B816A49E136CA085874AD11AB60434EF2AC5D46BD60ADEE62B1E5DBC62F0936666492115493
-F1A0CF78BB28497A6B08247733D722318A93016BBC6A1ABD41610553C5F5EEF2B445B6017244
-BC2B3E3AD2591A49E5E2705AC15CF05D8B3AD95AE7640B4DB3EDF7CA5716908EF1243D1E1DBE
-710E32E47A54727A3512AD71C6F9B536F7C6A79F8DA80F928B845D85B64E9875EF921A2DF2C9
-D3A98D2A5A877F563DEEC154F88049260E75116BB01CB8229B187045218AF267045E21CFCF89
-40A4EE09AF15E1C7CF26228D039698D5B1A5E06AC84660848706F09BCC82B84D8E02823E3106
-5B45E2FF7AA8DA41BA9C2378DD1AE5BAE47BC7D72B7364B57FBBA43B56AC0FB2094DE5D6B1D9
-9D1FA660F4C263FC63465BBC3D8F538A2828A172DCF2BAD5F9DB1E10B7AE74383C803B02879C
-368205D1C118A84A8FE07D19B3AE4F7CD62B888F595D8C9C82DCB9B55E6BA973F2FCB645FD43
-E563AE4B63B4CE6E80C11677C849FE5174B0D4CFC04843546F8E4237C69B039DBF0450101172
-6777BE80ACFE93D17CB57426E1A290B95D8CD00D33C00D067224893F98672A8EF073F8C6E99E
-25FCCFC8C1012CC0B6DEF0D30C64A3A31BEEB8AD1B465910A50C9FBC4A070978F17F0F1AC1A1
-885A0A33D853517B8388B39394AF872E05AF13FE519BB706E366F3817BF69811B04943AECD35
-EBB36E409A00ADDF747C6E5233EFC0B94C38C8E97AAF991DC41A5E4C76E4A668D52936B059D2
-4175FBD51C2923F0E2945005DF2A5CF042928EE87568F528CA2450480C8618BF965B8370B6AD
-A560D3B4E1EAFF2263D2E2BD6F06F2F5CEFF44B435CE87AE6DD6EDAE5340A0322DF7794C2E18
-0F963702AA1593D73933DD9E36A2E79D6B851344C87905F16D252059ABE078BD01CD79C071FB
-3B5B355B20CACA81D8207C8E0F425A10E563828F676F9CE062CB6FE1E0063B7B5438BCFA2C97
-E78DD8291AF17D0E7C12E5205F7FDB926AFDF0B7D9245C4AFA2414326257E53FBA0FF7918D3D
-201B63626F798E91D2F2ED1329CF87B692F4400707A8CB0E02D8BADFC849974E94962F07A557
-B0D85E17FE485662AF06E2BEE67097398C1050A01814C80FE109D363D388A3C718C277A72165
-85D330B547DBC337A845C18339559262925A9B240467C45C813AD98886FF930CE28E5CBAA425
-C009172945FEB6FA43216620223B96F0C0C12F39DE13DF117FD899F0F83BDCEEF807BA2410AE
-97CADA1549F0E505D952CAFB99FC7F233248627A23E87EF04330D544E136A6A5A8952D4438A8
-4F6DFBF354FF28165DCD69988D12D184141345FCC8A70AB393832C5CE72BA581B2011E85508D
-9972E4AA90796E6F375B3A00D9FD54807803A5D6E31D081A11454B1A0FE9408ACB583B27B686
-B3E1E7C396EDBC063F4375718427549BF1183FDC8DF87ABF25A6F9AAB05CF3201DA786B535C5
-959BA0C4FF6D57A90B27FEF94CD9AE94C3B292F2BBD8A42BEA336E24808C7635B38F23F04C64
-2A64B2C3DCDD868C4A8C553744035E239FC019A15DA741D45AD2DE71FBEC727B55F1A8A4A86A
-9686092CC00C879ED59C1F94567A76E729B91315EF76C0875F92F37487639BA1E8C2C693E802
-D5D6FE15847970F8DE9E616E7A6A2D898276A3BF95B1E7BBAA0646E1B15BFC1E4B8D5F5CEB3F
-BD9DA21A695460E2E41957FA614D4B7E50BF86A38EAE329811ED79AFB8A4C8AB38E517663C4D
-E0B4A9D7DDA0F8A6754AB908982562D5160C58E9D0C5810AA5EE7E745F141DEF2FD5871836A5
-685D5E6E3702BC6A24F0C8E9182284ED6E16A92A5EB512F4EDDEBE55F6C86E7CB8D71383C0F4
-9830B3EA2FF75DD2956F8F3FCE3363DD49FB85D51A62644F757F34AA86F353A8844127870215
-81D2B149A70DF60DB556B4B321FD839CFD80D257A1D307904402AB5C3EE16B501F9430C83431
-96AFDE969444625C3ADB1A1734DB155D54CCF9FE928A81485E266BAC53895EE70732D1F345AB
-645A4DA959C8A08E3EAA8A968097926793CF91DE5DA8025B1FBDA7F323562E5DED86D36A0BCD
-7F84E3D1F9EBFD8823EA952C1DAE9C0944C13B3BC31412C1A5A44DAE7401ABC391F7199BFD7E
-8BDF43FDC052D499951C6DC3481322C6C7A86E271C776ED12FA4670072706EC9CA98D761CD16
-271D3566712A56A2E38652318B280F947C7604C9B75D2F88837F2B06E3C945A618631F429F4B
-305F6BD26BD8AE3BDE01A086F2789831F8CC8D413B17A1CD82B36C3F1BC9EDE24D0F0FE43B49
-62E95A9AF21AFAF89D01A1E245A6DF385EE527E881F9E73537A93AA68279A579C259B409E460
-0635DFE6A166D014D719CB03F4CAF828FAE233C3C7ACD57274676A0968449BDE0A0EA6F1FB15
-FA3C68915FC3DDE2FA71C1F1B15CA794C9FD7EC9E7A216D2B5F1DD49BFC714ED7B3375D65C46
-3EE12E32CECD625E0B8E0848C507C0689AFAF31845F42648676ED1F8E9472B68FF91D3A7C8AC
-247BC169C738CC1D2D2C3F0FDC9FFC01BE679107A018D2C85C8D457B5443C24CB5310975E5BB
-8846DBE7B79842E64A0DB4A5BFDDDA3A1E2D9F401938E621723771E2E313F5039D99C429A6B1
-D92343FB2C63344F66A6E35218F9468913126C4969F35561EC734BC83121A21961B790D0964E
-1C6EC73E4FCD066EB57EB2BC769A77E70F18CA8B06121F8CAA554C3901296CC89559AC92D6FF
-D4C03E6660E801519CDBC091D6477C152A6AFA1CBEB2BC8F553AF3A56114AAF27AC7B729AC80
-E5592519628CBEEE1F1DE9357496B60C082F77188F59D55F95116E86E917B6FA25E8628F6D47
-3EB17433F828B1DABA1DE045477E1DDB96643143D99B98B4429212739002593B1FB82734356D
-90CD70105FEC564F50BC8823A9C6DC25230270EEFDC43C254B46AEC991E0FAEDCEAC965DAAA0
-353B324E791E46794CF0467C5E81DC79E0DEC212F77F73E8393FD87169C27A078DA0334BEC6B
-1069A4EC4912D3E9B206F78F6B278FD65D5FB96671850D1F61C1920D49AA034C4F1841EA9562
-2B71E1A77D54BD776440DB63316DC376418467D6629774E86D000420EE1DBFECB8E21851C688
-6503031F842FDCEB041AA63D55148D15E31AA67BD9079E6F6232062F7B189D9A4D2A13D71281
-A0768E6725957EF8581BC9EE368322050CB857AF2097166DEAD851B772FB8967F36B0ADA7E1F
-99AD039163B9BB02304D3189978F33F39052097F63E30699015F83607D834D8A3AD42A163C16
-11B4B85F40DFDF3529A2126D467AE8672EC4BC285F0F1DDE3F4A96A4EA6E224E8FEAB920C794
-87F30F06CA5B5906945574084D23605B98BDC8D6577C9D926BFFCBC5FF9369F44E80CE3F545B
-E16EC21D1A5A1AC18F546BEE5491C50FB9F8AEC5FE2FE555344C829EED36F5287F9ADC202D2E
-E3F0D24C9735BC952407A645F9C9552C1AEC230F9267ECCA1E3718ECAB42C531DDDB91B482B1
-18E3BED209C282BCA4556F867EBE4516C2323C53539A83592CFAF1F97200EB25772E91D5A4A3
-4A2E9250226D2F5C7B81C2E72782B9E3B3F1A88B06ECA767251D8F763C62F424948E7ED3DD05
-3C8C00BB3057C7A3F2ACDB7D23549A95C14190595B31A544D4B2E50FAB5A8268AB2C1A7DEF7D
-6E40528DAC5199421C8391BF042B1E136B057BF5FE8163C0BC71BEB3BD1A9C8CF13C0F163AF6
-2B1F3C87316710A417E94AB5FAB692EFCB0907A361240A13FF321E6B47319CFABA3973E2BDDA
-715CF2116A4EAE6A2904A599F2253C8FEA23607DFBDE8A89A3B31FCA441119AA551BB8D1188E
-648519D75B196F3EF3990E3FED01A81615A0CFDDBF0A46B62503AECC0A3F3BD25FA222B545D4
-C9D75265A8869E7148E7E992327D439DD63EB1962760EE72BEB9E16BBA6ABE24EAFEE7DFA35C
-91AB9FA03F8DC6AD75565F7968B3F6B0E9AB33C8592C08F21A11032956DAE46D9DCA09190717
-1C83EC3FDA6D81FBDFE84C20D28DF83179DA79CB6196347DC9A4C7AEB6E65F43037A29F5AAC2
-4BED20B21A83C73BC785439DB31E1E874FBA140FDB12DEE62C6330D1A7C5DF70AF507188CA5D
-74A2598591314B57DA93EC600EA2AD3DA8AC07BFBA27BF67C563659CC3DC55904D25F890B890
-1E782FBDEAD982087903A48BA89E8907619EA907AAB079681BE74DB150142BEBA6BC6C72D2B1
-99829262360BBF0C7E657180A96E6B81CE30C05BA04A073B0151F4404A7E13E369A80B71426B
-D8D3B573DE2BBC9599F67BA1A3050181D2A098B62B5EFEA2E6707F4C06D93B456605BCC402E2
-635F991A378897B9120095ABCBA75C08EBCAA90CD677827F2EDF87EF214656438FB890F1FC04
-D90B0EBF43C1DE41AA6367CBF206233B5ACC257B25B93A12717B25DEAE5EA6335EA8796031A7
-C5529D0065D3EDAF3296ADB44BFDBADCB4CD3BCC8BD5D718493A6155A27748BC268089B2D341
-B19A7B346B4E4FFEA44E015315C157299FAD4AC00B24A448022D7952DE7E9D845403CD0C1ECB
-EB8D596A359C215E4BF2368CFF9C71C03A907AFA1F8921955751B2984071D0238835BED2B2CC
-DC900EF320090FD5070E2DAA93F2461F03145D4072E620D12E34E4F6A8B632488C5A8ECFDFBA
-425388EA657BD19738533CD04FF3526FE99FD3EF5F276363F687E2661ADB87606E6CD82D6F5C
-B7F4E23163D146F3DF82EFCF6CF46F1AD62DC12768979E56C2919038BF95F4CB110F743883FF
-B6F3A7E2DE13992E529A8D137C8A0513C0877B2E7F88191323EB70A97E8AD6DF9033FC8EE538
-842EF62EAF6320371A2DC4A00BD823D3BF704C3225D4C37CA4362A8B4A4CC47B65F619EEFB25
-2154634990D0381BE07890B5BB2402C5951F0E7584BB071622AB05E1B052469FCB949B486D75
-C93CF9A80E8DB3F03E7139A6C8E711A5D4984890F9AE03088D821A5D98DD8EEE9AC91BB71330
-24B4F1CD1CE8F540D75152EE20F45307D83435525386CE1FB335CD9FC643ADC1B923B50E72FD
-64004676DF02BD9DE6836084E863D102F929AC6CAB1191451C9FD3846E607F7BD0AD1DB2A8F7
-C971BB31E74D2B35FC3F00CC8CC90B0B4BDB5D3BE91ED6CE10196EF2410979D6D7F42A499D3B
-35F715BC063B346F5C6A823159F7C083DB0B27622672B0EBC96CBC2B171AD45A2E23875525DA
-3A6B801EF8A612C7DC7ACE0CB7F0DF292B9492963C822B6C7BF0149F4AB4782C64F10FD21030
-FF2AF361188A0E3FDBFEC64778A8DDC84448AF0DAAD005D66641267C5E654BC65EF454959274
-E4C32B8FC592B93F79731560CC34054A4B206613E4BD2D0F14968B1305E76B3C2ABFDC0CC4F2
-724A3FD5DA2F2803609A1B5D2103D71557663F979D5249D077CC501E8E0EB6E9F470A5FF5BB7
-1A20A4FBA492E235EDC554A30FE3BCF7F9D78261CCD89A6EDEA838BE2C88F6ABA32E828A0CD0
-60F202BF38FD6712A10BADEF3531583E2902B63031BE0FD6AD28FCF74614FDC00268CDBC2391
-67A501E216C331F21D2302C0647C800ED058F163DEADDC7EC270DFF5EEE405BB8093FA5277F6
-633A21280908893A570F22CE34B59E192CAFD8034348B2CEFDE3389107F6F6DA293CA5BFC0FF
-204D350DC642F44CFA2847CECB3F5BB85EA2E268538D1740ADD2947620A345F46F97280B2240
-8DB74A96B5BC1BE5F8D81F963521C194B641FC66C48250317F8D0058D406D21A20C03475BDFD
-0785CAE59D1D61BDCB228F9A8E935C52D9E3767A2B6D38F8744E82F1F495ADA21AB62B539BD5
-F2E78BA890668296E64755DA92259C4FD6A65F151BA7CB78330AC29E4F86DCA57C78D667ADD7
-B2DBE89E76A8F0634D4441F4249973F2188D74B5CDA7360F9AC8192AE04029BB421E5696CA86
-2E1513249127A01904A6D6381D6347E99D422258A0E47F170D283CF47DCA2AB54747BDC56095
-0950E275E3E69B2DDAA7989654F9512A45530AD720BCA7F59FE65099E1DDD8A4D90714D85736
-4A7FDC9F0A84DFC86200AD900FF633BCD1F8B985495AD5EEEF776299AA3E74D11D2C81066A0A
-2AEBE05CC1BEDF2A6B3A5A6F6372EB30E77E0575E327A5AD23F188D16BC7D1DAB20AAA3F0709
-B8709CAA1661236443EE9B6145418EC5E9E27B73CFC884D6DE956362FDC0C254992557778FFB
-6FF0367DA6188502F524AEC46857D9BE66AC05E7702284D6255DC5976A47A976214AE793E081
-DC36AEA6D5EDC84E116F7A5F294F006D34A68FB8D0284E1D9955EB619D2465A21DFC027680F9
-CE451F5705AF47489AE0A1581A9817A33F2A6137017560E9D2683030CD567E4433415BD822EC
-285F7EEC241533A4A403D811A1E01D3BC21C6D0A28F21A44BD404E46C0546C59F6FB8B9F3B62
-F11A598FC018758B3B62291ABD5C39E03477B7EFE8DC8847DD258EE6F1CFA079D831BF3B2893
-1EA51C02BB8F3C729B4BF2886DD7CE05676235B3C63C671CAD6194646D0BF1BC5349A3A31D0C
-90FCCFBE7B84161BDF8C09A46A5A715A7C7AC7EEB2715543CDFDFE0088755CD5C80C6ECDA2CF
-30DE170AFE855F837B6963729E9F970BDBE326D40FF910A483B7A46A65355587E07BFA503590
-A208E0679486BBE19BC36C4F79B692BA39BC8B2CA89C8D11B6F4D1BB8896790E5FC205D5D2FD
-CA560E11A0F574338288DA041598C746D807AE064C18DAA834B63B55F2ADEE166C85862802DF
-F3D9D7334C0A06CB1D4898B77803B0D9B7E9B7695BDC15CF460D844D5FD2AA02482C7741CBB9
-338BD67CDEB37ED910364C0B0635D8317854AF661BB4D1FA2E60761224DF68232B718A6E2F25
-0350C46D60AD84DC329AB46D64E6A4839E9FF57C70D26CFB7C07AABC6CC0587AE8DEDD4A9E39
-CEA3D17E05A74724B115BC8E5430860EBB2738129A46D0AC5100ADFE04E9BDDB6AC4FA2CCF8E
-441161A3E8AC2051DE7C04F2B2ED0080EEBD5AC348BB6F1740B38A220FC4F5B2A110518DCAA7
-67ABC053637F24A76F1F314E21A41217BD403420A1D6EC046C5D7281486C421AB5D64241D21E
-533EFFA79670F96241A9C947B2E11E38D6382A345B1AA679DD4814112E69DF58B919B8B51506
-17DE7F9FFC3EA80CAD25BAB053F5F201BB34D9F5ADE6C1F9609BAED2C9344B009772D8DA7E0A
-7E9F22D01A9FCD8C5B020224563A631F5B9AA7257EDA6DFF1F8D9485744CA2DA626624574DDC
-7D0C53C59B4484FFD604E5FD6A46E1E854AF1CC647BEF84707A619C225EF2057ED4ED522433A
-C11E21ADDD0B6D1E9A898A39F4FE2E547298B8F3725DBBD16AC64AB5E6FC963A9A65B2663239
-5D28D10150C1F17BC20539B3F35E6F518721CFF54DAFA1FB36E686C66A27C552CBBB04D74F15
-C19BD418BB38BEB6E83AB05F48D9EB8E11100DF90B4526CDEBEF6E2CAE7F60DAC0D284490ABA
-6CFBF1194EE8F8CC56ED7F714267140FAC5B852D68E64ED79C109D77393499831EF413DA8E4E
-233F5DD49231F585E148DCC1A4623323CD710408FC6B8D55FB5D3F8B51FBD7AB4E59DA3D7F09
-CA64BF4B48CAA2A01C3E04791742BEC5EEC021EB356AEA9434D24876004268217352226EF1C9
-C4602AC7256EB1106088692F7EA099F8ADED5EE6044B41F257780990B7AFC50382EBA95F76BC
-A0AC595FD9CDC5544096E8E31BA5E4C313E631CEEBC515745BBD0913B2F9AD8CC79CDD1B6881
-2410785C4C795925A2563F8AFA996A8E6F4B0D069BD2B0C7E7757D22FA823072391B1C3BC235
-C77D9C052F8C8474B35ED6FE23925981B249B624F5F42F7D3F6860D09C65E105FCDFCB8F773B
-F59B5B39E37F45C5FD51E3654FD5F4107D738036A27A200260B8BD84A89764AA74F5797B9C66
-69F58A00A6E66767437BC37F978212B0FFF04D2A41C97A4B15F4E6A50DE388BA2187832D3B14
-6F0D23F5ED5F8F6E48E82F30998375BC3D172096C0E555F8B253CA73F003A6E12F78F90E9BC5
-6E94907AC59BDB37CA268E7B10B047F6A34D17B10A48E313CF90AB4176A0D8D3243911FF3527
-3A1091A793ED3B164284356A3688B0F4A3FE855FCEC3D73BBB2970B090619C594395F0E2C3E6
-A10AE07E24D3FF12DB03A8B410430DBF93F4124CF1F0E1B41144D1A51D377D160341BA174861
-9D9793E5899968992F4AE3EFDB7E68F21084A1A4B53D317CC138CAB8CD34AF674FEA0F03401F
-0888806BD9393B988B7E580D40C46C30374CCDBBE1546D569087FCA1A2D7580C1D2592B397E4
-7B28A7C30BD2062D6333FF3071EF6BCCFCD68758BA172071D7D8A83DB37F0304E20A349B0854
-EC6ACD98B1BDB5CB913187A9E3176F2954484167CA4D9A8E13F7273D843802E230D8061E02FA
-81EF7FBC93CC1EE297393E330864DD19A71CD5F01AEC327BF6C22D6A47B73201222CBE9F7397
-44E14B9EBC1233D086408ED92F9F8CFB11AA283CCF77B1473DAE1E00C5890EC07C68F5CAA1AE
-FFE937DC9C946E71B60ECA6D05C75FDC436793D6CEDB8458A32870E4D4D0BF6015FC9687BBA5
-88F3C256C3795AD9EBF710F2BC0AF4DF3AD34622027F412DF3B0CA26C8D324DD63DEB9E5C490
-D70FE2DBF83107C00285D3D1860FBBE4BB7FCB1DD37AAA2FAD6DFA2EA2608DF1B29C499B66F1
-62D9FDE39C04FF948647234E2C182A9DF05FE1A8CB7C7248BF261EE5824359AEAD1E37CD4561
-2066E517D7DE23DE2102334EF2DE91D69A646F1FA0FBF0B1D5E48874811CCD69D6AE3D82EAE0
-4BDF48A0CA9C7AF4CC9F4B9AED930BC66DE3E48E30DA7B6FEB29320690EB62AC11BEFF1BA528
-FD4E582CBF447CB3A4E92F15593CAEF1138392516EE452ED4F859AF83077B4047BCC17C7D297
-169E43E75E306A173A30B3A3C8D52652A0EE37A65C7A32FB81EFFE4B33A2EE95290BFC56390B
-DAC88FA47B1DEF4C92CBDEC3ACEDBCF188DFE57B1CA90983B4DF7852F9E4D70096F780C297DC
-F83D9900EF0D76D09113FDF1CB98DC4A4C623CFAF4900834C743756FAD2D1608735DBEC04374
-28894A19BCBCA91B757F5FD8E3D8673D8527635C31E8CDA6DF46B28EF4038FDD8D09B4324EDD
-2BD1BB0CE858AB8BF8C535CE63737AC9F156F91AF82E8F3366A661D7CBA84FDF0BF60B41B089
-C323F47B570B061E908F9755CA24047C937F11FA1133258F8EB958421743F10222FCDC3F5CAF
-3E0278797CC9806767348B9D0935A5053227F3EE3F9787046796CCB3BC63DE17EE5A9791948B
-9C055F3EBB14C68D2211702AA0CA4856E8AD5BED532D728E0B9CA42F03BE18B50EA47BC6725F
-3FFCA8EB2838F78A26B397E73DE2DCF29D739ECE2337D309D66E12743A8F77BDEA9935CB1499
-863D972BC1E2C094B7A2A8D8C0927CD587EE86BFA9CBF7599CD79759910A414C783C9D7AA6EF
-5093EB3183617A44882E3B47A67D1A5EE4D53A6F451F1C1DCEAF935538158C4EA81863EA1638
-3C376B302EF74F86536AA9C769EF507A60E35DA6BEDC3BEA269D8A8E4CA75C39E5B18B0F69B8
-E78255E9940B00E53AA4CB2724FC22553486E618F5F8C852392081CDEBE8FC7650036602178B
-84C8AD55375C20D9D5476A2D62DDE2FDF259DD1ED708285F87B8F5B69ECEE72594996DC72FBA
-1B498A40190F045A920A3219B9EAB7C9754EA7C864BBE72A0F95D9ACF539400D08029F363CD2
-9A3050C95F80860701B3C399FD047027FC15333CE9069F8D4EF6F354487C98862783BEE86EDC
-5D073BFAD403DB4C1403EB31019ED23C96C33890FE68D75355F35C0820FA08257A09E736BCA9
-9CAAA9C93E34351854EEBFD4B80B7536FEDCEA115EE8E5D910D0C95565317CD1FAE76A20FB43
-151F7ED5E48AD888FACCCF41B907676D290DBA1DE329553C393653B10D62BDF590CD73F98700
-1A0A1FE19041C1A60F92648C02E643AA7D20A5F27BF8F848D669E167D4167BB1687C148BC45A
-6AF5FDD144C56A3DE655D0F6F7D217CB33D2291940AED6D1A684876D216A647A4F665B59FA33
-F4CBCFE1E13174A1B8856BA8800F0282C68B72BA7CF545DD18C452FD6A769B00B38860C1408A
-BEA2B06A87E1B8D9C6E08B3DF98DD16FB3392950F6F0D11763D82C7D6BBF3119EEE58A710397
-815D15F71D2FF3484F71E8142DCAC57DE43290303CC7C16957FE7D4E2683D5303CB1CD30E61D
-8D4C5942777DBDF2A9D564EC0526336F09C7028F2B8CA0BD988B3B71A970A47D13E3600EA390
-6416B44A7D6138EDB9B7D7E83F824497E4ECD18CB7004B23D02314FD30B2E6B732712F976629
-5AD04947FDF616BABDB1EB68B8DCBDAE3B16A829D6316990FA726806329D34011B147FAF96DB
-5AFA4624B03A6A9046911277C1F509619592CB5BB517C4AF96B8DE83F06E03DFF6BBFA560063
-CF8073223CB81A6541487C6183737F0C0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-cleartomark
-
-
-
-%%EndFont
-%%BeginFont: MSBM10
-%!PS-AdobeFont-1.1: MSBM10 2.1
-%%CreationDate: 1993 Sep 17 11:10:37
-
-% Math Symbol fonts were designed by the American Mathematical Society.
-% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
-
-11 dict begin
-/FontInfo 7 dict dup begin
-/version (2.1) readonly def
-/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
-/FullName (MSBM10) readonly def
-/FamilyName (Euler) readonly def
-/Weight (Medium) readonly def
-/ItalicAngle 0 def
-/isFixedPitch false def
-end readonly def
-/FontName /MSBM10 def
-/PaintType 0 def
-/FontType 1 def
-/FontMatrix [0.001 0 0 0.001 0 0] readonly def
-/Encoding 256 array
-0 1 255 {1 index exch /.notdef put} for
-dup 161 /lessornotequal put
-dup 162 /greaterornotequal put
-dup 163 /notlessequal put
-dup 164 /notgreaterequal put
-dup 165 /notless put
-dup 166 /notgreater put
-dup 167 /notprecedes put
-dup 168 /notfollows put
-dup 169 /lessornotdbleql put
-dup 170 /greaterornotdbleql put
-dup 173 /notlessorslnteql put
-dup 174 /notgreaterorslnteql put
-dup 175 /lessnotequal put
-dup 176 /greaternotequal put
-dup 177 /notprecedesoreql put
-dup 178 /notfollowsoreql put
-dup 179 /precedeornoteqvlnt put
-dup 180 /followornoteqvlnt put
-dup 181 /lessornotsimilar put
-dup 182 /greaterornotsimilar put
-dup 183 /notlessdblequal put
-dup 184 /notgreaterdblequal put
-dup 185 /precedenotslnteql put
-dup 186 /follownotslnteql put
-dup 187 /precedenotdbleqv put
-dup 188 /follownotdbleqv put
-dup 189 /lessnotdblequal put
-dup 190 /greaternotdblequal put
-dup 191 /notsimilar put
-dup 192 /notapproxequal put
-dup 193 /upslope put
-dup 194 /downslope put
-dup 195 /notsubsetoreql put
-dup 196 /epsiloninv put
-dup 0 /lessornotequal put
-dup 1 /greaterornotequal put
-dup 2 /notlessequal put
-dup 3 /notgreaterequal put
-dup 4 /notless put
-dup 5 /notgreater put
-dup 6 /notprecedes put
-dup 7 /notfollows put
-dup 8 /lessornotdbleql put
-dup 9 /greaterornotdbleql put
-dup 10 /notlessorslnteql put
-dup 11 /notgreaterorslnteql put
-dup 12 /lessnotequal put
-dup 13 /greaternotequal put
-dup 14 /notprecedesoreql put
-dup 15 /notfollowsoreql put
-dup 16 /precedeornoteqvlnt put
-dup 17 /followornoteqvlnt put
-dup 18 /lessornotsimilar put
-dup 19 /greaterornotsimilar put
-dup 20 /notlessdblequal put
-dup 21 /notgreaterdblequal put
-dup 22 /precedenotslnteql put
-dup 23 /follownotslnteql put
-dup 24 /precedenotdbleqv put
-dup 25 /follownotdbleqv put
-dup 26 /lessnotdblequal put
-dup 27 /greaternotdblequal put
-dup 28 /notsimilar put
-dup 29 /notapproxequal put
-dup 30 /upslope put
-dup 31 /downslope put
-dup 32 /notsubsetoreql put
-dup 33 /notsupersetoreql put
-dup 34 /notsubsetordbleql put
-dup 35 /notsupersetordbleql put
-dup 36 /subsetornotdbleql put
-dup 37 /supersetornotdbleql put
-dup 38 /subsetornoteql put
-dup 39 /supersetornoteql put
-dup 40 /subsetnoteql put
-dup 41 /supersetnoteql put
-dup 42 /notsubseteql put
-dup 43 /notsuperseteql put
-dup 44 /notparallel put
-dup 45 /notbar put
-dup 46 /notshortbar put
-dup 47 /notshortparallel put
-dup 48 /notturnstile put
-dup 49 /notforces put
-dup 50 /notsatisfies put
-dup 51 /notforcesextra put
-dup 52 /nottriangeqlright put
-dup 53 /nottriangeqlleft put
-dup 54 /nottriangleleft put
-dup 55 /nottriangleright put
-dup 56 /notarrowleft put
-dup 57 /notarrowright put
-dup 58 /notdblarrowleft put
-dup 59 /notdblarrowright put
-dup 60 /notdblarrowboth put
-dup 61 /notarrowboth put
-dup 62 /dividemultiply put
-dup 63 /emptyset put
-dup 64 /notexistential put
-dup 65 /A put
-dup 66 /B put
-dup 67 /C put
-dup 68 /D put
-dup 69 /E put
-dup 70 /F put
-dup 71 /G put
-dup 72 /H put
-dup 73 /I put
-dup 74 /J put
-dup 75 /K put
-dup 76 /L put
-dup 77 /M put
-dup 78 /N put
-dup 79 /O put
-dup 80 /P put
-dup 81 /Q put
-dup 82 /R put
-dup 83 /S put
-dup 84 /T put
-dup 85 /U put
-dup 86 /V put
-dup 87 /W put
-dup 88 /X put
-dup 89 /Y put
-dup 90 /Z put
-dup 91 /hatwide put
-dup 92 /hatwider put
-dup 93 /tildewide put
-dup 94 /tildewider put
-dup 96 /Finv put
-dup 97 /Gmir put
-dup 102 /Omegainv put
-dup 103 /eth put
-dup 104 /equalorsimilar put
-dup 105 /beth put
-dup 106 /gimel put
-dup 107 /daleth put
-dup 108 /lessdot put
-dup 109 /greaterdot put
-dup 110 /multicloseleft put
-dup 111 /multicloseright put
-dup 112 /barshort put
-dup 113 /parallelshort put
-dup 114 /integerdivide put
-dup 115 /similar put
-dup 116 /approxequal put
-dup 117 /approxorequal put
-dup 118 /followsorequal put
-dup 119 /precedesorequal put
-dup 120 /archleftdown put
-dup 121 /archrightdown put
-dup 122 /Digamma put
-dup 123 /kappa put
-dup 124 /k put
-dup 125 /planckover2pi put
-dup 126 /planckover2pi1 put
-dup 127 /epsiloninv put
-dup 128 /notsubsetoreql put
-dup 160 /space put
-readonly def
-/FontBBox{-55 -420 2343 920}readonly def
-/UniqueID 5031982 def
-currentdict end
-currentfile eexec
-
-80347982AB3942D930E069A70D0D48311D7190FA2D133A583138F76695558E7AE9348D37CAC6
-651806D08527C1BB4A062A4835AC37784CC39AD8841404E438B4D52D3901E47A1DE4F7924E0F
-B3DAF442499175BAB1226EDF692A4956739F8828E80592F450C5D5C22AC88BCFBE9748F61D18
-243A16F4A4467F084E8E2BE46EF47FC51C3A8199E3CDA62FF9C4FB73956DAB8B6683D2156377
-808CB35026073E80523F59A30D195FCF9B9FCE4FFAFC6D56491BDECDCAFDC988206C5A457A19
-270B37D0AB776E03EAA7EB568EEAB6B5E79DEC03B0DCBF923A2AA8E4F4DEDA2CB043858D8B43
-0EFEF0ED914A80FFC9818ABA0FB30DAE3694E5B31DF3855892D59B897F82FF79FCE8444C7926
-FFC4EA008E63BC518936C9098AFC2D1C14030A2C15BCC9285C1B57E80570D2F6301009E63D9D
-DEBA44E3251C75CA9616A2C5DE9ADE661A15DB37B40B67C1BD71F77343912C62F47DD24BE0E0
-3DD833EA60D2EB020407EAF042FED30F0C2F01956C83D46D982B50DD979C2026ED5F71979DD9
-8C9C1EEE25054000A5EED5CE9498B2CB59A7EE55E0538AB47D86E69B2AC79229C8C6BFC93A24
-E822A2F1C9964FE630344579A25F66646113F73D1315C447F070C575DE76266C6773B01137E4
-A039DF5E4C2F199E8E3396EF9FD001EEBE090E04384C4089EC9F13A07AD8733C6AF775F432F1
-54EF50EC4477E08760C03C006E4D1EB1C7CDACB8955CA4B0D62C7DE046FF58D431BB9883D285
-86552C51809E3CED4DFC2198B97AD9E52EA33C91B9E335A80D9B77D3430E0487930F99245966
-4ABCA5DD9177150B0FB502FBB04CCAA94FFE2697ED07906F95FD98E7B325BFBF10EA6B39FD76
-C1E5295022CB39E9CD471B7B0B3F9D72DA99DEB0615D35CC22F42DBF5CFB70DC6279FD520D99
-F32201ABF25E621CFE1D37961D9C0AD36EF3D9264C547638152ABE30812B1B98CEBE66CFB9D9
-8E4EA24E4C0D4D77ED282661248DDDB6B550692891E6557B1F6238BCA064DDB13B68D37CCC74
-EA8110F7E0BC4C4F903C8E43C772EB210276CDB3B1DF2815060A0D29C94FB764B5BF40377AF5
-2248D706E117DE497B5426761F4D1431E38CF3F9E386D35F95636EC78DB4FAB1CE3CF7064458
-2561E1595EE6F32FE27C6A759857EEF72F29CD64FE015C020AFE7C7276C5FC4EC8F004722411
-F5386C1FBFBD3C78A2E6815DD18629B47E09ECF6E7F3706A4B080FACA3A7BCF65F7C2A38117F
-20F2B41529AAB9683F6128BF3EAE3714F5A5E5EEEF1306D6FA586F0EBFCCF49C7D65C9DAE335
-66FED98F0D64D0C3E3513B45FE6EA1B18049EA3E98F6FBEFFE3690375CA7F5D844F38CE109AD
-A4399FD08FDF5406C0316645AA63413239A45E59D1DF5D8E4A56DFE3004D7BC9F74C91668EBC
-429B4A4136DA8838FC897F5B265474691A86A2FAB71E295C704F6E8C7C6D88E38C2A0052DE7C
-7BDB3CA7F6FEA9D5A7140D4ED0A09077938E8D0E295FDBD36F18C23435C173F5F594C2B31CC6
-F519E1BEAD65BA0898B082F6B184484277AA5BCC8350B40A0495DF005EA5E0DB96366583D1FD
-A37242F1A6635CC6384B3941DC07F625CE8DFB5D94151B32091CD42D74A31B65D4DB66BA8116
-4918688A7AAE6274147CEA7B60DAAAA3D7EB494867F406CD2B349768C3A4D97CC371AD690820
-C31ED465EBD31DA12837829903FB947FCC3C952C9634A5AF421F9516A1BE4D6D3EE77E5ED175
-B8F45BC58735B72F30CF733AB137AA1936E6FB3732C67EC769A0FD213E3DE004A25808192969
-A6B0A02ACA9E18E905598DBDAAFDFF5A747F172983582037BD2DAF435EAF41138C95CF5C4420
-34388D3C6FC34E0A56E857CE703618E7CC59961F29776E7C37283F3A63D82E2E07A4821F6CC8
-B7732ACA1A40A22459A4D753039250A1D8286D24A72C6499DA36065986DC72BEDC16AAD6EBF4
-CBBF8A7D8F71C38C547C21CCBAA1BBC8D9541E13F92B8B304903D8FC9D21D2BB5A4346DBC130
-E2D4F590AD34FB8C3B18A7446EDEF65E1D9A5B90E1006E559E5EE90B7F2EC8DB355C3BD69E1C
-A06F4E92E5FA226A922F845C46A3D78BA3643CC0F87FE47039C5B018D0F1E53E83796C53899D
-A49B9B36515B8096C3423CD63DF7061A2646DB14FA15D66E57DE63D83B296937E90712BC39D8
-C1EBBB146F3E554012EEC00A6CDE5B900631DB48B189F7405ADC992F6F5C10D1CF8907E11830
-BDD7E7D8CBB7CFE1BB7878D3625EEEF8108CB65D1B6B740F9824781EB7F12C8F90D18BFAD539
-F6DF56EF3BB8B249769DD70BFDEAF4091ECE5A75A7AF9F690F4829684143DD95AEC471EFD607
-CE75EEB182E1D1E2136D26FF6ADFF7C30F4379E0E729BD1DC7F6D3915ACEF2A765F76E80349F
-85ABACE85ACFFE774CDEFFA87AC4F6AC509D4E13554F06C1C097E9D46A9EAE8EDEC4E4E67409
-C52A27EDA5ED00A5DF0897638A91B9EB420121F1F9FCA481F0F4B9A0E33B5F0D07B202077688
-8C555C9AD63E76DA475797CDDBC6352130F1207F70235D6DFDD5CFB4C733BE582124DD326AB4
-D90BDF98ECC6240F1B5048BEBAE1E6F389AC2222CC5E6B670121F74F20D635F64CE6DEDDD54D
-541A1D27B3F1DE23F65B988F223B60E3FE02D3A625411F263B2397D9C8926A07143675AB71E5
-29A3502A41271AE4051EA2CF27DA2AC34E3C6B446B1F7A1CCE0850D4AD71A5268515FF9CF430
-2D15EB4BF58A3B7E446ED411072BB620A14159DD3BE2E025CAEDF27417234C315B1A4372E6ED
-EDFB47D259A9B387EEBA5C73F26D866A59F6141C4BD3491F84D8C83B17B9FCB14F356390F041
-29F7613E63F857B67C3DCEA0C90AD82AC2EB8C09A5F5D4AB13244DED536017A5B08C3830E58F
-2967C079F84CE3A4426DB1FAF796B508B0AC52509A032EE76E9F048662EBE76B38FAAF702CAF
-83279962EBB4A3B78C6643F006D97CDA50743DC1588BD9A56B42338826199EC6172A8A38FE8B
-541CCE72B38E53689A96506C52C2DDF8D91EEE308D9A21951E8303252AFD186FC7DC42982D47
-FEFCE3D5CF2A2CA31CFB9A78EFCCA4169DEF9A50674DBD232627A6720FD7F6D568A7B93F2C0F
-12D8FE7F8C18FE9E0191B0BBC99B8AC5E48563F52C10E4D57D937F2FFC1EF71D54E0BEB4831C
-5B3008D8AEE79726C1627CB8AC70BE30640BFBEB7AF398CEA463F9AAB513D09594564DE02F33
-D6E88341C1613F47E4F99269F27E00E503D867A88AB0DEB6D0247DAFC6EED2A9E0221F49BD20
-160842782EBB5C7718E27E051BE48FF1B4CED46D0E6CDE8E133070A4D0080987BB197F57B762
-841DA3D750D0975F5086C0D993B36E6D802B740A53DA719F4CCB7C36800536E503452FBD0E75
-A86BAF0B1186681496D3B5576721F12DC34DD13BF2E240A89676507573AE2C8CEB3647D2584D
-EC0A970646C7089B8581F528D520995285688D39C05BEB9D574D64150ED98C79BFDB82B002F9
-D1EC37A089C0904C672E6E873950862713B05FFADBCDE299CFF3710CBA6A9DAEC2F479F31E21
-83E00BC5E90FDD0CF4EEAC202AD254E4F0836D69EC56DA6283206F1145B2958A9729778030BF
-BB570F5B4B24BD7F4C82F38813A8F2369B93D7B1D6AB67EA1F3CAE838F786781AD0DDC0C3A9E
-1B654330846DACA3EEA7DABCB3788A2116927E0196C504EBE144B96297A48110E25602FC1CD9
-8D7B2334AA55753DDD23BD7135F98715054EF7F74F6E4482C505896245B2AE4CD4F8D1F14DB3
-B6EEEE9D553FEC4E09A53F40B81498E1DABF8073B1D8DD5A564DC137700FFD1E0BC5F20346F2
-C48A05EC532870B026F38421D34BF43FEC5A9C3C17706296001648E93B7F73569BFE6F407D2F
-A9F183ACF09EE356DE476DC2210F5CBA7BAD155B0C4B6204AFC92D5847E44CB4327E2143F1E2
-8867EE5B800C9BD7AB01494E2A8F6D0C1D7A647F3C1357C0DD74CF94B167630E7DB2ED3075F2
-901A3FF1464F0A5902B3D3820F4D20E3A371A9C1592AC148CB07507E38113557CD0612D25AAE
-6C21BA4CDE6FE4E79CBD23B0F5B3A95404E2573AA2E37216411D34045E7148956AD8FEEE1298
-F453A8FD25B06F18370D0645192ED11D39A1EEFF0721753F03F1E3FDC895E916BBA7A92DF9BA
-0D1F97CA6A96D269A49E419095071FFC70BFFED6A2F78C5661B779B8528A5B01FB3DA091C39D
-64FC74A03E0BBD5D6BAF6A18F893604B6BAEDDD789A85F69A42C9CA4CB26572BD469A9AE1B25
-BB6608AB4EA534548D67F89190CF5B11322405E319536EAE5564D55C165D9DDD35F714E2C9AA
-9335A2B83203C3A09731AACD472390E86FD66CAC07B10EBB01C1A8B3F038C856D348C7B377BA
-C544EC0A7982E581039DA0C166440D12A2CD88608AA4F763C00F1E67A5E95F75FBD851AABF4B
-752E116DEEBAE5161F9E1D5C50450C3D94BDA9CF4CEBABFF8CCA8A3104884A6649740CB9619C
-60DF3B7B3663105700CF06964687BF0D17B64DD61F882AE257C17A54262B737D402F43CF5E78
-F16B9CD0D01FFF8009AAE76CC8E3C3C66C123B22511052C04044B9D8DC82BB054D8B68022A52
-83D85DC707F760D65F69D1BD3B6CBF07A088DAC2FD84EF6C441A4B0702071AEA16CF4A04EBFB
-CAC264A0C8F03BD90708170BF40C429F5E8CE51AA206706C631C3A00EE09A8CDC0752662CCCB
-D09CE0C1FDFAD14ADC8FAE05D7B6558A3DCF7993052821ABFF433EA1C93A39031F1CFE78B6C4
-72D0D8536D8AC21E5FB370369B3434DDBD17152722ADE0F7DF62A67B523DDB38C9BC865F728C
-6C509E95D5DED87E52CD209256481F30B69C8BD609DDBB99ECD320E2D341999EC039A22660A8
-ED38FFD7E0F4771B66EF931B3E0B307DE14EEA4AB841F353DE1F334E81F6C0CF0C7FD27741DF
-68BD164F1F64A61395769F2BC5BAE99304F2548DAF234F79CE8A02A4144A56BFF23780A67DD6
-924FD03AD6C28EF268CDA7D5B6F4623530968EE9E345812D7E47C5FE01030A1FD1691CB74878
-6831DCC4E1BD8A595105F09BDB1E3B2D13B5DF72ADFDB139B3AA49B836D950AE4B02F51F4E23
-DA240F4EDC12A710D743869741A0100E6B69FAAE750EF1B96D2F92179B99D095D81BC3AF770C
-C19203C0C7E9D6DF513E8300EDD8938A8C63DE3E547800674AE2891E12B6317A81CC2FA8522B
-FEB3BE9EF1001371AB147E7B65337E4C9798E9374AADDC82FED512D45318431951849A7AA52C
-02CF7F03EFD521218BE2D47F111390B7A31D135B8112CF6063D0121DD97EDAD1A7F9E98011F2
-6669C6F3E679858E6DFDD81CAAA24F0F45534C9EA134FC92254CE387F2014E160CA144638B62
-D2FB9DD62EF3F68C7DE8B4E43FE18A3C3F81D812C7735CFADBCE9CFE404B86EBC3FF2DB257EC
-25302AC7C9CCC987D59BD91C8A9DE0B243ED089F106400B2470FC96CABE6F90ABF48A1C06AF6
-B1C804E19E34B0329F623B8D5ABEE7A7BD85523448B1E5FA8BC410F54D3C178AFE83BDF9FB4A
-6A539879563CD437D30A84137E696510E84556EE344CA93D1F9148857F86D94B196C22F76138
-9A6B0B90C4A760EF44FE2ED9FBCEE7D701216CFFB1BB8AA530884CB22BEFC9D4543349E5EF21
-4217597F32103006CA78FCF3E9227EE98FBF864DDB73D148BF3B3DFC5289EB96B12D7FD1B73E
-FDC752C73A164DF33DEA3D3E980D69746256E849EE45CC6C333321199588A948CDFEA8CB7534
-05A50C74D3F4C8437C84A6DEACEA4BC08884EA2F9F8702AB01E9ACD86C98EA9024403454B080
-F8AA2071A9DEBE778FDE63B707D6B86F941553EB64CD0187639E2BF409F957150547E6C6E6B6
-7CDB829BE98043FD1FFA308D28311C326F8C8D4BEE30460936B6C57105B2759EC57B3C4E04F4
-A4BAC17C2E474B34568DE1AD40E338C13885E7D79A4E7F8917246902E372775FFADBD8DCCD7B
-DB18CF273CFB1E313953F8A5CB55573D0C4B98E35D32303F0E363530CA86CA552CA6B47C3482
-70C2ECC8FEED16F78AE5AF5841A8B8D46AB5B4514AD466356F342EC4D6C5811F86381B5C2F8F
-F2552BD6FD0E4F33EF6C87D8D7038039F4A16829F3FB1896B19B6D5AE08389A7008CA54C4375
-3F1CC5AAD548278734C80364E0D84A9B0E2F59545422CA0FBE0D55B0AA4B7A4BEA2BD0810384
-7CD86D4E9C3A142344740F27E085A0511C50BE76A2B53A77C48B26E33D30F90F24C3117DA2C5
-87BDBF42619BAE4649E923680046EC3FF833636D8649BE335C8937543F2AAD6503ADFA3FDAC6
-A1BE62C47EA9DF11B10B05F456CB815ACFB951F18C4379199B5841C7930ECB62EA953EF1DA3A
-7E9E7084E8174B7F84C174EF0EAD958882D0F412F2169017179E74A093CBF8ED69192B2B55EF
-9CCA48D6E53981C89A18E49E7E9C77B981B4E717BC7577F5B13FBD27CDA05AAD50C9B97843B8
-A7369F71356437F3E6A2088981C733BA0C3313E64E4E767C8A8A894C42013682CA4316E19D20
-9022D2C32628AC0814F146594E33C7172BE9D340DFBDC432B12D880B3C9602BB7B44389CF632
-4C70B33FEFB1ED67F2F887D657AE72D3425CA29B480C57A68AA41747EE4C2A940E4D1A9E6037
-FBB7FB6CC142AE5F5C67BBF58D4CE1D0D45E6C841917B5A1A5F223707CD620787F50A8210973
-EF4585DC0970797214307A2592331E1F060ECCD6945D027C1E8CD241DAD955B508B0EE93BC2E
-16A0AB299AA8DE8B586C13D56E68D1FF578C5E6D0CA636B37F11C6964F6DE114A7F54022B815
-0096C1046086D951FBA237711E2ED96EFAF09E36C03AD289242CA7593ADF982D666E55F9BD08
-D8A4ECEB1D0701E237FEE22A81FE68CE1960CC3FE1494EE3AB665A5309D3727A120BFB4319CF
-AF9F5F271D654EEFA4CEF84253EC342B6720292E27F89B2485D792861920830668596B56CE07
-DB3A0056DB8641921F7B23E6BC23B8D285D836CD5D0878303F8DC7F09075B09F6A13762AF912
-DD6DC2CCA9570EB683D1A16EB9107BFDF7B0A3687D2B11E0F15635C6DDB2C4FCB39EA68F977D
-61DD1842DFE9E3C2E45AFC04D32F6CEE5B882F878367A298D0C8117F7EBBB23B6F3789F34B8B
-6A19D07C81B3D253B6A5CF3A08EC06C8687473AE28B300AC914CFD3B4BE2A8883CE504E4C853
-C1B1B2DDADC2D6F2BFF34D1A239A8A88161D40E74C82A2A385DDF5BFDA51FB62A46DCB3E71DF
-EED1013F98EADC850B720E1B78C8991D926022CA8CAAFCBF697B1A1904546EF4909516463E7D
-3508EB851DF3E5C87F73F93561A5BA073C1290B2DE0C184B0B89439F0BDBE6F56DA7B17C95C0
-868AD087030D7E8D3106D8A3332FF4734090CF6F19B1C811BCC8E6DC7DADA877FEE422F65E74
-08353FC4E32561C1B35EF14C3DA10EA7DC8F5D60465FB1BC0453C1F60C17CC81261508A9E173
-DF7A1801F4FA2AE53D78F95E3CD48C194491022AB8DEE328D997EE1228763E2FFC54AB82380F
-2916FE2485B50C7068D0E015C52C0A2D50C0C3FAB63FAD955E1CF19E1550E78DC67E8C6C5D19
-005CF1A90700DF156F8BE1F9F1ACD61AFB7C56E5B7656F2A978A3EF8A9381027FE04E0011C8C
-9BE10DE276BBE7DF4E507934F752AF1DE58362D2D594284EA89C71210A5CCCAABD0F821532BF
-DB83833F09203691DB6222424A18B9FE06FED7E8A926CBC72E500A59FC7D5BC5A82458D50008
-F19932D15A9129E21A044D1D4E66609565F63BE4E8931C9E216616B11CA2DA444FBE5C08E72F
-C5F428E1DEE3F19AE0CE4EDEF9CFD6E94E6143B7724634FFD0AA25B1208E83A3FF962CD07CD9
-E34B3EFDC0146DACE9041E3C2B193686583597A4FA77BDF5E9AC27C87E4E1BC6AF0A9F0B96DA
-DBAA4FBA247D57818AF367789A394C7BC45C01AE146FEBB9F08A4572B915EBF46555D8FEA8BF
-A7088C1C1F0152A0811FB745080589B853DF5F29CD3419FCE4D1E649BA02322F23B5E037DFE7
-001D928206E0BFA0981B57F56BB600A47D33A97863A94A3E229D45DC33B1C19FFAAF6E15E816
-1D3F0A4E01167C364E987CCD0F6DCE6B45B2DD4A92A4D7EF1115B540D2D42DE22DF915D014A5
-C0DECAC74251C29F65C63CFC6683FA73B95B0D57068F6BBCB8F293600A682B3C5EFFC8965712
-BC984D495835B89D115C71B9468AD81D10C6F285EAAD6BDDB59E492E1D6D830D45F56CA6ACC3
-0D7E0D3D8BAD8CCDB068121DE1CE940BFD3836B17AC2E67A858DC7863B7CC1C7ABB2F5AE0357
-6D9EA73E840D06116097B1F66D7D66F1755153300C229A293D4EB394465C38839D634F62F255
-B0486D54847EDE769E26C78A3D6C2226ADF652D10C87CFB4703148B571F6F389C0E327C805D0
-24E8E20B0780A9EC5FB84385D2365E7F5EE7477F4352B54D23B8965E0210C18897B59F478C46
-57EE0DA73E1DBC1FD6DEB59B9F906C20886FAB9BA798E00A0DBFAF8899A894867D2B96F28472
-0BF6B898FBE2CE16B80941B1D227866713C2D20AEF7EE9D90AF520DE9F58221C78234EDA0588
-ADAD10986461CFD6531FB8BD8DA50BB38873C8BD788CAF4A8D8AD1AB0F103F0617A1B03FB6EE
-C9BDDBC770B23892876DA8409E9DE3B0E5083B36AF276E276AEA88785C4FFA018E82297A78C1
-4BDAD33C211AC9613C2C57B07F2C5A9D578833EB75620739FA3B1CAC2762089FEC2E8EACDC86
-6994FB0D7A2EAC25A7BC451E414C4D1C34B30DA616F46E57D439052C70CCFD44D30BE545CBC5
-3E129C634EAAB155AE80ECFAB6F6DED51E18B393D6D4190CE894710B1E5D0C6AB3CDE6AEE21C
-638F23E4AF4D0745C9AC56D6BF2D95640998A261C5FA9060FBE789FF4CFF17AF248ACA03A63D
-817D5EB4DD0091AAD344DFB4B6ECC465DAAE7203955ACB02DA28DD5F7E323E30DBECB7BC6277
-C60F17DFE5FBE4146150CE2242CD77F78B6332ADDA9BCE1DF4FF46EAB6C6E661251A83A7DBED
-B35234E015E73BCD55A0A59A1D5211FC1D683A2D75DAF8845E95F9D61ADAD9F97805DC8F7E06
-65905E3C6306577D54C4963E9B26DAE47A0166E6875B9672D5C9220DCA4765B3DD1A39EAD347
-42320ABFA2B9D9A0F0CBDFDA624BD119D23AB90F3F3B8817EEF783B7B5BA8B28190DF2F69BC6
-24C13D019E96D6E9086E203E48E2A04BA7D60B997500EDE14124839732D4FE99F19B9219BAC0
-9B78FE25D7CAB67BBEC16F9CD5CAEB66EB185F51BD1E7C392C3EB3A3ABC2E5F59F407DA6BFAF
-F916C2C2FF9A2ADE0AFEF0BE6D29E35BE6DE297ECF2FFA59C4B1D20F33B32E78A2BE495B63E9
-9F21033E4221915835DB9A8326D4AE6AA571B6B816513DD6E101A768CD5A5432233802D19188
-73BD6C0072FE20C695AB45E3DFB8C663B115964500FE5E22D59A05336140D1D2BDBC322B72B5
-1F28FEC02D269DD36312897EB0EA745D7B0013535A2155B8F533F67FA0B09260E03E31924D77
-645A75A89DDA5F49CC217FD79CAF03209F54B1E8211AA82F8650C1945283ED371FB026FC11AF
-CDADB84E123693E7632545018938F05D0DEBF1A2EFEE10ADD353517BC45B13D3C8513E6F5683
-3A28475D48D74FAA2B75522ECF1DD95492A1639B719B01BBB42AAB990189F15D6134DF8B183F
-02384901E075465789915D34EBE0781DAB7A61137D5329DBA75EC6E96D35CEB28460E9E6A29A
-51F955887D56BBD3840ED13DDD73C3ED0C2ECE63FBF20D2CA3696AE844A61126FD861299BD06
-188790FB6FEAF3800722551810759D1B4B85DB80A0CB353CE1A1A4683A4D96AFF29B4F9038CF
-05F3FC06ED4A900F344F0749024C11C31C0C37546DD9CE3D040F63E3D6DAE02A8A8F6946749B
-281378707928BAA4BCB238F5655DC68977E844C072D189FA57880E337273CE8D6613743898AB
-E3E0BAC266D271497287BDF34C08FE7AEB0A3F0A3FAF354F33F5C8845F5CFF43BC1EFC7DDAEA
-21E52AABB01C46E7E503DB67F68F6DEC1EBA3CF285C68364924BF0DA1BB384B77A07E9F12B74
-A21AC30D70F9A25FC3E324918297B22FD4D68F2131B7E59668864B8CD49F9C715DBE79DF308C
-FE2ECCF613DB62A2515C57110C37A2C281A7CA97FC08EC89E928D38127E9179647118F547E2C
-9667ADAB3A54E41B65FCA78111C7BFBED640E106505FF6BB268040EEEE24370707F21A66062C
-A9E0564587C8348E0D0F5583F19A40CA8DF4ECAC662012D252A261A211D2F6212702887923B9
-05DE5B53C7FB9038976650D6AD009B4E95DDCD986F85A3E9D6BC5296763CD466E59A5A7D3757
-24F8D1D0432AD60FEB4BED07742196DFB37B5C77B77C32B66DE91A4BD49FB61453B2F61CB92B
-EEB91FE4C12A7AF5AD33E97FF2A3E3D1255AB956B37BC47665970982344C3AD159D593BA5D0E
-C71B1BC107B9E323D906BFDF5E3BD90DD30D7047B8C44B044DD1ABCB87926DE2678501CF9294
-823B44186FA08E66D653B540A53427BC28E1F9EB2C223086A8E4BF9B99F6C06004168B78817E
-AA91229567FA3862AA5BE6148E4F9314316D207A7AA72AC7B41B7B53C0BC98A067F974308B1E
-B6D5F2908459FA57BC3E8E2E4FCDCD01E40BE70B23CEE6F2FCA644179AC3A9479E12CBBE187D
-56C5408AA95EB7E1C03BC49FFB2C68744E64040C1B9B5B3FA1A28FFE5D837E3553E7F8D7A6C5
-A7FDAF4D5C9B3C001DA736F5B65992F2E6BD37BD0F48F4CB0DB4D3F340961AC3FE596FEC65D9
-B9033416A8FB76C4D4DDB10324101D7EDFDAD4C3F899443BD4EF1E5FB3D9D6ADCC60EEF8E845
-2A2131280385C9BEFFE45A3C98F3B150436FF490DE1905C9F8C5A4A3C5F0694ABAEF2B0349D9
-E39C88A967FE207C3E50ED01D1378CD714DB7294743FB8BD30330ABC98D8BDB62D95F5295AEC
-2E92B192091CDD27A2825FCDAAAD53592A66A20A87419F3260ABA13A25AE6E948BCBE3E8A70F
-0EA9C92BE982DB4C9CA3EFF9AE1C7414E9758F47E9B3A1E956C35DD5FA88413A549EC565E339
-5A657BC69077808C6DCF4972A9A169B365670B3971698FF688C5D16C0E8FC33A17545A63241F
-756CEEE915D0411B74EAAF05D11B30585D565767703F2EB7F64ADB524432EC7B93B13F93FEE7
-C1CE8B07409EDAC54EC922FC4C8A5267E099FE74B6E4ABB95DB50BF96D21D594D6B8195AA2A8
-BCD32238725A2E958F036150188201C0D3CB444821EFD60E23A3CC6A7DA6F7DAD12BAB0DAE8F
-FB3A7A8FD8F2F0B22F967F9FBE68DFEC12F65878D3D3B40F288681A50BC99EDB869EC0C239E9
-EE251D55B7BED08BE2BE3A1328F51B62FE7B262AFE6BA4C8F4B9078FE9A165CCD64C333131A3
-3DE9664142AF556DFB39DCC69675F107A6461E57BEE4507F97117B79E7593DDF72897B3BB16F
-F00228FEF4A4F7EE2A8473501EC5A4ABDD5EB93652C03443CB7C96335EDAADE0CEADFACF3A65
-FA5611DCF00ED78AF0052B5373382336F34D41A3D8BEB767E59C47D83C69F567A955F62816F9
-08C1E9A45FE2FF2A356A38D576DAC757A5A7804F1EF40F27BFF2BF840E53620B59F611649923
-0E49719EBBC57C4BE304DA6332805328E5069A2015818E49B53935AC51C00235DA3D5B6326F8
-15964A52921EE4176871B8EE545E1D53A4CB952F624626D8F18850B80E6E9F3592AE5A768BB1
-DFF634B31488AD02854CD8EEE10D9F3F599954E7CD4FDDD4FEBE130E287021A2DF503821B9C7
-68368B26ACB7EB2B0D88F150D8749557E4BE40972743F3F1900D0DEACFDA766C0200A51BD15A
-9DC3065E935BC1C7A18AC10FB3C418E56449B4C9F810D469BB87B424D54BEC3ADE985CA8B8B5
-8295ED54360E90F2030755F735CCDE48E24A9E454F856B40B5698D5B5493A28F6F6926B3B237
-EF3CF15D8599F54F8D4394A831438C8F6D8C347A74F7F219248F5497E7229BEEEA667D6BDD2C
-4E94138676BC5DA90B2BDD8375F735369FEB5B7DD6C61E737747E09DE5174AC4CE0122463C84
-80D36DCADCBB1B35DCD1A94EBBB22F7F80B0E30EB27F3D959CE813A491217ADD5BF3905C8361
-180B23D044CF303FA87F68C14C9325C021F7172F83883D7DC50295B56DD40DFC1536F995CD9C
-BC98A39F446D9008619A091CD1BC64E9DDAF782A9C267C2B917931BD8A84A79F24575F73224D
-0E4754E11E8769303BE9E26538446C6AFFF619635B52B5E276E2C6262FBC25DA02D4D43B9298
-7693ECA561C5A845D3865BAEE7FC4A664E9913819445F8B13F1F2B1A542A94CB06115D1E7FF7
-638E10331F495B4D6971BDD9BF24D2B28EF40C8C82519065B4116AF8730E2B84CB6E5D8816F9
-25544F2FC576505ED3DF6C45A14BBA68F9F54B111E545072FFC9D13BE8D9E6F96B4ECC10AA1D
-641D13AFFB05A628E90F08B1ACE76B2CBB601E43A9B8264B6E62AFFB32F4812D5F2538ECEC90
-68CE46B51A71AF50ADF2EAA81BBF662C0AED7A19586C165ECF9D7D0A2AB959BA6DC8CE4E4A6F
-164AC9954B9EC7B5910C77FB93A008D9621704A516C0661330589325C30C7F0402319A66AD5F
-53342603D97F21FCA7E4887B7438DA6377005F6CA4D87E720E21001681967118DD5E7F459803
-D3100E3A283B48D860F15B37BBAF50FC9903132D2C6A011EF1AE16B95C8FA9F819857F972645
-E7CC11AF62DA297B2A6E231143BDC26873BF1A6BF23EF0C47C1ABEC894B7FBBB71A7664A7ED6
-A59B16C2888098BE97F574E9D487CCC83E5850761905A91E2FF6F66C0601D5DF0115B8FC095E
-C66D862DE36492E21D91449CFD93BC1084DFE0012D2A1E5843589849D1A46B630B3E0A122CD9
-DDC9239D473B86FCD2B0232D3C71A5D5A3C7587888A037EC9BCF09B4EE7D38FE3C0837E578FC
-8A473428457E1535BA3822EB4C7B1A5A8FD47B453EEE0D620B888AF7C48F5021DFC2D201CEE2
-F7DB8A50BA78D5D56EF597CFB3FA5D98671E79D5E070235B9CDEA6FACB0D15AFB96785E4AAA6
-62CEEFEFB3B16CCFF17B54294C40C69359B644C44901205F9B7C3EA316980F82CE0F1019F45C
-6E3D00439FC1B7B483E9B95BFD995B6CF59F05AD03D96EE331811144C386909C858F93FECE02
-7891C42288CB9DDA74A7EE15F415453C3355671B80C117FF5E19D82431C6292F27510226E829
-0CE956CF446361FFF1899E2A0C0A2028DB91BCC0241CAE66336727BA1415C514A818D1EEC8E6
-874EACE31AB3A8A9EE8307B5E1BF53ACCDA0324C9A10332F683A8B845183F93610586CCD7EAF
-F88CF4B2A0386DB108F2A330D25F4C4F40E85908E67756E9D2AF2AD6BEF46D862FAB8231AD07
-E30948AF6B853C2AD03BF066FA79DFBA62821F3402C32F927A25F0840ACF2A82DB3FC394ECF7
-E73B677208775D3582B7257F9B7240D098EA700558D4AFDC87FA610C52BFB0D710A91E14B303
-B49F0DF8635B11B1BA1E81A60161DF7C3068994E38BBA09C613BF83C956C18D76B443B407AFE
-41911DC1EE2122C08B0D4D0E395201E24A2C2CA9D8BCE2AA76C71F4ECAD84DC09AA17FBDADE2
-6744F64BDDD6165655539CE0657A7657CA7C3BA33F2BD9724A9AACC0C04DC987741F39120D75
-1B8B5F8E141A224715DA45235B1BBB2A8575E546B33D410A4689359BA4465D654ACEF599DF63
-7DFCFCE1D5639C9C4EFC24614301CFA172BE0EF1774F31BD2DDD0FFB680100507A0FBC9FA8C2
-7BA4AB15C762599F76E2B3A2FA7037870E88B145C5B508A08F17D01B7B55F294DB709C25B53F
-0E77A0746162D3603FE10F6BD1580277B6845D98D08A172F47EC51A36E4316331844CB9E1BB5
-DCD8B9D1F119599AB8182289E01A6470F2504C2E3278D36A5AA5D40D5CE64536316D39AD5EC4
-927644B9CA05AD31946C184B1A13C586412E115C3C970D13D5A4EC5DE0E23D10E5EAD8E78A6E
-84C3BE1B3CC75D26F6FD4C84978F7A15B325535EA1DC674CADE67E758B489CDC7DED5BAF2F08
-08A756A342F27D959FFCB92F29F034C7389650B422A69B525A3FB1435A7348C36066677ADD54
-1A0D3D0A3130F380372D9931B689E562194979414A5C18B22F839B695C103348ED9229413183
-F5DE7EB61FA793C144FD6D7A4735AB091A33DA0BAFF9802F2051F091C91103C06FEFE063097D
-DFE22F2D922491AAD5C3D47DA4CEE911A8B1F76776699AB7647D05BC4A202610E35D738BDF05
-69379038A22897DDCFAC647C2565457F7464A6E51D29B7527DAAEC3774565A9859CD0BA44EB1
-1F89A8767409596B9CCF3033D7B2CCA2197E61C28FCC700D93402E70AE490D2480C9233E94BF
-AB55F3B01A41C551008BAEE55CB82520808EFF11FA9F6E2379A17A533AEE2935A13909C834F9
-3D7A93862AC7C48EC151129750FB6F8101E8E4B216CD31BFA3D9C828C3757010D3760633EEB3
-F4CF6C5AD8C21970420436F0BFEE037AB3EF43145265D233B3BAECADD4C47A0636A01C40144F
-168AA598E3FA0F2E05B3DB4E234B0298F34E7DF1FECE36461F2FD6D8CC8DE4FBA989F07344B4
-2AE5839D4998459B65B0CCBB866D1E9F9263DF2EC8B41A4B45B110B390BDF67517B7420D995F
-30918CABCAE15D325B4C372E05776327AB08A0D77A4104EC629AC76842DB7963226FDB6DCFC9
-50E329C4F94B01E6BDCA33CA94A114C9037EEC9F933D46559499B7A14E46B6903683519731B7
-9E6C40DE54A2AD88664CF5A14DED3DE8BAED61AC3AA0303DD87F9959C0A989561C37EB1FDF76
-BEFEE80192B8E955969BDEA625D66D516F6725725799180719FB51D0C0C61800EC11FCB25DD7
-FD167A411B4979DEC317590927DBE751854F51D3FC155F3BB31EFE0BE81D51B72DD525B14B23
-C5682394BA9E65C5F284E51A26216A78E7A73D54AD3ADEFDED2368650C411532972C5CC8070E
-161386057ED0CF1EC204186E5779B74AA82543892C84D4FB8E4F28A94CEF830F726900BE1125
-53680C3E4B0641F124DA412BE45F6BFD08241EF4E7AF33F4B401FBFF395F2AE4A30F5E617ED4
-80D7331ECE4C4F017C5F11ACBAA2B5659F83BB1F70A9287123A5864D08296D25D25C6F0A6241
-DCA50AD824AFF5DC27A5D43AD8D13F877D768116CE58CFDDF19667BC6A7D43E749FBF61D31A7
-76E602039EE1BF04E386E5C25E00997C5C230598E38201C87D46841BF25FEC5A4767D43A494B
-C5C74567D98F9AA4EBB74783725E302E98C2571FF17E85B324B7CE26BDDBCB81CCEDB836063D
-E5DFB105F28D39D9530C5289CE330480771EED2D389637A1783BD5B25864C7CECC0B717EF79E
-FC8387D81A219781D25CED638CE4E66BEFA42B7426955B808B35C4AC873AE9645864FA75C3A0
-56D1CCABCF30941BFC831D5D8930D10C4D49AAEE0AC1ED8208AA10AAC54CAC0314A520B0AAB7
-4D488F39C382F316E6A1ABAFDC20A542D6CAB2B87E3D1AA0403F8A23BBD230373812116FB341
-3BBDBEE188ED5BBBBBC4EEE0C7E1850EC100D384C1AE3D801C589F98BAF5F82C8F7A37DDA543
-669AC0DE51D1DE0D67689E5BD80BB9F5CA640472E14EA223B0110647EEBEC13DF0B572418B63
-C6D6D8613625C7A37AE9E67B57F9E2347113CC234AB436148F45A952624DAFB993454A70B4D5
-CB2D719B2E17C6408DD82A8A570EA34333F4EC602B6382FC13CE60D3F495B3E2F12A425A4FFB
-62B518A8F12FF8E63757696E60639178B02B7AF8BE5C25287EF37B9C8F396308B13575A341AB
-87CC3851EBD094B688244D5706BC504005A29C8ABE1512FFEC08D8A7642659B1A6F0EEC82141
-A32029FF608EB4714D0D34C5747C89D58A5FACD98687BF5DF3E27E5E51F81C1FCF4EE23ED40C
-C0BDC75721802262D6FDC798FF6CEBE4DE75A7ECDBF0F15DE6B6C2153C5424892AA6F32E8968
-CD3F55118EAEEA574E6C4C18AF92DCD1EFA066AD3C0A7BBB42FA60B62E910399F4BAE30FC939
-0C1BB4AE90D098BB6E0D0A106CEAFA3E100633AE84DB1AA16797C7D49305642C66D5422DB076
-5BD992DE8EA728F53C73F186B37F6F036F22FBD06EA8E1764FFBDC42A07D96D430428437A978
-B4C3BD12BB6D76547DC3B87108A258A6B4BFDE93C7B22F27A6BD3EDF8555BB0926DA2FDAA8DC
-BE37A83D55A8D592B0A41630476E1327EF14571BCFB6D5B92F71B272A50D9DE533EAA88BD8A2
-4B1EBF68118457D5A14225EFD9B8E234A99EEF668089357333A22F4631680BE9E2B1638D7F0B
-1D1B8C84E39A1892DC40C16839AE8CD7F86E5EC8D7559158EFF38430B86B65887417D1F03C68
-BF82E6C15AC0DB8B58EE1600462B867DFEFB20878DCA0D1D38FF679337DFD7BBBE4D737EE240
-7E353DB4BAD5EAF81DEFA8CFAA9C299345449F8C42AB5498DD8A606AED7ACB86E0FE64AC861F
-18945A956BD3DC71DDF0A3C4C4BAA217117DE0F01619CAD2118F269B2BA1A77CB61F17E99166
-70E51596D5CE56FCD74322BC15C17F502DAE1F2FF719F082D0A698F72B8C7BE6FBF56BE1E025
-3D97AD3F122021ACBCFCB1504C856FD2406CC8F0AFF60599DB1E4CF989F5B28D7B5AD5F22D3B
-2B6458D460AC2D5260B7CE7C0DAE68942CF120610773CB03C219A44ED9F569B2076FBD4553B3
-1418613988A4EE39E91C0CF95460C4D96B443299C515BD89BE72CD9568D158CEFC75AEACF95F
-0A0B125C5124A020EDE1A42EC8E737EA7CE9D9FF5B6F457966D409654A9963620D4C195D7615
-C5507E435666C6BCE835973A830772223FDC012A1453F515AB733F81D7A442810AB2E3A8C144
-EE93DF45DDC05E1B6E38967AB2203B1C0604C38A18DB2FE937D1226DC9C6B198729CB3C02047
-AD65A19350DA0E019C6DFBF866E94E6F409D4B76F5D2AFF15583AA0FEF0203D015FB8878B5EC
-B35D5F34A42267C6DA0105B31C825F0B7DC26A1A3122F63C6676018A4AB05635306E51CC1270
-3DC200DC3699A5AB05B744DFD6DDECF457EECC2F9DFC1576F4C5F758EE3676C5A4EABB321077
-8DFA2B9366A7D7F5EF0CDC0A1E1DBA51645F8F72E8080496CE8517EC7BA1A77CF22EB451EA32
-D9228DE6E1F19A55405860E2D37891E85C0908F2DAEA89B9517419D4A76DBFAF8984F9EF90A9
-126AEAD55E7136B549B25D5C829D1C17950692577DA8B9CB5208AC4D137263A6713878A94BC2
-3D03469A5BB6A42822CB0BF4C284B077FFDCFFD91DB65407A155C94767E8A0085AF0952CDC06
-5E04050AF7316105566A430E3C2F8C8A97C4F687E90CEBF8F72156D6F86455A5FB05C4FBC41A
-2F51B1C3B1B2158AA27EB96EF726FBCF6C6BDBC40074880DCEC3228C17B0B657584E6BB48A10
-2602D4150FA60F591B9F4DE04DED1D32556D4E5708E97368E86120E9A3F87582E9903BBDA407
-085964BF184A6C11EAE4C96AD39F836636EDBBD0F2E37E988AD6C30B03618E402EC752C03F86
-29C25D7280932A3187BFA478E8577299015531BAFB7FBDBE5DFF28198BAE528E0ED6B9DD7430
-F9B16B0718273C3B0BF29E19405A8536E75A08A4E9C247E34FC9EBADF88A680A3BE547B87374
-5EBC170BA982CBAE1D3963133203249B5FFD9AFCE6EF95162EAD87D6344F12A8BB16F5B14B9E
-152EBF30A1D8E22B3B75073F2931AAA5A04B3B808AAE81F9FD3EDA931C45A716DA2D1B298924
-CF6EC24C3E944F675E53A929CFD84C9917F9471844A97CA6F242A736DB9B3A401B8B726E4AC0
-C48359D1A85EBA65F0463C34524C17B596C26FA7BC8DB77A8627E417D8E24FBEE8007BE37B6C
-5B39D08D3CB4242EEB78A3ADB969327D48EDA63F01D211D1165F4078EA2B067293A9FA26B7A0
-0FE0426E702C6B080901719F7E3C2CB9AF82E9A5B3F7ADEFBEC8AF79EAA9CEBF07A1A8873EDE
-5E64333FA5B6672815D887881B705BD0837C84A7AFFD3CE49BC330C0A6FDB04EAAEF4E360020
-B65033FB1115C7E6CFE781EAC31F15E03614E5D03AB29F7DAC49F6048341272DDA6D8378562C
-0507DF81DF62D9F88B80AA0FC90FA1A5A5C23265F8E3635B7E66F336D30D1F977EDB218EE691
-86F404370A3FF65BCFA8F62CA5BEDB42FB34ED525ABE07514CF86541F182F826D91D9977AC84
-9AFAF9924F888C5D77AF64FCA191F831DF5582AC7CAEF95BB346E3BA965C9A0C66CD494FD3DA
-7015977BB7D0B882C5F080D2D62AA0DC5D0EC8B961971F1BCB6BD97184DCE6DCF91D91A54EAD
-5F7FDF7AFDE33894B084467A81DA5F415867A9DEB99A5828DD7DD2BA9C2E6F9D2607B5AB4066
-12B233CE306C32F658D9C724E7C6F31500A0ED8BEA38EDD3ADD82F3D6F5DCF92921F99D8DDB9
-C51B116070AFEC7FF60E7FD6128DF63767D31F31705DA4D0BD7856374B3C9A2AF2F128FC9F18
-514EA65F196EA0B949EE95D580A4199F1CC5C61590878714749F58435F490E50854F5794AB70
-30FA597760FD6595ED958139B733613CA6CED006B9A44BCE4328DCA478DF66880778241F8A57
-70237963FC975898144684773875CCC1340342228014B220D90523FAB8DB26DAFD9B8F64D040
-BBECAE820C83AF312757A2E59A1B1C15D58C659FB9A0405BC72CBECC9C582E196509DCEAF3A8
-7882A5C9F880A884D4FE6CE66D91917045E9EDD0E0EDE475D951C3070359EC90E5C9A60FEF17
-D3E3100BF00E86F25D6988C134C37B85B8615BC3671EAF698216949A68F830415FF4FDB07285
-A23190714ADA60C6814FFA9908D771C2F3D5F0E86FAD4139C1DE4EDD71D8ECD2B06A6DA31D9A
-8E1C8DE00A6AC9D9A434655E468FC9562874434A8E811E9FE54401E169AA61808B1D930E9549
-02C95FA423B8B087F0D8832A480592C5A46BA0ED7961791B1EA4938FC2334AC1B34734DD4803
-C1A56CD855512B6DFFCB6FF446C4839825CEC5D6F57EF330807841982FBB3B25703F1001A05D
-60FE789FAD9FC1E171C27373B8A1106437066475309A1519CFFF0AD34FA52A0A82AA7087A9F5
-DA4833B872FB8F61077B8D66F50A565C2C3A7F292CCDDF3A6717A43FBDCB92C4EE796901A9BD
-8753EC9B84465DA0E9BB837578AE9D47F6B73B60EB60357F77ED3BB9C75FB4A2D2D350494E90
-66741F7697E22F465616F01965A985219FC92CAEED20BB90AFD83DCC07FD04F6E989A0FDC0D1
-88978134E1B637AFE34E91B80E81BF3C05894AAC716F7BD514794C412A2C682B54E00557E6F4
-BFA737F4DC41A75874F59A8359B0C803BA72DD896E8284F37DFFF3538764CBC2107CC1B7F064
-63517666EFD66A090260BBC9DC71A0623FDCB4067B611CCF8BEFFC05F31EC52B5F5332CEA64A
-ED69A4C795E17A8CC53F071A89222E6595E26ECEEAEF557E4E73E54AD4F43B4C3AAA5DB67560
-313289836DC7257BC84A915FED0A17405B80EB3ED34F7EE8F77FA2F8AAB05747CACECC0C3DEF
-8C3979D1703AAF55D828D6AE2ED8F509269023D86A59F57F5998A51B8ABA96E7CF29A9096E07
-1D074EC6D07DBA5F6BFAB8B87117B16BC0B705457FA97272B70303DB02346EAFFD6A2B633EA6
-77B160E00E58CB12BDC02114EAD6638E409DCAB1FEF4C3A277556BCA03BEA50E8F5BAE8B30EF
-C2F4E6744F8D71FD59A7D5C40D5A5FBB93BC58CAD1CDC0CF57E3FA86F7536911170259763A11
-E245B14C95E5A0C06483A7925996CA105FD1E93442C07BD3212C13EBA4C71B818B2D96CFC690
-274A890202864959AA0AA9239049B18F0EA83C8F439DC5736ECCB54B4F1BA0E05C93AD4657BD
-56F554561F308E2036718CFF4F15C9D6883589B3D0A6BAB9E2271A5B2A8331BE2502D9724734
-2A499EB08E38CD54D9478E1A896CD91958F7A515194E01F5D63304B1A5A85AC17EB25911DD25
-372966278828C463B25CB6BDF5DE6788F256AD4BC3282FFCB584F45367CC370C04ABEA8E289C
-42A514BC50D551B46890D790F9386A30B9D4A47D41398787815FCA325D895F729D4EBD11ADEB
-EB8E7C52BDE67066721FC14D4D724EB766627923B2E6DC3B0C1716A1AFA2D70B6EA534269BB6
-3D1C6E7444BA4A7DD1CB5666223EA863B258B3C5729973CFE95495D228F5014554D2363869D8
-6EAFA9B5C12E38D3A0FAFDFA8D532FF291507A305B16AC9CDE560573E3AC42D5779CBBCECDFB
-33111249084A1A87F87EB6ADD8CB718FF185DFCCA6FE57678B9D00BB308986E55A3ED468EE9D
-2EF7B607FF43AE0A78A841C72BC4559E46E6AB26520CE469D8F9C6F341AC162C18D4616CDA92
-534A5FAC40FAA9E504971A42F30918FB19570E189C31AA548EEE4BD78B2B095BCE72A63BDC22
-E3553D1C687EE2A4B5071FBFF7CE1A6209372D8EC0471F77433AD0C228301C09A85FE4156070
-7016526486FD6C2A5E2D06CF1F65A54616087AABD0B48C8E5B5BE4314F467C3FC73A5DE6D02F
-FEEC369A73C80E253EB5D71445780FBF8F3A386090F7CB405E92D1EE91E1B3CF95EB9E960884
-56CE40C5F6CB2B54B57BA10995DEEB1744A79A94EDC7D446125718E3527C225A869AA5DF2B73
-3F66C4D4AEBB7D12950F19B67F3523CD8A7944F94FB1EABAC3A00C9E8B0B0FD8381BE5BDAF87
-5E4C344312C4ACC28192EF914FF8ECD7F3B7BAB5F00C0DC889E3E66A643A72DCFFB408A43760
-BF2143F7DB1FD83F767AF8031A7457F3098E4FBFD6E1A916CD677DF4F4CAC1D4804030F1B92A
-5E347B8D9EC40C048E9C87C61CA42B74AB6037628C585C08120A0511CDDEF4C9581224B15220
-A67EEC584490FD037C605BCA48555BDEFC022F8A88404A674DD68EA03902E29A584F1F392827
-DDF1F6D62F1B0676463DF8A141C779A04F3D57279CB62851F6067A36A457402B160BA48F1BA0
-388B716E2FC0AA891C651AA36349C3C97F21B92B12350E2DFF420CB9E9B6CFDD2AD4EF2CB602
-117245C76D9357ACA27DCF783552F26C67CFEA38CBBCB48F7DDA8770E0F6CEA145095BFE2453
-416C6F371C66BE5E25DC6DA77282E600BD616792A2902081CBD392EFEE3B4CA16D937552F439
-E6A026E9C28FD06EE0ACB7C367D07AC9F4932703AA02EE3EA97E179736E73C093E752B27D706
-20E8B5587B605A0E473CBEF0CAD582AB053CB640669F80DBDA3F85AC8D5CB655A5069E110671
-F3243CF0664DE09B0A89A207F132B796562C315ABD21B36AED6AA9AAED4947BB6273DE4BF15B
-A36973BBED87A6E00C1BCA35684D9FEE78FCB2CC733C52B11749CC257B58EB0BDD5FE1C21C2E
-DE3B445EBB456BCF25BD834F6EBB9014B9B8680272323A7F5EB1156EEBC7B82AE735D57680BB
-BEF0F31E8D886C46DFB8DBFC061493263DB6C6A9996B180EA6F8D1436087304D99812103557F
-1E7ACB0BA70CD1669051F51B5D44FC1D58B2D2FAE10C81964455E82A018B90B6F85CCA71B1D1
-52C1FDE71801032DF28A243972B3BDB6FBFDABD98A9CBDB95E9C855163CEA53FF7A5D5E47D9E
-82AE02311D185828B2A27399C93E5C93B05CC1E6EA7432FF47DC18399D9A9D21FAA7EC9370AE
-88F41EE1E639AA26AC448FFE9C81F6DF823E67CC0A0B7FD36BD60FD0C8AE4500A224ACDC0F82
-324E3707BE41179533E17649839545A24A78086868BAE76B1434C013F3C0A2FCC861248DD806
-C08CD7F0B963CA49D908098EB2AE675D6842079A6A9908F3EB054D7374197AF469306B6CB33D
-AB82314EFF1AC09FD368E7306F240083C977B2CCA89ED5E41B47C2F4BAC33F37B8461A89EFFE
-89C82B5A71CBB986F9EEF472590EED198AAFA8FAA90885BE4A55FE546C78B6CABCC5372257FD
-9C9821F85BBD6E03918E36E44CA712F8D7EBD89C66A9BA0EA7BA6A45D4203E0E2BFA64C12416
-DE64764153C743224AEB73E6BB67737CE9E8A847DDBB794B66C4244EF2C62BB09D281D43EE9F
-18602FE93F1C7C7B90765B345F31A1474B57DC8A13CFA5BA7395D2E558C127F1AE9724B1FFC9
-8D9C80C8FA727C613D94A36B43249F032ABE651FF83E173FA36F78ED3E03C2C4F9745E37149B
-30C6BC2CF5FB4ECDF33548F3711CB2031C82C41506A5960BC28AD8BDB7DA01073C8277364C60
-1F1DA2B4476B0616555F9C208586817E4EF8987E8E504289BC8CF6AC33B6F18C5BBA16F581D5
-2D008DF6107EC616DBDEC7991CEAD17BDEB467F9C6512A2B0548917AECE5D5D32F74AD93B7B3
-D27ED685F9FED5A957F90DE6C7E94553BD1E61C6A3BA36B2ADFCBA9ABE991441B51AFD9E986C
-A845F4A58E4BAE93EF57929A93AB4F61B11BDB0A079FE5F3984FF53E1718C9FE0A3F2A0690D3
-4FC79B89D1B4DBF2DF6000D5432A191CAB9C57A0E6F051162AB976ACA2023D6E01E1D4D80EA6
-DF6B6AC91535F8867502BA4423648C2FA0EA61FC42F76252345AF3940844E1F582EB9B20D845
-998A2D6C08C0E4D2A8F414B1AADB1F42B3614CD1696A621E134C0A2A70CC699F1A04611D7EBC
-4D67F989FDD35EA05B70F26CB1EABF536D5AA6B7C88934D2E0BBDD53DD953A14FB2B9E5D6085
-24FD736B6D79E96DF083EF23E0E3EF123EF1296CE0B61076FBD36E226988AA31988C799A3CC5
-71676C60A131416F01B0AB89BD8986681256F49D18B65F81FFE00C2AE1043E2183BFA14C3898
-B21F4712983FC6A3D4165250F11D2885362AB2152774C85BD95A056E380FC76ECC6A59E88888
-B8442645C5587DB510F0F0B444F28112638ECEADDA7A58C552061888245D6DFE75B8E2ACF8DD
-16D9B823B4E2348CC52A6AD4B3B49A0E789F180F2B3331E2B063FA37EC766032A05C6898D2B8
-D742AA8A388815C861038D9C9FC03A6708BDE079C6916E1249044D57CAD64474996D86A5538F
-34C07BE9D659EADFFDB7308E40FFBAC387806FCA0B99D3C17A1FD85074C31B6F46DD26035692
-3F4E9715B5AF2DC6E4E297E9D760AF91D5109A58591A2F75E640A6298BB70DD88321F6D2BDAA
-C9BCFEFAC06B075C3E43BECB57D13D882C8CE483C01CFB669BC6F8B6D0B68F9AC5E5742B115F
-3BE89A6F6BFBEED2B393AD2276F0AAA090B4FFD18AA8B44CEA8D4E3BD0BE06983AD199A5CDA5
-2F946F9ED82AB66CE87287761F3A9D4524F9D73CA7996E1E93E2C545888C61978E87DFBE0E32
-7099908C5FB2A3EF38279E3F3911164A2B0B992F79CBD1CC70B1A4A00530D9013CBBE6ADEE01
-D66523D73C88D978222E32F3739B824B5CA8D57611EF9AB62747BF609CDFF960D587F528BFE0
-9D097DAD87EF86FD09756626CF9D02D01A5CBA43F01FCFE11F9E578799C33474AE18240FEAE1
-47C69FE7CEB717F13787E6346CD0A61459CFD64B7DC40B3D93C7DC664B6353C1818613B3B5DD
-492087320A37C62EC64AD029C6C49CB22B7D015A792D09696981166241225D36234D30DEA1DE
-B9D971465CD8294FDD3030B0A2FC1EBDE69975E5ECA47C7AB1E06B5903A88DCE7723D9D67D1D
-BB248DDF6B5D45970345C80247AF5282E4F9D9518A79878CC2933D1E1D465EEA19A212DB186A
-D6040C641F6413FD1DA3F999E92819688A140A017EEF1ECE2D32858EE89510225C2B30D15324
-9DE7F77DE4E9C6649F0F21A51E6D544F9ECB7BFED6B29340BA3C73C7B5C4B0E3381D8BA2393B
-56731F09F2504CDDAB5215141A2C9EDAD16161C0C6B295B8070A760ABF872E7FC5854299108A
-879394435F8B5D057FE4295C5E5DF726405AB3B3849B785102AB8540D341F150ED63F6E43503
-E7CA5757A3A0A1E5858EFA672AA16DE1BCBA1B75F2FF0F285C8BE0632A1D22D40A18820F24CC
-CE7377EB648FEDA70DC05E2D9E36F39FEC53033587B0DDE00D492A9C9BBAFFB6BF70FF51041E
-1151D3678D756C9A8E5A327B824620F6ADB31C43424AA977DB9CAB16E7060D199C0E7030BBBA
-EA71BBE39FBD722E3B340D19BD6F17089C8CE0BE3F2EB9A0B6658C2A732E52718F7355CF902F
-07ED1A978DFF17961360420C8A8BB5190C334A3DAC4C3607ECC06C040F2CB7047A7466CA795E
-79B0C743182465E8FB49BD2A9EB2D6BD02637A8C2A462574697F56871C38301E9A79AFEEC64F
-1D2F1A3A75180CEA35DC20B4205F57D67D55C7C21E5DD55982146BD0EEF2AD0A450FBD37F821
-AF12E1CA094B19D84AFB70004660668067C85862DDF6D9778C8993762D9C56B1A159EDC4577E
-CA8EE5B6CF54B77BEBA0220D8D981F873A70B803452BB8992CA67870411A987A5AF976C0A79F
-BB9668062CBD933B3CDF88FA5CF665CD411457A81F08D85B6FC1C416089180189FA968E40C3B
-EC9CCCE68512C639B9293DCE4BA86A9FC78B5031C178901E81FD3F1B4520945688363EA05860
-595AEEACFFF017E2344828BF05C291BC464EF424423918DB4CB74AAA710246BFB9F03E383F71
-680623068D62A8C53D6A4DCE7F8534B52EF8C1A7311168D4D94452A66025ECC809CDC212312F
-F6FB04691675B1CA7E4C3066BE804EB42BFA3ACC47C11C2750C1CC06EDD9ABE042C1539CDD4B
-CD8698405C2602ADA651B222447BB3DB5B83633B9DCF7EDB246C68A45521B541D0CA137EF11D
-BAA438E963A4519EB0686503B5605B637C4DBD91AE1F4331B03AE7B52CA0DA8F54E43A15C402
-3DE87858D5E436DAD8A89FA067A9ADB46B2EEB54958D49D1D84DC1BD84512DC15D32EB8216B0
-8A3924D215382EFBCA0D448969DDD79C9030F0AFCCD337925325BAB6A97C7067AE96F70B4187
-D28FE69E8347ED3B8C73098615E032A5A012D9D9BCDF39BF8C42CAB9ED7FE9248F06206D4E6A
-10668C17DACAFCB033EF3362507860E3AB7C540EC53BB725B28C24A94D86CCAF8D51FCC2C0A1
-9B89AC78113F1AEFDF3A53E463079BD3ECEA106B40C2B95F3D6414CD9C2684952EB02E7C9874
-2DAAF5F1824D36E2DEE047A5D073AA73DEE5836C1D9462BCF78C837080C7DAEC9F33D2D6C1BE
-9D1042424B4FDFB0007119ADCBD9EFA371EAA1CD149476AE6B004EA47E5DBB9A2603A49F19AD
-0CB977455583F60534325D0CDA9E60A3A45AB4F03D336A79FCDB1F2D1DE0706616D159A3D261
-AF379A4E17829D395582F943BA93B66502A9EAD5F42C501A6D7CFD2EEEA17D7297F8985B480E
-0FF41233BA53E120FF97044485354767C65A84415240DF537CA0532DEAA1C23ABF3179823664
-EE96A488B0C5F1BA284ACF67DDB28B912C73018A34638B7E4200DBE758CAEF3CF6F82638F7F7
-87CEFB19DCC5FE6EC226A92051E2818D835622B92E70866303F49C9589953B67ED931C4500C9
-7A1B07CE250CAED036A2831F555A21130EBCD65ED599DEE85B3B2B4305BD6CF57BC690BD4E46
-A921C6D9820DF80D3C300B477DCBAD45E1DEEF467D7A64AE5A783BC75B2D71DBDA7F79FE651F
-1B9A38398000F057F19D875831227E1CA67408057FDD1C3AC2A814892C3265CC07C250B4B75D
-C0B47543ED8738B57FA413B2626E27CEDF5C52520AA2D2598DE04F868312D6CB72C945B884F4
-058CDD4F7C27DA4AFE166A9B3E5328DFABC69BB1917C85487F694E2213F018A0F22311CB9D7D
-504B29E914066B23E892D300F803036B31A329A2BB71FB2A4B8FEE77103D21EF7137D3D398C5
-2320825EC7B22A8202E3602BD2F4F65971AF5B6A01CB2030175F901CB3773355DEAC2BF2114D
-3570323C93004768613BBB553F43BDC0D968993E317091521061670E233424CE69E38A21A903
-351D84AB69B7625E6C52082EF004ED357B3CB5A76C8CDCE4DDD28C361E59B6CAAAD0B546516B
-9394CB06575E028F4CBE115135CEBD38FA60EB389388B9DDE02880128DCF5673B2FC1FC6F309
-C8BD6307EAD226E1EE855A11B0ADA7459571DD4A8231F4AD23907A9696D1EE001E2DD9A7B45A
-B8BF746FB9C904DEE41A95CAB9BA42BEC4923E68EFEE454ABC1AA47BAAF48573E9AC766878D1
-02E5BABCB26E84FF4A6CF6E9D45EE580528C256E536A156E46AA27630A1C342E0BC95DEE3B82
-2BF03D70763AE43565CD0FC1DF62D5AC7405AD00A79E6D44C7B4E125A13CCF7738EF45313CEC
-15D39C0D02A59E028C5456A6EF56DCE7E6F751DA4F16C21FDCFC055B28042266AB5D3FDEEC83
-4290DA6F6031F051B2F0BFB3CF564268CE30C03BA44970C05D2A4BAC6DAA103BB720A45D54CE
-36C45E8666705A653C76E31B9D345B221DF2D0F32115C312F5F236C68542784C8FD905464A4E
-3173B640C2AE997CC7B8400C34234D692B742872121D55BE2C9F7970110A5B98A94C39E37923
-E785767D2F0FEB6701B305038AD1D15C437E49448C8E0A5CD64865CBAAF4AE48FFDA3F3EA297
-278B7D28574025BAE092ABD8C20BF2C8B9C582ED9A0BC1B13A5B3C8D688371455F83D4D3767B
-5A1B3381FA9955D180ECA54FAA8FB6EA302DAAECE68A4CDF5961ED46FC593F17EEBC26E6BF1F
-E7B28E467F1D3C7527F0F45A0F51B2CA758D3BD98CC4AE00BF77E16D3763998D95829D1790B1
-597526244A6B9F7C8D9240C951675569E76EF318A92A9FC084AF2E0877D42D749CDB404B5511
-5ADB09B7B21E1918E9BE95C75FFB10385D5F3631ABA94F729F4C2B958B64AC746421D9C5C1BF
-9CBE5A987FE752F0255E065A9410DA2A3719BCA214E0267930098A6021F095B9F040B02213B9
-AFB034D8D393BD3A13F5D43B7EDE4B3BD16779D739A03B33F971AB70435CDC3AE6ABDECAE52E
-A875F33D741F0975FDEB9D6BCA5DC70FFD5C37E560303C905D5090D02FFCEDFB111F69A46EAB
-7D985C52FADC5DB2452A4224BF21A121171DCAFEA055A3A35AB94E2F18B6A10F2E6C672FF64F
-2A75C9DCE20CCB1C11DF351734448437507C232F5F714EAA95CD4E81C3CCC8304E546F9440CA
-38C84BE6E363F6CD6849A7B594C4BF040A6F0930239E8E9D638D759848FE1DD498F927C1B107
-F61FE195A2C299D5159C5729F8E40A5938A9C3B379C957E11798208E467401640B18D2702104
-DFDBFAF6CFA9D03E01FF8020888E26660E87668626F8F3511B1F58531A1CD83BA87D832197A6
-E8D8237229FCFF99C8D331C3B98AEE62C04F818A78B91129D168A2DF79FA9EF1E177CE7766DA
-B3D1980C066F9DA0638DDCD0D7772929C831E8DDBF2B5D49DE9EFBDCCA7CEBC31CD0491C95CF
-AB411034FBF3E507351BE81A263AA3A9F1D29DABDFB8042A00AFBF88FCA7DC39A528B03FE1E9
-B2F9223856CDEC8B25777C1C7A84C99B688CD70315F32930799F2431277857E7ABE0FA8D5652
-416B17701155017EFA45B8B4FD013EF9FDCAD36F3193D8BE66400216CF6903868B2047FA5489
-39B5BA92D84ED77AF779FEAF84D86721DC1A8195A21FAC05DF38D2BE615E51785751B6ACDB8D
-141DE66C3EDDE7A2DED9D2A43E86F826797CBF740FE94A76A518F4D9A44EF0873CBA9DAC2A8C
-938C3D8E821D9C2400BBECABE75A93CF23D419504C97345E2466F2AAE75C103E2A5F8F7116B0
-0B3DBF1D3159412897ACEE6AAD88E876E76D324855CB908CFC5B1090BB79286C9A4F915820C5
-E0D8F4FB56F5D21302752F2DC26E7E10D60FD3D0487BEC007E71549E32C66AEDDF127D2442D6
-FCE6536206A4AEFDA8C0DC38993FDDC61861AB92EA956A310D5A35FE2748167601C1811D37C9
-6E8F69914A10ADD3B46DFB25A363F9B82D67BF1B1B20DF3EBD975B0B2683BABDDD33CEB5D0B3
-95A6DB274FCBB679FC367730F74A316B45132AFB81B4527D07BC2BF645FCD83766FBD17FA168
-82273B9EE48CCF04C79104D493B5EAC14F7C7EA3239DA2A65DA4C279FE61773F39CC3AE4F939
-2B184994ED9BE7E762EDAD0A82B331DF4C4548E7269EB78EB2B893AC87B65ED6CA41B7956465
-9A0F59C409F96D71C72F45D8DEB64D97AB09CAE1F25C2339D30DB2276F2994F709C6E65C78BB
-FF17E83A8874616A92E24695ABBDD94DFEE84DD6D954E753B0C3DD38618F31303BC6864BE55F
-E0FC13162B11415046D54173AC128D56594061678FA2ED36A1CF6F4EB89AEE92DD7B7A75A52C
-4D2CA5FA19C7AFB5E99243BD582F59141A3E0F01B5B5C6A9E82440E606A9713E07EED6F6DF67
-8CA00FC787D789B6542575D43C6F50A168B9FA85A3E87C1F16F3EA88C01107BF98FE9FAD6211
-3736DFDE04D4BA546D7FFDD3CB49091FF64EFA01593896E82C8084352A6BECED8D37EF2B5478
-90A13C22B50653046DB21C47522846DDCD8A59DA8D838C200D435D835EAB110FD00FE2EBBDD5
-1955FF01D416CD55193FE697A63C0F2A2D0D8A934AE720E2E9CD98B7C8753B5E150964428AFB
-38901B2A793CD6B2D25AA29BB669AA020D076AA3690B8575805D121E09C9781C0D7A0C7B61E4
-8A96E32220BB969B3EE86841E21EEC4F1513102D13A351ADE72B3D829FB4606108B920C0FBC2
-9898E0F4175461D55A528AC3CC0C4D4039B2357A8CB1F6BE44E63A7FB3BD9F8685D137AEA474
-CB3ED67891110857C8B0CA1DB9283BFBCED04E9E89573D375F0E85D589E9A0DF4AAFE9113511
-8D3922393F622DFB4EA0649471E78AAB23B0EE990E17D42D7AB437879DD8ED4F60C471EC437E
-DD299CEFDF402B33ABF419E5AEE40D57EF7E80F10F5E40A0B6A655160A729C6E96DAEFA27357
-B917DD49B1CCA24BBC7CFF29773B16B10440995308405E97CF17B12A4BB23FD888074012EDC2
-45D91F3CED4DBCD476503064BCA7FBC9AAD94D4E7960AF26002FC3B29E0EE6764E4447BB5FE4
-B2E37A1FF9EEDC3502E8DA5766B18701F36BBE5C9D072AF87AA5A6428A4BA47FE958EC7084FE
-86F6CB85F768789D66775F4698DCED5F1B3A3907D5F7AF50493993DCA677FF01BE16526ABF8F
-0000F117739029B07A83ABF3F3D3B8097AB7D4CB21644EA8215FA089AEB3EA1156E1B7D132FB
-37BFF2621DB030756C01F6F7A1466062CCADF0E1F7138A3EC3892E5170E0F9073A2A0114F77A
-2CB6967ACF385B1CC55CE009D236A26E3CA79F6896A9ABB52ADF8A09A3934A0A5DE66F4E9DD5
-6B8E1E6599D1B40203AEEB01DB50915435FA6C24F52E1606CDF0EAC9AEDB3A56E3B8B6D8E9B2
-1C64A7CA2897C48F135B6343D27613D4FE1B50F7870D3F003FED5604C27206EC71E8E064E664
-76ED9090A232B92676936FB6327D84CB1D6153FE299EAEEE6DBD147BEC68D136019B04273897
-04189964684FA591DA67CB9A7A30330A648BA87BCD08F6B013E449D1FA1004E2A78266F0B3E7
-B2A25C7EF8CCC7EA2DD3F94372F5122754C06D1B4C0464173DBD28D1A9C290F038BF4D5E5089
-21EB77665C21B36363F44DDC8700389EA2B27256F7AD022F048B47408265DBE1BBB1138DB740
-34FE477A6548D809EA520B839AC464E2878D25688F22174C1855FD2464AB77DFD0596224DA83
-A8AA36B8F24A90A56FD65CB5CE456ECC7EF71FC20DE464CAB14785BAD3094A4F15CF90CB447D
-ECF8D14E3E547EACFC0D36C212E48133B7FCAA0FE577E3C972B01FB9D8C1ABCE833281691AA2
-709BAE479A0A7E629DC923938762360742F26546463FCF2C8981056A685E186B03902BE14940
-08BDF16EA82A5C8D776711A3E260820AD4C3498E061A8002F97891BFC5086FAC4244E3922D65
-2099CB62D5CB917B202CB791D201F72AB4DD3B4CC3883EE1F80A9A2B5EF221C7789B4DE6DF63
-FDF867336BA8B0AC148B7ACC0657D0BA694BAA0EF715E0C5DD22307ABCF6EFD34E8BB285D6B0
-CEE70ACEC8D7E7195532479B86588BB7D1DCA1A81493F1182D4427639A1E62F15E2BD8075BF4
-CCAA44A0D7E228EDCD7D1ED980B282EFC90020925D1257C466B39221457F32A3635A56C58752
-2E96854742033D5B1246B5AF903669B5411ED0B60F57204F9B21BB43E9CA51F10AB5310D8C17
-C4A24C6CE0C136E79F3DEC4C440B5A60F66189EC926B7640610F50FD39C4225EEB5DF4DDB9C8
-AA7BE015AB4BD74BF9A9B64BCA6929D2689F88743550421D4852F5BE355CD84498BD2CA8CEB9
-4500A329EE7D61C9F0EB79CF44DAEE37E7E0EFDD37C90A45434625C6ED4CCFDAB75FB93911B6
-7EDD130EFEED849174B0F0893DD8D271283D2C51FCC3D50E918A079F1B4574045663288CE95B
-DD648DD7A471CD35B0F283151EF57CD5F63F08EBE34AAB9AA427DE040C7FE70F66109DE4B6E2
-15C26742991A099E257F173D13CEE2FBCB24B732437B39BB4179A4819A6F70A8840A4560DE0E
-538E6B0366FA916C1D44A4629CA0628A0907399ED8C4DCB469D4F0E22BC600645F9CCB262BA3
-04303C81334D89B6C9889355041E386263091921F52A1504A3008325684A96AE21862F15FF95
-0267CA94CD5B821EC6E230B5FF0FC3D2D477CE778B732838147D0F0244AF0B4E671CAA94FED9
-BE21952B5ABDFEBCACFB1FE0E7482C01ED5B46F0AEC9E17386C2FA45E1E3D21AD2DB0F7A47DE
-9C63499B50948D7C47FE9D48B9449755AF6DCDBDD3E2C008780A7C732519410AC00D2DA4C7BC
-917A42AF0E52E00AC875152951EB150906862BE1B52C616747557E1D6402E4CB79F2398A9EB3
-C0DAE0B7D4FE10DA869DDCD73D46802736F6A91F60B02CE33978DD33450A1FE8D417C8FAD4FB
-AB3A87E8B1C651A47B603955BD64CEFD822542EB854B1142E4797E3C2D52DAD0ACB711F9F7B6
-59D82E040D19D774E487487D481D48B1EF9A29AAFD6E50245E313248B89886A8FBB209AD6579
-346DB8C8A8242206072887A7A7988B03792B18A05715ABB55AD1BA05FA021FC0F57347CEC136
-640E28144C51723BE05FE6234FFA586195E9B006D37DEAC26320A42893E5E50D3A24B054CE37
-D83E79C17734D682CB0A0758FB85F59B5692FE0C5C040A9C36FD929B590F6A76B31690F046BD
-CB0DF5C873D645238C9D4E95451C1120C4EDA4F218BBDC875C8B0823E17DFE6EDD2E9D603ECD
-1B24B461D14EBFAF8D39AE3D0E8CDAD6AE3CF30BC672332328D92AB4AB70323512A5E2C652E7
-C3AAED966AD123F1ABC487BF0371FDB349A0995E0143FC2DEC12857EC5D7EC3509E0A26D9C69
-D5D4FBE4EFFD6E494DBC39E15EA70D52C5DE10A7D6E3E4304D613AA639B5515F247C2B986A70
-4D799A94C663D2512A41E5EA4DC8669279274C9B056CC8B4C912C6109CD5574588EE7ED100EC
-03C4B45ED67BB06DA222DC081667CC8C6D88E7583F9BAE0FFA374343D89BF681FF7145D322E3
-3C03B370A67339079341AB0FD5AEBEBCD65E043A3F558CC2C24C4C6D9E728B4A872102A758B1
-368B1E15776C49C010F166F1C70DAC9656DF581271BA84E44585DCBB8DC8BF6900889CA7D790
-F7C65E82A8D3F7445C32FCAB7DF318B6196C6DC5AA4127C54266CF0904AE861C975EED1D64ED
-E5E7D4681CE3B6097658B28A6E287E0856757768C73FF2F450F320B63B668F681ECF6A5BCF75
-060AB93B79BE1C1712D738AA9E573D9734AEA44869CC09A489D127D222B0801966F28449D004
-237B7A2441E4B337864D8D7DABA93804723BA7389C91593905C561F1C04C8447E5AE6CE433BC
-714ACC635C0B1B59832175CB25BA302ACB8F1BE55A09CAAA61DDE770E86D60240D43DC68C2A4
-92E724BEAE9779216D676510AD5E574AD63A0BB1861F7355C4F75182329F95AB2532A2F2E6B4
-CE91D1B1F2FA98C779C35E482ED6B0FEA72F22B6AC70AE7BF5C7238ABD17CC95F434196C23DF
-E4DA439F085AC210611725566AE889459E16259205ECEFAB1BC48117C554B40396A72EA790B7
-C6203F6ED09EB7905C9B27902F34A536476C2734BFDF8D10A3096039905CDB15BC954288C2E0
-65D34893B58535C9DF43D345F8894BFAD402B6F79D0C1C28C7F49548B6ADD8A469AE00AA7200
-14A1F1DEEC776F7998016B494987E715D539652EBCAD1CBAF564489737D68810708D4776BFD9
-6781BE22A5FE0480295EDBA74A439E8B7D6D373267B2BE22BB064952F6CB595BC5947BD406B6
-577A5AD77C2D15306FBA0F8117A30B41A76C4FC3F815347C0C3CB278109E36AC0F2C16D46237
-65773C9707D436C43DD150E4A92920E60F255492082B8379426B394232C75307E43C375D472A
-D8E5423DC665C7729016CB9D1DBDD6A86B6182011AFADC8F3410B972CFC2343CE556611FC11E
-3EDC8AE467A795B43A66A0A8C07A6E863E619E6E9AF8E4E848E957F117B96F08ADD1E88D0A8A
-EC1815C17B6C05BC254B808AD7EE8E9FD3375FAAC26D3A4C1A81C2E227B43227EC74E9C329B5
-969E75D2F06679D44CC34F8EFB42344A7694AB1E280190FB4B5C833F6FB0E843649BA4B56AFA
-D3975C4912CBDD02DACC7D04DF5AC677728C4127D03962E5FDC17370DFE3597BD21D1F8D74C1
-16D9BA9D079C9235805D74F198B4DEE3422908A24F798CBC2B67DC1FF1B4E785B81FE1E5AD90
-82198E369095EDFB3E74CC015EF05EEA5B87B5447A7E72CD709F88FE9CDF01D3EB30AED171C7
-0B5449D32609AE8C46CE30D0395E78C7DE36606200924A1F6562E06AED297F6470EA0CABBCBD
-56B73E765469212C0141ACADDC9DA6BD29359488FA967511A43AE30DDD8B8CC21667CC2DCC6B
-6AB444E95F91ED1EC4D068B20546D517324A3DC8E0CFF5D2A8CEA448EE494CCFBF0C703D59E3
-065D1837625579592E3FD5B2C2EE120FEF398B3E0FA16E7EB9919DA1FB6C0E26A3AEFA2745FA
-7393F1B93C55C5337232556338DA70E698939939B38C3277522383B302814D7073F97BDDB254
-13E47A3F19BCC8A698AE70F787E317AD08535FCC8A1499E17C0322E6AFA5BFC57DD1CCAF9642
-D472023F8CA6FFB57A450AD8B0CD283AE66A099FB1820DB5FC12DC4E07C16FA9CB37BD39EE76
-FF4EA3227CDDE14C5B8D7B0E2077262CCE37F1813C2A0EDFC41B52EDC9C9050483CF263976AB
-9EF4F7C55FC25CA6F43E34CAA79ADF09763451CA762BDFBEBDEE7344337EBAAFC7B2D89C258C
-D69826B113EB46A7F26E081185575A8CE3ADD08D4A81197D6C9AB05BB26B3FA8DB4E632A77F9
-D20B5EC34022CDCD4B2CF77B879937B973803AD408C74D359FD0F2EA3ABA2E029156D4F1A897
-49DA865804254BF243DE3F6224ADB144195E4FE92ED67C3B2EA024DC5535997EF0094613D925
-CDDC28D1BE6A836B9151810BF805FFC297E7C8A9AD7B2EC822177E81A2A8A214E8AE24E077FF
-B164D1C68F3B5AFDE3DB0024C433B8953173C74DB24052F1E6B215B869AEA0900AD3C9E70695
-1ED03E1CBD89D0F7144E54D4FFC8719896073C34E4199BBFC5BA2AC53C1BFA31067638AC42B9
-825AE2462F7B87AE44C28FA721E7624A1F5D07A39A99355CA7D854100C19A6C52571A7A5716A
-CAA9D054B3BC41F438F5DBD3E76AB7F00ED728FEFA7DFC7DDFCF0CFBC653647E47D7B2A7F09E
-736D1A8D409494B75F0DE766729A77AA5DCE37351DAEEEFBA1E94D1C7D909B7F076D812793BE
-D6BAA55F0EA17AB272BD6921FF46E69629B5361AF26336C89B5F33556F3643EB0180BF3EB7DF
-9441821E8ACC3EA44089A2E283F4570074FF897A2171AFF48C2CE1E61C766423454A348FAF94
-1F7BBFCFFDE7FD212E4552EC0E3D10EFD796327C0B61318A6E4480052F0A2B7E236B93CC5C65
-755439E31EEF936E414A20FDF2764EB8BB24CF88A7192E923978A380B4A991D028659F27AC94
-1268853ABF6A298378CB585F0E102BD85490999AF2B5AD246354C20E0BFC8A6F94BF21C279F8
-74BF2E9613BCF3CC70E2798D2F064B9704E739E4C8EA88FD6C2495F4AB02845486373DA3BA87
-8B31007200DED291995C34A8E2254373166D08E1A154285A96D205C1AD46BE1BD1C7EC6834DC
-8F00A3EBA8366E7CD3319C9734280F847C4BE640BE3EA783F153499A9B1D0FAB11FF2065EA3F
-5391926C58716B869E84EF7E4F293FA26C89865E029ADB446DE7E5C8BEB3859D6DD1B323C911
-1286772FB90D937AEC5E579E57A3291628BEC3D2953760101AD9092A10C4E44EEECB342B8854
-6F547AB90BE167505CD11F97D1050E5044718699E1817BBD600C259B86C505C97F0935312D4A
-C1DE2A70E2F78212A0C2E6DBF710381F6976941104D6583E2D9463A471F758D734997D285084
-34234D32D1D632DAE32699EE3CD46C0E6F3167BC4A11551BB52DE011CD3A806043D9173945DF
-055AE07C4A0D82FBC3B8300697A7E9F5B7C0765F408FED473CAA5265CD8D65AAEA41B01B0849
-47C91633A492194C636016872C35387482CD816CEE8A8251B55093684319C4CE2604514C9229
-7E1C45760E10D674EB13E126F23C298F4E3D760F136351A11F9751FF1237E3DC9348BFE63E64
-89915663D5F79D74F19DED09DFC6280C32A1561AFA8261DEF097CB01094A4F351AF09C048D0A
-34475FF54DE3A6B7F7DB73109D6F62A702C1521FC46A7125153096878B3E39BCCF4C2A4FC796
-109D6F14ACCBFF8024310A06EEF4DA0A32504C7AEBD0595443A12C3F6F44DFABECBF07D49B9E
-A3E02149FA33C21C56A6D7C6AE41DA5AD9A5154F011557B8E627565C8DF7096B956235D70F79
-B8DCD0EF7D296EA5001F4989C8DA1FC9E6F69F5A7068A47BF8714CC75BA00DBEE21D3AFFED28
-C285B82A5F6EDAD30B7ACA46C18E13134089D772E5788AF18CE815CFA3BBEF58F5BBB13B9C62
-5BAE15C8B98EEC8FF7FD75C6DB4753F73B3A8B5CF829C8FBE25BC043025A765DEFD2D8B212C2
-82E10514F6D4528FEABB34343194207B095B98BFD90B65D042F2099C7112EF4867CCA1FAD4C9
-F2C7AC2AB9B22CDDEFFA98EA18C4AEB41709576F5585701022D384DDC880598CA805405401BB
-2EA02B81F2CCBE1E60DD3B5FBD59CB89F910C30F06DEB5DF0980F1C8528C81294FA3D2AA8FF2
-9CD6F344AA32FEAB9B115AF83F76CBC320BBD3A3C560C9C480252F749E6EA3E0097F794F0797
-7B8D0DFDA392E023118E07AD3F54C4F50EE33272F27CB52EDD29BEE23C887E8C7707CA4AB66E
-54C6BF795A402A27FE7F77DF9C86CAC7781FCDF24952DC1E4F135CD266C6119C932C2E15357E
-34DA043E4D4D71C789187243A9DCDCAD64A7C5B143DE2070A0D050A9255C676EEC668B854A4A
-6AEFD204615D80D953C50CCEC3A59452AEEFEFB7C72C6EA1E4A4C47CC3ED760807DBBCA68DC7
-7EEC07D79887B22B564A3554E9F603424D862DA3B1D811AF1BD7A9B475BF221CC9C688310EC0
-ED2D52D542897451A46C53145C4A87CE8FCB969C95F8861C184250F9327DB65590148A133C09
-85477D229BC36C336D5711615F5DE608F25E3FE549779A436FF856C7C46FB2A263B0B2B853BF
-3A4D6B9DC1009BA966830791D9C80E60935A977A7B94607147F4DC7FD2A46824ACC9ED7997CB
-1461ED6D73A0DAB61A98FEA14170FC2806631C19A2632FA61149A2A55F1F44EDF8D7F219B5D5
-B7D8BB98D25E67E994D9257EEB01ECB8D266FE53FA5BDDC2D947B84DD9C2C58118130DA79AA5
-E2739B4428EDDE707BD6D3D8FE7A5ABC1BD9B2D18CCD144118C220C5070E95D95329274A1DA3
-5E49D9ABAF740DB2341C54DC1DF33F23953D2CF52DD18C13242D7E69D5257A33C38151A71449
-638FF3F945BBCAA5BE1178805A47265B6529C1D3B5D99154852C2BC6FCD776D22E3214FA0469
-C1B50BC0404275DD61A624F7EF6DD40C74D54FB0AD73D4F184F5E3535EFEC89DC8B229E626CF
-10181E0144D2BBC3FB9911CF005D3E8EFCA1063BC78AAB37BFCCCDA687DB3BBBD38865AABD79
-D559BDCF259CD492D444950DC30CF846F4B080A95FC83A5089F85FC74FF2332BD82D656D12CE
-C0503E7DE2BC24D4108EB513D5275D83921FE0ED13115747ECB6BF97C8DC9EFC258621B0CBF2
-E1844D5FAC4D65101DF160C8B6C7BD470C30725BC5BC12430F8B70B9358B15A17ADF11900121
-C78BE837A37F21C0E85969E98EB86B162C93F3AFDFBA837323E6368693B022B90E7CD69056D9
-6A80131828E2E84499728DC78CC573DB242DDDD6CE23A65119E3398F7B098C52514A8B41A9F6
-8B2DE28F46ADF30255EF44FAECD65A3A8F58FDBC735AF8800069066117D6AB3D8B925B514912
-743279C5B34112B3A74CDF95465BF6C9927726C877840BAB0306D7EBE431AB3AA702AB1E70CE
-78450C9D1F5CC953776A5C6BF2CB8F84061C4101337E5446A89658F45636ED60962C09AE4A9C
-A299387923B0ECED4D3932541510A92E7DE24990EF7B046142E0576F570D146413CC03A0F788
-088F0F4AB07F7372A64394D87E18BBD9CDB48FBEEC60A705FD02986DC9D169405286F96E819F
-18178198765B6C2A0D6297A2E3C52C29B017DC840391C7A81ECEC4F3B61714FB6A9C0F43A715
-28F0C6880CADEA1B59811700F6FD8299C7A74AF69DE6A63124076EE2CDE284AC0C4FFC026008
-E8718B58CAAECC12F1213283B80E09047DEFD054BB04D07E8297589FEF1D797B3A8FCFAAB614
-EE1848AE998A9F0741D8514A669BA1B5C8663E00E1B773E9F16806462A809FB7D3F575944757
-6253334ED0F63B659FFF6223D7C2254584DD0E9BB6E92C75DC434905E9A8CE58F56B5918C899
-FC2DE34FA4EE18991E0ACBB29E9AF9EA3F5D770DD45DCA4F76A18857EB9BB6E44ABC10394158
-4A100572BC075C7E1B4F5F56C150C0CC98F8CEF9F60F9BA14F3A5557F588E53F918F29FFA80A
-9DDB03271ADE30F201720C2813FCF88CA1226690B5AF7E436AC43037D7DC4DBFD981F3257BD8
-5B88BA6F732FA9E611328E7FDCA1FF3358F6F5331C9ED37A05B78895CAB97B70BC3DA6BA6334
-9D33E4C9C1B10DEEE084D37AA949802ACB151E644C588BB3AEBA9DD2B091032C497950689A3D
-410D4D7EC73CE868CB5D5377744EE57ECE85D4F7E508D6234F215849473E2A10B44E85F1F00B
-7B17310D03C366337839C1A48B6B8CD6360A8D6B48183E4DE8C0B07AD2AFEA9039001A554E26
-C70467A906634955F67DD3F1D21B87926A21C8D7D37329187034B3BD8D85D347350D80DDD99C
-69A236E0553FA650D80712F2862976C1C6636D6564B5C92F7325CEEDC43B8F91BC05240BA891
-C1B83C6F77B832AAC3F7C878AB7236C089DCE374558084D472BDDEB1E1D0740D68B75649BC31
-2E4C9272BC82D9CCD37C634CE4D0CAB9C8A8604667BADBB179BA758BD7375E97838D28CC5859
-A135E8BECA0F5794449906C09DDA95D46DAD00F12AC04461BD805F3057398C73F4FE902EA3E5
-3E2052CFF9D9E3BC3E59EC1C6ED396E3D70A6A989C82F6B4A599785F81C7F0D7BA2E1DF1CFE1
-1055026D64AAD127F490B8D3389C98B41686292A00BD6D88FE232FEC10D7099580926E4F0C1E
-402F0FD8BFC6DE5F1ECE9518100088B8B2699C8B367201DBA28F8CD23C8D0B8DCEF6E9656CBF
-F25A820930EA923B23EF1402C9CD9AA9737B294B22BF59E2A6114C9E5CC418A104DCA45CD8FB
-B0889A583FB7B16E7860BB6F10E2A4B4A5662D1D5532E8121155CB4DA0E6A1E86BF65A6255F3
-043F222594BFFFF37A0013F0FF88B33ABE1051B1653384AA1E02DD29E3DCEBD813B6CC21F974
-BFC1DBC1D748FF2E41DE510B445D47085B285B4DCD2F129D72F620727636CF917D8B1B49BE85
-92CF0D52931C5156C3AD4CD5F3ECDFB10687EA9165B7C674AA0B5B6FCD44488070CC97900C18
-6AF9B8417A63BA19DC6FBA8EE3805D83C2DD86A68B833038880A911FDCF034F97DD498DCD359
-1DAD3E4CA0C09B8F0589307C56EE977138120BC1790128F45A1560309944299C41499FBFB9C0
-289C86E936B6B833AC32C77606B5C69F8B2426BF61C01B5CA3F44C8CCFFBDFA7FF8F3BD82B47
-EBD24357D98A2E993329B64D14B749A764467FF73DEEA612A307C0BEFF194460D2A3DCDA6FB8
-6B8425ADA9EA896A3A0F2B1CF63E221A426FD2CC11E459C2C66547F69DAEBD0EA480B8DFDE50
-1A5DE4616BD10032135C5F47EE013EA19FE78D5D00DE8E9C78F3796265F6671DB9FAD451FF1B
-FA9123A0BCCC9D5DAF877300A580555A57D79F26EAB6E4E9DC21219F19683CD2A89E6D60BDBB
-E89A7553E6484E702E1F57966E7709C754747F015C930F3892BA75BA9B368C9AFFB959BD36AF
-A9038CE08D4950726F7A3EA30931E357C2CF924872281DE365A60571334573E0F25EE5B57566
-3FBFD2A1191928402084EE9C179612B5EE417A3563699C2CCA50CFB9418DF13083BF6495C707
-85392CC0584CD7488C85646B22CB1EC66DC5C540DA2F109714BA2B4E75F260682906D69FAC11
-E22F35467E164A68BB4A3D87ED95FDF1F787091B3BEC87765407369F213990D0E5E3DB7440A3
-4B1883CD7A0DACC79B4AC3BF9D313C7E8DD8BB3E3F7678C1F6FBDDCB5821FF3CE594312BC729
-2E48BFD9E02A0519C8F75C677F5DBBF1CF4487A5944AB0B3673FB95B43DBD8E7E8D5B7B8C261
-A4FAD6BED10801BEA7F215CD0E6CC7B509DCDB9D00D8BFE5BC2B80EA2F5F4A35B36BF098AC6D
-B416A989A05027869468DCA7EA368620F8AC46F55C28BCBB3BEEA16DA7A46B9424EB726502EA
-ACE8A02738ABE6F78750E55A2DAB2CAE0C6D48FF269861A2D5C304CAD986E97F28F940500741
-510C98171A0FDE3F52E0A46355D6035419A03371D94B8426B0EE0750B29DC9BC1A3769EF8888
-1920185D23CE3A878A92EF6214101BE7B4DCB1C2042E11A9F388CACEC3408DDAB2A20EF72C8B
-ACAB29CCA51BA518570A5D3989C7772D8BA1BFB7B8EB519F40216D136AB62499ADC48E8B1619
-4C832F37B8A4A32BC00BD3080CD12A31386D6158EF9C4B7FC105F9133898E2FAD594F894C476
-34E2F57311BC31DEC69610ED59018119615223B13E6B9DB30F8B158CA0BB0321B6384BAAE1F8
-7D2D07C37BDFB0EA98CA580DDCFB1C6067B7F06C3E428B3F27D13F5EA451D35ADF23141ABBFA
-8FB0E117E3C3842471241C630B407E951BDAA919A0C89275823AC1E67B3B611DEFCA251F8DDE
-892FE09781D6AE7FD12D14C0555172C5B3FD84C2982D9F509A7A8C7DF18E1F11114E7343F540
-5BA165362038D91F29058A9A5189CEFDDA6AFD7046718368F90A12E011DC32A7F9315BBC24BD
-CE36CC7FF632940FDC80E1DC9BCF4CC6924110083541D6E52A6818817BDB9A31FB7485F58167
-E7C9FAB9DE53FED4057B9DA8FCBC2A0F391D779CCD3CC0E6E493871E23A01D5B86E7FA03C16B
-9BF1ED9954D2E84B763AD5932E9EA115E6466AFB39E2896D45BFA3DBC58AFD5D11EA4556EA95
-46EFD145D16A8A0BAED9BA235A448A6AE658BC70F615B8180A1B6780E9C4DA899A23E8AB05E4
-3F540CC404C2D3828E2637E6663131166F027B2B4AFA5E1336AD86C8AC66D5430D8D2A76B288
-F1D265B79407FB2FE466051FE11CA4E6E7B0092D8FEBDE8D2FC1D19A1BEF6D740A401C1A7186
-CB82C7870499A76640B0846A77D0959EA9686820FF3A9CE9605596DB780248781E97EF4CB0A2
-E2568CDBBAE19C7DD605ED7C02C046FD102159ED2389CAE2C543A31DF3E407895FCE3E8C7CF5
-160089FA84222E95D4F91473B18795CC31903DEFD7139CBB5D1079614CCE5FD27A6979F42559
-618BAE0A75EA891C54C2D584FD43AD7C447C52A16F1B7D54F59289B687FB76C3591AC8F0FF30
-08F33A96FD93011AB0D1E98736DF71E2109D83A2F10A73A042F5A5982EC5C68C8A2959276A03
-7319DC64284D3F659FC4CE81A12CD6A821F9BF2A1D3E0E86D685FA4EA6453FC58E6F616F0623
-EE4D23B150D542A4A7002AAB63683840AC17AF3C3788900FA15BD98F14C4299EA2769E170A29
-064B8F2D44BE23D77C2F5A9D2E0B06D6D08CF71CDBBC5D406652D9285EE5F5815187E6767967
-5502561DDC4A0864047C4EE04CF2B3E78410B03BCCF03B45A78CB6E3457948F5DEDDD658313E
-13C69346CB82D26348799213587AC43005C7318BB10D845874244566472C442B617EAB079091
-CC1E7AE87860C0DE4658D82B2FDE4152D3B126ECB3592D971F6A1FF28E8A1E260FFA98C5E576
-941EC7FB19423153B644452E39F1A287C2ED235A185A202D1157CDB72BBE9C8B187C91521A4D
-4C0924BF587EBB6D82C4F326F760D257E51347E05A00448E4C7038A59C2E44237CD51107A8A8
-64415E0922DFEF09E57F5C0C44D03C5CDAABD1C57D5E505266D842CC988222FABFE45BCBEBAB
-83E7A6D3E9D1DD7554DFC8C27376D907EBD5D7324C9F7D8D667677D6643C433C88983F71140B
-F113D86E389E0F3D9E0AD2A292FF9E29BD60088B0F62FA1F27199949ACECABF9975369324417
-C7E4F7CB9B37AB4E612584584F44CAF21A71A6AC5D8EA72BF112F0E88FE657578F076EC90698
-5725F8492DB5FE67E959DEAEA81BFDD0064526D868F832BA4445012BCB8091DDA68362A19BC3
-87C1825EA746D32F8DFDF426E52D860EF67168AC5B9F8975B906D6E8888CFEAA8047B8AFBB32
-6A36312B59E0C5453BE806A87D4404E2EB9DE203A78357F963AD1BEC14293775DF10F62033BF
-FCEA7F4F75B302C1B5FABFAA4C46751724B413737D75A3A16768DA6CF657A6AAFA10F18B9633
-7E49FEFCC4A491FC71DE38804303EAF732FC116A6BE70EF0EAF390AB57C0D86E1DC317431892
-C8600EA9676CF4549CF7E5A010CB6EB5C3E2B8EE16D117B15B1215C0B50B133D9EC00146D523
-5E88DD552A25950F7B49ACB9A22DDBB33860E4716916B370E762C288C6EE28538D64C69A0177
-9D268EDB9B18D595401991E683717D2B44A0AD13AFEE2448C069A858CCF61F3123452CF8E9FC
-F8853E037B50F72A7BC7996D499EBFF308FDB4A4474F9B481B129B2C39EBE840B02E817B4A84
-DEB38FE1018058B1A588F74A618DF1B18D6F3DFDF0A1BA224DC006F9F3EC4833649232B0DB68
-9FBA723E9E79D1CAB87576D52C15C27C6E3717B1FFC2CB118DDED579CCE5736BB6B5281036EB
-84B0D647F576DE2BA7CE4C9F5DB48A6F423E131906F470B848BA6719F8858D1248F27E64D026
-01A203B581BF64B9B2F93D34BA85463912B41303ECA2C4B7F145F3D2D8BF6A38EDE77DE4F637
-551CEE474F678646E8E7B9C62819CB748B9774D4DAB0F9BD50766360E8F0B86ADE1177FBAD7C
-07B0ACB71775337BCBADDA38777C623ADEBA21185A8FCC8D7A3E5BE040CD007231C994E8816E
-73E0D4F80BC89755DEF959A43B3D2F9E1B4DE70AD52B3C8B8F96A0F453C552B83D290BD3B6BE
-07F6D7EDF42571B2335A09F11CB3777F8A87DE46E0BF85BB95C988657F6ABE3DF552683DCC1B
-299A671B9828062C4A2BA1FF406DA90DB6CC37F7AAF27DCC97C28DEC1B9765901BAF6C1CF913
-BBFABB60746882F3AADB86302858E86CE5637683CC1EDA6181A10F131D7301669C7D08617C7D
-1A6D2827352C249C6FBF9E9907B97276A33E8C65706D79FF4FA5E15080353546A1D5AAEEC0D2
-AB122A1876FDD6A8C8A1B195D2B77ECFD799F889A43CA951E7D4729C24109DBE02195241EDFA
-BB844B9022379D80C97E7D5048851BC3C504480B3F2AE8AFDB39F811139C44ABBB8E5E46B31F
-40F92AA2130B76770FFC76EC84836E6DC1E5A085E436FC7D7EF7B4058C95DE9B99ADA856E55D
-3C3336ABD57E6BE3183AF24AB1AE510ED477D77EEF55606FDBFA7847DF2635319EBE523D7BCF
-5BB2FD8203AAE32492128EAB3A06BE2990649C77B65A43250CB6DE1CBF1A7C4297BC9C0ECAC8
-7EAB263A34D210A159CFFC61A07D398E9DA551EE43D983D7A7B190678727D2F8AED192495201
-0DC57C2EA5CC31DD1DD82CE33ADFAEE7A15066AE88EE4B850AAE3CA0079AAD0D65A5AC347471
-034FFC121DE0B7862D2C80E13E6D0652E4F974565B8F9F1974F9A5ADC624E5EC59C01D18779E
-A5E52FB9712CC2652D225E65EB4A50F26AC336B0E41DD4D6B6A7E5E7E534CF674A129D8B723F
-BE5D18EAAE3D0F2A2D233DE87D0347C8E974452DA5D12517C73632FFF60719E7E343D542F736
-0046AB05A1BDB87A021AF4A4A279817F50220514145B2E65F13EC96F4E97367DD0CFE56DFA3C
-CFCD354B4A654E10DD04D479B83674F790505C0A0588DF89A5406B571300011CB268CA95BA45
-5FC101E131F8B376A7AA4406D9913646B2B5631D5D08687B74CDCEA9E7355D58040E65D3E0BE
-0A10984179E872DE61A451038C50BE5624BC82D5A5229C4947A03616F5DBB60608A87BA98040
-0157A18D6333FFFFDDC039831E4765B4544138D6DDD736433676710F276A0C1664F62AF7DEC4
-D4B339D04D8930D404A486C1A5F0094F7D5EB1AAE1B079EAC648C7AC58476940D8B83258DAE2
-7F5CBD45F9021CE248A6B0E93AAA05AF4FB15494AE9DF1F2BA08053AF733B9BAD21A9945FBB9
-1A3752BF6D00611F18748D0D47BCEAD5496F72B927643052DE62AE38040A55F90885620EDB30
-B546CD76D985EBEB51C3CE24406728580700BEFA27C97181829D69390E549D7425CAD8294790
-D72F17E15341D0334EB8802F8CFE60DB68379402A77846E633E576913C7356D11405D2A2994E
-63F889F5EF67A5F8CDB21E5A38614B05FBB0F0701251DE673F55B83E77BF856573D13ABE34D9
-6758A73C540B936AC2FA4B8B29E71A4F6FEC539AB7B0C19B1AD8186B9622B8154881777BC4FC
-A7413812A86A6271EF9153DA2803A74C6E8F0AEEE75A52601E08B9B3E271D42D2B890F66C6AF
-B58345328AF86CA1F35355CD656A074EE6CD3639D6C69558627009B067E8AAE0F189764E4538
-1FA712C95D77622B2BDA55A39942FA18F2B3BB56412C5AA331037305021C2758D344468F1177
-537D7E48A2B533FBAA744AFE3FCF1082EE1B9BBB8C7D15617F5A55CB4A4DE1B0370042EC0344
-DF7C97CDB748F6B253E752B4DE504B90CDE9543CEE2A9E3FE5C8FB579E3C5A1AA1E173AB0C3F
-5191351EB6037C81E6A1CD39FC4BD25E4AD4CED262D725692D5948E2D8772BAC555D941B0832
-03D57BCBD6F3F326E83F085F1A4F45B7ED261C3CE12DD3EA39747F561C09A408B622156CAB3A
-BBDF04719A8A1F55A544788B0DE3EE7290CABB9C28982FD171822B586F290D9A6B8752FA5BBD
-565CA13B3957140D72D275C3BF623FF098DDC68FB32E8949F9147E4433D606F23524392710B9
-4504C2B08A98F06A068364616E18E61C2E1CFE9C315A6C23733DF7D7648FB18425BF265BBDDA
-C77C4F327181F7C5C118A6F6C8FA3C1E48F9F165B21D2B9587695E2E87952CE0B691F73AB970
-968C2409DF1D25F8B56E9B6050F903F6ECD9731256E9C50E14229525EF8C33F97D105A1D36A9
-862A781DCFD8C12A218E05A7CD3CD92EEF74788F4CCE3C2843EE394C7A4466DD6DBF2CB6CE8A
-23A70BF8776C6741FFDDAF36168BB472D16A22CA4262B9CDFDAA824F2BEB91CB57723FB75124
-40A927F2124245BBC8C34A41900FC92F8DFC54217FC083919C22BEDB7CAD9BA218D8FF9DF1A2
-BC76BF8B4DEB64848209553B9457FBBBDC95B923210C8BB8063CE90CB490AE34CB748205D268
-504388BD46E560DDAA0193DA98692E4B6F4A57DF26F6FEC0C865C9C982707DD9548AADC58684
-4CEEDDF34F96EA3B884B4A8247EBB18EC7C3C698E90946A4211D192644E1C33CB0663ED4B2A0
-6FFCDC724136F5D858433FA180C8CF0658665023FB58AFF86AB880FEA67067069B8581138AF1
-B3C9104CD413630965AC9EFF9D58E7EBA664B68D7D807EE0E2B4E7A25F06847227CC849D7F87
-7E041920F096666E434E92F0BAACE9CB1FD04B8807AFD988F4EF220FC421650B034EAA9E8036
-CC2C73D321D33E2DBFC4F9C9AC3D99C14BD573BD470E222257DE110104F40E17732126AFDEBB
-A7188D092345EC330CD7FAF1A1F8C96FFA89075815916EBBFEBA2C919E8D6A5D4B6B805F2273
-2461F2854EAED84C968EDC50A346B475B31B41CF55C019CDE41C848CF5AF5BF5BF936F90C3A1
-64B19AF696123FC4766CF465920D263499F9A3B3B42988FB2B9731A326EA54E13F71D13035BB
-035426AD5CAB6B761EAD8AF8BD553C8D493F2AC4F3B4251A2A73FC720B70D74755953284A25C
-3E46B0403B3D09DCC6B3518FA89CEEFCA7A9C8EBB373021D052276334787C7CBE3AE9EDCFC92
-AD5CCA883F4DE572454CB2A306CE9F325EFE13101E300154FAFD6B2AE11ABB351305910BBA93
-374420E55EB5FC4B48BB525FE8B13231AC03657FD10707C9E8FDD5CD63A1C84C7AEF5980A7D7
-0F586C3B99DF22C901D90F7D427E6ED644336E58EF39960E675DCF9516E4B9D954E577C88199
-B2526F37CC2AA9DA1D82C6C60AAE5131328FFA178005098507525B95CB469BB56D3065B73DC3
-989BABE961A7B49070544DCDD7D506A69E01466F44D18CCC25B1380AB3A66DB31F4B7AA374B5
-7767F875DE1178492FE3EA5C6A25826A74646519BA4FDC3E5148E9D8516082CC9337E531266D
-6C805A580128DB06C6361B2871D348C92172345243602C24F94293CD76851846AECBD1196188
-E4727075E7A6583CC711D43AAE5FE1B4817A8C3441122907A3AF4D88C4D4BDA0F850A2B18A60
-99C8D50E5D798FB4EB59BC83BD5A80440BC4223A25D75A52C18A2B88EC02D7DF3E26548895E2
-D3F45020DD67838EC668F544D3650D9F84EEDB467B646A4F9A22A4E993E6ED1B988E7236FF0C
-2708796743061C38B515D1E811FF618022BD794B44C3303D695B6B9D6AEB7E47CF9A6588AA22
-40C2DFCE3F52B7BCE5A5EE738EFE99F1D7F3209AE479D470F6EE9618F88F25D12861D28CAC1C
-9B433AF0F812DEB94BC3A950B014AA10918847B6FB25862F720D3030386CB8335D68ABCD3512
-D29FADC6ACF37AB20904FB79DBD7898B4EF73596E44E322B6137AC9E9241979FA798FEC9FABD
-B7F1E44E563FF2AA41DC414CE4761E2D6A0D0F919BFB5CA23EB5599C6D06521914A9178C04C8
-32604F9DF0B6F7ACAD82ABBF59818E5835B0A60AA4AA2F9799C03CE2EBC3FEA221541941B167
-73C0102EC9197D0B74AD67D49E6F5F8AF1F2DD4B3E7B75E9CBE4E213A57B4D6E3BE6A4915B64
-CCBE7F0732B78F9464541D0871B193709E35001E287243BF898DD660CC163DDC90FE28781214
-1531AE1557E9F7E91D3C51F73C9EDC60B5DA85B71148864B85FE72BC597A4B3440AFD9832AB2
-9E165E893F56CCD10FC45C322BD682F5209B495ACF4C57F0490FE4ACE13E94F07A5652B8B6F2
-D3908DFC4042579A7BA2DCAD968677A07C5CEE8165A90236DA1FAA1A045475B945207D723D89
-806A9F5A4DB06966CE2FD153659E5BE84B14D32FD54E073B858B63DA0BF32D77427199B6E823
-E48297F32B2AD113FFAAAEF5F99ED30B729AD4AC516C4740F2E793AAF8A99632F87748629DD4
-F1A71FFBE3567C01FE54E04E1ADE7203F9EF568F59FFDE3B138B4FF5ECB1124F77A4F804026B
-3D62D4EEDE127DA828FB0390872E0547E1C02E924B2E564790586C2F8B5473923CA8EFC39DA0
-B20802C25EB60FFB8D9B102A5EBC65CA9A4496891B0A711813DFF664D01CFD35BECA93E12A04
-35724550FE6FD2ECC7AEF330E367B60BB66C2F02FC8413C53E38C48ACDF8CC0103C5C32D33E5
-E2AB7BAE0858FF331417487E3254408E73608DD8C1B921ED1E6069D6CBD38ABBD9BBD5FDBA13
-0AD577CCD8763C061C956234C1906E6ECBA8FCAD31ACB6147D33868A09A4FB10412B218AA900
-FCFB078BABC7A58FED6EB2E45CF5F0DB3506609BC70393EC8F6482D798654CC171370F0B0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-cleartomark
-
-
-
-%%EndFont
-%%BeginFont: CMMI6
-%!PS-AdobeFont-1.1: CMMI6 1.100
-%%CreationDate: 1996 Jul 23 07:53:52
-
-% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
-
-11 dict begin
-/FontInfo 7 dict dup begin
-/version (1.100) readonly def
-/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
-/FullName (CMMI6) readonly def
-/FamilyName (Computer Modern) readonly def
-/Weight (Medium) readonly def
-/ItalicAngle -14.04 def
-/isFixedPitch false def
-end readonly def
-/FontName /CMMI6 def
-/PaintType 0 def
-/FontType 1 def
-/FontMatrix [0.001 0 0 0.001 0 0] readonly def
-/Encoding 256 array
-0 1 255 {1 index exch /.notdef put} for
-dup 161 /Gamma put
-dup 162 /Delta put
-dup 163 /Theta put
-dup 164 /Lambda put
-dup 165 /Xi put
-dup 166 /Pi put
-dup 167 /Sigma put
-dup 168 /Upsilon put
-dup 169 /Phi put
-dup 170 /Psi put
-dup 173 /Omega put
-dup 174 /alpha put
-dup 175 /beta put
-dup 176 /gamma put
-dup 177 /delta put
-dup 178 /epsilon1 put
-dup 179 /zeta put
-dup 180 /eta put
-dup 181 /theta put
-dup 182 /iota put
-dup 183 /kappa put
-dup 184 /lambda put
-dup 185 /mu put
-dup 186 /nu put
-dup 187 /xi put
-dup 188 /pi put
-dup 189 /rho put
-dup 190 /sigma put
-dup 191 /tau put
-dup 192 /upsilon put
-dup 193 /phi put
-dup 194 /chi put
-dup 195 /psi put
-dup 196 /tie put
-dup 0 /Gamma put
-dup 1 /Delta put
-dup 2 /Theta put
-dup 3 /Lambda put
-dup 4 /Xi put
-dup 5 /Pi put
-dup 6 /Sigma put
-dup 7 /Upsilon put
-dup 8 /Phi put
-dup 9 /Psi put
-dup 10 /Omega put
-dup 11 /alpha put
-dup 12 /beta put
-dup 13 /gamma put
-dup 14 /delta put
-dup 15 /epsilon1 put
-dup 16 /zeta put
-dup 17 /eta put
-dup 18 /theta put
-dup 19 /iota put
-dup 20 /kappa put
-dup 21 /lambda put
-dup 22 /mu put
-dup 23 /nu put
-dup 24 /xi put
-dup 25 /pi put
-dup 26 /rho put
-dup 27 /sigma put
-dup 28 /tau put
-dup 29 /upsilon put
-dup 30 /phi put
-dup 31 /chi put
-dup 32 /psi put
-dup 33 /omega put
-dup 34 /epsilon put
-dup 35 /theta1 put
-dup 36 /pi1 put
-dup 37 /rho1 put
-dup 38 /sigma1 put
-dup 39 /phi1 put
-dup 40 /arrowlefttophalf put
-dup 41 /arrowleftbothalf put
-dup 42 /arrowrighttophalf put
-dup 43 /arrowrightbothalf put
-dup 44 /arrowhookleft put
-dup 45 /arrowhookright put
-dup 46 /triangleright put
-dup 47 /triangleleft put
-dup 48 /zerooldstyle put
-dup 49 /oneoldstyle put
-dup 50 /twooldstyle put
-dup 51 /threeoldstyle put
-dup 52 /fouroldstyle put
-dup 53 /fiveoldstyle put
-dup 54 /sixoldstyle put
-dup 55 /sevenoldstyle put
-dup 56 /eightoldstyle put
-dup 57 /nineoldstyle put
-dup 58 /period put
-dup 59 /comma put
-dup 60 /less put
-dup 61 /slash put
-dup 62 /greater put
-dup 63 /star put
-dup 64 /partialdiff put
-dup 65 /A put
-dup 66 /B put
-dup 67 /C put
-dup 68 /D put
-dup 69 /E put
-dup 70 /F put
-dup 71 /G put
-dup 72 /H put
-dup 73 /I put
-dup 74 /J put
-dup 75 /K put
-dup 76 /L put
-dup 77 /M put
-dup 78 /N put
-dup 79 /O put
-dup 80 /P put
-dup 81 /Q put
-dup 82 /R put
-dup 83 /S put
-dup 84 /T put
-dup 85 /U put
-dup 86 /V put
-dup 87 /W put
-dup 88 /X put
-dup 89 /Y put
-dup 90 /Z put
-dup 91 /flat put
-dup 92 /natural put
-dup 93 /sharp put
-dup 94 /slurbelow put
-dup 95 /slurabove put
-dup 96 /lscript put
-dup 97 /a put
-dup 98 /b put
-dup 99 /c put
-dup 100 /d put
-dup 101 /e put
-dup 102 /f put
-dup 103 /g put
-dup 104 /h put
-dup 105 /i put
-dup 106 /j put
-dup 107 /k put
-dup 108 /l put
-dup 109 /m put
-dup 110 /n put
-dup 111 /o put
-dup 112 /p put
-dup 113 /q put
-dup 114 /r put
-dup 115 /s put
-dup 116 /t put
-dup 117 /u put
-dup 118 /v put
-dup 119 /w put
-dup 120 /x put
-dup 121 /y put
-dup 122 /z put
-dup 123 /dotlessi put
-dup 124 /dotlessj put
-dup 125 /weierstrass put
-dup 126 /vector put
-dup 127 /tie put
-dup 128 /psi put
-dup 160 /space put
-readonly def
-/FontBBox{11 -250 1241 750}readonly def
-/UniqueID 5087381 def
-currentdict end
-currentfile eexec
-
-80347982AB3942D930E069A70D0D48311D725E830D1C76FBA12E12486E989C9874C2B527F092
-5722787027F44470D484262C360CDFDDDF3657533A57BB16F73048BFBBFCB73A650484015441
-FDC837ADD94AC8FBD2022E3EC8F115D4B4BB7B7F15388F22CC6198EFE768BD9FCEB3446EE4A8
-DC27D6CD152485384EF5F59381FFDA43F2D20C8FB08AA27AB2015B774DB10DACFDCD33E60F17
-8C461553146AB427BDD7DA12534BA078AD3D780414930E72218B3075925CE1192F11FC8530FC
-D5E3038E3A6A6DB2DCFBAE3B4653E7E02730314E02B54A1E296D2BEF8A79411D9225DAD7B4E6
-D6F9CF0688B69BA21193BF1495807E7A1E67ED7E41CC25ACC04702F68EF703E3D45722C1A91F
-DEF7100A48631A02A6F02A08C6B1F9B4DF8310385B868632718FD87119A233F219D9411383B7
-FA9F3E4780D8C27E2E89E0CAE883D664C3EAC57A3AEF8988A2E9F0F8C7F53E0A80BDFC4620E2
-1287D0390E19753985447F3EA66401024BEA75E1B4C4437B7BB188F76F96B918AC7C6AD7E8AE
-7F21D8C2790F08CCCEC904FE48EF39E597ED4D4237C1D1F596F5906B19EA308020F7A35C168E
-327EC3246B1DFABE912F6B6DAAC09974876D3996E57D180261110DB05F15E3E8EEBBA3D90B57
-64C03DF3033A1ED678EBC679569A2FB297378B25434C0F205313ECB8A952F07242D3EE731B0C
-DC086A4481178A3D65129C47C09B22E9C431E11B3747B94C26A757C38D06001798C6A568303D
-541385244B967D3B1786EDEAF65BB53C4C2FE75E4B1B15C2C78D930B4296C80F08BAD8601245
-1EDC8E9F0854C3B390A16E27B11B3D45A9F72EFF8BADED2242DC928A61685D79E09681C97425
-5B90A498614CF560FA5B1718981388268BA206A96989E6D0B5D485D9ACA5594AE67DD7B34D8A
-369ADB06647F8AFF8814D6D9CDC04A4835918E557174C5BEECA6D04D7F46A92178834CCB8306
-EFAC4D4035CC601E61697E492BA6E5CBC7F60843F9A2F5A5D4E544B79A8636E512C9CEC06824
-D25123FF28CD58B87CD0844E5616797257C5F4C4C4A733E9FEB40C089105A9865584BFEEB4D4
-C3EA5573B0381BC64DBD4223DE702D60DEA3E8767434E9546B4B9BC7982FA35897090DF90FAA
-9EC1188709A85504074B825124DB9D95D5D26C7F155890D75E0D6F94482D2BB1C6BF31316152
-1C3BE48CC4A8F06CC3C2B406995972A9BF6F57636E60164CFE24FAABCE38211D66E0BE881604
-8505934CC2FF08F09727B454AD12ACF21F9A09E45CB6D949C41F495EF289C405A0E36C8B3AA8
-B5ED4DBFB452472A5828667B1F898A89BAF75FBE08A31A58A36777559570E9D7C7AB812BE48F
-92F1B8FF6EAB7EF43B4A66296E4F105A95C602309F0E2223BBF1B4ABF2542AFBB99952DD80A0
-8D1719833E515B76EDB5E9A6D2D98542325B1DB66B909E2833420585B3C5B58362C8045AD7D2
-E0CA50F0A439A5B9D32490A25A7F92B69888C61696120A9A113219EE75865290973533384D2F
-F987C552902732E6384AFC03E91B547D0CEAB4E03C7EF841A69F5C707C40DDE858E5766EADDF
-D25561904297C5D3193E7681B038F03D41CC5D2715E98CA356785F6BAFC8CF731BE3ADD3F6D0
-EF40D4149A917481B1D4825A81CAD418B4E4D9FE42739D1E7CD4DEB93496DA8008472214C83B
-361E33E055B3A4E5608B5A7FAB39766A8F85233E8FDCE9C653967F40AC11C342918C2EEACB08
-F9EA3A3A602B8DE3555CB26A1A25783AEE6E9EB323F487D462069594CEE229E750A061FB88BA
-6ABBC3C9360E06627F866596FE8AAD90DA7648CD2F360A0F953235F6DA11DCE635A9E6566116
-F63620433DA40CC5C7D3EF8EA10CC1068CE26982EEBAE7BE2A65619DBA174CDEA64EC416C6DA
-9558CD51261CA2E39A62EF879581C495D4BD835B3C45AC771A571E0450C3E8964DFA366E8F0B
-6A82352E70D9AB9CCDAE2CAF0C5188254C4D1F64DA8948D8583DDC7BE7635DABDB0914D23C8A
-9DB588EE957D2D1C03C8F6A4BC076FE11DA57B1DC67032D30C3E9E9369B545A2717215EBE2ED
-DBC63DC364055A0247A3F185D78DB812686B302EF6598862AD87F194F644019C0C58D2DE9379
-7D85105D2305586A11107FBC75EB823303C10F9319854A8A6A257333E07063050964EAB5F828
-79FBA218D4601BC43E1F8FE446F6FDF4F9B811F9E396FF58E9F424B2569B90257029C4DA059D
-775CCA5791C236C5C61A107A8F792DE2FB4BB71D0576D5EC2FF716A151B06AC77E495A39B68E
-E61E2C1002CABCEFF81F7EED5B5D14D66EC08365491BBB4D7263F757110B7DE0D8869DCCEE22
-BF74BE388AB1E7D7F7E950A49E24AE6AE3C517066D046E0FD4CB43DB4051EFC835BEE6B69926
-7F369FF8866B0CD0A2F68072FB7EE03C9105C68A7A486AA2E2696C8FB756B9D2A8D9ED8C60AA
-A7FFF06951F0EAF8B75D028CA978915FD8747F31CA8D85F52CD6E61A65B8FD3251F341E21297
-42191B7502C29602ABCC8538A0E5577DCA79B28A320D8A74118FD1874535A56D96247C353516
-D540BF78F49397A63FD2BA217D34DFA234F88D94FC94797625D7AB9A2C9AA9B240672ED4B96C
-550A87536AD0419E275AEC350630626EBE8D5DDF03C79907671946DBF72AE29DB11E8E4B44F2
-1E14073C01252D1EE3A5FE6280AB85A4EEE353DE0A1D48D0249ADF8C68A3A6162E8EC9E71940
-41A03F2D770C605626C9C2F9103212CEE55CF90F436993D1FDACE3E3F710F13BA99E4F3D4D8D
-34D24DC49830F48D38996805B373E9A4829BBB235F87636EAB0F17EF0888DCBCD47FE0804EBE
-E7CEBB4BEFBAFE39D11A9D21D321C670408F4A531194CFA150AC56E57F4279E46F0135246884
-43E6AAF647E835F0CE2AB0A1CD0390A1FCBE0EF93876412A440550CA0878DCFEDABFE0C7F572
-7B072D479F301F7422A49833617D6551F4F3470DFA0C2E05B782B20450CD2F0FB73C776BD092
-78C6FC7701BD809C66E0BFCF2607DDE1AA7600642ED25ABBF6D9CA6BB87BA1D3EF300CE0DA34
-782CC5EF92F767A4A9CC1C6CEAB5D38C844700AB16F2DDFE7F5F3DB4046E437430BFB44E400B
-DAAB4BDBD0E2B0B3CB23BEA964A624C82D846D96A2654EB150C9C93F440E4FF6229A32F58581
-E6941089F0C79DF8FDE6B8898DDF0EB5DF554D1DD2F4D396E7F74840645477EDFA322FEDB111
-A034CE810C852ED2EE0ECA5CF5528E29502BADC59ECD9E13418E3EBD772EAB12970856576365
-9B74B8C0AFE379B0F4FF5BB6F716BC152D5A4B748FA49EB6C348DF6F32696C67BB57AE1F80EC
-97556D458C6584E6EC61F5FF0EA2BDE3FE0CE26F7DC3C4458A82DF77323B1B16E0EAFFA3F0F1
-35FED4C12814AB623D8FB605245EF2182167723547CB1A18A02E338CF0171A13CA0BF324C73F
-4959178F8B08FE6DD6843439CCC2B398D2510A6FEF4A9675DACA5A51A4525E86862369B15A06
-E3A30E409E0C4BCF0B4ADACFB46219EAD7468C90ED860EDCF90571EF7463316F554EF9D89678
-91E7A4162A5DC9D1316C5B7CA963796B426BBF811D7ED60DB55771A856C7880287CD2B7A2405
-BD247D5E1EF8C4A178B43D2A6C87793A6359424AEF64EAF1F824AFE9C09CEACF48BE262B5D37
-FAC147659EE99E89D1B6DBA8546DE7ECF61014E69C9AAC7615863FB0786A34248FDD95CAA267
-D621127221F899947304B7C86B580154B10E148D1284080FC54F4E7A4D65FB4D2A12D0C225E7
-559724FB219336203DDE4BBB19C420A57D1738A14D3943F94576CD4996405438384D98AF09BF
-A0E31FF77E004DE3C86C5125091182DB692AABB5BC40102BF76CFFC253FB70F664F6644985B3
-208500DD33803BB6325E7324B9EB28FEC94F94513CBED2AC61DF92537D43A378435DB4E23A68
-29730BB5C5D1E1A7B684E83FCD76E68AC15548220C162C53E15DED615CEB34AFD0692CE70702
-51A1DB9A5EA54499B196A3373CD566AE9A4EC04AA6010A6E1BDCC617F304E4CF358367ADE16D
-539EBD5B8B02D8F7E28D74923515005749575657A3024DF109E16C97E91566D49B22CD271214
-77D2A4CA33A7F322548B48D54DD4DBFB0FF956D361968A4ED2D5AB6631E9E07B0D20839F8688
-FC93D18094E35E92A4FED41615EA9E9DFFDB80847AD3DDB7A4FB1497F9DF269336CEFA90ADA0
-78D58BA571FE8D5452FE91819AD62C69283E6FA95BB7BE1DDEA3953E0539E1CAA642A2BE1B81
-6FE0012153135FAEACB51364429A6A41B5BC61AD021827AAF320CDAD485ED084FFFB5BE9FA8C
-B754FDDBBAC7987FEFA6468684D5C9B8FA5B36D4A6F86E759D94CDFFB2310ED8D60295253D22
-70E7F3E4E0611FBA50286857379C19D3FF452DD4EFD0E382FBDEB2A900A61CD15C25596225D5
-FA90D70B511655B7EEF3836FEB03C3AE86005482EB96DEE289D030C8BFE378AA4912E5F60BDD
-C99A771C61910D429A1F2FE0E921D921577F9D5AE558B74FB6A57D8AAC9416F084DABC82015E
-CFB2D4DB1C409796688D05BD1B2B284F4AFD9FAE9ABC2E7E93CE7A555156522D4EC7EC1752DE
-571ABC44908C904348CD87967E6A417C87E1A4F702BA8D77472D7B11309F86BE5AB259F6D235
-81DDEB581E423A7551779042915317E49E2FDBE8612EBD6D27D442A39968D3BD2F1E4597E84E
-DE951B87FDB5C7655780FFB444F31D117D20BD47FD6C3F4290ED6F2B269448C2385A615559CB
-DFBC3AFC9E12E6FD01C357342A6E502B93FC34D3218F13CB8D22ADBF840B2E16A009B72EA153
-87E0F80083E6876486500A0B9F7275E8733F3AAB41E94A8628C35C52A62262FD22A897FEA603
-5B7A9B507307ECB9E2CF686BCD3D29EA3B8151D458EE68B9D5D93214F7B51AB88ADEBAA044C9
-7A15682067AD2FFF8445003917810BBC03508F0792540C47F5E7FC7C6C731FD1BB057CF8C403
-9A688A304C1ABF73FBE44F9D1B57B753C2D0AEEBDE041183120C72E048266736E2F109212933
-36EA923191F0D6E6F2ECF5B35FCEB54F3442CB501A4385088FF0D22CE5DA5FFCD7DE3402C8F0
-4F81A5B9B229136FDF0ED82DC5A46BE0B59228F129EDC99FED0304F618C8E69394C3FE1A8299
-CE353B2085B72730B679FA7733783D312F9C84776CD5A604C26C975163AD88E86190514F3024
-7D19EC2823DCCD48E4A1306B5FB1FD010FEB72375527579719B33AB727DF06C17CCA68C8A8B4
-30195AF9692DEEE69114F8374A0903CCF2F5007F7ED10AB95C33073D1C039D0677B2192397F4
-48A39B860587A3235E27B66EB84D6DED90F51A11E936F450CE06C3A19CB5294BDFC62A55E563
-58E0B24AB98E00365A27761DBC4BCAD98845CC15586E8C77256BF7AB99D6C9B1FEDF664F06CE
-D9AB97991F3D9F6001E3DEF98A2982BC9F1422E283686EE8C1C2812DF4527A8AAF122797E8A9
-972EEB7BFC8BA7E023DB7E075C6FB09D728C004147BFDB5E63232BBF4645E1BDE2F43D4BCE1E
-018E689A8700264BBBEE13C53CA085EAD5ED92D39083576DFF74E298A35BC9CA7BBF41B3E9C7
-3C5DC8044558B10068880F4A7D4E3670E4A81B6592218AF404DA284A46B2CBA0D3B00EE70386
-CD34A68767F11A81FF3F600B810278B529890FEEF6D59CBE45364A35AD615453D351A78B3470
-DEB26B13F0D64FAE5EEE32BEE4E55DA71699A69A709BE49A65A45B969DBDBE1A748FBF2455E4
-C814D4C36FD4CBF83F566B801A860EB3F9E458862A2FA7AF339B555B95B8AD6833EF87BE35EB
-0986F5D3AA52E7F5E90C4F6C4FFF01585E76F317795D1EC969693F06E7ACAFB589263220B207
-A7E67FC72F2B01BD86E88B0F75B13601FD5D9B0E8E12732B3C134742BA47872E61B768B01A39
-2113A78D10F16AC123F47A17FCE7C8AB707C34FE72E917B9B88E4FC3E80AB968A368084ED158
-3DBB47E6D57EFD410090CDD7925FF5EB64548A158025B0145E605EF7BB850C646BF41826E4B3
-D4E1A16AD19AD24ED1333DC609084FCFB0171E11FCB2F301F4492011624644D478E39FFA178B
-58E715F7C4496BB93F3F84B732557371CDA8BC73E7AE1FF946E3B7337482A3281B759542F683
-13BCA915DE865063634A4034BFFC5D1790802260EC845FE0C95F0E368982C1D88D3B5C9089E1
-96A9BB3D01D3B6D3F4222429C60656958F61B54469CF4FD2FD20D6952004BC84572B553CC5F7
-D7CFF365EFC34D82C9E2D52462F57B636C17C629DB4693DD62E49B3F3093031F7ED866B4779A
-CEC4B73771AE02407911FEBF0EE3406EB7941F6ED0F386A87439C8CC2F136B1A1462747F8ACA
-CA28196976D851C057A8131CCBE9236154032BE1C8E63CD5CBCB0DC424B2D79A11452BEC5CD7
-A1ADFA1D737F05029EB9182F4504DFE420932C527FC2A7C940CCEEA8923442307144C255F34D
-C682E3D4734426DDAD8F587CD6BF17DED8BBB8406FF187AD82A1191D545C7912157159AC80E4
-3037935E6FDA1826225AC37742B09725F982E8CA4F91DE4227DE4D4A8BE845F9F876A80B8B9C
-FCA58DAC05CAB2F3B98833046813F28EFBF9CFB92A6ECA4F514549C9923E3EF18CBCCA48ED8A
-266B748AFF1C8033D5A7CCB6964626EADEC47C4950216DF0D32C3B157D3B13BA03CC24B8D3CA
-31BB9480D021E7C164EF78FAD1C807EED0DEFA0A8B7856C63F83BEBEC836C4506F21F3C9371B
-C71ED43DB27B743C7BF51D461703534034E5181AF2A251E6C7E7FE9F2BC7AA0B50B7D90EB5D5
-2CC5100F1AE7458BC010D6073FB4E8F0377774F23F70FF332242B421225A42566B28E9EB49C5
-E174F60CF8D46D8B783C62A0D9EDE612691D423A2F4288C22F9F47720BEBC56B6969C3303ED3
-9733C7F4071B2A0E5D9BB9F1D51F723DCF1FD80F38F0732628AD360DD96F22B805AB8F4F2F5D
-C0EA3DF0C38795A9B3DD22E87702C8FE3CABC9AF906E3C1786E76E12B11EB11E4B8EFD728067
-B71E2B1ED858AD9949C3FAFF2D8B0A944586FF6070B890B2213FB3D69EDCC0B9AE11AAC67DA6
-475218638FF665371772E73A8282B949433C4BA3013E287FAB279718393447C03D23BFB6FDDC
-1832F11FB13AA4CD0B62A0E70B6AC299D1EC6326AD9C128C056F3B98F22266FC32985529842C
-2B977239CD48A75E5C10A4EED841D23FF09418D6F3A848A612D3C91475CE701D0118D682F28F
-3CFB34DB0ECA9772B245E0D6C10CF2A719FB2D76C2927E80ACA872088D41421511F16623075E
-F62DB5CE45782AEC0F7EF7884014E96B20CE71EF55764E7F9FD3CA65E497F60FEEC943323FF3
-F6C3962B7AD75EE500F0C64B71A3147EE7572B5D3EAC0167BE27C97227052D16B4A64D208A8E
-701B1C2DC0A475A1DA48B5A4114CC4C5D93C077C7B4B676F48650B882ADF8633510745ADF684
-1123782E5A7CE8AAF5BC177D3B3FE87C7E8029ABB206B643B7432C254A29F3910A447CFE95F5
-73B2FCA76E94A9C28481DF2269E230E7F6AD57309C72BFD7EB6B96DD3A7310BB24504A9EE881
-95D04AC6CF82CC23CD0E9E5529D9182548E92BC8CA7AD6FEB493BB9D1A71384DF6FF62D7B238
-0C0AED3AB924E7BFBD5C2B2257B86D300CCF2298636EA80D291807BCC7185D6A33B85FC489B2
-7A94AD38EFA6B9727D154819274E8EB82DF8DC8FAA9893E5B045398E2F08B0C59785F1046002
-5EC4E1D6B5B0A138CA4294402950B39617D46083FEA715534644924F52914F0404A79788CDAC
-63787EA1AA66712ACCBD850224CB6312562BC13DDB603C2C0479DBC2FB96D10C27E82808E9C8
-040F5E770AC58D3B70F4CCBEDFFC59AE08D45B455A621F55DB785DE6BD35D9A86CF4F5544413
-EB4AE9154811905CD798D3A75BC974FC8693A8E62ED3C8C8BA1B80021F8D0A8324892B5D75D1
-5284DCF1EB045F85D769263141295D0E7F72ECCBD49CCEF183FCE7CF9F2169E6A0FA9D19BC0B
-F5E2EDF47856E60D26C5C83EF5F22E3A55C1DD1A328EC29085ADBA834006BCDBF94EC2BAD3A5
-57336C6CC39B022BC8C570DEFDE3DCC7DD0C77FB143708E6D8AFA52D319030068FBB9BBBC2B5
-8DF2359D1BC6AC1019FB90FA5884B149C29B3E334F5D0DAF7A31618DD3752BE1E5057D0B8C48
-A9F51F3D20367FFDC9235591E22A7F45CD7D021D7732087B9F33C03A2BEAD60C6D4E49BFD322
-3A0368D033291783BD6EE1F59D56DE4BE17E19D472F8DCA1A9841E58F25AD719CEDD8DE4C17E
-F9D894B6D8F10AB53933DB7A753A69230E1C9D58D3CEE639ECED76EAE6EFC7F87AEDEB09B1AA
-25FABF51B12EFB77D6D827015FD30B54802942212E7828D7E9BFE44044EC7D2C63D77A80EF8C
-A1AE3C02EAF7868CC02CB5CF1BA1E1388248D1336E0DE78E2FBA09945215334C1DCCB3A84244
-B0918CB778FFD40164BBBA7782D48DC22F27C8A097DA3A311D819416FAD2ABA3F821060ADB38
-ABF35E6A59699A368349C8367F152676CFA0725CE664BF0225BB0E703831F4B51E0585FB5792
-F8E22586565F24F0BB10FD5DA91F997E79515E19C45F3E3E850769DB8500DB69B105A192262E
-B4441D7367BF02845A98A25C5A7C37BAF62E5891F87880EA5953B43AC809F5170E2E987D97C6
-0CB0F0917D30FA44250861D6C267734A325F4FEBAB0B0BD59ECAEA392140853C2F2986A698AC
-8BA219DAC15600A13E0A2A09AC31E44CDBD10D344C506E7F7AA02E5D82B9E669D57997773907
-BB6B277ACC6D132FE8660B193A8F7660735215C519EE693FF15C04155A483B6C63D0D73618C8
-5073628968D62619B53A69752D02529FF4009AC6865B9200FB6D9557664C0A564553F488DE5D
-309965D9AA2DD24CE41E0763350CA0D17EE334EB854BE53E5A7966743745064C2DEB4A25C882
-1C9DEB1EEBA5EE8A323D24C2836A25742DD2D62ECE9208124B19B67CA3834F29424F9B5A2CA4
-226707DB2DE09D34891E78BC5FADE7511BA4ECB683BA15CAC1630D18382A8DB1D49612B1C94D
-2DB3382A4294C2047CFE4A6BE84ACE0163F89545C2739E350C40B30AD4FA0F05B9E428C02787
-835E79FFAA2C4F950BE745BCE8199CB81B1AC3D896401AD6079BCCDD15649CC36FE2CE12DC1B
-5F6EB780367CA8F03035E6352D860D8E549231CD47B7673F5EC5C965ABB77A4E1268D65B8AA8
-497AE3F4E160E917536FDB66B6AE23E573F26A000B77B613EFFB4FA974B6921CDC961CFB17E7
-18FB3F6D825A3A2FFA4D3402F573F807BFAD5B3E857CA389007F92DF287EAA3F669B0DC27EC2
-88FAE57E73B4FB679A7E63D362A9144256265897ADB557DE0BC239E96264158A047B333F04AB
-16623124047DEAAC228BC447DDE6CC14E5058B752FEF2BCD9167A0756CD1631E9A3B2C644ED6
-43E711E19CF087D859BC3AC787C62BD1BB94951F43297762A326BAFA18D1290B13B8FD430BF6
-F8526A6D707B3B2943BCBD3C0424EA8CE09341C0DC0066C84B3D1EA3554BAA88EF1880C762D9
-A20A8F72CCB6918BBBA03FB2B0AC4D01E60067F5DAAE62906862863752BBE67C214C909E0330
-24AC975C268E69BB9FE2A0E54136AC5F8E0EBBCB4B8BABBA3D70A93DE665B2BE117B977C6116
-6812432EA60C682917E550211BC5F6B700BF6AFC734AB213A2879B9CDF14B2ADAEC7E4314645
-3B596E49CEE114947C7E53536778A9A6EB898990DCA8B43B3AA36828EAAFCD7A9BA47DF443F0
-F942FA5235D22EA46D1F1B41911BC57885500E538D6248E0FB394745E5CAD647E1B621DD6BD1
-6562AAC964656A716B03276623DC74658AFEBC0737756CC528EED19CC6FFFA53503E8FD53C89
-8A1B6D66E43FDB516A9A7ABA0C59CF702A6A1FD1EEFE975F49756ACB34421845571AE257EACD
-BF984E0B530690D8C4CAB01377C550B3FE5F184D8B6D3668BEE13567A705F0830F0A578CD325
-04CAA6CD41680B328EB975DB45D512FE6C67BDE90A28ED1F049960E7159540B5F957340DBE56
-AACFDF6E110D183865A98F1DE74B6DFE58327A916BCC7AB61838444368E3D96FDAF2DCF1F3EC
-34A6CC506F7F4BAF613E4A719BE52899BAAFE2E40E0106C8769B5B194C34D4102078532C4F1F
-A0EB8F7434B6C5F7A70989448A22E1B6D5C06E081BEBD6A513D6CF4E6A297CB820836259DF23
-DB8DFEA605F76CB0383CEE69BEAEFCA9CE935D7BF54FB37F03A4351B0C09CFB8A65119C06B0B
-35C2DDA6B043CFB82230BA082589A74862DDE3814FFD80490D6DA055C2A0ACC5AF84C08A9DE9
-D53C6F7B9029E9C40A32333528C31C61B1FD46D43D9321908DA6B583ACA876402ED4B406C67D
-C010B2D595D225107BDBAA2F3C2C86D4F8B6870E0197A3298DDA0E7C1B54B37F049CD2E3499B
-3968D132F109C53DC7A8EE8F11BD8F623C15E5F062FF7F7CAB6DA8E5A2BFC1EDA7351B8EA7D9
-9094040091419F30EECCB249301CBFBAE27A9949A05C8BC395E7B178D27954A4ED2E39208A5E
-228665E24DC8688112307217A062A3487804D7DF29050FEA32B0FE34921B9D90C9BFD5EF019E
-574BC40B56F900DED36F945BE6B72FA20ADDA25F1BA5392CEDDF4F248ADDBAA8E793D6666648
-F805C0ABCB39C6EEED68B6B502D2C2C980954C61A7990ADDD54AAD4C5188870248C520A3A5AA
-261972722E5CCC8275BDA179560DB290798A6E424539E2497804D21F13FC8F81BC51B287C3C4
-468645799060239ADCD518E0F4789BC5ED7E14560B26DB4EA65C7F03E3A793847A1719EF3C0E
-F973D617A026B1567CB6679ABCD84A74E77A544F8D25B020516E7CD9F17C4F5433612B190980
-A5DFEA6E7C1DE6162D988911A4AFB02FB1E8959C201AE6048E78E990358C4F28248A28ED1DDB
-1F6073C81284C62F9ED8363FB46354D37DF8B318828A0C655567F0649578958B08D60C0D311F
-D5DE83489C2D66127CB935F2DB8C3D274FAB71A5EF07521C9AA8F16FEA28E985B8398733CC0D
-4E7BBEAF994B27C5492B63B6C8060B7CE73B68D71915532ECB3D0EAF6D50298EF2BF24D5CF58
-8D93DC013B26D58927737CF449B627A9EB4FE3BF9B69D3E9E4007FE9599193F7B93A136758BF
-BE941DF7C3B276F65A17CDFC0B2C698E51B7BFA28695AFBF67BA9B38310E3E4D5D50D1EE1E86
-3A9E0A42C714B4EC267C36D1DFE4B8D39C871B39B1E87BAC57664443E5DAF5CA706B28BFF19A
-4E946E8F6C864FCB53D7CA529FC187729EB7124838DF5A0C303F14D1B7EFEBB41E07EE3240A5
-CAB89F90F75628D754FCF6113B9655B3A94ACB34C43AD5A2EABD9337CEF3476DB565BD402C0B
-6791CBFCA3F77C78D21170E737C20C1E73EF507C9BE065209CB01A64ECE64DE97F8C5273AA6B
-F65017B74634736BBF7C77CB492E619B34EF8B045DFFC086346674BE946946DEA4F546E51EE0
-1E3740A8303072EC22268B42DB1B256AFCDC29649F342A5BC66A3AA5B150D1AB1F4E2FCE7DB5
-0FBD9AA57547E89C9B222D0C72CA2A5CA9F78C496A31750FC3167CAF8D7DE7D6D2B7596E5E5F
-F48E206105510E9173A942E05F1260D65A4A2F7632CA5EA7D6453F11D6697538094CB74E3DF5
-1DB6DE131D95EC051A50C20054AD0BA315BE89E3FD558CFE2CCB53199E60408771D775019095
-6E10EB01B30F70DF48AEE8187DBA60AA229C4EA04B32B35C4C901267A6AE8FB8065014B08D09
-CFBA4E65ACB1208D347A17CD176A8B48E9D748B80DD87461AA1195C993AD6A177F41A10094CE
-DA6F972984306C28553CA9CA94E48A7F3F8BE5930083373DDF9ABEBD27297C7BE851E3E32B39
-D3DECB3E51EAAB500A334CA59EFB52D367AEB4FAAFF01DBD7626896A07AF2591EA7388C71EFE
-AB3EAADCF313BED3CCB1BE40F46E925D8C02FBB7751A91BD77017233F4BCDD41DDE48B6E652C
-875728304F5D877A150716EF18308135AFF04EFC69203E5757E0F20A2CFABFE3AC3138AE2C1F
-442624E90A5EEA57C5A7CFD0AEB415D2FABF90B435304458D8C160AB9DBFDAD14416770746A7
-54B1F76BB078B3DEF549AABD51B82CEE9A6B9E097EBEB8E5C99DFDF9CAE025B408DAB4497290
-B29AA98014AF1B8C70165844D41A6745380708A5978F22D0DA3F13D79B57BFCB323E30234665
-568A41C914C9F1C375276CAC0CBDEB7B2AE2C00891EA560125C57F378EA38738F1670D7105D7
-59515D7955345B3246D27158184C8243319F0182CD8F449F61805B9BFEA54331329665B44E25
-5E5FB491A11627CC2D3308857764FC090631107B00F125044367DC5ED8D5BFF20747B46D187E
-46DDF45564B6693D7EE2EDB79CC42CFB98AD11D3A9A3E17D6326AB8F35D6E0F80CBB9AE8C476
-56C297554AE9C8081612F82A74E2975F6A0B99920009D6F943F37674CE32948E9A3E8B66D895
-EDADDF6063F722249999DB33A48A9113E150DB6A6097CC93CFD555BAAC5AEF635DB1F10A1B82
-280E203B23B58D8611A82F64A2E85105A609C0D5698AAB9C0096D68D3B5FE15374690D484050
-2DF00D1E81DF1B00C04738BC3EBDEDBC9E52D813B2FB0C122A896BFB8ACD8E9CF5BD9BEEC5F7
-499366991F2A19C85837EECCE018DC27B0DA9CA5345503CA5A2421A8BC50612ECEA06CD37B9C
-4C05FF3187CEFE42239A8083952E27A9A68BF8E3127B273EA563BBE9C171094F9649ABC55764
-8B99E59C4672B1BAD8C69AC94E8A83D85F7FE31B88D45282DBDF6CA6CA24CC8770529E54ECC7
-AB07B125FF4C5B5631131545EA7F9F8C3EA39C2BA6C9903D7E47BC78A10A6C8DD9F3C0769863
-BC65823D25F9E01D3D678C3DD4F6A9CB82FFBA54F07FFC2CF696526EE7C1FF53A24E64ED33D0
-49DD841FF6F9A73C6FC4602111654914FE93B4096A96045058F9369105DFC359444A6222EE09
-6FE3C23E7D5679F623F65152F30866934CAD07CB0E9EB2D664EA37BC809535370CC0CFD74C1F
-5088CC4578EB78BA7A72E387A5CD5CB6FE8FAEB9ADBEC6F1B4BC5628EF518F38410D75B00426
-66184E3BB0FD5ECE42B6308218A7B4623ABF288D62D1A5CFD954244A9770C649D0FBA10BF789
-AE4C8314EDF2E8977D601E1454A1EC77FD3E8AB18D07F6E2DD3405920D039ED79AF6A87858FB
-3112D32B0297D811E17B2BB9C56F8F80FECE68DCF7DEFD9A7E0061D3892BA1681F01CE71555E
-5374B9DFB4E3EF7BE21B5462B49B05E309791358BDCB9EF0E8267705677184503E72799FA2F2
-3C2201A0C551E9ACBA02144A597BCAF934E8F3DCF07AA22A12488DF558E83FB19A6ABFC79357
-82B99F3493EFF689A2DFD5928712CB0A2837B2DEFF51F075A5A7263DC61BB1696B7049C8E720
-12586846EDB4C38B93966786107F99542C7B2D25EA19A0C96C613A7B109513D4E893B46A2333
-78B6098E03AF4C94AC0DFB7C31A7555A90B04189A48C7D56814DE9DDB0468E67BA8E939891F7
-DF64F5514350D7B1EF6D58136C4BFBF5F7FA98805D09076EA100721CCAB3C2DAD6E747262977
-3D079DBE2CA257722389B97D57DB3A1031FEFEB169FF3770CBA9A1272D2D55B1EAD4A3AA4F15
-F624A12010D9EA3150C50663816B1C00F77E4CC52F53B288D3BAC9F52A6F596C05B0495C6FE2
-EB88660BD25DF4488DEBA049FD1CEB3FF9271E17E4ACC11E872166686611F8F60C7E7C01AFEB
-421AFB63554B9AF8FDB0C81E501E4AD4C792EC9A62EACDD75600577E996BB6826A7EC34913BD
-56CAD25466E9FE70181DDFF055EBF4CC32E83F2F606D20991A4539B75CDD6B77A9E2F041A71F
-01CD8876242961C04A6135490B448D17B40AC17BC6745779B684282454B0AF95F9094CF29FCE
-CF6FB82B6CBA5B3AADD1C1BB64795367DF32AE44C520F0AF5EEE7FF395E7E17C5874CAE4406E
-4F02EA95ED1E088EC71B1887FABA2A18F888B81B5238DC495CC589ABDE435634AAF641094C3A
-DB7431E22227D9472737AFBB5E834513597840F42136E219D93A799581CDAAB03CB7AA70C330
-F3E84FE21142B45C8E699604302B4E2C074A81267118D8FB4D44F3C9CEBEF9EE21C33058E915
-1167EEFAD9A040EF1556EC9BD48B8F195104E98056485C556FB9303090B473F7482D5C5F5F97
-B18AA96F49EE790FE506F0A9C6A3A2F5726CBF45BD5BDE0096A85B7055CC32A0CF7D9AC0663E
-96E6889A996FC4EDE1D8E9D868FF9EE20B9FDB100B4D00F2C06AB3851BB06B81EB88F55515A7
-EAA36C602ED3C02792C4697DFF25639DD76A2EE8B0DF0B67B17F7431FDD29091172F9AAA12CF
-EF673D31058CE6B8B5A5185C94033DEE05D4A8A511F3F019DF5F1376E945E3BC254F7781CB64
-0B95711C737681EEC2FEAF135C5B400817BBE3182FC4DC5AE537E85487F8C14182C92B96B5F3
-E80DBE1490A844DDC247C57196BD67B2C4EADCB895143D9F887037902621F5CA92B9B39D1140
-81D71054F19A9880E8EE9D3CDF69CF2FCC06D72A03000807FF109B6B3BDBA5DE94F64D6FECF6
-2E1FF5659E3D07770D9E18DED6064F71F1D606C088510B2437E2D917088696771E56ECAB1831
-54E9388EB8624B4AF70C7169AFC8F22AC394D87AE30BF1937981D586C33B3D901482A4ADA0D9
-AD3E8E53E27892AF58E3C184C6BC935247F0D8C202D10587CE3BC034908BAEC8B617D7BE6F3D
-DB6C94896EAB3697928FFA8B25522342003FE9E640C1AED684B19322BFDE889DE74A7A6CB3B8
-B1C0F19C98AB06C125BD3440ABB06A6AAC0F186988D7A77CC738174A0069D8B50119EB1FE2C9
-B6CC0EB0748B8734F71A480ADA4052620C202F0F61016A05A894B58F9FDFF03DE8E9F7C0957C
-C05D91C6AE73D9A4F22534274FFC325DC83FDE347C1543824E51B8A30B676EE6C36044B602CA
-9B7B83977D7187093FF98447B0D04B80B8A9565C14A5477ECF2DBEC2398DA5E6E4A2CF337F6F
-075CAA383A1FA18E8209E10C0F2CD9733E5AA5105A3A1F6DE8417E79BD5F425729340D0BF77F
-8CE558FF4899944D1BFDCFDE1AB12469941463B95A6CF186631B904DC1A929DAF48526979682
-96A838B98B657F6CED5E1251057330190F6222FE603F7783DEB7EE9E9CBC69A84AA938FCBE03
-27C76F1F349D8025558FEF5657B995CA19F8246DDAC948AF1D5D413CDF56B0BC5B5C473773EC
-9C577BBFB1DE635CE94114B5726CC122DCE433D05108A6B6E153E67E179D5724A0F8AED6F4FC
-F0817AEF63327742E3F6CC2829EF3D95F2F12D5B1F3ECECB99797A5948ACBE18B99A02FF678A
-8BD2D25EFA671D0697AFB65239225FDDD380C96A72826E29AB16D5F2C6AE687064D46D782454
-C561F5662D6C5CA7E3C2C968D0CE8828A8ECF1CD83DF8512451E8B48EB8DBA7B1E1BCAB27C20
-0111A0F2514F3637ADD070D6A18F96B5283AB06703F4A0ADC09CA1E76302C7255AF08F58C44A
-A42311037A0CC08FDF0D89B79CA1BA587795951E082E1054E4BF0D10BB048145ED6C01A45FEF
-A9D1F7BB38904DD70CFA6DD59D4F9E009867F6FE51FABDF1F51577F26441FF4C1BCFAFE7EECF
-C3D9FFC4FD7179732E54487359DFF3421E5F6E60480EBA68151B3EA918F9952C88A2898D0E8A
-FAF7AC95BF3509C404A8F8F64655201FB8FC24269FA6C3812B8BF01E06EECCE7A4B9516FE0EF
-14EAFB23EACAEF67781E24B9C1D6D521E9DA93613A824F709EA3B4D869F477C96264E214D5D6
-96BEB49E68DC6B5D3692F3F0433F677547188A10D54E8493B58BB0D976E2D611A2748DB42B92
-02CACAA04C3A000D131E4C8987F776DAAF46B03F65F363495011A2E3A0F4F0146CF8BE4484B4
-A56E77E1C4FABB76F092E930E454F4F55373AB71A173B3A00407D31AE22F9770AE07133684D7
-747E4D9DAB05B24A4642E81E3529C460DF534C943B89A773F7A28EAA77DF7FEB3AD2E2D357C5
-6DFE08F2059B714955DB8FF86F96BCBC66762138D3A0C5489826052CA1778A2F214B2182C229
-BEA5841B789080951423042041050E8AC580D89B830DAB0B45F50EFBFF8448220A8210121B1A
-538BD5D882D0E9880908E42FCB3D4730205835D00F9BA46BC774F8CD04EB4D9999722E15320F
-34FAA0FD3857D3D345517B94C3EF4E96DA81F5A678DE9BEA30DE1730CF37950C6E69DEB26F67
-52E62131BAA9535362034C1CB53711E1A2C9010CC4D572028EA0652A0E750FAB6B8D2480AFC8
-C20F78BA1FE4ECAA4A5121FE47D802AA80E9899902483F0C107C6F585C5D2895652F7ECE21ED
-57E509E258C174DE09D2B6CE9D6C51BC20B75020CA004BCF716B137E2E779D03ED807CCB3597
-DA0E838FD70EA9DA638A702B02878C8448DA37CE796899F7F2EF8ED530252448FEA81540807A
-AF1E4E65E3E13819101D384356B54C58FEE8B7CFFB5860085549C58820A4966C022E9F7BFA6B
-CD83D0F3ECE6B9AEAC85D84B5C4A324ECD208185C00124747CD18E28E3A4E7208217AA31DF1E
-B4887E60920F43795FA1C7E99B59B8178BB298F90E9710CC78E67569A2F3E1A4DE1F111413C1
-D9C999D09961678B5E90C59949B4FBD9104EC8D9036673BC5D11DD097ED300D66EE9AD98E6C3
-2C9F332DFDDBA3332568E1F4BA90F4167656E1C38B71756F9A9A0FFE011460E597D0E741C1BE
-4CF3E37A3673562054E70724F5CCE2977CC13DCE366F774D04B5B27DF6971A08088A7A09203A
-5EB7182B4EE4AC6B4D39B097B1A11019FE0138700DE0727590AB41FCE5C145809A141CCA6712
-784283C6A101EBC171DABD68649F1525F9576E1322C9468F2C5D78ABF1A3B7292ECF89888A98
-8219A761E9B05510DEE4F04CD71669629E401F164636F3116303D15E41B3A606259C040EF142
-146C0585A921AC01723FBC89D44FA21C2C1B0D18B0A41043B60F1CC161C70E7A1E68281A7987
-C02046E0E3C95C3235D30C6BF8D089C01C5C837315FFA54427072B9124332014C34940BB1A69
-552CE3D716CA5CAE0D260119D7EA88C8EAB5FF55B26F5BAFD9AE98F953DEF3798804019455A6
-318FE37A0DD86A2FB61C86B5E791385F99942027C6EEC9D4F6C53206AFCCD6860F288E5A7497
-0797DD7C0D4A0D9BA3FFA40AF699B6BC825FDFC3EF8EF74B53CD3C9A2BF7472A11A86189AC48
-40FA39D55E1AD6D4C49EBEEEBAC2F5C13FBDA719B961C2E86E7CE5EC0A480EFF51437ECB0ADB
-46C939F0343DE93749554085AF866AF9AAE9194089F3F465ACB40ECECA85DBFC0FC2470FBCA0
-0A0CEFF682598564630AB9B631D49F0A4D94F358FCC3769C92F915B5E50D22418DB17E12B21B
-113DDF64CDE6B924DC01A99C8F2003A4C1B9E0D0DDA2453460EA11CA7953CA90848EE04C50C7
-8974AD56ECD5011A856F131F5E918CFBE9D68F6C81D8F7912FF9986CBE7D2E4BA7D26504AF41
-BB53CE07E3A983AF069FB3514DAA2755F47A06DD668B0D91F5A26B0C140495905A3DE0AB00B7
-3D932C53352225FBBA38CE398663B3F7582F22D87504C0E5B4A9EBF2DDD77A46731146D57E07
-1F0A5AAF19284BB3B74DE2843406E4C536B8142DAB8F84406A1A03F5E0EF1057D7B39CC2A7BE
-3EFFDF14340469327D2BB2B8350A8F8FE8333A2710E2442A6C8EAA022773CB141C35C479944C
-9AC0A1DF9BA341C7661C9688337CFECDB641900A80E415401FE7A34C224521BFF05E5ECE7F92
-478563813AACD315B65A1ABD00E6663CADCE921B4E23F01EC1820017085C4557019BD2A4D7E1
-4CC924CA286B7DAD8AF312AC2344C24B29A221758E0F58307D7BC6521B0FD951CC9370CC5848
-ED0DF1CCBB2CEE3A8B4A706B63065BC8F4269E70A95143C20DF920CF12E35809E749ED8C3BB0
-FCD49DBA7C88670A080E3C366432A69EF37835291B73DC65ABC6DC95FFB0FFDB8F2843CF31F3
-2B00508F69C3C620D88E4A8112E4074A7A5C0DBCDF27D8DDDC4E43AF437F87FE75FD350CB7C7
-B0BB268999A6FC35DE36E9FAC5100721909EF026D4FE2A0B8771AEB3C6999F968C11E3F086DB
-F1996EC272F550C7B376D7C72F1099F47D2A8438EFD5B1A7D71236B14875901DB66A741174A9
-892F09CD664FA8E1E4C8CD55CD39ECE4CD581223737D114CDE375EB5DFF9154795D3A45CF788
-491FA9740BB15A4B936FF87E574C61AD755E7055E0E32B0EA12041213FB397F7E557296C90A4
-C20ED66D10BCCC0DDDBC4CAD0503214707E3D0A07459E2DEEEB20E65138521B1A6FC5A78E3D7
-1D511E5A8F0A6DF4C7BB9B6A83814BA34B48EFED99CFB7231D70FD9B20246552CE4A4AC31291
-1A1F98847026C08A32D076E57A8A3C8CD129824575B8AC69D9AF83DFFC44362E92756C477FB7
-A28159758AF4518A74607FEB5DA77837217D0133FFB1C0C61465E42D120F2DF2857FD9F7C4FC
-0472D0E201AB4CC716E103DE39DE4997A1A3EFEF1C815D7F9C595CF4F38DFD0AF392961EDA38
-2508925AC51A5B399B29F885DAED0BD03CB0F687B568E508B6CA15EA611912313122AD80333B
-3D39817D365FCFF4461140E31FECA21B380FBEBCFA7CF1DD360A370B4A9A0EAA69BC8677B3A7
-41F404DADDAC15BD96EEAAFABF66875B89781BDD85FB3419E973966341D0D077062BD1797B2C
-8C4A4F23E8DAD9B1B0E78D38AA2EA100FA4299FC877F712B85348B01181985A4F9B6C7F9E7D9
-6C177F9C84EBA9C76CB45009ECD038571B19E6291EA243ADACF2C43CE513ED1640F6B8ECD6F8
-ECBDA1DFA6B80142A3BEAC03F0C835BB19E2E6DDB190A636353709A14696A7749E6AEAA923E8
-E44E4D79FA5078FF32E5967E4C1129CD03A7AFB008371143A9650056DEC977F9700A6E34F66B
-F55AFA30C9F7E4C15286E6A39A0F8534B3ED1B728871699AACA31B87F969374386DC67723748
-D79DDCDECDBD1959A9C88A31AA050DD5D80E49610EDDD88BC7DDE7CBC0BABAC3DF911DBAAAE0
-46FF8665516A0221507C3ABB2BC7352267F211BCC04C43E4861ECF8F12AD85515FC7CFDB6C84
-C01EE3BA2D9C1E1C11ABC24F8B0AE9367A7447DD7352E99B07355B571A1E2FA1BE11C5A4AEEA
-C975250D6CA40F2BDBD13FB6DC511427D1B783ECE8A587511B2B353E44351BC0FE868B1B90BB
-2DAF1F83935B76545FB2D951A919A4A8B5D1D30703246F8B31870FB2DED22858C11A81A8FC9E
-195C9361CC6E0C7F0A6BD2321E8D7D52101052D95A7913E9298A6678A95EB956ACDC6312D26D
-FCD9F991497C9D238382234A0799F4EADA8B2038E71496EFE9B88FF839A3823A6B5CFD1F0F6E
-E9FAAD07A8D82051B09F8C8409A8C7F1C6F725E47272752A270AD5961527C05758C6F3BEA49C
-8624DFF57C0F11359EC1D43307E0F40A33D8B6B2A9D57472EBB495B047C894F170F23CACDEA9
-915385CB4258144C06BFF7FA3712C45EA2E0EE1122E7CB644ABBD83A3497360A074A72D1AAE0
-E76D6565430E3101DD90367F87DC5761511C05E88DAF8C666AB08A5889B1DD00D02259CD57FA
-A9B8C58EB5A85537F6267A74A9279212AAB2018149FF2A0E291F8F5F77054D067C9264D5C227
-A823026770114476B240FE1D6702023265F95DE76BDCEBC522A7E4C1048AFEFC8FE904A7091A
-703F0BAD4C5F97C4C28FA2859525C4A6A999ADCD39093BBBB91D5C9880A2B4674558E64A8267
-C53FCC40D4F310562737637C72F6ED2A171E9502CC249CA5576C0E3716A84C088E8E3A4FAFDD
-FE56E1023C63B4F46CE54777B7AECAAB9BB939E0F5F8B1229AC60ABA460E86F6D26E05E150B1
-741915AA0EE489D656E665844BAC393ADCE9CEFAE5BDCF54AC75ACF52A75EE006B504351EC4B
-BCECCC4124C0D6C3C9ADC16B20700F8C789ED13323A83B4D3B3D0BE8490A3065AA6B472A7795
-B25E16749992726226B0F76227A7F00547CDC2FF21B3FBF1F328E2333745DEFCB9C26675CED8
-2DB7630169951536A1E9C51E928A1C7D7CEDF6AA82121969A462731B4474578FEFA779961D56
-BD962F7615871FCEFB1B6E1B37771C755E3002D96BFFAE0AA34D87C5BA182EED9E5C90D27D5F
-7572CF0DA89E8BE4239AC9F1AC3C56B348A27AF7E834906127617732D708CB1252EF119BF6A1
-4EAB7502D2D8A5358B81EE20345575F815781E373E98169016B9FBE5D940CD7693CD19924BEC
-9FB27463F06699422FD7FB6299A673CA40750EAE58064B1B61C5A45CF13D616A041AD396B3EF
-D04323805DBFAE5818677E28C6CA91235044FA765DBA93D5F0EFB8D0E9D5722ABC6BDE296091
-750E2B568178DD452680E459E7E990F6AFA86D356F8C09A81E70652F428016C8537BB6EACA70
-BF4D51763DB9A5E7E84CB946E4595579B3B0724EED0845FA14419FBA3508C4522B49AF595DBB
-41FED045A236A183AB2D8311977A661BF10FD02E689F34787318428CE9F91B4EC76E85289023
-AE0FEB08A83C479E28C906AAF3338B4ED0C5A81EEBCADE06E0D2E3365414F46BBB6E54C4F9D5
-1D4DED3D9A66487457F3105FA8B4B1A3E1617A5A63C1863ED3C77F1BC2FF8C16E867FD4C3BC8
-E0DCF670BAF2D216E787E8587422F66EA25761266626B08BFF4C7389C074BB1785388B06E55E
-749C73F0E0D0736E16C756BE0E20704F3AFE33B0AD8B09FBA5BF9EB1F37900541BEE65657B29
-4E7D60894F7728BEA4D40B1B7D44A25302FAF6DA2143598744E26700350C6A681DCD3412920E
-2EF2E2B2EF62DC2C9BF590AB82F0EE6385827ACDDF1AD34C6C5057A6273F7B5A5EE85A501756
-AD39DE32CB1644DDE997CDF179D38A639692568E0BD48B5835FB37C21AC55D9C42DE2872EDEC
-DE1AB1C137196ABFC7147E3C01C6CA32A2F0C0FBC7B37ACE0BA1244DF378E931F93068D02547
-CBD6DD46431732B182C71A70CBD5A2B59B25C4067407757AD8A044332EB2A9A3B2980B1BE8A6
-EEC67335D07030586314D0450BBA630C27B4A2E14457AEA350AA22210457915E698E84F667CE
-4A16B5AAFB1EABF4C7FFEF9860A8ED5C3E2574275FBBABCEE3AD6A343E8F7335CC9E95CD9A91
-103D0011B9FCA006AC82C389A2DA3EF2BC7B2BCA665AE4799522849CA69E54E708CC74E5ECC0
-A06D1B348178F358BF5798FBB3AC105D09AB732E2F11748C32288C90F90561481AF74DAFC524
-25CDBE97E2097DAA663FE479B9DB5174093FAC40FD2C26738E68FC4B5F27E7DD4150D2C3D8AA
-AEB9C57D3D35A3810AFAEDFB94988D633E1989F34A880BB7D102ACAF099D1E37828824230BD7
-BE130095B1D69A616481A904B9A84A3E4C87DBDFE051E3544635A3F51AB632AB1D151C0284BE
-71C1A0FC23348B49CB66457F2CCD7DB672A529C46990017FC66E19667C3CED5008E15B7D6B1A
-550C4FA3D3FF50283CF9F9669866F47B00012E30EADD2CF8E9BF40E6B75335987080BB813557
-A82F786906EE527EC6F3964EC3112C8C4ED4F5B0256BDEF7F6EA44D2FD8E7A2682C9CAF7B34C
-9B3DB56CC1E314A7049DE9BF674B7565A6A98E1E2A95D18D8836B09F1F18A89F8500D09078E9
-091CC22A35C8F949ED16433BBA123CF427C6559B3F9F18855909DFD4E0019F5682773C1DE28E
-55562A5EE1A6D8DDA2C1E681E9CF58129084A509AA69D5410C8A0925C44350D920243D3CA503
-0BD7174ADC330EE71F81A2A31B437F392785D9F25CC118E4C58FB9C1902C69E05FEADA918AC4
-DC12A4EF9B1AE0EDF33582DD4829FA8808263A945FF7B989553F20BDDDC758A2D2A97FFAE02F
-2BA2B29AEBD676FCD308315DBEA5C7B75296FAB66FBD4F64C4AD66FC66F25902447932C05AB8
-B8D7B0044B1214F300926AEB93D490C57777E286F45E4C087827AEC3A88236221765FE2CA39B
-B5DE181DC51683B15ECECFC7DE6356A6BDDD22D9A8242A67E091E16B9A2B689054ADD08D99DA
-D3981FE089966E739BF7326408902378C31D1BB51F00548C344A0CF5B5F51E6425C341307805
-5197DD67429A1B5101B5CF81ED04196211DFD849D88974833C3CA9469220047721F959831806
-9E3921321A3D997A9B045B36548079891967198152236FE58C1AE7F28B09604169F989967F62
-C7BFFCCFE4F8B01C0FA76AD04934A1891A244C1E1F4E8577C25625930E5406AEA0F8DFAFE958
-9BDBC8872757359789FB703DF8B4983FD643B05E075A8DB19A5A6550AECB593B843BB214EB3A
-721692594CE1C484F128E130F1BCD8C2CA0F7DAB8268590CF26943AA9E1314A79B2596A3E080
-4FC185C719DF92572E74B007315D2189C0D73A553486F1619A1C5CDBF77E79F63333A462BB16
-D538C640BEA05AED1CC00B0FA61B60A4EDD2C4F9DC8FD26594653655CAC0BEAD5C26BA8EEF6A
-241BB3094D590947BA27107AFD335AA9BD2980F86902F7791A80E15F35C6675E5C3D2BCFC0A2
-8DB3B8414F64B7F6880AFD4A43583840058F742D1D3C76ECD2880B8FFF77ECC17145D3F7FC89
-3B32A19AF34958EFA7302B14AC2074EAD30EA5B502DA7ECC633ACEE0EC57F5B433E80AB06A1B
-6C3E5E63488A3D9DE39F2BC5BC5F09316ABDD75D7C293CE3067E7498CE4A7996CA4487164555
-05DFB63F281533FCDC04A786610EF22C7C3096279597A426D5FD6DB719FED05290A092FC0C2D
-05A4D8C95354ED08EDF77B8DF872359FC121B83E600CAFF5D850F4DA62A66CF105B1CC724872
-88154CA6D58C5BABC2A6AB9A898D7D022A47005EB11DD822077C6630BDE9FC1420B2B10A9496
-8ADA17398CC96AC93A86FF63FF64C9C774571B50E9A35D6920772B1C3B351295569FAA16F3C1
-FA689A26EEE4BFEE08A858F93FC95392C4707CDA935C1A39E6705602DB2CD0B11FC233E34D74
-7266AB8689C77A59F0E5B355EAB8A42FA07706D794C4B1CE6A375DD83FD70F1398A08F0E42E2
-4295996F54045B47D435FE601400AEDF678968EAFC64E652A670C4476C1324EE57AE67BA63FD
-F88EC7E05035D5940E9706679ED9DE791606DBE417424CD1931F475D0E005479141D447A9697
-AF0CF8BA450A9E9BC8E3DFB1EDEA01EA1BA738EE58B9E7A1E842124842C2A6A972981968DE49
-CD0A551F74FA76AA842CFC2A9E130D8A4F9D8CFE99310BAE082A467E9DE6C075595A3DC43727
-3832A248648DC2679AF3A8B10E55000FB78A20ED6E1063179F790F6CD51B117D4897777A8D2F
-3F9CE9D51CA3E69D3EDD1498D89D45D7F4350298205386BDA893CA2AF84A1CE0AFC770BE249D
-6F851AF8F4E484131D5CA8435C3B608EDFD83C7EEB0259A036DFB62C09B1DB70E759B81F14FA
-00B8F18253007826A61E8A73E79C85713AB391301D965CD6BEC909FF77A253A42D14B7854309
-1BBF86DD9233BFF8CF21B5A6B66A98F2F515030133B9A5BAF23AD19DC4D16CA01F36010B2E6E
-180D6AC6B76ED4628F219C44577280452480027E55E134257CDC9049C813AE93EC848426B890
-738D8286B664DD77EEBBD29C43E2D7C892C54EA36B9C81C8A72AF0D92C095B03DEDF5B6D563A
-8F9645D2FE6514664150FAF8BFE113A91E414D7F554F041206298201F6430EB56069BD0BCFC3
-D2B2938D75BCBB328AA87F3271B73CBD878E3B4028054370378475942C592D8618D8DD58C8FF
-DEECF2E08927592965E00E895EB2687C67DC3A07D1E8C6B8F27AE73007D6AF3D83F16719B6CB
-01F4B57B54D052B3F79EC8E4AFD01FE8F246E9D7DB2E6FB1F2A681733DC11E1F5943758C345F
-93A043263FC97E7DD8746B2739E340AB3AA2B54AA86241E5F3CE218531965BE76FB68F734724
-3A4C9DD808EFC7A5866408F61241036A233499553C5EB871F5B78C3D1688B615B8AD295FD31C
-F1D03126D9CA5F067EB92AD55E203EFC39C62C33F4B8D3029F57094611150FB3E3466CBDAAB7
-3702C666BAC042A6A5CE2CF1C7C164AE75DDB23C62CC24198AF6BF6E667D7B64514B3007911D
-AC202771D141AB4A1D5097AEDFF137C58E81387182E0D58B19B583D2F76D1BF6327FFE1CE792
-8F515272E9E93A08F895B202FBD293A7DC6E82FBAC9565986536E58A4DCEBDCACACFC4F2A5CF
-0F5E589BF96794A3758E5E6FB86647AEDC686221102FFCF4DDB5EA5D0DDADFD83CF83CA1A56D
-F196430539A0169651E947AD6146A689F819041CFCBCAB5BCE4F50BF006F0EA18CC8C9925239
-BE6C310FBD160725264234CABB3F2F4FE81FA534A21C920F0C4D5D4EA4EC2C50C3D89A5B7992
-6A6BEB87B72296088CB965BFD4C2DA4382482413D54BC6ECA7EEE44DABE98075FD821E35958E
-51B4F845EBCB80781CBD62444B419B2FC5F686A24047180E8AB6BD5219EBB776D3C405487307
-A9414246B86F4695C67EE5ED6813334C83076EE9882291F95A283D5ED245862258E8D121D7E7
-5727E69B0739E57DDC0C97EDBDF283115A32FB54526E82C3EA395255CE7E3862289D19B79428
-A9FD8A43701C00D0D9ED5A3C653BB47F9D30D4792023DC8A2D6A973F1DD69FBE954DFF1A3F83
-032E1BEB766BF80F729A98B68AB57C8E8E5ECEE9C12C2272E95A46AC06E6D59CDAB4CADA541E
-E87B4DA51C3386C9A385A54C1B41DBB81D083DA941C8CAF3E7917634BE5BC078B92CA2102962
-BA0D315E2D07BE034AEDCDB19290069B07844737B417528FC1CEFA49241CF6C8BEC39E4A589A
-F3902D5244300DF56489EC63B09AB0FDBB91B51609CB2079B5277437A8379A037F33E3E278AB
-81FE819C676F7E2BB1DFEAB982ABB98A4114108A04036FA1A7C140A4052889142633F981237A
-A00A7FB881029C36365FCCBB791FDFEF586D38A8CCD8A00BD20701C35CC5804E81B8B7EDA8A1
-768843B30F86286B1EC0A50AF308FFAAAD22587631AD221A36A38F7FB3B354E9AC7289084111
-4ABB90722EF6C425E648B849B5B14743B16DD9E62CA25E7D109D650BB7B3E43BC109AA7EACE4
-D2CA8E8F957C522CE561889F8C71E452F61AA01056E7B10C217D0B2719DBF502CED6118463E8
-C03932C223A9FFC82E0BE29692662F6D76DC42D8381D68C84A396F1E49E3F500F5767486B7CE
-6BC644B6DEAA0160014332B7E91205184083E6E086C416FB6F21EBB04964DED52C398068FD60
-49E2CA1EF0A802863066D5DD8AECF3CC3DC0BCF8C367B92BA88B24C393149F261C1C690BFBDA
-0FE8500617882E115DEDE487F795EC4C372753466EAA4BCDA268DFFCC171B9EC5FF352FAEF55
-AB4A644184C87AC82FA199C37D3413E5AADBF99937575AECCE5899D2A4FF5FB53FCE91F444D9
-74B53C18E362B6D327D9C79F30C5886F3272A00B1FE8113299B6013743234AF1EBB707085F1F
-40EF88A71BFE0CAA195753576A6F2FE131C4D24B237C320A48D3AB85EE7BF29DA5D2B678ABF9
-E3818EE2954BCF956057288F652ED5F42CE5DC09FA730D1820691CA2B97101B3EF588B3B6819
-3400F29D70540DD910DF4FBDC20B41B63E0034692CA5BFD63E95C2762BED8162BD691EE1287A
-CBA6871E881CB33E3EADB1AC8C1A34387DE8E816A3E81EFCB086D1B594A2EA6E54044A4AE220
-DD61C5D167C1BB2494234BB3ADE17C1138031ECB7CBE6B9B2381DF49484238912AEB5F7B385C
-1300EB13C2EE6DF2F7C32B46ED02A0488D428968A5F97C9E95E4FF7BE2091DF79EC79188DDAB
-F7D6B9AC911027248A1FFA3EFD8CF2C3D3D973C942120D390A6CCE4ADE888BCF393AA2C8FFE0
-BFDEE8D5CBEF16470FD1CDD78F5FDE80F253B156D037EC050B75C5F0943DEB7ED179B31183E2
-96CFC7D4EE82DB5D33E12CE52FBA0FB630A35C5BEC3D5A9A457AD8AB3F5851AFF19585817A43
-65F4DDB18EAF7DE8D2E5D27C8E6E8E5979C6521D57102088C13C3998895F2F13DE3213868FCF
-27B2C434636D6879A3B191B667471537FC1666ABD530EE64AF5930689E52717C8669D1E3F9A4
-4340DE7BA7E9C5FD4156E6874442C54E327AFB09D0778D16B83988BDE11F4FAA11473E65AD07
-B7EFC6D08201EB5C3E5EBE9C742D8ED6869D9C29923714C40649866EBB52A5FF0E5AFE799684
-8F238B155065B9D7410BB07297C74EB3A7E1378B11F68861A5B63655C72742B484FF7C170BCF
-809C96A005DE9C05E8919A9E37062A80101ACA472EAE91B80CB08DBCF6EFA68A7E6806D5F8C1
-17D6304DCB4F0C753F4C7DE92633025699B3F60C874A63C7F13A430B1F3CB6C573E5D8665936
-3680EE78BA6BEBA767A1FF21952603DC6459D214E127C8DC718363D07768C8970DEBFFFA4C81
-88E4D109FF21A80770E7C272E726DF9ABF6A4057B8F0827A60F72BBCA1C040F3C73B5A547986
-395AF437A1A5EB6DED1D2B62CFC3739A60BD5A66DBBD6A6289D3D6F890C6746587BA8979DB9D
-0CC9E19DFE54C0F1602E054BA170FC65EB7477CD49394069F6BFC38AF6A883829551BFBF6986
-70D78030CAD26377BEE6547E64CE8E4F02C0A4E2FE50AD03DBD6E5D9DE93D68D5D4224C39097
-EAF79755B78D5E7CF561DB1DC391170CEA8F2E68E00DD48B57B4905EC036E8D560A2BEBF18E0
-6DD28CF64BBF3E0CA2C0906822B655BF9C648EC063FEBCF253EF15996FC28F884EEFFB2D5930
-11A95F10127B4D0990CD32F93E1C56E23225DD2468F6DEC105B5AFC136C4E455F76CD0915F99
-E749C51C3075F70B5E93BD71DF3CB9A75980D2BBB7E54824D1295922DAB9AF0B045861BB2A12
-F132D7AB71F8F420430D2EDE096D86A20B81130050A3E13E9F700F517BCCD5042A0639D99F05
-FCD92F3E65B88F4618B3B35E4998C49F30FD2F15D2187DB5812053FA0F3A4F197811585F111B
-67604A2E674E8C331250875C713FF0827F713C09712F40973CB8331715FE60C5A06735D0205D
-EDD45BA112C53D5D0D8FD68934CC9BED8C6F7EB60D6E3E039390E5C5D80F215F59EBAA86D2B2
-E93BACEA0331A838850CE827894089BFE982C341C20EF24EF5847D43A178DDD2EDCB5098C3D2
-7CA3793B3666DC1274311998A9B692064C116A38418EE30D663B400E1FDB9CB9B259C42FEB07
-FE59C27378941F03F51E61911F07D51C19CAB6BFE5AB99C82C77CA8D11220E5AD824190F14C2
-F07D8D3BF7C7521CC908FE33ED3CF0974A0E954644B23326B431438B2D9131F7519559F74983
-C5243582A5DEF19057BCB954556FF3E3204975D28D8D3AC524C9DF4BF73F91EBD62466DC528C
-8C9CF5D571DBB35DA632E229BE0066A35E760412219B2721F97DB76EACDEB179CB0BFE577E83
-6403D5B337235699AA2B297AACA5A8D64E872D08D879CC0FCF66F16C290BF6350F2DEE328D15
-F1FDA5761558ECE82949FCFF004B93E572D9028DC43D9BE79AE5C592C2621B62745468A728BA
-2674AA7335576E90A8FCE8F1C92E34912DE1F42F503FD0F02DA06DF2F3C1C7452A41B373F1F2
-988C8A58E6959A960DCDFADEFC5E5135DE9F920E9947E4A5102595F95C257DE1B3C49811CD59
-560C3181B04C70EAABAF5C6CEBAAB81D23A29A4D6BB01E9A82CE18632763506F6929B936BDE5
-DC36BAFA738AA1251957E7B5F807F8E19C01511BC277BA09602FB261CC95E67483E325F01DFF
-741789AA8FFF524870EF3FAECAA9D760382A5B99C61113F7570F093B32ED9897C4E9E493EAA5
-704A340D902935EEA1A32486B95B9233C2AE3E73E639E002DB7E9DBEC548ABAB62C8522317CA
-13246C381AAD15A988C740F199D57AC6A484005FEC6DD56215EFF4C9027572C51693B3497BE4
-20FAD82F298BF22061A0BD86D7A55781CD941DA1D0F3F7EB9C7F2733DF9AD378B4D426662BDA
-0AD2A46DFD177E22E1F2F0E2117E5687FD82230275EC0DB951CEBB10098498CC9A4CB2132589
-747B444016BD05364F662B082A05D7C39B22FF1F95125E6AD5DA22C8990B6F0B62D6100C93B1
-C205ED2AD813F63D4D2C6F70BFD072F23E49E8D2F6FD9B444EA73EB6FD5C6222DA5975C6AB79
-AC084DED8980AF0F8C5DEAAF9ED34C2986754BE20A0E379BF61EB9B4226CA4BC861C991ABC84
-64EAE5E7A2582B4E2DAEF97B2C9531A57F73B1C0B98CBE627C4A17983F79E0E3BE2A115403D0
-C7DB412230FC039C951B20CFD096D59E35B528160739B164A4642C7E59F18B976F4DCACB8D98
-6936BA31BAE9DA3D0CE426647775F7ED453FE50CE658A333AADDED9C88B740AB6E5FD3CEF7F9
-59AC52D33E9E62FDCACCA35A9685625A480924A64921CEB4763F551D254AE1820DA9932C97DD
-7F43E37AB666FF111E49809BC18A6227D59F06CA1B0AD06DE785C1302D9B24572D3F4BC07733
-88170685DF3301A76B33726D3785BFE3CD0E34187C4F753AFE734B1F91C4CEB846599DDC42A5
-C9EF8431090E667A25ABE90B240FDBBA04B4687DEE81752AFA00B298354217E2CFB1B2285D87
-0BCA42458A135C959B02F5759A7BB6A78C50EBEAEF1E334022C2F4B7CD610754CC6EBEB00E03
-2B87974471020E3E02AF14E1B09E51C0BE0446367DD3509560ECC8D93652A31126A6B58BE871
-7174065488CC241E45240911BF02E6366B806236722DA482303903A675462F71CDD62E14C553
-61C1C087D2FEB2CF4CD6C508A6962FEF8CFB9383B1EA51949A622D1B09D6B9B9C93E105AA9A5
-CC200BDB5115016755EC028E423C35E3CD86D2DDDA72D025D73C61F410B0127A99DFF3DB09D4
-DC5687D6FC7B8B742447269EA45B60021EC53D3162C0A8AAA15C406CCE397F4C05831407767F
-4381CF59E964E425F84F9E5A0BCE860F4EC0E4016886DFF6CA58FFD6479E4AD53817E1660677
-B809E0E0885C77D222D88BCD39CFE62632908D058B7CFADB6531EFD8B0398E8F3FD94058C62E
-B489582C3FCC0611BA3CE75AB3CF38A9DD59008D268F77BB105691EB7C03CD7CFA238E84AE3C
-9150BC0AC8CBE10E24E31313866EF901533D9AA8DABA9D4EFA97458E714984349C7B41490388
-A2F15E7F8C7B7ADD5E9F994060F9E7C7737554A3BAFE0415D2B96DC62898737CC43C51C52331
-137B8A2F1F7AE537C34AA9D56AC88FF0BE4AB89513AC373CD6E80424AF411BBAAEFFE0ADCBFD
-90BFF13B502B435DD12755FE8CEA41FEFB00A81D09353E840A5FF18A40DF9B9C46DD6AEDE5F2
-F7B59D1FA9C7A167074793A07ED04E03C8874F02261DC4BFE4301E42C10499F82FEDD545D49E
-AFFAFA5A9AFDC0D45319214CB461426C27425249843FB0BA0614EB80717BE0D1534E115EBA96
-49DCB53563478C7A3A52D422B4073189C0FF1D70ECD3A69E196B1452408EA409A733AA023970
-094A1980636EA4182B8FF9D636E773C3161E01A3FCC23B3E3E0BC151AA3512075EA6C232B424
-224293B5C1B7B554B8606ECE0DFEDBB55216406D2CE51EE4242281047EBFD3E19D07749DA8A7
-1D7E5DFFBE8FEE9A23789606622F38C4D91E457A36A42E4361320E27CDE499BF47C7D6914BAB
-55461059F75FD8CB1D9E5DAD7224A22C4D425EFF89274A8AF0931EBFFE95502627D0876C70A9
-A027AE0D5370BF6BF55BB928CC594899D47F1854870BC4296E8D8E8714020DC17DFFC0B0099A
-FEBAD36F9BE600E98F71FA787DE3B2968E7C7209787989C7613F053DB726F575601C94F75971
-3F373E844B707EF60089FA6425B65F10244B81316595B0842B22E550C85B935DAD5474D4B6EF
-274C8D6D3414A7AF48F20AC7A61149F51AD4EE9D4FD71F10BF8167094CCABA72660DA3BAC9DC
-01F12AC0330F06E768205BF7F020F8D4C7F2674B1104AA3825BAE123BDD3F0A1B91AEF608EF3
-4A79541A2C63A725D4BF7C321EC9C456D7BF40BCC38FD640BBCD3B3C354AA98F5F446126CAB2
-30198CCE2DB6F908C6D1B245307160CB300B4762F7F37B39C7B335622C2AF6555CB0751662D5
-9156C86B0834C172A8A02D4AD5A370E659B914560D9C875CECFB6AC9A212950EB14E876700AD
-F60B5682964F242E3C635F90EE6EB4DEE8EA55948325D0F13D63750716873C6A1AEAE7434B8C
-D94747D67F12AD02E6A3FC3005E3313ECBFA45E7B0B1D81657FB76879A235986B43DEA1BB8E7
-E424048AFF68B69CF84DB6AB6C1CF42FA94A56D8F3C32224F288631A9B2434316DD00EA1924A
-3212C5A78A809E0E35EC4694601A04E4CFD3211B6511624F9215E69C344F09AA312CEEBD6004
-CE8FF955311666BD5E0EF6E3458547498AEAA8A147E67A52E1F82B82A67096CC76C804147240
-38B59F19476609DF98619365E70261B8CF7945B25CDF2A91DC0460E36588C1391E1C95909749
-3D83FFC79C48ACCA8842C6AB92D85625C68258049C4FA27FA79E8F6B1C5C7D2978631F7B69DB
-0790C6C2D6909F97D98FBF3F4D7F8714D242CEC600100BB82A6AF564514A3C35606CC05164EA
-4B907FF473E4F85AA819F932E475728696271B47E4E7E93EB45C6A33A7FF6D163B5862FC5588
-2D0EA227D891E2296EA9E9CB67DC7803A6D20990D6DF541831064B38902B36B38A5724A847F0
-7B363A405BAD5CC96EF4302F171B0F56D7A9843A3DBFB14264D7067BB1CEFF88B94CB5BDF134
-BB20A2B32EB9E6F7E62A5A94271CA9E93E2FCCDA781A77928777D4E05F85C5FFD83CD88B96C4
-D40BA13E3B37FD47D3A2B65FF982ABEDE05BC56D6F56BEEBBAF32C30C472461E101D18185E11
-C72234A7115375D98B8FB77766EEF9B2CE25CCACA9BBCB302B5B1078B516A1B69EC6653F827C
-A47039C439B21EB791968F9A08569B4024D3CCAB0DA8DC8A7037E1438A6C04FEC40C625E13FE
-AAF858D36679CA9F8DB6D07744FEAA34CCB6F9EEAAFE86F1E29A865272A25FE921EA01C0902B
-1606F719303E541497E583C99C472692E117FBF046A9A772F5E917D7D6F5757D687896F11AF3
-86F038EF9EF55628272FEBE7B0C18DA6877BD74D5A05315A088E2316F999388C4AD22BFF272F
-A0662CBFE52BDD00A4AAFCD7176DC8D3293FDAF91BF97069EC7FDEDFC1AFCE42A05115964B89
-D8F32CF77394C8C659AF35FAEABFD83F713CDAE091F266A721AEC77556526E59C772DD9B280A
-80A55AB20CD4D80724100670AECDFB7FC6698BD3943CA4D391A148F152901CC8A56E7068C5C5
-C0F9B3AD9CA43F21AD7EFBF25C6451E420B5A204107931B4B71E141484301B38936552BA4CB2
-37CC8B6372B351DDA986F613B10CDB09CC8F211A20EE1BFED07252CE9FADCCE4425D8B0788D6
-1831E55FE491110C6864BBF8B2D2559FE9FE9170095AFC69846FDEC69156F5DAFE2E1A496F30
-B8DF312A2B1441BB527322881EB351DC3044AE37E7F7DC307769D0E67CE7E03C8EAAA2534CF3
-F14228B41D6B414C582D145DE880A236BD16BEBA58E6B8D266399192DA034E28C97A2D49703B
-EC65A3B4AF9B4FA2D263DDA0FBBA7EF607C80DD616AAE0B699103AF98436E1B9F5967742BFF4
-678152D97A39DC3506E187EFBB18533E38AE02B9DFB95604F4100BEF7E47AF038DDAECB6B2C6
-9BC6540DB2A9F8878839840A4B64CE74BA953C4BA384ADC925D76F70A99233563AF59556BBC8
-837618E7228A294DDF9018B130233D15B19CD6629C85875EE18B4B21D19E0F2D467A9E2EBA52
-02A0A468AC96360FA93B7931125AD7D6D23588EE644745C8813838471DA250ADB0EE5D0AC974
-02F7491029F765938D79F1EC1E5AC99A5EAB27DC354F82BE44F45B07672C9EE71CEB6CBBBCE0
-8B028E9161C03B7B94ED1B06AFDC7A73FD44437538C13AA502E1844B282637E8D2CB1B6D9E4B
-64DD0CC7FA0A088AB08AC3395DCCB3FA16BBB302807569603668E84323EFC3AA0834B8B463E4
-B0FD4E0CFAC2C6CDBA69FBC3C0FE9F3F3908FDBD20FD67CC084F5EA39EB4FB6C8D4A937AE24E
-AC6CE48F731D3EE6D5274F8A8FACC2EDE324851E868B2F3AACAB889E9C1002722CB86F2B8872
-F13460EFD6A4955D42E9A8B8F60ACF4495D7922D481531B83C2FAB0399C3DFD0BE115F5AAE7C
-8ECCA864F0585BF828FA5E021F03DF76AF0C7EE7A29C6050180AD9AAB88B0095118C82BDBC8A
-7F11C6EB5E58CB9B795523EA80596926805372D5926C2896EF1A3EFD44C5D818600B57B61D8B
-9A87B455E06170585F589EAB2ED17663E3B35D3BDC58A2D2C0D663C27D90B5DCF26FFC3FCD72
-8AD1D597A2E82241812E274F165028AE112E7554C40ADF077E5215C9915F8ECFD688B6E33A13
-1C5D08E4BA91CFBA9CD149C538A18C944AD5A1ECF8F5BC2979C343A8AA88C86961AB05BB73A1
-38B70A2FDCA7C369C93A46B6A7EBF644BBBB8BF813A31F50194344F9DB40203FFBCE2255A185
-0167999BD24EEA951161439AE06AD517B83F443F168922209D079B164AF643C050E9D0B8886E
-DCF5E5350E6A7C33EFA1837CBAA360F8E1094742AE629E7B4B00DA13A9F02805FA11F16729D6
-B770FB8475ECF5743ACBBA889C3B59594DC6DE520E359BF45EBDE814A4BA56DB032E7F09E37D
-AE4EDB4DD8DD558AEA476403AA5B2A8082A2EE5402EB12E231EDBA442FC3653CC566CD92A524
-53ECFFDA7A4B2F9D88560A7F2754DF7D4739933D8065A94CB6FB2A456508939EE43A574CB144
-36FE96149999E135E75621F5CA0354BF8793E0F2EEC8927BE22CB2D18092027468121030D60E
-649663F359BD00D9A1178112E9987EBBDEAA525F242BF6B41E6CD7F0C518BCB8AE5A03A87B1D
-7A763825E97C8A1D365DDBEA9A869C6548D69A759A97E977920B783699E8823E3D3406E07344
-04D5BD169DA1F165B6800A16C5EEA79440D969E5FE445AE3421C68C5F9CE35E647F0D5C6F38A
-3756CEAA51570CB799A1ACBD598A22BD71B0A5EC6B05E23AF845574EE27E30149F2A48823AE3
-8E889D17B547BB395EF9A1D3824E53FFB99D27BFE9FDF0872FD5FF557F78C1E043C7C130FEBA
-7B9E9C3118397466F3021F1CB3693C711D5025D9BD099FC24FC6516688169026581369CE334E
-6A643733EE0FA89A9BAA3FBB8E196593C37F4CE5FB7A3A4FF424DB950083057CDD0E82B95493
-434D91991C827C7A1C4165004C6415661C77FCEC2328D667F6FB744B53D93BC414BEFB36582C
-5CE05E23BC99AE0EE96DAE1B5EE8E27DEE10A7770DFD265279B5CB79FF054BB42383C5F32E92
-97D2F2A0AC8D25B5505A4341540240EB55D57B5DAF406A3160EAED1394D5B5AFAF270636720C
-5E1A36FFDBAC50ACE636E7E2AC87210D271EE5A219E49FD060EE71D6D51AEE5ACE8DE2E15477
-9BA063676BF272CEB31BCB8F58C26029A8B9EC77D1EB7FE19F2C9486008368419B7E844CA6E7
-0BB1CACFB7D48A96B88CA3A8D8B19F5891539F9A2C52F5E8BBBC89F7D12D146E362B5056D666
-39EB71CBCCA5390BAB0A96C413E2AC28B871E926D08B29AA8D638DBD51F3149E691AD8AB0DEE
-D4A0D68D16B945C02F36BFFC3AD91147C53CF645B3069B7B0BF6C0EC18AA483F8D91DAAAB9B4
-20D19C5D81735D25827B14C8D56C2435DB604BC012D92241B0C5C664C1EA6B49BB0864C582E8
-B98AB0AE9571ABBF5D8CAA29DEA39C81EC4F0234C18FB80EEFB76EF35ACBE2E576801BDE2CC1
-028A02461E105B2D4DCB3B6264DBE5ED7EEA041A76FDDEB0C365A4F72B80347A5F968F5E260E
-144FD6A4F125634AD1C09B06064B6F3F745F8D8DF6DF8DBF96C428228647976A1FA2843D94FC
-1AA5A80BB92C2CD93E82FF801053BBAAD90A39CFEFD4F29A21494B860081021111D366DEAADC
-1ABF9CC0098D25F258F292AD7854B10CDC6F89A790267C6D060B532DB258AB5447E6D28FC675
-8D2676E71D67B1E53A59904EA8FBAFA635C8E69804543A6789C203FD41582C3602B05F6193D6
-312B000FEBC2B3BB7EB7488434340DAE624274ABCC678C8B7509524D81BAD38CE2C5C0CA8A0A
-EA18CB4B81705EAF3E164C92A48B2D545FB45830DDF2736C48A46A6737403643C80A017CA970
-E5D89ABC0B880D9C1C6D7AFC54BE26D03AA6C1E00D0F1CC13B8D713BEC2609DB7883A903F644
-B6ABC3FCE078A70FB69D8A4CF2BA0B7DEB21D9DD52125DECE217E616224CF235645731A0139B
-BE687BB1C23499F20E239B8493C4802CB8C452C1870FC48560F75A8AC4ACFE9957A3FF6CA91C
-0AA0C27678A211C12989A1C68D7AEFEF79CD45FBFD9535443D3C4E39C7E933DEE44E80801A7E
-FAF757D90733B1DEADD32DC8ECBFCC35937713C5F7B5C9AC4F8A00E5D887555E4062CF91D1C4
-10E83A313E93A82152186F45E186A8D20692A5E352B2A014D2BD8F2D414F4AE75E0233E8CE22
-EA2FE01FB8D0E37D290DDC09E332F00699DCB0FE17FF077843B5E906BBBEF1E94C9044FA886B
-4187A9DD43D97156A50290C9EF89ECE2E976832F87EDD665E6904CF84426085457049BDCF8D9
-574071A6B401D734034004BBF6F95E93E04B7E0CAFCD312F5E66BDE680E5EBFBFD874773B5E1
-EA1E6ABE5ED766399729D38083E9C10B390944B20F25CE7D725D283A263C7825AC55DE7DE2BF
-B946827FAF02E9AFF6C568F8A7B6DDB6DA10B6CA9803AF5A268F75C636BC34A2A1575FD58237
-B4DAB2EAD7F334E9F989C7FED41592C9FED9389FB008E683505AE58EAE1EC6332BC7C717CAD2
-CCB932511EDB9DCA19FFFDF3EFA77972C92069EC577994B900C92E8B82B1121AE97CF579EAF7
-77ACB625113BCA51B424355531B22F1EBD7D1717B55933089D4A6EF2B676C9C1E81D400BAD0C
-E760FB5994EC467A08281C4B48FE8AD9737A7BE4F9E5892CC18C5EDDE4C20A86893B24DA902F
-179302D7213087219166C5950D27BA451929FC3F5EB3F66A4CFC5DAAE4937C60607C38652EA8
-A36C964A27A1F4CB06469572575C1F3A76E98F4C433D88F307355AC4EEECF5F89FB84B2E8913
-07AD7452FA3982BE0E9BCEC0A75CF27DB1F49227A2A72ACF7419563056F4BD11BD426F70E2C3
-AF2B4D71B198094305F0E283C8002CBBF2D0A01475F1BB4101DCEF32673B4E001E2E8251F762
-CE8537A0360DDFF368976F58A5BE0F7E6863AC707BCF95AE9F26B352ECD810DF75D8D3C5C6D8
-88FCCA41A447943B313D3EE026F4D8164F6888545FBDA5CC7F90772E80817F2941FFB300B2AE
-0A98A3989AB56A5383E6C56C20219EB776F8114EBF7CEDCE603C9858ABA9482AFFAF5AC6C1B8
-88B69AD4C02633FB3ED24ECBDBB4F612822D9BA28A15348686DBAED8E1653386F62A3B5EA975
-BE17ACE48CB88054E85FFAEB74EF1BA0D83EF47C38F9D1948E063238A183F4AFF7B11B654FF5
-4D7C217A69C3F4B22E33E752A51CEB904830F8E1624873FBFD419BFFEDD0C308CF4AAFFAEF8E
-0C532AB28A5EE2397B3583E0C2849C1371537E5A1C480A8CB45DEC915756325994DCB7A71558
-E9A1612FF316E27D3D9AC79B3A2590A595BC5A493572D160A18203266BE62CFB96F3B5069AFA
-AF48B9F450916B02F573269C50F970F3D186652FB0CF47C7BD5E2048FF3A6BA2849906C609CA
-E6A5A017D60D3CA4FF24EDE9968E660FD3AE3B2C04142CA0BD186B1648B218842B180AF14841
-17BF4CCC800D9F5F76585D6AC4F64D26B9A458A4945762372CEC119F91870BE7E6F784FFC9EC
-385735228620A6E75806AB4F04F4EEBF3A23964F6DE114B5F16F23C1486D6B060C548E8F3E06
-ED70E1F47B82272A642B45AAC02010D0019A64737FC92CFAFC84049F782EEAD531C3DCB9B3A7
-2527426D1CF7694940AEA4C2681180DCEFBD6413C58D8E1A6DDDC75B4D0A39D7E4AD3C0D08C0
-E0A264885632E507881366596FE386D4648AF430474828F98B186988E88E5C374CAD22C5DFAF
-648B80F2D4DD131EA6926B7BBCECDD78DD89614D21C35B6A62BA02A33AFFC02A577658093AD3
-D67E04024F19458CF8EFC3A6B6D99C2B5AC89D87E99191DD5B62037CE448D8434EE6A86DB340
-606B1B49B98E75FB27D5E2FA2160E8AF157A4B751313EA868F1D564E829C77A31FED587ED82B
-3E0F67BF16CB82E2074E1BEA101E0C7F05C159750FB23A9103CF83BC150200BD816AB20B7242
-963D6357C903ED09523BC339474DA38942CC46F446758722EE9506E560CEB4AF7748AFA005F9
-31FDA425E0F1191EBC9F8D88CB68242E7313FB90641D9638210C13ADC9818B7398455C2A52E6
-A42077062A4612A6C22D02BEC006B0115B55C3E27231AD41655C8DCEFC028AB128126C81A381
-7BC1FB2C1CF12C5F8A6E241A4D62BDAFB9DEEE7D6DB4B94DAE0FC498310DDFEF710B68068AFF
-A4A5E579B6B8009BD6F96E8D33FD1713004BA60B0EA187704B91873891CB96A8907535C6F467
-6934F9860A9E2DE80726F27DD82FDDE9A091B74B6BA20D1C5B39EE060400B57D007A3D21F89B
-C29EAEF953587923DE8E63AA79860BF4B695F7491780FA7E82F6CBB5F1012F265BCD9A04BE63
-01B7F026D872EE431261D3934D3DC8E4F10A7A045B03EE476B735637A8F34FDF9C8CEDA2AE86
-25D1BCA5A05F63742ED90E708F4471816A493963D0481AEF1B25B002F0D438F639BA7C90F7F3
-E87B940667DA0780B11F13F37E768EAAD5B943811A9327C9B9C0752C757C4E9E14D693CAB5E5
-586097A0155BB34114B5AD501E9CB443B4EE3C02A183D794E01A018D23863468793690386DF2
-15C4F9C3D8FC4CA7A2EB0CB3A34BB891BFA7E61ED7E8562964986A5D157519EBFFBD96B1B736
-1923564FF16C5FFBA49551E6CC38B46482DEB027136C813EDC8BF9882797EBD7403CB7D7E7AA
-88DCA2B8AAAB63A73F67893A955AE23205E77A6C17B0AE03E66532FE283B65D93BFA1E7075A1
-99A4CBD83770BC1DA19C5BC76AFB7E34B70582585B731D91CF3F51D45D039F0039F0DAEE2E1F
-975B74499BEEF4BB8A6EE97BDAC55DC752DBE35F3EFAA01A6879F45677B31D9362C867FAE757
-DC3E91B896496AF3073F3C2ACE015A069366F29339DDCCDBF7AACB94EFEA2F51A58181F8DCF9
-94676F2D0EA7D9B1D13837BCC940BD880693476E38DF3D76550379E61D636BD1DA787BF5AB4F
-1C46D506D7AD4B1A32905D6623CA05BAEB9313B8CFAAF577D89F9EF32EA8821472DA3770BBDB
-38AA975A91085B50D3A2072CC63091538CA76EA2E595836D99D57A066C2D224FAD79583E86EF
-00F7918CBA33238C729711C1EB2FB02466EA74414A6C3E24FC7F0F14EC55EC944D39CE9D95AD
-119EDBB9BE9C4F39646DBE84662ABB9D20092DF7C0DE3CEF76B5FD988AB9FFFCD6EB0DB3D574
-DA4713E5DE444B3D98C79F36C89FCB0FBFC13112775134331FA1F80855AA3BC5520CE7477B46
-1895E033DBDB429E7FE57C490CDEFA7D424F368F7E6E7F8F471D4717B7DCAE94527002E6AD96
-26A2EC8449F75E77F15F6D4B07505123685DEABC4A82EDA62A868D84C98F6CE4F550AEABB195
-8E06DC768B335DBC779B972651817B8BC7D2C706E4A527DE48C68F4ADC8E70E5FA4F60B7BB83
-7AAF42059DCC491A0420C81829A035423D5334D65DEF1F978F99887FC255DA8EBC0D31822241
-866D2F4D227A8772F646D9890870C8FB101BDE5A21BA0BEF4599560DBF1C72FD40B288F35D07
-0E6F26495113C62C65014EEDB07C321C94B701DF0969624B3C014D78B3988E3A995599FE8F3F
-4DD3AEA7892567CBAB1D812DF63100C8DBEFF95C0A5E29751DEA2D2B5275072AED5B792E61B8
-BC0C996B9F03A511DE02D0FB2FF7F464A0DC6F9F15B5160A77DD29EF66F0D73DB1BCC57F4607
-7E2E85F296DCCA94FFAE3D6AB92EB5FF331BE303DE4FC72220AB1C9C8A03F21D2C4FB5AEA7E6
-3A69333829970EC8494D2E3F1B2D0E99C966D38DF3AC6206C3E9A69869F9C5F47086EFEE29CF
-873A5EC3B7F25E9D4365FCD83483254B1872B26C6FD0E21946037C40E0F9FE38D59B5EE8172B
-1EED2F8CE7FCF685D6C5D28A439DE7C9890C4E675C0A924C6CC5C7CB1C2D7E5503879DA74FBF
-95AE097161F80DEEE75AEC1B022E9CE2652A56A6AFC57A06E750EA4552D12352193801994700
-A771CBC33EF0C15D9871D0AB85D0B96FA9CCA7101753635B2395714BB74A24B91F2B7F13F446
-98CC203F97FBB3780956F01E19BF03874B7D1B4DA0D56BF1B5C14C22615221B8CF09E533AA46
-0E5CD9FAD42E434AFB3C9E65643BD19088C7BFC79552B42575632640D5F05E8D2D5A80292385
-B6AF05C247C937071AC381ADBF2A0266DAA5509AC233645B8F9578819D76203FD98794AE5EE7
-CBA2E543D9D742F26324734033E81F2CE0070B2C153B4836434B48458AD5D4C40B1418B2A53F
-001904E75C9944F383A80C05E3F7DC160C442741EBD239F55B7C0A9C1B3730B263E797010421
-25934C0223BCA4848E8150B31BFA7300FFD377F009493862EB8EC7505C5D7A2213719A6B4806
-E8B385E1B4A542DED24240C722FC4E6AB969F7D3750F2F76E02BAF9BDB7D6DEFF714032C2F50
-E0C3A18377129A09B7C827563F9FF4CCDB313345FC5880EA233E89F8597B27FDA798872A3595
-466BACB8EE61BF79B56E3082C8AE75CA708654F6031CBC61E1C3BCEC255B19718F557CF4E6E6
-A7BB94DC8AD7ADA4549F0E908739B3F28BC949082833A95F211D1A8D30393647E09EB5065159
-CE80E945D05E1234485F74C55D6A47D1B431004B63480AA63E6B9BCA1879E4D90DC9FF7E0E52
-461CB1037ECFF1FFEBCB42E21127A9F98D51CC5ACA430928BEBB39C2AA8B2E810910742439B1
-63A86C032D724596E1BBBF5809032A6E0012B2BE69B3AC750E874B27047CE4F50E2D0FBB778B
-ECD49A677651B2A843CDF5A953A8EEF4590641BC6FC2663310303185DF05FF02571839BC9865
-956BAF99A70DC8F6907321C721AB3FD6B2F7AA3F8D48C6D5A7F6E5B62B162BE346CC72054355
-ECF830BAEE80AF84B6202A6157C59DEE52E07EBD580002579AB13A7A9F48B8170278053CA1ED
-4A345C0E04582705D0BC924DE91C306449CB8541A355D252085CEE5824BB68408AE62CE2CE6D
-8BA05C7394400D64EE59FB23A5B77B4DAC2F30A5324886F4C9CF3C2F6A0BA347DACD4388DAE9
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-cleartomark
-
-
-
-%%EndFont
-%%BeginFont: CMBX9
-%!PS-AdobeFont-1.1: CMBX9 1.0
-%%CreationDate: 1991 Aug 20 16:36:25
-
-% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
-
-11 dict begin
-/FontInfo 7 dict dup begin
-/version (1.0) readonly def
-/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
-/FullName (CMBX9) readonly def
-/FamilyName (Computer Modern) readonly def
-/Weight (Bold) readonly def
-/ItalicAngle 0 def
-/isFixedPitch false def
-end readonly def
-/FontName /CMBX9 def
-/PaintType 0 def
-/FontType 1 def
-/FontMatrix [0.001 0 0 0.001 0 0] readonly def
-/Encoding 256 array
-0 1 255 {1 index exch /.notdef put} for
-dup 161 /Gamma put
-dup 162 /Delta put
-dup 163 /Theta put
-dup 164 /Lambda put
-dup 165 /Xi put
-dup 166 /Pi put
-dup 167 /Sigma put
-dup 168 /Upsilon put
-dup 169 /Phi put
-dup 170 /Psi put
-dup 173 /Omega put
-dup 174 /ff put
-dup 175 /fi put
-dup 176 /fl put
-dup 177 /ffi put
-dup 178 /ffl put
-dup 179 /dotlessi put
-dup 180 /dotlessj put
-dup 181 /grave put
-dup 182 /acute put
-dup 183 /caron put
-dup 184 /breve put
-dup 185 /macron put
-dup 186 /ring put
-dup 187 /cedilla put
-dup 188 /germandbls put
-dup 189 /ae put
-dup 190 /oe put
-dup 191 /oslash put
-dup 192 /AE put
-dup 193 /OE put
-dup 194 /Oslash put
-dup 195 /suppress put
-dup 196 /dieresis put
-dup 0 /Gamma put
-dup 1 /Delta put
-dup 2 /Theta put
-dup 3 /Lambda put
-dup 4 /Xi put
-dup 5 /Pi put
-dup 6 /Sigma put
-dup 7 /Upsilon put
-dup 8 /Phi put
-dup 9 /Psi put
-dup 10 /Omega put
-dup 11 /ff put
-dup 12 /fi put
-dup 13 /fl put
-dup 14 /ffi put
-dup 15 /ffl put
-dup 16 /dotlessi put
-dup 17 /dotlessj put
-dup 18 /grave put
-dup 19 /acute put
-dup 20 /caron put
-dup 21 /breve put
-dup 22 /macron put
-dup 23 /ring put
-dup 24 /cedilla put
-dup 25 /germandbls put
-dup 26 /ae put
-dup 27 /oe put
-dup 28 /oslash put
-dup 29 /AE put
-dup 30 /OE put
-dup 31 /Oslash put
-dup 32 /suppress put
-dup 33 /exclam put
-dup 34 /quotedblright put
-dup 35 /numbersign put
-dup 36 /dollar put
-dup 37 /percent put
-dup 38 /ampersand put
-dup 39 /quoteright put
-dup 40 /parenleft put
-dup 41 /parenright put
-dup 42 /asterisk put
-dup 43 /plus put
-dup 44 /comma put
-dup 45 /hyphen put
-dup 46 /period put
-dup 47 /slash put
-dup 48 /zero put
-dup 49 /one put
-dup 50 /two put
-dup 51 /three put
-dup 52 /four put
-dup 53 /five put
-dup 54 /six put
-dup 55 /seven put
-dup 56 /eight put
-dup 57 /nine put
-dup 58 /colon put
-dup 59 /semicolon put
-dup 60 /exclamdown put
-dup 61 /equal put
-dup 62 /questiondown put
-dup 63 /question put
-dup 64 /at put
-dup 65 /A put
-dup 66 /B put
-dup 67 /C put
-dup 68 /D put
-dup 69 /E put
-dup 70 /F put
-dup 71 /G put
-dup 72 /H put
-dup 73 /I put
-dup 74 /J put
-dup 75 /K put
-dup 76 /L put
-dup 77 /M put
-dup 78 /N put
-dup 79 /O put
-dup 80 /P put
-dup 81 /Q put
-dup 82 /R put
-dup 83 /S put
-dup 84 /T put
-dup 85 /U put
-dup 86 /V put
-dup 87 /W put
-dup 88 /X put
-dup 89 /Y put
-dup 90 /Z put
-dup 91 /bracketleft put
-dup 92 /quotedblleft put
-dup 93 /bracketright put
-dup 94 /circumflex put
-dup 95 /dotaccent put
-dup 96 /quoteleft put
-dup 97 /a put
-dup 98 /b put
-dup 99 /c put
-dup 100 /d put
-dup 101 /e put
-dup 102 /f put
-dup 103 /g put
-dup 104 /h put
-dup 105 /i put
-dup 106 /j put
-dup 107 /k put
-dup 108 /l put
-dup 109 /m put
-dup 110 /n put
-dup 111 /o put
-dup 112 /p put
-dup 113 /q put
-dup 114 /r put
-dup 115 /s put
-dup 116 /t put
-dup 117 /u put
-dup 118 /v put
-dup 119 /w put
-dup 120 /x put
-dup 121 /y put
-dup 122 /z put
-dup 123 /endash put
-dup 124 /emdash put
-dup 125 /hungarumlaut put
-dup 126 /tilde put
-dup 127 /dieresis put
-dup 128 /suppress put
-dup 160 /space put
-readonly def
-/FontBBox{-58 -250 1195 750}readonly def
-/UniqueID 5000767 def
-currentdict end
-currentfile eexec
-
-9B9C1569015F2C1D2BF560F4C0D52257BACDD6500ABDA5ED9835F6A016CFC8F00B6C052ED76A
-87856B50F4D80DFAEB508C97F8281F3F88B17E4D3B90C0F65EC379791AACDC162A66CBBC5BE2
-F53AAD8DE72DD113B55A022FBFEE658CB95F5BB32BA0357B5E050FDDF264A07470BEF1C52119
-B6FBD5C77EBED964AC5A2BBEC9D8B3E48AE5BB003A63D545774B922B9D5FF6B0066ECE43645A
-131879B032137D6D823385FE55F3402D557FD3B4486BE79011D1F5BFAE5C1F476EE6F05EB1D2
-CAEB269958B194521197B312FCCED4867F3C8FBD030BD715D8FFDA1DCD454B174E7A1A97B59F
-E770E67702519D9D9B23D61AC08424D55FD4DDA249CFF0B56B9F3AFFE9D0DE215C02A52A6BC7
-7155FF6B8CBA5CD6646BA331254AC58ACE650A967D3B272331B87B6DF06D5AB9D80FABE9F9CE
-AC10139B61244814DD9FC295ED42D1B5CD11C2E4BEAB318B20F51CA2C2930657E5343AB7BCF0
-E8870A0D12DE4FFACB6CB3FDBDDA481C2FCB84408D3D902E9A32070B2AF6CD9317A33A42A438
-57C114B3F4C3005CEF9401F1C2BAD3E69150D7145B79F95C9CFAF7A335B277E6435AB374F6A3
-E78E124AC1E4615511F743AD65C5D778403A840310AC4902985F107FD33C0049623E4C496CD3
-53D2AE2A4804110C6A420A38ECE5D5B235C5E35886A0987D284E6110527B653B1B2C68C636EF
-423536180ACADB954843EFF4D9A82568A713B682C574E30F793F093AEF755C650E0E7175C2E3
-138695212D9FEA7BD4B5805CA90985B134C6BC15A19AFAE0BFB3066363CA3E2568C2AC4559A6
-9FC32FEB443B61243B4508B7524D974636855C83C8149D489B68FE18BD497116A953431251DF
-BD09FA1D2EF66C85B997AB2F6B7C6D9D096CB43931379FF9782FA00C5F3CBD0604265D16891F
-FC4F66E4A0307765D384C32AA83A4E4F89AC79EF8E76EED08450E3466FBAF54F85324F0A4F04
-BD00CCC22D32E4132C53D14C043B47034450AC4A843DF14590E5231E492B11D4A00011C3D6E6
-F37E2632EDBF6CFF3902D64929BFB8FA083A2240FCEF3F733839C9B0EEEA1D5A2F357435AE71
-B7E0DA54985D1DCD073094F9738B4FD4A2EBD338C28CDCBBA1CF3C6BA6E9937D4433EA6689A8
-1E54101843E8730C60052FF6DC6C75CB96AF37AD445DF0AC49F897E7856DEF63A850C6605695
-615950B45A664E674063A99114CF1E0EE2E05D79D26E7FEB81CEB509A2F3FA8D94096CB0EC26
-FBBB135C83A55301E5AE52A96AD6152A8A84A7D2AAF0B7766AF3C67E5A4E3A6078A65541C9DA
-FFDBAE251854E29ACA9819D5D7122659C4AA49F9CB5CB445C3B751507E431E61324C3F99776D
-811007E6AE830DA6DD283B06B2BB817E2A8A25E315A881F2C5EAEE90F2E09ED012BFE04D9366
-2204AAE6A1E3B39E50860DB40ACBB8A8EE89EB773C871FB59D43F0F5756C535F7E5B87C1B6D0
-FE9E6C785F2006C2A1FAC9DE611159DE3296A2432EDC180285D3AE485E5D87B08104CE271D97
-4AC9A0D5D21C14CE3FF9D4AF46B7708454F4CE5B4FD5A6D81547C379626E78E48998B2993603
-6565CC536CF1D538E2CDE8463FEC33642070F6242BB1BADC7679F05D1EFFC5ED912CD7B68785
-6CA126080C246873CE60D093E3C7C39B9850BC2D2B9EC18DB277E0713F03308D56AB50AED4B4
-D863D9689DB3F0C430AA0CCC5CD7FD74A0A6C7D45B5B874080E20E23E5B11886B6D9A0E3AB0F
-0299FDD011997614990FB8A21B8CC211ED8AE5BDB814C2DDD263A7B75F875B540790247AD09F
-D8C54F4EFAAFEA7017CE0CE7AB81A7D539A307262ECBBD06433F31F689CABA4AE3CC7B1B89E1
-88E2EFFE3EEDEEC80DECAE4733BAC7CB2F926215C5BC6E1B1AE50DF4E50A09802E7BF185D7FD
-65EFBAB4A6B624DE6CC28EE8F736D5951411CCAEFC58FA23E9578985AB61252217F790B063D4
-68C705C95AB56247541508AD6672BE6EF7928B72D624C1684644E67A654B4CEE27A3CAF8EEAD
-4D02677A324DD60B46353E71E6B6F2B1E456622904B19EE31DCC9E514FD0C05DE43CBAD782AD
-9B3F6EDD727B4E8C9BA72FB7FD7EC06F2C175AC8C22CDB9EBFEFC3D7BA96233923989B51770B
-C898AD7B43347A67E56FF1D9416A010991CE146CED6C7FFC15A4E1103B270623806F64BDF815
-0461CFDC522153C6AEDD9B47C3637270DEF1E5E46FFADD86E6E89BD69BC3A909120F1C7253A1
-836391DAAAFCA520E8B96A4AF920ED0D3C363602196A93F5B320C9F39A01CBA90801002FB02C
-EAD7E0410D1BF8D4EE0D1F471E68CD00D6CF53457F97C539BCBC2195C32ED2ED9354E0EDAAFF
-448E0CDF54E5074C4C00706C087E88D6ADB9792AECC35D3D3364C8E1399E48D1E43F8BE14A37
-6490F513E202A068B30FAD0C6D7333806B95C4EC1173E64F9A3E4AB730EFB2AE75CA9E264B94
-13E860352DCFDDA8EB4B77BFDD8ABEF55B4C27EAF4E03144ABB94A95D27F6DD36C1CE7B76179
-96E0B8D595011F92ABD6B2CFFB38303A765634147195D7895FB99218D4C85FE9D204A67A460D
-B5DB7E22CC1A607209C73169B945F63B1C170370155479F0CC287684ABF485E89C0AAF26E8B3
-7DA9E4BA8FDC53B7FDA67D1CB5D016B97797046D7F75E43FD01E72FF2A395C5CE5F526464A2B
-2C62F4636779EDABB868D5DE9B9B6E753C35F87553A011C24C9AE4BB485A5C75C0EBE42A4FAB
-2841A4030BD94ED18230A62D3CA211D2E646C73A12D51EA7039FEEF84967FF2286B727DE882D
-9F078739BE443764E6863BB4D836A0C5F4B0CD2C0F2ECE54F4F9E58A0F680BA71F67EB135E4E
-112D574A39FF1C81BE0658420286709D94625C4CFF0103C31EC7516577AFDFD0E4EE8F2F53C9
-0FB5837560DD5EFCE5B47469F5D078CFDBAAD117E6C7925E3B4757B1C1A798DE23BB34AA1F3B
-6AF328193552EE29096005F6FF0BD8C09C38AA7E3FE9C148AE05E7304CDBD2177C1C24525274
-477CF082BC3C4E315094D089A6E595FC5F809E0E7390645D0F2E99C58D8A8582950BB31CEA64
-41506CD7714194B36F415F76B1291633FA0A780F641082BD769EA98B8979AD38B76BEE38F6D8
-93A201A772172E3DF12BFC7553029A51CB204A79C6FE777B8565F06418CB3BB8C5E95D870152
-31C7BC825D752EC1F82F1F2D60E99535B5B4D4E11FE94A07E7B16D0B8D1FDFFECD00C2AB1021
-D18102130F9466F88F80EF1054E7DB53FFE14FE5354C5F312C10074569FC74E97E34A1806C79
-1A9B4EE5D9D793AC4A045C3AE5CDE8497FD743DCCCB279990E5895D48C1B72D1D0DEB85BE395
-20AECC1894AEEF26E3BAB27DB1DFC8E2FC5D540691EC194E7C7B23C19F645C6E0255DE78DE64
-C7D3D7D39859700EF28FEB0F6CC67C26D60B6159161650F006AA47E388FCA7D7FD276E15F12C
-8D821BB569378502E91EA55F8144D25C4D99545813EF1A42F40192E744ACAC70D0387C548EEC
-50A559A40A7E397137B492CC523D62E21C097A623436D5BAA02FD0C2B0D6895CBE4C47E5D98D
-3D841739F933042C6C44D3703A3C8AF669EE96175F9796BCB708C977FDF83CE0C9E5166ABCC0
-D98A6B4D924117970BD134989CACBEBA177829FD58C9794CBF73C32A77F5F54358445192E538
-CF7595974F2D0AF714396D11AB5868CA3D7DDF6A6418D85BD233B1BFA6DB0E18B443AB8BFDC2
-932C5E07E4FEE70DD44CA82D8E437ABA2FBC6178CAB0A59BCF5952B83C2A0DED09453A5BA8A8
-5D7A22C02507A765D7CC579C98F6C4B2712E0957C6A06EBD1FC49FD0FE70CE802D1D83DCC897
-84975DF899021F06871B82E033DF39204CDF6A11774BD825CA769D979D493F99E24922E913AC
-45FBE5FF0524102427BA92826728A76353F124DFE3D07DDB77A6101C677D7D76FD9F9F4CFE0C
-71F5CDCD0268E8480D88EE5EF363BA5B420D8B1EAB7438272189726BF07621AD94C46EC54FC5
-00110AC4D1875D66DC17DE3B0B7BFB6492B6FB5DB31485448FA1E267FB44E7245C5A433A522C
-87EB839E43753674C63D846C8F4E53D4873B487A9301B987F3028F38EE3BD8E404B52F90B5BD
-2DAB08126BD94F93EAD5DB7300CD6CECF92B79C5CFC1E9771215600203AF9A9ED7C34CB64DC7
-84700CAD29E813EA7B569604BD3B4CA4A964A1F08BBB4FBA0257BCEF05EDA0ACA3B35BD9D3C0
-16AD67D6CCB32DC1D55CC7AE55AC2F2793A92D4BAB9F7A372D157C7761751CFECE46A587760D
-9E8FC3E23123AB2B2709824865FB55B56AB42408B462BE454D0D4CE17D8967203C9F1ED409BE
-B34CD6D0168E78815BEEB777AB5AD377CCA54DD20953B54DA7C737186EC59C2E725D3BA5F831
-DFB0A72424079E4E819FD235FF07F4B5883BA6A24D824E512E6F5E2ED5754572DDCD6FA41FB0
-C1AB1BBEB6DA8B9EFC2E4A364E74434A534AE5F478B4673CEB316786662E402B796294C7E604
-5F101143D9C1564E19D57E677E8D9C5A972474971715872055CE3F61419414C0A39198BC5D2C
-4D639FF5463E882773E568054EA94C2F52E34B189C70F953F40D52A8E81E6433CF769BE01B35
-957123E16D9A1B2280D0724FCF385B407B637BFCC90BC1C53CAD285C9E7801DE3417E77CE996
-04E4B8493B5EE9FCD7E428D4F5EBF70FA8D388B84FBA322CFC45CB4EC1F05D614DCA3B299116
-A21B01B8A0CE18A32A98CFD6EFE3EDE85219AA23291C10192966EC2B942A28FB3BC0D1765482
-93CB644F748C27A0C8F433F94630F16512899185AC74D8F5F64EC737B441B527450EA4A11E68
-D70D3D73BEEF3DE41810CEAD1536017400A2EB0837E676FDF951AFB123CC0B9CCE8EBF9AA625
-912E330898C3D43F7582FBBC94C8EEC6DC5C92BEC51E84D9367C4C532E0081BA348A75F6AFD0
-81EF0AB2DDB6BFE67DFA6457366BA2A996DE953F4F5C064C3CB75E3FDF8BCE298B96327C0091
-69A7E93D24243A48F722257BDD152D00D610928F3C854BCAD10CD3D32D778D73867597D0B0A1
-4FB901CA84E27B982486E005F995216B32F37375273DA4D7EAECC5F27D97B183CF033713E26F
-CC8C1C3D82C900C48B0CA4AC3C6B26365D3D12C891E4973A924E2EEBF8A863557C5AAEEF33D3
-CB83357FC4C4A17BFC50C96B0443D39E1FDFD9A02E27B7303E4272BAD588D8D0CD7FC2B6559E
-20CD1DB8FE0D819349F5CC8C0C122E43DFD8C6EAD727FEEAAB941476986B6AB401B42E2B8272
-47279DEE6D1C3E9705BDC03023FFB2C59719182CA876380D9D7BC4857C001DEB9A8178252C9B
-8CFFA07CAC0C728905F7AFA909627CE9F8C31D49ECFEE0F7BC14C5BD2EA81E1598E9121D8D35
-620B72920D7C3C8EAC6677748B54D0D7A70D80D48150B3759A93075D53AD73B89A8E246A227D
-C11778037311022933EA89D66D3D0047D255ADE787331E41AEC3A05239D1A2E06523E2F1D331
-6367CF5F52D02A60F5015E879EF86FB9AB617DEE2BBD26CA15C2AC5BDA17D7B34A093EE9905A
-423611FCD6E6E1B6F944FDB3849C154557183603A0542751FB5261FCBB62E54703D8F9B7D620
-5DC2EFE6D4E37B3D7FA606FB7A80E5D9A354C1020133AA7D7088611BAE0EAE91F024DBADABE5
-DF2B25487A4BE9B54CCF153053A909417FA5D429D4974D44EBAAF1727FA216F23F29877A043B
-D2C39FE0F57098B3742EE7EE3299F3086D683E7D055BFDB86175467FFFD6BA7E51E17AECF866
-6280B7E3405E0AE6F78AE2A38A3DC0D899059F578CE6BCFD581921B6C6BADBDCED810E82AA48
-50BBB773E403911415CFE98D5EC2C6D9BA8EB2B01B0CD31566FAE71899FFCD16E95CF015701C
-99FE6E3D9D05DE10757B264075952070D627DE00B70A9FA6EEF03450932E97376F31D0340349
-696693A1FB8B922FED4059D4CEBF798BD9B59B42283DAB9CF5739AD8719D68F6C67D68A65206
-EB748FB709B01EDA72677000CF97C581999D9F3CF8FE125D102F4B63B72F98BDC3C35742453F
-C184A6D27F965853FA307630C7D4B3C1ABD2DF23402E3224964B103A5650EA476082A4145B7C
-1BCBA3AC8FC5DE65B6B5478499D9F7196028E57BC547352F639EB55EB82298DFBEA807D88597
-8532677605A26FDD8B4E3C5181B614CAFDA4E0E275A25F5157801EC8B94B0E467BE19F787168
-3D56BA2E44AD07B04EE35308399FC004388D0D608A58D3306D678E6EB5AE29D2D8D24C4CD8D3
-9889CBFC004E49EB78805223E74016C346D3B479EAD2099398B0E1820B0BD11D5E4920EF35AD
-6432EB866685ECAC7CBF6019A068B4528253EE4184889AF5E2BD01A40DC70D5D49A4AA226635
-585DF8F7B47899AF9B319970262ADEC0E4B1677F0AE7C7CB547B7470FC0A8BFF45D77265F595
-98E6FE6593D073677B9D12F42F124281035D71666F425E6DD00149DA51888B3260EA0837B707
-6AAA0418E6AB821621926EDDB19AEE609C179BCD50D8636CAF9723BD7FD3E9D40247F55D1356
-9BE45E4F31096D69CCA44E19CBD0114C8B3B44700037BE6806880D09B4F1FB1F5D1B872D4F4B
-A44D39E20007BC04A8BC6F1F9CB7D8DF769F12B1C07A9EB18ED05AB7F202A01455956F46608B
-36D555FC662C632255F8F6B5966F367483A47D382D87440B1E51EE88DB5AC005C96292A49976
-666B84BC78524FEB14BA2C8334E59D6FEA8E7B0C9DF14EF32CC3F05F9ECDACF9A95B9A00F7AC
-CCD013F8F147B90C06AEA2DBDE2CE4C0DE45685CE89ECFA59CB8063B8354D0A807CE2B9169C6
-8058CD6E1F537FA54386E2025F7150580E8F3BC40F88C8627CEA8458C605B4875EF7E1F96D4A
-DC3B384076250FBFE0DA308758BDD810D86B29B2500B4C01FFE9C6605CA02A01E1FDE8F2C331
-E4B7EAE9B75BD04C281B78DFFA503615FE00EC094B134F254DB8967A77AC764400FD2FC481F7
-8AF06EEBA1D3E37C9BA375232796B048573E3CFC0A900D056EBBD4E9E124674F9D2A8F3320A2
-34D024D0102802DCED346AB577E5E067A3B1E06B1276595100BB46B316030692FBCDA776CBB4
-A457D48B2BD1869E2EB528B51886F0CE2184A19C591002D912A72BD076CA56CC366E7590EA31
-C19C100880FA57A0860B2BBCB24E8A8F392D118DBBD8ACC95FCB5D89CE9AFF74B6516360E193
-26FB55F55C56325A52F98B1F8EDBD4899E4F49CBEBD16CE5D99CF35F48BCF3487461F22D7B30
-44CAB0A0C994BCAAF75E5E2DC25AEC7372C4E6E1147BBD301002596E52CEA587ACD04E190D34
-EFD4A0D2EDF8B1D77770318B2591D84C530C56A20FFAEB6B98AB0B388CA9EC8069DCAA504664
-F37572724DA03A449CC88E2FC0616207F5DA7736D5EA601B1FA8472CCC4E98711C17B97FE944
-5969D571AFE0900F2CEC78AD0AA017C99BA494226F8D39E0B9A0974BC8DFB0FA708FE72609E8
-18D396E4A4D312CA1842036B8BA9E81881113AFDC414549F668DC71D1552C941C9B4E13B71A4
-534B4688B4710C4E77FE8EE85F67FC6F277898FBCE887B6EA58364AC0D865A8B137D832A4BE6
-3A24435F659162CF81E052474B58450EDBBC735D50D82DCF303D64B599F91AB3ED49C8CFB165
-C1073CE21158874348BFCFA404123691D9A4971B2F97C5756407EEB5A80A3B69536EF1ECFB20
-0776987BF012091AF3FFEF95E8238457C6B0737DCA4A1B775783EFE235484457F6B9198B34B1
-F10C719E035E5B55E603B82E7028A8F594DBB9199BD2A2C21C256F8CB0F42A6883F2A7B87937
-D8F27058EF91056186578B7AB58F306127C0F91489E95978C5B8F3AFC6036FA15FBF7F90CF7C
-744FFA3FE2152408A73B76C9A788379E60F6555809C4E9F06449AD043EF5E391CB22AC137DEB
-8AB7A992CEE286260E5C0FE273D1C0DDD7703343CBDA9431876E4217C5A0F816D687DB4EA72B
-1F358EA0E1EE16A6F346F88AF4733548B3491F676F0E880F6A68AC25DC7181EC2F708EEDF9DA
-0BA829D13DD744133D6EB3373F9FDEEF84271B4A925405F7D5F6C0094DCE449A1B225BED1BF4
-987F7266266567C006E6F7853EFF71E3F130ABF8212D2F8F87C79D871584242BF451DC071852
-CE63C1A16731701AE03C1AA570DE687F6848E3A4F15660BD8C1E2C515A9A423FE09EB50472F5
-3AC4D8591DA47F1C5DBFAA5371E80D3687996F555F305C82DC41420542F0D958A658708BD8D7
-599DBCAEB318264D481EF02FC4EC6FDA82F7F83EECF18B75E86287566BDC9090C83A6BBDA6D1
-73BEE814C4DA731E777B5AD7A6E25EB5E8FBCAF89B3136535EC65655C2C57C70505B7A591BCB
-77F387DC6D51CAE11EFE6638BFA683A9A298337CA99E772028CF84F274C4A1A9897F8AE703DC
-E375F33CCD311A6D606F28B410FEEC117982768CE59F0E637F0313EB6FE3A66F0C7183ECE7BA
-00429F0ABB65D4667690BDE070174206A32E84BE8C14C96F6DE5E534BBCDE17D94EBFFEC646D
-FB9F88096174B21A190165988CE9CFD75F8E4D704A18810D4A37A5D2B1508B617E25D1201206
-91A5CC49F99FD0289CE32FC19BA24F4F2ECCBE43A2FC2CDC91C55EEC85552EBDFFDFDF9641BF
-EE74511407794C734C6EA6A96CCC8F1078D52BD4FC4AEC99898C6FB807695F5AF52E25C01441
-F153BC15401B84395207F419F4544F039A0A8A693FA8281530ADF3CAD98F482EC90EC6B78EBA
-051040F7F71A5BE4EE0496BDC4E65E25C9B0275E3EE91BDEDFBDDA97C4A2303AA8F5B1D5CCDE
-DA4390BC637D8BB59E613B4894337514395A4A2DABE070DE0B28F7A5991ADDC42D6450792BF7
-A64488A0A786B5A193A3DCCCF39B7B23A83196776AECB4046EED1068E9101D14C1BCC5228578
-0D78F126A7C33E1F4BCE4D30C8F48E1D747D2B0A874991643D6C473FD36039E82919CA2CF0E5
-805E697121EF2A572AAAB1316AB9F8D300A3AD9B19EE58112C13D87B72DC6CD30B509468F848
-B82936A3E839BD26512E7A7F09327040DC3DBD67932486B1A6B5E21016BAACFD7E2B25110FAE
-B41E1C202DA88891ED4C5817D0089762BFF4FAF7C166ED38BAD867E1069BFFAA9E31BE54B66B
-50EF953549C611E25A22252472708CC06E092DF1D16AA474B8E019E755D4925AAF21C1FB2737
-6A9C8DC63504F0E4A676A1E1D77485FC8AC88E2434963C12CC8FB8A88BDE10174FAD988EFDD6
-027A68FF736867275E4664AE0A98E1A835DC218D3497FAA6818B6952546FDD40FCAC71E26B38
-4CC50B682AB0ED702947607BC124737DF00910F4B2AF852AF3DE65951AA4B8CB4E8DB66225C8
-81AA5CD3F464B387C70C01C022D16BF7BE28CCCA58DD3C5D8153633E76DC9C698F6FE93FEB1E
-53B0C4BCBBFD594C5A9CD857F4BD3812359E6B3BDD62265D78D6B46844955C416FE17E75E7D4
-CBAC35EE3671FACE80B951A6EEECC98C603EB482CE96C358177AA0CFCABF72054EBFD67AE93E
-BD6535B6E90D0346D58583778C3694EBD85F236903BE439C1DF6761A317C333FF4F22D784C20
-DDB9104C4CEC9147FE6910F805C1EFC03381C6E12B8DFF411CC46DB4C74FF5E95995AAC8F6A1
-CA1CB7362C73DE5475E849A1ADB7854AF29516AA3FB55993F7EC38E87287C6B53798F6E93953
-B4B43CD58FE638F00FE188EAB0EE60A75EAE5B83D08A39F860D878A6F96F65770656900C426D
-B6751F93F4509793183C0FF94052493176FC1DF37AA5CF5DF5672C61829E20EF5E3F17D499A4
-291F043C7D953BFD4133D804C25FC86DD66263C83688AC3C0ECB95A1F328856E5A0C020F3AB3
-27218257A498D0AB4A61C786EB15D190140E3BA94440AEB9FC90008BBAF2DDC0E7C5D0953BD5
-C7EDE559628C4246FB13E8BD83D8A38934142FBBAABE150C3BFDD4445EDC1889EAA5EB4B2C6C
-B44B22BDE3C7313CDD998A53DA219B76E67141F8F3E3BF348AA98C1E7B98F8E1AEFF57435D0F
-DAEB29BB76109AAD8B1BFFDF1415AD5EC2C177A41E7D845B55ED3420DACEBE071DD7BB86F503
-2839351A98754AF026D21603C501D43745E95175370674E9D6CB49BDB5A096B314E1822C19A1
-2EF4D32D08FEA6C4DB9CFEB4287FD6B3AAA408A709FC40B49057CFF8CC07E717642F7299D21D
-E2CA1C80708ADF9320DE6128B4774A2DFBC728D93F3C696E595191AC91E17BA9B1967498B343
-D4A5F0D95FB24FAB21464D845D9FE0290784FACE9681F2018BEEEB3301E73E6D2F65949F7AB0
-C7EF82985DA75F0DD4F98ECD63E03E2697CFB22CE00FB7775D9B6D26B0F17341FDACBCB060A7
-204E944179BF5BA6CEF15FB357BF844500A13C099C9FCDC84F1899A8109847D100487126298E
-59264F7A380E42E483C28F9BCF7548A88AFF826D4C7E33AB55B6B00546E272D6C1A3236F3B6C
-91F700097FC221BBD8E3398FE377B77E616F04A3DA9B4693B90B1896B1BB23A4074FFD240B2F
-0642B3F17946F013BE5485077D04C8CBA954BA18750E77B5D76C173FCD6ED20B368C72196C9E
-C08EA9FEAB86199299C024388C987A5ABA2232C2715284BC50CEF1C397857FC6B64378E51883
-D68F06E47857E7BAC5A3E09E91E0EDBA77CC8105FE2DF69D58285358406108C1F9B92C05F19B
-4F9D9C649278473175D6FBBB20B6262DC21C12AA0D87B44FB7A0674527C5FE280E5727F7EA07
-FF7AD11ABCA00D64576D9BA7E1EFE6BC55EEAEB0FC7176C1256B5AAEF87E1736B6336F9588C1
-42B06333F86C8FE38E907DF552D480C6A9A517A5ECD2C5279A4EF9014977F905C4FE97DF84B3
-A1AA37F5E602397665A784B50F4796340CDB26F12236C613F110A089E8EABC83FE10E4030E82
-2859D65740B270FBB109B348A22373865928CC5BFD38FCD4D19AEAF1C22E4A7621F1C7870083
-D8D076B2F6B840E43018C4CFF83E2B643309B1E151D5B21C6EFBF25FDE417B5C0800A6B7E0BF
-49A62DE46EC3F2D6CC13844389DEDF1D267E215C2225A3FD0C447FB7F7C28F051DA34F872950
-A845CDA1D7F6180CB231BAE3EEAFD29421C88411B481776EB2DDB5A1221C2E68E967BA5F848D
-1762F84C1A9B3C170B9EF2A7AAFDC749693720947A7C5DAFA3213F8D610A1F7EE531D8A55E68
-69AE5DDB1555B2F386B5A4740C024F57E6D1A35FB63F950C1AD02FA6C5C111104364E7A8EF6D
-7C98723EC5DABEE3A85487441DF4D8A3609A8B56EB4734361C4B315AD056FCD90FCED54C3E96
-6E8D10AB1B37C1EDB2601DA8D536D96DD1A36E5FC92CB8C4AB0AB3727E79F5C0D730A3903A2E
-F5C468E7BDF798A151BBFE7E4CCDD9839C1B23B25822CC12B7CA5C9629FBFDA20941C1FD7D54
-9835F63D295DD84CD82318EEA112817F70C6946ECF5F76956CB9346808C445C9C9FBB5FEE5C2
-A3166AAFEBF438FFDEABA9FD0C713DE74B905149F12C9DA3558841AD636AF404161BB95B0F50
-C83E133827B53948F5688A0C784A066F13E485FBA27DFA9574327306C9ADB3766664F3802775
-73EBD84040A9F8F3A5BE989EA439D2976B1F89F83C598B36707A0A8A46ABE2495935DC0780B4
-2572B8CE3DB099F808E55DC8115325626EE1C9889DE216A61AD2D4CC5A3236DC99D175B6A9CB
-B17A9C3F4F769CDEEC24809511A94289D3B5F71E9CDB557006ACE81414577D4DA65A33AB6957
-B2D56264EB5BE8E70D156A456CE77F610E2C4FAC7E51878EB3D8956892D4A9016C29A06AE0C2
-4E04549301F25CC2DC509F0213A70BF63E379511DA2FEEB4E6FE4A9A34250EDCFE73E18A9DF1
-603863A63CB3577A02134B0CC4B3AE4F1973010CE00010F53C9AAC4C1042FCD6DC57FF3BB324
-5DEF5201BD5443AB27A25B9F4115A2D95E77BBF21ECEC82A2B573B4704183578CECBAC19CD66
-C8E2B341AF048847BF74B7FAD9BF86AA5AB192D5B1155FB6E5C6ACECD0A572303ED5F7E078D1
-BE931F23189469CDDB1F1D0F778C57868419B8951DB772DA8E24E85063EC1FDC8A9FF5F3D18B
-60B1636083334ECF35B13609ED325E07E5BCE1D90D8C42B79B8BFB43E27E076A7ECF537026A9
-6254AEB5FDC4847E515E356C90BB4350546E5BF865E2E015CA12CEEC765D18A96EBCFAE77A49
-71C1D3BFCA81CD887570DF14F98E13E7873EBAAE474CE7D986D38BEFAF76701F3C1964CE3569
-90AF831015AFE322A357FFAC74B081A05F591EEB2D9E19E93A01830E5C0B0CC896C91C96A5DC
-D658BA4440C909F224993440315709870EF7AE8EBA94F29081A795CA1A4E948A5485D4FF35BE
-9178241157A547F3F406C7E19CDC67A1460654C39973CB719C56DBE0628D864C0935A32BBEA9
-604452BEBC7D21C42718C3E7E4F7883ADA706447A3930D451B1F23DE11DFA2D6936AA34C789A
-73D3FD2B0A4A172A7EA5303B6C10FCA0D8BCF8B77DF7969219C6D98C866DAFBD309745FFD1FA
-6AE150B7AA5DFA9E6B197D5ADAD971680402F6B4FF642B7871C51B439E6ECD41FCF1FFD060A1
-6C867D020EDEF0872E5B69B7E7B6E0DB7B66C1B5963F20CDAAAD3980AD4DDB742EF281A72A40
-5D8DA221425FB22B35D0B6A2D2ED35134D547FC82089DFB966B125773EE1A542688025BA2AB8
-B1AFC05B127DC9B93262AFE5FE1A5A3D96D403C563EB8E84CA6AF775E89BEBDF9E9574A8B8E6
-40F0DAAF1E8F43E89C1924EC2E8CF27375F77444BABA2985CD5886D8C8C5DCDD31D2BA892CC6
-5E4889792BA81F593A5AC3891D02DE4F6E49B181562B3B0A045DC269CFB580C66BCE4B28F692
-3F9DCBF8572CEB4EFC0DD7480962ECDE9AF2A15BA92008A95DC48FB9EAA25E231D48218FA429
-0D5FDBE30F56C6F8E2C7F398AC177B6253D84F50441BFD587F27DB1396235C371B17FCEA8A46
-622AC2E677A28E4D8A69CC3E179A9775C5CCA7C323DE352F1B8DE7A61C7CA4456132188736B4
-05B2F5AC605B5567E0B5574873B3D5B17D53CB5DB167771DC17102477EFB87947A4D8FBDA880
-8A5060523BC03304615111BD80E57B4F0D1E41E79C040A75668522B3A217F5C5D463E8453217
-C6963E64530350BF411A93973BC96043A15765401A764704BDE6FA2D3048E3068C4E18BE9F97
-389FC65BDCF9ECF1959DC62B447D6A700411AFA6D91C1C2490B797AE751A6E5105D8BB6883C8
-F93F23817EC6B476C62DB4B161A9CD7BD824F17DEC3AF0BA8CF291E6E778A4648AA0AD108E79
-F9F79788E3A7EDB2FD5DFAE5B729B440BD5373251889300AF640087BBDCDE774758BF840FD8E
-AD08863E81F6203E19B70A1E58C465F48DB338651981E1CBFFF02A9042D0324A207E79E28F88
-67EEF9AC2D5D56262D5B8A33333D598F67A81959055FA9BCD265F169195F057C741869D3C673
-71B2E22D154C3D3FCC1B53B98F4CDA65BF55484CC4EAFAE1171D9FF659EFA759B65DC133093C
-4B032279425FC745EF552E5F82703CD1590368BBC78166A20FE480E82F9E0B9D793C334B3C5D
-D7122300B82E7C710AF371E70E0426BEE0CAF0B612E142C9D1F3A6B20B7E0F7EF26219045E9A
-6B237B2E7A898CE256EBFB45C144E2D7FB0EB023309E622235B3290BC722D7F2FE5FC7DECACD
-BCF15755ED2E942174C2737478F81AD3C48DD89E9FB930FF25BB85F755D39DC6D55E6FBBA72F
-3414BB89DD6D73C77A4C93557D25416970EF6C0013684067AB780E3310AA2B98A9708A112A0F
-1DD101CC1F8DD8E6F7071C2EC8BC5775C7A3C6596DF24E79841D756A0E5FF538EB731F656526
-5CA05DA718DB0C30330BB765A4D4E9FAF0D639B20635C06466C51972846A208F169801BFCDBD
-7A8B6FFE1E15E360453C9CA21A9DB212D0CDF69F04FC550B29287FD177951AEEB4797073878A
-9FD1657F8167081764CE5892456562071A711787436AF61D7C30288D7BD0E453444A5EDC1D51
-5C55812AB9F10293B98BB53FC20957CC295FBBEFD65C48D82B530A78145DEAFD06656A9C6031
-980ABF0C55CFCDFD0A9CC751F0ACDB23142C8E9B5F3E80B5C99D1B10787658A7AC78270D2FDF
-6C19AD1481205BDD463F9305B36DB7E8A3B1C4FA6965688EC100F09A6C0248F203834747FA65
-94A8F1D9899E48A587299FA2EF3E9DEE0E89E572E04D107EA32D76AEB9AF06D01FD7C859784D
-994D56F7C773F4A11D11D6015D5FEEACC6696E30A4CDD1D22BD7102C920E876CC89F5CC7038F
-D9F2009B312C87840D400ED3BCBA0A10AE0A98A3988C25ABC988D77E5B16AEED350F2AA13650
-E5121CFA98492970FF8D3E058E4FAEBC6410AD41CF3343853A6575E1E128867A0887C77CCEFD
-32F991D7F3CEACE3D6050C42158F567BD2886012719ADB450C9E3BE17F0BE1388827C236D578
-6D70F57F3F5EA6E5CD23623B6A3148B33A04A002CCB0E32FB0EA58851B97A23E85EDC42699A5
-090F5803C7518775B23C18AB8C590CF1E56DE8878200B48061ECBC629F5577D3547BB3CAC7B9
-21B2EC80BE4F483BC11A0F807EA413779C7DCFA990EB16100F6325047CF687A3E9A3B7D77117
-8969CD8313CC2CE5A63221B25DA9C6B84DF9ACCA5617ED3355E8296A91230373622E30AB16CD
-50908548EC2A2AE6BAF6CF7A0B683CA440981B8E9F9DBA573CC26619620E035DBCDF158A3A95
-1AD2A42FC30B8982A1B2D9096A7074632546A0D5F6D9510BF17A69299FD1B63E59EE53850552
-06F39B159A744C180E7D2ED2EA2F23D92F4F8B28C4DEF8500E6C81CB416ED4600C67E8C3AFB6
-7AF7A8EC474F4959B4E41F9876CFE84C1BD0D82E4983C17A8D6F439804F2B8C3C707B4A3AEFA
-1DC46D18EE730A995A15E9466287C68E263E318F1504BA13730B9B0D5AA3B6F9D8ED53E713EF
-0ACD8A12A81B440AA8071C38DD4D4B5BED93FF6C16ABFDFAC2F70C09376CE914BF9E3D4D08A9
-B11105ACDBA227811270817BED33A8AA0670DB212D9FF4464D3396696B7EF095482A9ECF2D02
-B87896E4474014A5E1E33CA651C318DA372E85DF035FDFFD239C7B8F717CF54AF760F2907466
-3E136559FCDD34C8022C99B6A27E7A7983F00D34216C21B6FBC3B4E0672291B7C0B1FE2D5651
-1351B0C0352502C403D5AEAED1D426AEF8C7853CF7A5A778E17C3031BEE0A2DE5E44AEB0728F
-D9C9991C6643D0982F897A8B03BC772421E32DD17BAAE27303C075832C4C49B6408C879BEE27
-73325614316157CB39BD5407479AFED13E02F4CCD56B7173B840665037A4CAF98ABFF3609F61
-A1B3FF2E9730344C59522DF862062C649326C2AB67F5DEE14CB390A9F1465803457DF40C4C02
-449451B474D4FAA670D64483E868F1E93D8D4BA7908F2B06268E2E7987E9A6292EF3E28039AD
-1DABA53F9355AD22F0EC83B7712C7DD9D690A286B74FDDA8FFF281EE4D0245BFBABD03ECAF9C
-35FD0E43336B7FF4FB7515BAC959EE060975653D9A629C1FBB35696F3184EE386091CADFF357
-AAD69F764B5255558DC69FF209CD258AB52CF8B92962EA9D4ED1A3C113F775E80C3CA005F682
-3FBF1AAAD88A2F2F1DDE80C82D557D1EC83E8EBBF183FA0E42E3C2BA003ADF3D5FEBA0A4059A
-F15E2F4A5E30FF4FC3E306615948A30A1A6C41248F3C592BE489996AB8C77D93DD7ABBE1DB26
-9EB67462AAB92718D6C425F1192131734C240E70D154D3106881C3088BFF926C36BFF3282631
-79E32721504BA02C32F1680A17131BEB2218BA0D21C3A337A3298F9571F3C7120FDABEDD6AFA
-91F26AA33F7B99477731927AB90B2D742CAB545DE428DE79AEFC160AE1680EF056CC0AEEA30D
-60C4E66C7B8E7507DF5DF6A0DF6B57B3029BE7012D559E03AE16617BFC7075FA646A3EDC6537
-F7CF987F4F5B7C337C6DB259C4C72B0C8C20E6942D068AD26F5BAA574A3435E30436516000D7
-1179EFC69C96F0F42B2268BFD0E0743740A253535EF2ECAAB5E933292C2392438CED5FC77FA4
-5FBB801BC1A4FD629DA05BEDAC312F8A3B4BA5A627CB0142E2BB8B37FE958A9084D26A829C50
-C2617A0765DE70E88B58D51A9BB7281EEEF68F9958AA83276FC4401CA6C96C0D8B75281CBCB6
-A9E9E91941F3EB68FB725160B8E4FCA470B1AB6656242E7350583DEE0377DA6225DEB1D10D1B
-133164A2812B5C39935C7AF7935A669F591786F6048BF6F2B5508646A113707FE4646113FB35
-523A789E3888AD3892692DE714088EA35888215360EB9B045C4F5B0F7997E4F442454953ECE7
-EE7FC26EEB39E7B93D37F6F1B073BE2FC07072D96A750F646BF8F6DC0BE1AAA446B6D60A69B8
-619C03638FD4ADCFB17A06D8EC6C5AC26DA36D141B19B64902577312EC8F2545E734F56DCE3F
-0CF72425D411F320E38E40CED1A0E588697E4DA951711E6BBC07419AF75D6B69D7601751B75F
-44B98E39522D31C23D08DA41EB5CB5E779E7B69FB26FCE06F80E843413B17524A57A1BAA1D6A
-935716473E937205DEC335930831EA045D7BF1CCD1DDBF35F0E1727946328253B924CEE788F0
-06DC4C7929AB337684C0637574DAC41FCB49AB48B6A2B68632F5117C3364374218FB84B4F0BB
-DC99FA2B31353B239D58BB1B6E4B070ECA0988F9C9613961AA819C2CC6A41023818F2E1B62FC
-FD924559611C5E5496679641669C3A297866B5943C6C55253ADD59DE35F735BA116A68462A63
-70EDF3399EBFA41BF5B311267DE66F4C7A1D6DE312C6730C4FC569EF156789AC3D5BC3E0E0AA
-FDD3D3B38BC03B5A11B9D66D548304F71C737944C8DF3A75A93913DAC18EEAAFFD303EBC6444
-84F65836D78629A5A62376E58EFAF6960995B5251E0428DDE30C8CBD0DDF39C0DBFA26EA9FD2
-96CEDA16307F0FA6E5C737EE4749313511FA0120018575ABFC6BFA05C2690CDB44C4A817A59A
-291569CC347F90315F950EC7190B07A62A401D16480CFFC36B7A6C2FD17C54F2971D95296D15
-5113F06C9240DA01E128151C0E878F447CB3C12640D0E9DE8BF20190AEA64244095927AA2673
-9E78A805A45CAAF5535A9DC66A03D063DE82664148F982D9C095211DD8760866335DDDCE67CF
-3BA3C77AC3E2CB332AC4A23C14F3AE9E645C819B5A5D217D7D02BCE0E485DABE4128F6D29F9A
-703CFF2A3E8253E7C02678E984B623C58F85C90B4A383EE7D54A9D0D48F9BC8107676F0C8C87
-1F77F45837CCBE1FBDD9B3B60C02766CF5A42C9BF3F430C4450E37E90548BBB6CC65A9C15D24
-B5E6375799203EE74451D3545E9BE918DBA7DE337DAA89C8913689136340F97882EFAED95B0A
-72142EB2404DE5284319C13EE52398905B0CA7FEE3674577E1A35436F01FFA6C2D83EB342F76
-BAF48B9B9FB1C0267AD93BE405D505ACD0396BE9F08AD5241999BA59C9F6E1CC0CE022E6EF6B
-9510D354B9A8397FA4527552E6A9643CF47EDC1CB3C6BBD8B8B7F63229BAC2C13F11E9F08618
-EFB6EFE63D78CFF1855209FE240DC0C3E02137AD65582EE49B6EE71A8C779597C37DD0AE8D05
-84750C003D8A6E1FEAC3423A9A313E9689A690B1E66CAC19440CC2D033F0A5D1EB39E634C267
-A058F02B2F5D0051AF51DFEC521A2F740CA909371F670853DB59526DC1C590FADD133C851DD2
-ACDF2AC5497E792BA7016A5213A88557C0D9CCD8822EBCE9A51C11435D6E6EDBF7D3EB4F8AA9
-309BF2B6B2611A1F24766D6B85CEC5A3ED27B74DF514454896949C8ACCB110D2618E585EA9B5
-7311C7C5FEB75FCB7B675CD0D4D6A47E54C2A7F260A636ACCF74485EFF528EEBF42399660314
-C5C5FDD8C7FC043FE0DDC0ADFB485A2B098D027B3389B0AB6B6BC25D6AF04199E0F9C2E9C115
-0B5742DF5B8D6083D3F0FF254F560651B8BC5DC069EA959E52B73695050BBEF5DE99A8B9254D
-0D7E5FB00DFEB35870E3781DC36002B390F5F8B495DAE3044B4369B4AB117B45761661441A2C
-DE3768E8428A9B6A4D00BD3ECF9A1222ED291D4D99247BF98B2EC4317BF13048B66D8ECD7F3D
-CB2E78CFAF484D0682B95BBE275048BEFABEBA338D02342B6A587C86962B167CC2A8FF814D46
-D731FFE53FF589FD437C415FEBD30B36E766AA114FABDD16F3FB1E556931F497858EB75F4ECA
-83599E88624DD5FFCE39951EB0976CCC2CCE87F9848986D2B3654BE993EB7D170D34647DC8BD
-724A67FD5A41EF701E73FCD60E6A18BA28FCE2E5263C21FA90B050B1F45818B6589148EAE441
-297C6C9312A1EF6156524C2769E132EAAE49176DC01193A7641615276EC689C043474AE3E480
-CCD47BAF2CE168D33F57EB98C0C2B05767112DD08F32866DAB5CC7D7DCD3599EBE45FD3512AE
-FE28C80C7FAF1D87C344D59585493B500D83A06834E466A9C97DF3300BB063A22334ACF1F782
-659B5DB9E3FCFB8123AE6F2D270839C7B7AFF2FD6AB5C11E0D515DA14B03D85511DD097E9002
-FFFD3AE925A8FB527A23A86E5D3D5A2E98CD8AC71A2EB8027509209E6575F2BDE70F365320E2
-5D33BDEB84A4D06E7DB348E06F48EE7051EE3CEA94337089F73F9D25BB8CCF3D3050A5987004
-C745CEB6C0D1EE7E112C0CEF6D3C3D06D7D464E8F57A2674DAF8E64FB168AE37ACA58BC64289
-68FD86ED14A589BD712527DD11C82BB626B371C1A5A0834D6DFCAD69EF196A211E3234AD94DC
-AB5F15D4B3FA1182E08E5256217C5DBAFD116A66F897C4B588A8EC1F66B4E23E2838B0C04A48
-358EB0EAC2E9A7E0F38E55F88AD9223365D368087DA4CB58F540DA6A563153B58558DA604A8F
-7CC1CCB3223B97D23F9B7D27EB7BA291545E032EADB1707AAD041C257C460F6B4A6A0854BECE
-A1CF78C70E03EA07CFD506D44F645D2BF1123B86E58DA69E18C32BB6B2FA665FB1177EE410AF
-8EF2F4E85A32D90B89DF7ABE2AC5932CF270C0399192A08879D974CC06187AE7D71287A963F7
-D677A661EE5D5F053659575FE6E7162DE1E006186D75A937A13170665FF4F75C7E252AD967D3
-938B4EE75C4D29F3C20A599E77FF1A09E71F2F42F14D0549DE5ADB6919E323B6CF909E4F0ECF
-2F84FC72D7DAC5B5D6B84B81040C29E7FCB2D54599B049E706075AC6EFF38E3278F63631C02B
-A5B6D2A384D9CBF4E33985DFD37E77B14148F49105F40A447622BBE7139151B53A080E55978F
-9A30BE03B0D77220AF6DBE9153A91CFFF211A45E620A3FC527E80F090970B4A501D162C4D885
-116E473AF6B5EB5AD9345C77BCF42BF9E900202E5E90EEEDB3A81EEF9F2D2D43285DF0DEDA23
-B33909FCDA1105FEE7F1E416C31E8DAED491346E607426F151A743D13479882F482CC5E347C5
-D6605283932B296516ACDFCEF6021F6C441BC7D4E9E766519E583D1FC3DCFE20788DF6B3D8AE
-2A92B860F8D0FE8FEFFE8608B80DDED0C3CFE914BDA32AAC60151D957963253ED528D7F57AF3
-51268823F937AE52A697EAAB5A14D8524DAA0751CA6D02919F8F980D30D0BF065B0692CB7318
-398A294FE220E13CB75709CBE215C64BE3975F0ECDF550FD48E714FE77FA68034530A6B83E6D
-80318984E554F0004136E495C6BF1054A5D3C38A29DCC8978733FAD1C1829E68861C928367F8
-A1C9813857802B6A2F716D3953187CA8784EE421EB4714B5C8DDD60D6032B97338A41D12D505
-DB4CBFDCBAAA78474C3B75F12EDF8D4EB9BC89214C6EE1748C0CD72F5C1ADE69A2229086FBE7
-D205B65B2F2290D931249A47E9D7E302D497618C193549316F86984291A3A5B85DD0D37AF323
-F7C0F6A8F618526885D73C09290F42D4E38C2D1AA4E85EC6690DBA5F8C2A76A429F2A4979C04
-328A1856E4CE0D547E9DD94D2AF3243EF6EB50D680C223247C24975BDA4B2AF5468D453AE8D1
-7932850D7DC544A7ED202069F685BA4FCB75667FA05BCD04993DE2E4D2B92D9286126419CFB5
-D966C6A45239D20DBF701F14E003E66C6CB41B4CB4668A50A86A1212217826E09D391673480C
-94505C7BCC8268F91FEC1368F1448571C25BEC6A8E786164C55B8B8545E21BBA091D2A34FA5B
-B034BFCE18C8E5B521A6DB186BB4149AB0FA3FEBCE3C9B91431B21C661E54960239BD694B014
-CD410FEF9E58014B48C8B746AEF5F6293AB7AC35447F2F0C881BFEBB12B6497630A44179AED2
-57E90384A4CF377499B957EA921386B803710F3F192D9ACE9B5BD7875DF317DDF5A3B0D59170
-2BA6B0DB65AD4771640ECEC04669252FFDCAE35CAD3B5550F31F7070FD1145F6D6814222319D
-6FA189FE76570F5E30A1748AB7D7EA0A7E758B14EC0509919D8A550E42D38B538E4F0765C780
-1DEEAEC5BA234679A13362E407775F932E1C67F7EE467E43415EB1F88D3666A98E6D7C63BC1E
-5674C975B632D1726D14D7B48C0D2E4B2FBA4D1A75A861AE382DEF1F9C055CFA537994F7F7FB
-34F22B353736DE821FBAE95DBAC18D737F5BDAE3779081F6D1328EA9C8FB129D39739524591D
-5271B6C0584D2E8825B995AAB4C518393B594A5DE1C6C63C0ABE0A299F01C1F1D1D4247166BC
-0610DF812D859F82022C037058773F044E03D24610F54EE133ED7498B011DAB47E63420ADB11
-5AE8071B7D011773783B709AD41226168FC17262F98083AA14F841CDA768156B9635A00BF862
-A4B493B1521E035F59B6D03C9B23BD5D94D491F4BBD6B152C59E6E6618A80F00AE817E9A1D9D
-9C0ED134C91F93A3907417AD76708BB23E26E44C9C7F33C774966EF7D21ECF005FA6D9D4D199
-369682991DD47C10AC5AA8475136BAD7C504518F92D6A8C15310CB87E0DE3BD20A82D62DBB0C
-30AB1C9BF4C98F6939C66F9AA70CEC6B1AB91F2B41290CD5C1FBDAD66C5374A7F14CEDC39026
-8BA628AEA41FF1D2A5CE6B8B3FD84629C6E0A687263FF9F69D1FE8BE49DE1EA64193E5445133
-A97EFF00C749D4E0877CAC880993D00784A54859B826AD73BA0F7D5948806DD9D929E066F036
-8D1B4CB24C54E9BC2E74C3B6E8A4B6D836EDAC77DB0AA0C04D63931BB99E1A0A27EE7848C4AA
-B5500C89294F6A917004415A7B88557D2356894980489916FCFF7EE9424BCF9FAB30D9D50055
-8622CABDF15670A270C2C2366A719777FF2F5780E84773A04926E4D537CE72AB8BA3A362D8AF
-AAB532E652D1E847E0BB055B1CF3533D6E1586C3CE54B03FC2C36DC5E0D1B2902CA303191928
-6D2888AD297898E53CEC326B826C0630E527B1171CE3B4B81BD8FB3A25DC4F72C3C11F5712C1
-76B94F24FCC17F1E9A0AFDC2B833D91B1817439F3D0504530A5CE965FED812984F99858FD158
-2CA47E04C7931423A760FFACBDD8183B69CA6354D43ED79EC0A55E5B2270E4C11ABA5BD92B00
-588E383B044745E877B7C838773F5F17575CC324BC03F86D8D781766591FD3721D8429423322
-A0DD560823383912C21ACD5ED7637D194E8353EE7FBB409752ED7024B4E4776766EE0DC23DC4
-075E8D244A37C3CDFFF9C476B033956AA73D12D7A193C3107C3AE53B4ECD7B76C3C4221608DD
-FA773BCAB288EB898B97228A4B1424E67EADCD8C651060EB99A92D72AF3D7F79D4BB26375E07
-3188BF48DAD0782BC0762D4D5D277B58F43E4D7F4DC5D84093E1A7E1F0CF1FB52D85C88C6F3F
-4D5E48DFB2E3458C5D81418E7DFEDDA16DC5B29A018DE634F2354C2CFF1BDD3AACC6AA99A93F
-B92676AAEB6E2539BA28A7240E81CB9E4F8B5126AB1CA38B72E1D2A94A00C933C3CD5A0803A3
-B4BA7C72D7C7C0C4FB97E52EE2599C3E13389B1158111159ABEB300D5FDDCB173F9FAF8E643F
-6E45FC4858CD5773ED91A4DBFF12F22808D63176EBA6E0E405F46242C89218D55A36A9A882A4
-311C893AB6E7614F4DF09B93A06AACF5E7FFBC8D174BEB84C8A7856ED651842A5D86CF2377A6
-50527197AA1570DF1521CDD0BAC06E4F599E4C7828FDC787B5841B46395B4BE221DBC59CF54B
-39BABBDCAA35D4FF29C9F91660832D72565BD92042E968097CA426CB148D85A8A8DAE64BABA1
-C312E7D9C6C4E5BBF36E74978923549B0C72513A7C7EBF89723792AD65902DC8A0AACF042F0F
-111BF07F926549ECF1F5EF1C7BDD27E1A15538E3822ACE4548A96C3AC14A6058914CE6441041
-D80697C45299ACB1A690A13E7AED19B321DA21AA313219E11D3E941DCE8E9A969E8C6548FE30
-08DCDF6D7A4776771D802750DBF30427F14E133AF8FA2251E3CB03A5D271B1918679959777C7
-E3199052C95ED630CC2DEF0526EA3793FDE0C06C8A5BCE8E6E7F0375E23B349A89034C0AC7A1
-3238C2375DE73ACF1D3B0FA9118E0FB7133F928542AA7E51902E00101708DDF600E430D9CC9B
-76F793640A0B7521283010E08232BA92E206E1E55814F00A563E3B7BC90498069657A1082C1C
-F9F1C64A852996FE3771CBC6F622F1CC0281159F9CFB7A5944D26FD07A6913396D05725C5025
-8F11BB252EA00A5D3376C95A9719A0E1A055DCCCC41B9646B57401E0984F24B6226481B45881
-4D1846E64742CF4700E8FC5079E3BED9C370AB805548E9315AFC32EFCF13689647EA4E584A8C
-D45B91CE9186AA9BA44D10A31DDCB0C1BD7B1D15A61221A0729B0864B1073446C8E5159642B0
-2590E26C45E69A5E7B34461D84FB0A849C1B8E2F3B299765C23A14445A6E99CBB881D9F502C5
-7E1F9ABCD532C1B6FEE965534DF5EE3C4F4EC4E3EAE3F43BB57A2AE92D62A9547DC4EBD1AD93
-D04BA97C73F9B66A44B9657C79A490C4BD016D6058AFBE844A1A33837A804A32BB1A01ACBA54
-6F9E58F5F3E56378F78723E01864E9E871FEDC405DAD24662964583C08545AA5CD72BCDC3ED7
-149BFF490C70DEAF9D4A92F192672108C1AA277C630F378AE59B8CD944D8A83E284F875AF49E
-0DF3BB33A0A8D182F35F662F271D1353D98B3226DC7D69A1A8043733D6B3E8C7F934357B61E9
-9381E5F571B108997F8165F259FEC4369F7AE6D5D99C90F3E105ED4CB98358A7697E08531282
-B49734FE4DBCBAD1B3F174119CF17A8146CADDB80A7FC0B4D48A15548E3F1C5B1F5E752B7D27
-02DB3F2E9907E8362314D2FA67EF2325B97A5F3E99CB0277DF76D33E09888523687A18A83744
-82578B601893CEB51F13B034527F3427D0E631A6066D69DA7A00C4A6E93B448456FB0C77D207
-753A923FD92178C3EFEDAFFDFF218AA2A253BAE8CDE9FC1F75122CB43915F2BF2D9272A6C72D
-AB3794592293D160CC88141D206E87B8F53BB90DF46CC298CD8F0296F28ECDF1588A7F7DCB3E
-3174C128CE165EA008D94E10B6037C22A592B83DAF1469EFF8A321129164879713E962304363
-60A494A0439F9D67DEB8909D4D4885F2AD30B4920EF3D43960B8BD8AEB1B1B6D5135AE5B86D5
-D7B14E06314D1332646102951E747940F9D30557AB14DCF9A3E5A40A5003A3E4735D36409202
-0FA90014FCB0A01A8915BA554A7D2A87E03CE38AA0AD0DF81C862526FB3FB811E66C566A4C87
-2819D4378A6BFAD128C43F1ADDD936252B8F7F49F3525A03C6183E9CF41F76EDBDA3E451E669
-C603FE4F45A4BAF69C23F1E26CA34EA01DC31E38357CCE8A86C427001C13EC748683DC39F94A
-9FE04DCC020371FB12E3D57299768CBD92033FECD3ED7726ECA4FBAE87A8BFDA1C1E643B3D77
-9697536FA13BF1C53DAC801ACFAF4428D2EBB813AB377FEEFF8D85BE5A0CC7C26AF5FFCB99DF
-5CDF9CBC2B55DC94535B100651704617C5B35BF5C2840744ACCA8ED0CD0EBCCA25A8AB56E4E8
-81BB9D878AAF7C526BD7595508F08C5A5ECC1762CE4E19E8B3EAD539AEB6B6A534BC0CC75552
-233A68E541DA9FAD073EC318FF54E58EB409B4C2FCC3C6838A833B52E10064877D51065684AA
-F65C2CE5326FE229403D6938DA946546545294BD512C5D04286B77EA119FADA1CFD513D481F8
-BC84F54C0C2B634826CF4E089287EB436D6E00555F59A203D2C76DD3EC587D7C949067FD5D2F
-C79B4C1F325474C0AC6EF84355E73E99F7038D2232EF662299EF56E9EF16BF576249DFFBA819
-C68BF779DE6957ACA86C35DBDBC5F386865A3010067E9C44A0A7271CF8F43ECE370996801350
-8A7FC0ED6319351FDD2A668E61B58307C8004CB45B196758898EC4DFC6390C73EC807C0770F3
-00898123461907A9CCC9A620047FEA5424237347C116144E8B772988FF01B0F350DD7279ADF3
-326550404E2BEEC3E492D6F41D1A00D2809DD0B9995C112BFFC0F9673879FF4D45F00DF9006F
-CDE4D261C3604C6CB2655D7861170997426D105047447FD4FFE92C6444881A8EF7D85A72C2A5
-B8905ECDC9B053BDF4717E062BE6B62C48176EA869495673DEC4EBBE85DBD39C185A9866C3E8
-BF0CD2306015C327DD11565C1974F46BE16D0910F61E574CB4951D8B7ECE3BC15420C226231C
-2E88C77F941E55300BFF7850D9BBEB504CE229B67B9C3C35F9B83992869B98FD5F3282AB9D05
-25387C6E488A5AA7A9A7B603A872DE4048CE4B12859B7670CF6DC00588A87ABAB811B17E846E
-498BFF3C02E2074684C347876ABF57F0284C7CDAA4C8C382B54F5E6ED4133EF271A5502575EE
-18B4E1FD80CBC1E934995928A96627440AB5088181E89B43748FE57816F8CFE853D50A21F680
-8D53940B91E7F1226F8DB9FBD562F9173BFCDC3884F43239EB1202DAAD69A9B7D65CB317C530
-BBC1A868F3D4B3145C30BCEA50DB7A3737E8F0DC43FB91FD0FACFE7C00A16731D202D2CCE518
-55FED733BE69EA6FEFA57AE0054CF999189FACB585BFEADBC640BA11E78B68E7380C7506BEFA
-C6F2C0BB1EABC08C88D25B61221B86C952275CF5120B79E23D0225A5EF8192DD01CD81A84CEA
-506674781A48D16806780241CD093D71DD3C23BD02704248592860353C53773F7AEA81D065BD
-AF6D821945DA7C7F040D4BD2E87B7F0F08F419412DD04A51D672C2919ED51081F8F91C30DB05
-35EEE9653B4639DA7029B4B78225A1B049F0025609209BD24CB573AF7CC475FEA13BEB67F254
-9364E471F23DFC752F70EBDCD762E95531B18A3AC7B4618FC6ADF0AE34E9923B36F8CBC135C5
-226E4B118DBC1AAAD15FD85F72F166E14314A42F3CD0DADEB17205BD2D5C549221DB93A9413A
-013E82D003ABC3D72C1E87921C98FBC2ED9F21E227E6D65FF6281921117EC43B9ED262CF2273
-CFDC364622AA8251ED6FF454EDB636546A1BB037494A1964AF31E806771D1DC14661E71FEABD
-F16FE7E57C97E9DFDA3FFE449A9664F362B69EBBD5F35073061318EF0ADE2260131F0416C7E1
-6A6826E52805356A7E798A27575EBCC503CC3CD57B5C465D50D7EAC5DCEABD52A7036D841943
-0086B753FDAF07D5C806BA52D65F7AC19DF53ED003DD403D22CE7E69544F4476C3DACB13D870
-706F474FB2F41C45816BFE60A28586FBD35AFE688AF50B2A6E0A98ECA4A6F10C6ABF2F17CCA8
-D056A31C21EACF58AC680E840B3F6FF946C584E9C993D25D22754C169C6709F556B5389A77C8
-372B5F40B3E37CFF113B6461A63AF0FB4A7508D151819400FB30E890E4CDE55465072553BC4C
-7E46F3BFEA100849843469E4C659A3E5CC7A8DA97CD7E82BDA11835A2ED8EDD61706D9795065
-926BCC4B6B383E0EA77BD87C92EB0A90FB7CBD609EA40F50E6CF5176233361460A235FC646AC
-D1712780E19ECEAFB8E5FD38F71B8C03335E20972D522DA13B6D6200598AEEE45143F59F73E7
-776070F7D2613616A1322E961BFBF81216D408B122F41AD79D75A5ABAE59B6A5D885D6FF0DD3
-84C219711C074AFA1714408C6C127E937C1430EEE9D3E36F2FE3EFC64316FF89E15421D673AD
-B9B402DC2A2066F8FE4CA8D6F73567144D497EDD8F921A34F3536B724A578CABFE3B578FAC97
-24482744ED648251E78A582968BDD92DC2CBBE80FC0FB454560979E7DB0B95166C368C9BE471
-4C26B018D6F570E8050EEDA594AD4E5F17408CEA67125A113EA83F7E1EB400597544C3EF576F
-104B5FAE4541AFE9274CA18E1B92E2E0286FD3DFA8F93492539533F011BC5A867D7A61F580F8
-A39080F84AACDD28210BC268760164AF019465C828FEFB9FA784F8EA3861277948B88CBDFCC0
-C87A61EBDC7BD11A9442925FEAF03F452BC8630994C4D68B82644BCCEBA186A97650BD2F605A
-196FDE5784706939A207A4999F5862908073245A614E44B620F92E0ACF0CB0A6CEC0805B6E29
-F24AA93646A4CA54DF808A78992116520A43309DB4E20611CD2F8B9F223BD6FC3B8753D0BC6F
-55B5124E09C8B595C0DCE1E07DD90A49176032F37780F144D547991349CF5F94C0EB91A34F83
-16FFD7FA683DB4AEDB96739D017D4500C1A21F104A1612C24B127442FA3540C27B4A62895F89
-1B0973B55B7096522AE6290605DB9C4DC35EFE8D7967167C9A33E635B289B879CE56BC968260
-9A18F95AC9E35EEA99967FB01CA5D235976EA3D417386EB55F444F638D86B8D41F5DDC7D6E1A
-B47E34DADE70714DB4C6E6715E7FB10446203D0423899650C24B5A5338DCBC298B71AE8BDAC0
-DD38786DC598FE4965A6AFC4FADE5F2CF31DABA2CD973EEF151B4C8F044750CACEE9E5899B65
-594EC3025EA6D528F9C3FC5F9BEED56921F64F2F5D7C6345A02BF70AF611876FABB722FE6FBF
-2E1B734F45ABB9085E41A3D646078C76CF2C5C0246DD4E0C76AB7D1AF20D8D3E64A3BB2E5AC4
-2715C1FC865517B9C75CA4B3558EB9443940A2352EDB4CF4A1458BE569DF40A486CB9B325D0C
-3622C932E1D6D9272AEBF2EC1F3336681A2DD1CCB851DA04549F7F04763AB5BFD55C4BE900AC
-5BDD1545BBBF16B8559C7927AC2C11DB9E20EFE94C191B4CAF727D5E19F2F52D2BA3398F3DCD
-A95884F55EC18C3BB44694E4632EE59FAA7D1909D589927462FE5F72F180B7A76B254EFBCAC9
-4A7A946F67EFBFBBCAD1025C89236DE2D4E2B05047266AC3619A2DBA3A7799D4B90CAD8B460D
-11EF06D3EA479033A786A5223EA82C6E7C6B78E9856677286F043842FDF1439023399F1FEB67
-599B48A4776BB07CA78C000A186B1A1D89F4162B57E45982B68F73EAE9B025A604A1E3541BC7
-C9E9E86FD0DAE4FB3C63B40EAF9EF40678421E0D8276B6D61BCA6FA5D8A230D3738E8B699059
-EE23BDDB42CD2E0D4105FEF31AA4DD10F7C667C001C6191948DF0CC876F1246FF9BB667E8C3A
-3EB830DF8BD380AD3177F4DFE88482DF39A1E6B95381CCEDC5D5B6C6EFCB128CDCB5E027CA00
-80E8B8546DECC13ADD8E63C0AC25441FDD4478BF9F4F7F35E744672D99942AE1675A5DE339B5
-CB7D297EF8BBA905FF97497AC913EC362167230F827FDD1E0A5060C50C198BE7F84089E685D1
-342B57DB1078E46A6E14DC9E71E8A89D9194B7363BF946F498C1C4A84069FD6C930409B94799
-2E08DAD262DB7E788C620AD29F0C87A943B2AA58A32C66E68284751993DF94E74C269AF75582
-453115D2E86E8C0D4B18B3B01911F1443C9DEB3C5A77DD0EF50A43D68962F84E21D539BDF2E3
-4C2C48CAC59D8F0C1187D1A2422D3BA7BC5323735208496BD277914F6EE0AF8D614C5D8B1C5F
-F6FA8B9FA69CA5A1E545DF8CDEAD76CFFFB5418B0A97D19DAB3FB7D8B29020702B27A45724BC
-508A58275100AE822D9CC77D7F577FE10FBB419DFE8ECED2D2A3FAF61D57846B76584EF820B6
-CEE6C52587E101A9296BD62B15787541B2125E38307AD2771E3DB54CF22948B85E7492AFC0B6
-EFD85EF615732B9A4A3678A5828E68ECC53A6F1D214C9B1DA7F0BA5465C06A59F89B6544AFA7
-3A84C5522BB5A36507A51AD6FDA7E40EA152BF80F4421FE14FF8C3FBFAE4570F2CA77CD4E9B4
-D226FFE77DBE6661DE331A5FC9F9B7C855FCEEC97903EF41F6CC310EB7B6C6A81694D2110E6E
-4C30D83F227A2C72F02E8E3EE3CD28F5DEB39547EDAB7555D51D080663176530D5EAF85DA93D
-DC0F8FACEE59E0F80D522B7275C78148623BA2946610DD0B9A59D7DFCE2F9AB508242CC52170
-E6377D9F022356B0E6DF5A02B9647F05546F8A73F40584E02D07FF8062014BCE5F4A0A0A82DC
-32AEAD1DF0565A6131C43A6C8A7CC3EDC9733E1FB9C3BABF31EC434F8289B0AB884B2A50EDB2
-E5DF7E1D9E2EF27B8DF4C9B41CE90895384635FC6D0DCEAC02EFFB6F064058D03DB9853DDD0C
-8B65DF88B4305DE94A439C3C7D11964FF48F4DF2E5EF68D4CFC9012F34BAC4FBE7570888DFD8
-54C205C05257BC0C17A9B9C1C75D3F1F5D3AFDA82642AE5AEDA5B4FA0887110DBB0A05954CE0
-D569705F1D8C18281AD96616CA2A0FCE545190B72469562810824EFA8520BDE7D175D35A6AAE
-E26CFAAD87B1D227E2385CF848668F18B5FF00244B1570924FBE414FFB2B9D263B4C7013F85D
-363769E8C6536F85A386E62DB6856AA900A3BC8F4ED6E60EB91630BA2A2519910A21900DCAD2
-DBBF5C47E34D7C8B1280C660F9B0A8060B90D3FB68C05DA9173E7EA37BE7DD8007EAAF34D539
-B82BD72E8C5F051A81933CCDB06FD81DEECDCE23EB45267B2C01DDC68AF70AD0BEE312D1556D
-0A095FD1E8F844E6B5F50969D94A47BF363C1533D265945B6A3191E006150F2BE1E6A8C828E5
-0AEFB8F06FAB6163EE09EBBB9A36E0EE6535A8E302E46F3713444BCDF14A0B9812EE566EBF7B
-116FF46E7A0214D4B02214B7E5E7166129CF3286B63C3AF887F3DA737FC1DC3A85414946F948
-117D21B4A5AFD41636FB0FCCA9859D1F80267DCED9986CF1764C4F952D727361E431E84C78A1
-852080DED9D940469ED1ACA97FC69A9362C1F92B4D287BF38CC2DAF6370FBD81726A47D7D7C3
-8FAA20C17353007ABBED7B614D6D2F12EBF40F9132613457897EB19D0B416EBB2231D1D496EC
-FD6D69F75251ABD0D53B02BF3E81727C37E40CD46643CCB92F7B2D4D84EF9D29C9ADCE2C49A0
-4553CF8938A2DDC712069171FCB8D55C77C2C2FFECAEBD0C8D1B7DD5DF6766ACD516940D5F8D
-DF544F16416B819AA276411926319BDF8D3539B6F2B0C11D4B5EFDDA9106178148085CC581FE
-6BCAA1235DA6A1A59786982873E6D3ABF454CCC8A5A57D75882FF12B01A652EFF79D2B27B8F5
-A06E5F4E319B8F4E396FC67DCD1D078E17DFF3496187C1A71DB732312FC69FE9B05B2D0D8621
-7ABF3383C0FA9016AF5A229CF814697A47DFE27D39280B3356828A74676ECDE78AFA206521B4
-BBD41A82C2D89B8A1CC5C850A1FF6DC677F35AE7AB0BC010DCE6DD20E915E3EC1F3F08389021
-33FB3658DAFE41DE26BBE105771AFAF1D979199F77871AC2E170E569133CA63558198F0E9B72
-8C046BD63F12FB30794F2D111852A896D4851FBA0E3F6C0E5AA57C13F744DF9173851FEC81BB
-D7E1E961CADF3DCB7CDB34A02A277F52C805CC6D4E0E0337D25778D7CD8DA8CC2F51FA4A02C1
-B79891C609F7AA402B322A201DDC1AFB8C28B46C3326A6B95A8478D5DE4F9C77CBFB81AF47BA
-4E9907118925E4B5AACCD999ABE202171A5BF5CDB8588B85BA11C015AA894C7B6C13A46FE315
-4F7452DA01EA0DD833B7EE9EF6F56CBE8FB18CA3706825982A5220BFF150B2092D3CA8B3B188
-05BB84B2AAB97B76073D461EC6F2784A69F230F4DB3763157C12CFADC53D463DD732F2934B46
-271DB64DA625559B31C7D422EEF1ED5A35CBF1F37958EA0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-cleartomark
-
-
-
-%%EndFont
-%%BeginFont: CMSY9
-%!PS-AdobeFont-1.1: CMSY9 1.0
-%%CreationDate: 1991 Aug 15 07:22:27
-
-% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
-
-11 dict begin
-/FontInfo 7 dict dup begin
-/version (1.0) readonly def
-/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
-/FullName (CMSY9) readonly def
-/FamilyName (Computer Modern) readonly def
-/Weight (Medium) readonly def
-/ItalicAngle -14.035 def
-/isFixedPitch false def
-end readonly def
-/FontName /CMSY9 def
-/PaintType 0 def
-/FontType 1 def
-/FontMatrix [0.001 0 0 0.001 0 0] readonly def
-/Encoding 256 array
-0 1 255 {1 index exch /.notdef put} for
-dup 161 /minus put
-dup 162 /periodcentered put
-dup 163 /multiply put
-dup 164 /asteriskmath put
-dup 165 /divide put
-dup 166 /diamondmath put
-dup 167 /plusminus put
-dup 168 /minusplus put
-dup 169 /circleplus put
-dup 170 /circleminus put
-dup 173 /circlemultiply put
-dup 174 /circledivide put
-dup 175 /circledot put
-dup 176 /circlecopyrt put
-dup 177 /openbullet put
-dup 178 /bullet put
-dup 179 /equivasymptotic put
-dup 180 /equivalence put
-dup 181 /reflexsubset put
-dup 182 /reflexsuperset put
-dup 183 /lessequal put
-dup 184 /greaterequal put
-dup 185 /precedesequal put
-dup 186 /followsequal put
-dup 187 /similar put
-dup 188 /approxequal put
-dup 189 /propersubset put
-dup 190 /propersuperset put
-dup 191 /lessmuch put
-dup 192 /greatermuch put
-dup 193 /precedes put
-dup 194 /follows put
-dup 195 /arrowleft put
-dup 196 /spade put
-dup 0 /minus put
-dup 1 /periodcentered put
-dup 2 /multiply put
-dup 3 /asteriskmath put
-dup 4 /divide put
-dup 5 /diamondmath put
-dup 6 /plusminus put
-dup 7 /minusplus put
-dup 8 /circleplus put
-dup 9 /circleminus put
-dup 10 /circlemultiply put
-dup 11 /circledivide put
-dup 12 /circledot put
-dup 13 /circlecopyrt put
-dup 14 /openbullet put
-dup 15 /bullet put
-dup 16 /equivasymptotic put
-dup 17 /equivalence put
-dup 18 /reflexsubset put
-dup 19 /reflexsuperset put
-dup 20 /lessequal put
-dup 21 /greaterequal put
-dup 22 /precedesequal put
-dup 23 /followsequal put
-dup 24 /similar put
-dup 25 /approxequal put
-dup 26 /propersubset put
-dup 27 /propersuperset put
-dup 28 /lessmuch put
-dup 29 /greatermuch put
-dup 30 /precedes put
-dup 31 /follows put
-dup 32 /arrowleft put
-dup 33 /arrowright put
-dup 34 /arrowup put
-dup 35 /arrowdown put
-dup 36 /arrowboth put
-dup 37 /arrownortheast put
-dup 38 /arrowsoutheast put
-dup 39 /similarequal put
-dup 40 /arrowdblleft put
-dup 41 /arrowdblright put
-dup 42 /arrowdblup put
-dup 43 /arrowdbldown put
-dup 44 /arrowdblboth put
-dup 45 /arrownorthwest put
-dup 46 /arrowsouthwest put
-dup 47 /proportional put
-dup 48 /prime put
-dup 49 /infinity put
-dup 50 /element put
-dup 51 /owner put
-dup 52 /triangle put
-dup 53 /triangleinv put
-dup 54 /negationslash put
-dup 55 /mapsto put
-dup 56 /universal put
-dup 57 /existential put
-dup 58 /logicalnot put
-dup 59 /emptyset put
-dup 60 /Rfractur put
-dup 61 /Ifractur put
-dup 62 /latticetop put
-dup 63 /perpendicular put
-dup 64 /aleph put
-dup 65 /A put
-dup 66 /B put
-dup 67 /C put
-dup 68 /D put
-dup 69 /E put
-dup 70 /F put
-dup 71 /G put
-dup 72 /H put
-dup 73 /I put
-dup 74 /J put
-dup 75 /K put
-dup 76 /L put
-dup 77 /M put
-dup 78 /N put
-dup 79 /O put
-dup 80 /P put
-dup 81 /Q put
-dup 82 /R put
-dup 83 /S put
-dup 84 /T put
-dup 85 /U put
-dup 86 /V put
-dup 87 /W put
-dup 88 /X put
-dup 89 /Y put
-dup 90 /Z put
-dup 91 /union put
-dup 92 /intersection put
-dup 93 /unionmulti put
-dup 94 /logicaland put
-dup 95 /logicalor put
-dup 96 /turnstileleft put
-dup 97 /turnstileright put
-dup 98 /floorleft put
-dup 99 /floorright put
-dup 100 /ceilingleft put
-dup 101 /ceilingright put
-dup 102 /braceleft put
-dup 103 /braceright put
-dup 104 /angbracketleft put
-dup 105 /angbracketright put
-dup 106 /bar put
-dup 107 /bardbl put
-dup 108 /arrowbothv put
-dup 109 /arrowdblbothv put
-dup 110 /backslash put
-dup 111 /wreathproduct put
-dup 112 /radical put
-dup 113 /coproduct put
-dup 114 /nabla put
-dup 115 /integral put
-dup 116 /unionsq put
-dup 117 /intersectionsq put
-dup 118 /subsetsqequal put
-dup 119 /supersetsqequal put
-dup 120 /section put
-dup 121 /dagger put
-dup 122 /daggerdbl put
-dup 123 /paragraph put
-dup 124 /club put
-dup 125 /diamond put
-dup 126 /heart put
-dup 127 /spade put
-dup 128 /arrowleft put
-readonly def
-/FontBBox{-30 -958 1146 777}readonly def
-/UniqueID 5000819 def
-currentdict end
-currentfile eexec
-
-9B9C1569015F2C1D2BF560F4C0D52257BAC8CED9B09A275AB231194ECF82935205826F4E975D
-CECEC72B2CF3A18899CCDE1FD935D09D813B096CC6B83CDF4F23B9A60DB41F9976AC333263C9
-08DCEFCDBD4C8402ED00A36E7487634D089FD45AF4A38A56A4412C3B0BAFFAEB717BF0DE9FFB
-7A8460BF475A6718B0C73C571145D026957276530530A2FBEFC6C8F67052788E6703BB5EE495
-33870BCA1F113AD83750D597B842D8D96C423BA1273DDC63EB43F34FA90EA73A234AF35FD9B2
-4EB74F19626E9124DF7ACB92629CC87B2CF81C4A89B27447806DB34393DDF0402958E2DAEE0B
-4BF315F9AA0A2650D2E11EE289FFAEC940D0AF160FE7F29411C1C45869D1AC3BCFE2AE4DA724
-22549095F1FC9FEC606289316F106E719A4AE2FD45EE868FDDF07A80E9A626BB43CDFDAA9DCD
-208B88FD99AF20B27F9D6B5CDA6A3FA88D5ACFCE8B5EEBB38854378CECC3BFDAC526514F1BE9
-FB28AE628918BA942A3E05B78C3AFF808CF416A5ED01BE280231B75A7AEE0C66B0549713BEF6
-B3CBC188E625BA827C23C9F98E9ADB21719C2782EC86CF2F3A3B9DC4983DFE6FA8FACBFC6688
-7FB1FEBC30181988F5CA22DEA95EC6AB2AFA527C7FCD14B7952A2A0BBC8047DD79824D485FC0
-696AA1DC16819993EA7CEE925990CF208B1663ECCF6405381FC2D0BB26CE5FAEDF0ABEF0633B
-1216B4F3A138742AE422B81B1589EAFD2F0BFEE3444689461F0A6336BB7377830169D5F4277F
-3F231C8A779CB242DC73606C74AB947F0A39CC1EA1FFA441515210CD7E40D8541A96E5C22526
-194D88CA57A5CEFBE36576F31FC584EBBFA8EAE365B545796EDB64080D4CAF6D77A2435389B5
-F41D3658CDB2A77ED6A52E4C678B4FB66F9D9FA73F7EFB5DE11582821504D6E019E85ABD2685
-6D0B038528DBC0B361453D4CDDBEBACB41DD6C99EDB10E367D10D0A6A945EF3DE4668B5D5473
-B499ADCC4E19A65EB00BC4D901FB58D32F597ED5FAC20FF128BAAB7CEACF55A1EC4B134B727F
-51CDF9F4872302A930F1F1F10BA883FACB0B4961728F77A80D02D1F41BC197FFC7075541ABCA
-84E46FD5BE0F344D9BB27F9363347DAB71E5F7662551D0402470EE97D77AD369E6255C3567CA
-3ABA026DF7114B20A2F1808E053FA31E7F5B59716293FB106FFECA39FCFBB39419B33EE8BFB6
-D8C42CF02B13A51EB0BE99B23680DDBA5B0561FDFBA96502D5D66AE0E40247A1996D6B498871
-D21DE813CE776490E30870BC563E94214EF315A119224E0507BBC245FE0B59535D6A77F009C5
-20D1E3BE8D88C38BD1BCF5066233B48C108DCBE9E9684AE8A799C0B5EEFD6E4F0737F553CBB4
-0EC5091C4666218137E22DDC2093D4AA874907C89915E9825063FC68AAAB12946721025E4AAC
-2EC295E27E3D35844A4B452B2D9BE5C24BC171744E3BD2333DC1959B39F6A422C1909067755E
-148E0225781320B0EF70058D9038D4F5A12DF8E1A563271F2C55CE87A5BE074F9B29F861C651
-D6021BF253B2206DDE18AB986521EF272A4702DD3CC285F0EC3476D2ABB40942CC8FF0789180
-B6F4027A7F6413651A729B305099225B6F909C020E4506776437796E3D58F041FE75EF5C2110
-D9A51C5921F6456095124A293E244A013B70A870A38F52EA8D56E5196E3E55A6B4EF08F84F09
-923A77C14B3AFDD86AF386BF3761B371A1FDA2359C0DB444920C01DA74423E0DABE58895E7A2
-B65F443540713EC3EDC7A807D56F36D4B7B6F1FE12747C032F0887FE814421C09E368EC8C740
-540459873E403A213B0E9CB411A6E8EC67428F1D2EF568F10DBC94BDA1BDBC39F81F58F7FF4B
-30786894856C8F8D754B5CCE9540D3C15B730D0218F0E692DF1E92D445428F1B685FA4A5B36F
-6173EBBA9E90B932F66DEA4B38E59374054EC2323B9332BCFB84C008B1442AFF6D1BB28991AC
-4689BD163B9E79BB1FE0AF42D6289F8C968994CFB4F4655903B88770A00327CFA335745E3640
-F6F78A66C3C5011FD2C86136808C61D62EBC76DD420111E45509EE5EB3E1CB4B3F9A7FC4BB0D
-2734BE499512D059E191A77EF684B578114A53D06DE66A9B1521716A552982E0190DFCF61502
-4BD859A882327E214F26B1B1243E68E8B93F1C143EA913CD70470606C639232D2C6E9234DCD9
-2F0D303A45599BC82D824B28228F4FD4F16229AD8B2E27DE30C5608EF7F6C40DA0A7705606D3
-A06E0020A7C5D6C5562403CA468E4D66314456290FE657C202544E53B63950D413F3BB4E1D20
-6ED0B244F323B047459477B7140575A28AF141B192100B9E765CC2BB27D886BAC426546CDBED
-EDF0B7E06AEB18AE6D4499E25629366DA9275A2CDD531211CCE2DC40451E3ABCBBC3B859982E
-10A6832872DA625F10596C05E8FAEB24BDA12180F22C3C7C52CFE9D6F53144778FD300B8B1E2
-B886A977B9EED87C49CF70AE2EBC85F2236094DD73ED14144D85FE4E02E24ADED098F44F1EB4
-B5A9E0E67D15FB58408F85258EFF2F06A4E239ED4EDD6E8D91DEC719C8145DA7E533A8B7CC0D
-55AB16A9B517787470E1166AC2F5C899CC7594AB1651D4BC0EB399750997D8711773E609CB89
-CF1C27FC4A9154697EFB6CE1616AA105D0DF2DFFD30AA7E89A9FB633756BF609DE90506C3DBB
-518F42A6066048E12D2972E708101AEAD629EE8E6CDD5D164F9E1ECE4C763580EB6D9A83D6EC
-6F5994D658E3CF83E81248932577B91531288667C219916A258F27AC81DE34F1B5CB6AF41E07
-CB833A0E8BF6A99249CC434C34746D0F6072E8779E4B9583297B1C16E19CA066FCE82BBB2C34
-D821598B2208EE678770E95B48A4F569664FAFE7FBD368D1E5D09B69D1DF73099205B93F2174
-D95C435A9DACF24071DFDB86F6D1E19FF71DC06C48561538E068528E30CB949801378860E2FF
-6428543D7BFDDB1C8C8A392745AA37DB6E341931D4B99CF9B8E22E31BE25F6CAC27F2CA4C4BA
-E74FF7586D8819D6F81A986A5FB512150AC9E4A6BF6FCE5992EEE58191E7283FDD1BCF67D069
-3607CEFB968A361498C783D28153E1DE5424BEC4032349F61036547E37BA99E59E96F8962EFF
-06784AB6E0340BBF0AACFE04E12BFA5EB89490587DA665B466A0A2F3EAEE78F1D4C0EE8DEA69
-7CC4DA7EAB852159A707AEB4AD5D46B8EB3C7A52E77FBB683A8DF45662AEEC32E9254696C259
-1A27AA65A48F9B609B7FF686B0C516A82F2F3A856F9CBA2FCD485259B9C4C71F9A2DC8FE5459
-4E58941176A7B4714DAFE21361697FDFC349430EC5AA31EE3AEC18C40A1D3CF225BD5989650D
-05D6FFBCC63824ECBA52FABF1AD9C8466CF1254423D999B1A6B290F4D68D4058BC1484F04D45
-1AF4475ED15109CB3B3AAD319D60181B1D03602BFF1D6C1C132AFBF54E1B91007778F7044260
-078C8FBFC7E6FB22B9B7974B3B535D27F434CE90F29E29E3A120149635595CCC832F5AB85385
-A16037C188C2E7C3AB229DECC132AA637FAC1D6DEBB8E49AF3B30E416793FE1A544B36DD5263
-28D44FA684B2D3E092FD2858C77E9273B157A0296D72A564CBB6B09F9BAAF13F79EA7C6F0230
-E53975CFE29A317A0E346ADDC61B2BE668836DDE0AB038B10D099FF2734065473139FC737FAD
-1DCA597923D2A541F672A37BBDC20AFE015777957DA771CDEE85D788FC320753B0D753EC6FB4
-2ADA7D6C912F8385BAF497E861D1C3E27B422F69894D7BAB83B17F9CB3AF83B4CD813DCD6B29
-B61BCA6D780695BF3FD41AD0C86C081A2BE931B1BB038F0B720C1B7B944BB5A2343C65CF0547
-46F17BDEA90301968D294BCFDDA9B996AAAFAB2A56FE05B9FC6020D35C465C49F2DEB486DB89
-9AC323295A7EBF6A6BA54A455AC9C40778E20C0EEFB2442A7BD9D7C423CDD0EAD44845EDF403
-EC3BC3D6D2CD717734C95240D3EB7C7EDAD099F11920D053E81D887B151BF10C59B3EEF406EF
-0BCBC9764D603B80A7373511003A1FFA9B411AF8A7552FB0E830D4BC9FB7019CA7F477489FD4
-76B4E5CC71DC01341F93160C8951E622E79CD953DC8AEB016123BA71D5104525CC03FFA37310
-57FA0FDCEBBB97C0CD8415080EF05459BD0AD8DC0D03878DBE98A07F0C55D2EDA9BD2B8C36FA
-757F94B538ACE2E5D10FD0CBE23D83FA8179CA8333540121300000283E9C2D3CCB5C6467EED4
-95F927EA7F77552E22FFFC27CE2B4BABDD9255469BD13BB81CA5E998D7DF014FA8EBB880AA21
-FC586DB5BC9E8B722AC26D730F74CFBDFF0358AA7FA75A1EDF7D14E9EB3AEF23C78E52983E36
-F98FF47FA2F897AB454A569E1EB4DE93964384F119CB8914568FEE7D064B35C42C01B66DCB05
-6EE6C39CE49E81FEBD6793290886E836F8D23E89D87A220B9BBB45AFBDDEFE4D1C8766E86386
-21E6BD4F785032C15E1E74F8F5FD2612AA6917DB64E8BC7678B9EA2B3B5EFB49EF34C829D497
-95F9B6101F394B142B54C5D7C128B58A9126E1848225DC27185F7BA9C3DB09333A4D5824EA9A
-F8527CB48A315FED48DAEECDAAB29BE2D286FEE06E4605171D5AD9DEA414ECFE793D6E303864
-72499CDD88A1758BFF59FAA4577F3DD75D2C625337A33C7A839B50C4356714424AF153A4E5D9
-43A9DB4849EA5A768F8EAE150CEC5065BB3AD73F33E86DBEF065E0D268E98C964D8D5493C710
-90195274C3D5C13E58D9A3DA34C16E0F7254D5BEE3A59B99C36E261084B4C5A7DFE7EDB357A2
-B9A3E6144DF5FF533774F22E97055C23FDC200C5EBF55A4E1C9D4B31376BB703C802F1A6E72C
-A288D241C383F09238C6B5092A7C4A4881AADD5CCE18FED9DF4AA25D7EC8C9BF79F86AA2A5A7
-6FA09BAAFFCB1A538491BA9B74865F5C3BD715B6F8E61E1CEB54CAD601E3CE91EB09D3E9CBBA
-3D8080C0A70C43D04758021834D18B3B1509B6204ED52B3408D2ABF63173C48B137862A62098
-AC93CC52029855A98C676B80E3ABA918BBDC47C9A7C75CF4A9EB8C160A417DC0179CFD3F9643
-80828E1CABC9258CA22888B50A5E92084EB60F21862CD49898794F5510F04AF4EF7BD021E54F
-AC2E8E51245ECAC62E4673AD8709635FBB89FD2A1B67CFDFA65F64074580415D98130B430A05
-57C44B565DEF83534AA0B389FD8919E7EA31DB5905187671DEFA960BC43E193D402865FF8324
-6EC96137B4C8D54BC69E97DD3175E260339390BC5DBB8669AEB50FBB52C4A6B8D56FB982F80F
-DFD7B1802E179741A021043F179D5C867321FDE1820C9B8F56C0FDBF12A9B5B0AD988887EC72
-943D8E8AA508521CB72381846697C09E46A082CA77CBE2DB556D4D24D3250BF5F4BA10BBBC08
-5DE7BA6FC161E18380E671FC15813B1ABB7B88A7DAF87C705329469E2659A56131001C4CD3B8
-8B62C43359FC38788C30296F08751BD034109E4931E395E5A6E38126ADEDC01E12B99D65070C
-591AE1D306A5CB814C6B6B98DA11CDC736FCB95C2A3C393ACF1025CE94DDFCB9F48091815385
-7F0A5A5741DFE09000D0F6DA9D95CF1CB72E1D04362CA54F0E982EBB8D0BB7304C5DA7AA6374
-CCDA7628038D1C3FBC7B49B04DA3274F47E0B630B57AE019D0FD4AA917CD46562A6278CC43B7
-5C695EE04F3A1F318DDEDA671A1872E6DF5E55B45611394AADCC45269BE9CA76BF2C44CA39CB
-0F8432E5FBD6037FF267BE21B2BF6558877EAA1B3E54B668021FCCC07644FCCFB0B43255A94A
-48D626EA2BB20098C6EA743AB752332531B5E8B1BC8F6B061DB8CD7209A7D6CE79EC1961706E
-5FDCBAABC5B55402C018ABD452C2CB976B71CCCFFB4B57401A0160066BB3ECF412C779FEE337
-F4B477E1EFCD4A0F7338E55FF93A9472051F4CC2F53CC8A2A694A861A033DA19DCF7DBD11802
-50A442E9A27B8A2EDE87FDF9A8F349DB9CD22B66EE990EEF9EB1AC364900C634B5FBB9C5C6DF
-057E254F2E95CC8261BE32A771D197DFC18F9088A65158B77F8AD81F60C6F74688B7107E6CB3
-C426446EC9101C338119E5C9A0B7E9B17B85C4C25F0E2DE47A55E4813BCA52BA6F53A2CC0B3A
-E230939AFAF657FD85313A26FEA45B3D930CA24E3DF13E43CFE1393EB0CCD3888A4C1EDAB4A9
-16BCAE548281579A44078BDFE4A81499615BA7146EB71E67687F01245853175AC0FD2C740AE2
-3FDBE55C339ECE05B2EAE5594F993F0D127D745420C5028F20365EC46A1F0F6D136FECE236CA
-5D511F8EB576FABF0E9279E812A42B6D57D8DF698EC35EA53562CAAD449B048BEAE2FE1EFA98
-A9BD531671E6AF97CDB8AC81A472CC87FD4AE92ED2019E2D8AD7729AFE01A4647791D8C95846
-6ED4DBA14A86E26AA83351F1C5EFC20174E4C2371F1D0C5A9768BBE26E218E343BB94C35D4A8
-F5AA51A18A75005B583DE5172B842097AD8EFF20B866B0FC14CA9C040851DAD663DCE4EC7297
-9543EE0CAD8F8BECAA466AF04432CAE16D348080573E4E43B244AEAE5ED1A4F411FB6631E73D
-FB837579CAF434210BE9F6F4AA5AF350AE14933A6D315EC86C75DE6EE4BA25E8EE8382BABBE3
-00F8FF48B5092C83AD9EE9A02318C377209551CAC9CB575683C30A3F1B0B8E153F860DD5BF09
-7A3630A2B5E16BB4EA0B3489F6DFF3003F7941C7B6D65E93BD22C4328DD2C318C9FE2DBAB677
-2EF861FB7FB7DC1AC4958E942203418734BFECEB0A95F6E798F7D5292EF1BE205FA592B52FD9
-BD8BDE4955907ABF1B26DE31A09CA50DBE3B2F0C9B1942AE87AED1AA7518B271B9EB352F9C46
-957B328D9BBE2F0DC9F7C6B4329316275FA35BDE6F904B5624D364D23DCEE49EED637E43049A
-18BF2D8C6A119FAF128E71EAD2832C3BEF95710D82EEA2D3D2D9FC2A77C79EDB311D0388B651
-89492BFEBE4146AD39B0529E78BDCBA33F779CB5E3416A80E455BF0B9F9E5428D98C011A52F0
-3DA1442C0B9F6DE21827D004852442D51ABD2D4A7B17862FA6F375FF5DC3B94E2E88B47EFAF0
-EA84C00CC032622BD4EAABA30B0A9CD9511CDBF0C4751CF421CF74BEFB8B1DD81D42152FE5B8
-CF6D987AAD0FD44C3FB8E25732F512A10A23522D2F6D35270BD5925AAC0D4AC930A7E2ACD799
-3185A2DD084F6AE3C41332C40EDE1F22DFFD008819FA5AFAAFADE015FC821FA36625772405BC
-7E995CE0551E1B0A0BEFE8EC3F4E3E1528DBC4B7DA5842561848570167FA2B20CF99176D9175
-2A5D897216ADA21CAF1D0B954FE6A631CDB45684888FE366C3C4ADD8B6A2B3EF76543AA28B66
-19BA29253A8DA942081A63FF830F2500942C65A6F7248ADEAD0F19AE0FA1268A4FD68E33BE0A
-C1D8BE4223086043A96746C7E1F2ED18B6B46FB5775B3BC2022C56935040E6581424EB4B2294
-912F8D7C5ADFDB8438A5C57D7DD4A78105DAF771C0F6F445B96DF66A33B550EC3D1F90851DED
-DE5681A8826EF6969192F99D4B3715409700CB201F977344D84484A688659C954581423EE028
-855B18FF246BD585FE868FAD52666F26556CEB3E506B844D88D42A16DDE434AA4E31B9128443
-8A9F78F8CB488FAFFBE492E66EBDB881B67AD79AD11B5035461A60246C45210A983A097161F2
-24EA8D0D588B0C82A3FD1E9D76934FA221E88C61D22D17193B22D0286EB7D1CC280603C089F7
-EC1E34866EB172F86470A9368F6D12AA403A5F7DD429C877E91195EE19651ECE85C2B27F7822
-B6B36A34D0A33FA7F7E5B5BB9401058F04AE30AB700C2DC95BE029611A0933C2FD41E1FC0017
-B3A9AC906A5229E4E55FC97FE0822F2291F21CD51730DCDC55272E786255BD2636EB76C86B82
-6E65897064C776C21472F4C2846647925713998FDD8EB8025D6879FE259EFB46DD34AE5A5863
-D5EF76CE873D9127CEA0260453DFDE8BE0ECE48F75EEE1B3B79245FA5EA93E2FBAFB78F794A8
-E21C7A89BC5E9A3ABA93D52F7AE3EC1872F91062DB372C7AF1EFAE6E3B5ADC4CFB0169710D3B
-CC46631B9BBD58CD1C99C01E2CC95B74E3418631245903A6CAE538C8609F3F9773868A95CF08
-C5EDB36695AB8E8FD74204EF34F03EF530B79360C34B4A359B8E17401BB2115482EA6CA6929D
-16D3B06883B907C800D6ACBE41FA4BA2223A3EDB7E2C90D0179E9E20B4A18C13739553436019
-7E83E759DC15CDE695F073E9E96CD131CA87E44FC8E45E049B9AAFC8835482570DD872D19DD1
-58DD646C431F60D5C5BFEF48012B7D89618A94A33E6F49FF72C058C042C1871DC552D430AC51
-079C496A60CE005B187DE662F1E78E2AAFFF90DB6BF6948FEF33BBFC38D56F8B104A28C9A1FD
-239F52EBD02E4AEAFA944A0DD7C5ED895CBF11D23C860AF0BC482E2CEF8CCE58857E28B6263C
-CEBA1A6C0533CC8E20AD3D07ACF6DC69698C8CED7C2915173714F6B7A6CAAE546FBD14FFAFF7
-0CD63471B9B9CD660C89AD66ABFF923C088F20AFB06DAB331CFD642D02003AEFBC3C5433CBBE
-FED40AE1F11D6D6379DBBD5E10EC42A60BDE3BE5C4D904446EB0E17E1C57AF3D4FB9F0666FB4
-5678D598965A2F3232118FE53BCB6FE4CC10EDE002A6C3D1F5248ECFFD853364034C12C5BE69
-8C942800BFD5DDE2E9D80B521C4AAE502FAD878A03907348875DF256A86C38CF5C40127A9274
-F959DDF0A604FAB03933FF7C5DD967EA143B14C0414F2646D3D96026C31BE5B5FA24AD4239CB
-65667B84167C4927AA6793913F138EB1EB68AF857563042CD300EB157273FB4FC64976FF9647
-5C900AA939C59898698E393462AE4196D32CCD2E688FAEFD515F8869C0CAB9708E12C4AF9B90
-FD3979A568E77282C015B3A2427F7B841239812AB33113F68A77138AB3A43DDF7152091BC622
-38EBD112CC5258290FCE54D1772794A7830369DCE1D742AE56F926B2807EBBF14F4035BD6F47
-DCFC1823833F9E9A2C6F6434F0CD3C3BD875286350ED70A236ECF6D4ADDD4E9C95024E75FAF8
-D16DD22402CD95B59C99F6B933116FBA8C7F8E1C110032913658330E3C283E97AE5D621F3F15
-E48EFB86441C33693225252B6E084D0A3711156955EDC564FFDC2E8D4AFA588FBF44742A9EBA
-7ACF477ACC90272864535ECD9A38229713CDB6D7A34A1B6997936F85F533915199A8AA1821F9
-6284102BA39C719E9A34C2DFFBBBDB05D2E23810C209066317A2F470CD759FDEFFED0F0E9A9A
-FC557B76A094DD659BE6E35C61DE2CE9D56E4024E97169E3988F42F1DB4504796DD56E10F211
-49E65EC6BFE7182EC39A312E1D5870B91AAE0F1D11635D80657ADAC08B885E2E185E4657AFBC
-C069C337D1B03FDBC3078A9D39FD48BC5B12D08DC5A1B4D2E3665CE5A15CB780FC7C58F0AFD9
-08C6C37375C726EFC5C4C6EAC60AA862329EDC01428E97D7512F61BE628027020CCAABAD392A
-3F4409EBE2DF498459E7B60656216AADF9D591860C56D7C743AB71FF85EE652D9B6B2AE96119
-C627752534C42C5C781A76BBE0C8CB87E4D7A3BEA6FD1E56A845590AA930CD2D02B895687A1F
-5182ED826A88CCF9EFFDA2A7C91AC34ED7E62B895E694FF123D80ACE7DF2ACEEEC3253DCE5C2
-4262BD73BC1E12253CE4611DED198A5E075C540099E351A335E0EDC37B07202ABF457EBD4D59
-5610C99DB34C8959A462B132DB3588698CCEC771DDBA29D0E708190A4BF31C5E88D8BE9D25DC
-0D035B752E5B254B32DDBDD23C23DF47CDB65B6A3058828EAC63E54D0633898905898C0C696A
-3DDD1F8AAC9E251C5020EF4216C93B65A712BB19CA6AA941840ED9D18377DBCC82E3A72FB085
-6C88F41DC6D3062737D7002DC612E901B6C46C63B04ABA1EE805CE76AF11B73DA47460467D90
-1448125C8CA5A72D3A6808FFD1F2EF62A17C3FEDFBC57E2C850A6944E249A47B1B8958025848
-CC8ED2D53227D37459F5C5238753D86E96AFA1DF0A67EEB79E83689C4BC5D935968A9CC68B03
-F1C19F8169619330FE1C2B16EC36A860C3E1D0B22D3C6384885ECC63268D0EDA32FC07A2C1AC
-7DD406925174B46CE8326134EA3B934DC19E9B5A7DB355DA603C369EA27F14387BE3AFA9B7A9
-55A1DE7BA12780BDA66B2631B5260C4BF644FCABB000359241AB8DB95E89ED8D90EBFE561B1E
-9BD56A1F1776F510D20DCF760F72FB2F75B771D1D5E1175F8B53A96EF584888B30086B249447
-50C06A1D374639912492A8DFB9EA7D8E79BC02B864F0F5F2B18D9BF7A94AC2F72B279B634FDC
-DD9A98274F1C0C75A4B8EC98000530CFCEE58AB4EEA884CC9507AC73F6F21C627BD8554F3DCD
-725807ED8F88A2CD55C596906C3DA61D489F4CE9E6CEFDF06F3ADF4296B52922099F51A21325
-A524A2CD0EFAED0BD84694A09568FADB9A64C4959D8EE03D3CF37C202E21468888E8ABD0D773
-04BC831DFCCF2EE88F8D97F6477732DD6D16DC66D161CEF06472E85C2B850D1D7C3B63ADC14F
-C7C90ECBEF2A862BAC75CC29250E0722D7831E6AE0705D90F5F893A7244C9FB6F0E3BFE2506C
-B45FED22B9685A00CC20E217707B82108B47AF78D8ED4EBF4A08DA80903ADC1A70CE142698FF
-37EC752E586FC976D14A97BA9FB28437F75F09AE4737E72CD67CB7F2213C1D22CE2F8CA3405A
-D78FE376C07A7C26DE7196C19CA6FB74C76C80EFC21A41A33A8E76C7FCD254BB5C2D57131B41
-64C11547B68569C38C00A4454EA0C205D5CE20AA3B706D2879C9C90470E796B232C04C93F163
-99D49DE351FEAF6D62207871459DA540EE220744B58584FBCAF4FC6E5BD9FFFBBBCAF2BC2AA4
-76CDCBC2C29C319603B3C09970456F4693FE516E788740BA2C3FB8C415B5F4A110765A1CD435
-EEBE56CA018276A74A43EF7647477C7DA77E2B9ED6FC18E7A41351FD36D2609D791778D29D94
-9D519468A6A890CCA284B668FEFA3799433AB8B863864811435286C4F1BB5A951E7FB08DBA1C
-0B1927D21FBEF7A6979CA66F3310C709DB2FCF5F73BB6FA08DCFC1F0B2971FB1C3089548FC26
-24C292FF756D29BCEC8372E1BCDFD5A01156465B9054BA7BDC7019201FECB2F6BDCB7887AA96
-CD30F2D24C6355A33CC469D625C249487D5A2C99BACAF0932ABDD72167A20285D1A4B76CF7A8
-71C3D959682A3408041936AE6B906F0BEDC4F9EA4F076C198DD4A81536B37DC78C63C6A0A798
-72551796371ECE49B3AAA60D3ADB54E2E3C914231A572ADA0772AFD522056B0BE34822E4860C
-367A146D66EF2514729F754D507108FA60D995A5FD67F4F23E600CBD9324786E0D22367E6F03
-EBB356E3AC5ECE078FC7F05E10CA32D45D5FE4669D6E1A3F71A76648F72012129FE0E453D04B
-9F0526C194356E0C6591933FE3FB5C79761A5F7CB06E73D48C0EB224947C2BB005E09FCD76C1
-5EBB2409520FBAE004DE49EC443C8999FA55037E250845684CAC71B4D044337D72504771CECE
-C8CCEE0034F22D31E6C84EC1F8BFC221F117D141CD183E5F70E8BAF43DA3E8641015AB0D147F
-9A828CC1204515974F351584F70C5EE742FF30D6CC59784FA8B59C3F32115487A36A6C85FA26
-E154F4B5753A4E6D8EED81003B426A7B39902D092FAF36E8729E419E2A40C25C659CA6838366
-F974E1EB30D4CA2E52C1D282D2C7E133B5D1C4AC0E4358B710084A4D141F8F020BA1AB31C2BB
-D13D9381FC1B961BA11809173DD4775D60F2793ED94ED57D9D72002F8E77185478511CD9876A
-BB61C0457A428FFB7178E13A4019C8B20425154D2E2F42ACB079D0F385E7DBE259C5BB3A2984
-3AE99197B3D8841E80A9456E7C8C8C2E8BB5B49D6AED06C69E51954C8AFDAFB52A785F5EF962
-B8FBBAC4E7EA21480C344B9B8669E0E9FB60F7AD6E20D902FAB38665679E43A03123B34E5F92
-629A739E569E72AF9707C8B06000E060A45B540A5DE6D1E4AE9C2558186B28077744769407F4
-F42EEE8AFB718F31BA0EE9DA255A4428FE5225451BEEE7C743677D86590BB4CBAF0A728F34C6
-6D0CEC029923076B65179F37FFEDC1DEAD6A678F4D4A6DA47B66BF101CCE6A90F7A7A8749B0F
-93968F2A5DBE91CBB374732C969631DF81A4F1E51326D31810614268EDAED254E4721D59EDBB
-38E1CD3138931F95ED8A2607FBED22652AFD0E181C40B627F3ACD2D1CEEE74CA61A601BAE64C
-567E6814D204EA628C109023A7149772ACDC7FF9FBEA97BE232D016A1F1EB6F702BB7EEFC069
-853B8014D9C33EB84BFA60B6A1FD4A419CA54B288E30034B228F5E3E846274073F5EA3DC684F
-555DF2E038CFFDAA6E4041EF7F5A7068507BB98AEF63F8BED04D4E2F2B88C07FC7E80F3BA0B4
-FF9F583F8BD6FE45311ACA1076AEB46B391DE8F45C1F74991AF796666A0B50A0CEAC343D598C
-7601667CCE477811C450ADFA5CAEF0D659781553EC53B50D736F214363C9B0EBC163F95C5C49
-CA52B9314F031707CB9B0CBFFB63995AB995CFB3E65FF76F42A94BC3B397E449213B77858E2B
-BE42044857D1C7058F23456A5EBFA36DAC299D0C64DE21D97356017D4E53EEEFB9DFC4A5A18B
-E8CD7B5870F9756CA3688875C8295042FA9C38F2A7999F41AB701128C675AA400C183D67CC63
-6F98F885C9536CAC182DD12F9077279A20D589AE641E4763F59FE68D1E3D496031076F789A31
-B92797810F9ACDC0E0A7B33A3A88F3900530B5A02CF37D87D13E01616F3381360DF4BAF00175
-1B1E8782FA197BA66AD408AB8A3BD38D924FC0F721AC6596F0C14D3CEC9D875BE797004B439E
-84BAB84F43B1A12AFF062E3BEA2BC5618118BB0653AA27324FE8A3DA281C0DC71DC7AA17A4E8
-7299E97BC95AECCA52B97EC6D57AC590DFB9E52E7CC5E0BBEA58ACDD75EF01B932991D5BF410
-1ADE7EFAA4EC44EA551C974F9E7119151FC071D167944B23023CF7051C909101AE3114A718E7
-1D325392DA31F75C88B11DCA76ECFF14F1BDBFC695CDC4D089294DFFC13A57E34A72E454A9BB
-E8D482436E8A788854287C69FB71D3BC0B5CE125A9E7FC7898AD9FAEFE19BDD664E291FA5706
-E9EAD8C77CF5CED09C9274B2322ED7B6E2E8F9C3D1FA9FCD2154E0019B535E69A69251C866A6
-9351F8F24F3860D3D2B819FDEB7B02800E59E0F5197B3AC612F9BDD5DF719A1911C48435514F
-8D6690B920C2350BF14F0089A9BD92C2BEAD5C2FF30934B340EDB90D0C36889EC4202507BF53
-2C91E1E001FB9FDBA89D5A9FBF9C83228B19154297E3B3F87AC2D19A0FA9CC718F60A81444AA
-7818D14DCAE8DB9F35BEDE20330B1F2FE973B21717DA229455C782F533B5A21758F4DD11BEF6
-5D41380F0557FC000E43C76292C3243D2251C4ECADB82867553CCAC7CD7906E72F10C785BF2D
-884E80397A34C4268A80B909E9CD10328D2ED4A336F4C048DBAA4A3C67E9EEBE083BA7C8AF7D
-BE50354F22C9DFA81E24F08B2496A67FAF73D78B34BD3D9CEAF8BD47A1111976927650E72109
-5B1C19924FBDB76BF91D710319FC897EB76674F8CFE5BC8AD95DD86F1801FB56E1A94B0D1C19
-31528FDDE30D05291617B4E095F2ADDF2A4E9FD339FBC3F2A710402C32A503DFE0670777B1ED
-3B1DECCCCBD2A36927EA497F4401A4173615A644D639D6786B7A194EBCEF642BD4A69C7F11FA
-541D1D5E96C2328EE5EF5D3182EDE03E9F33CA1A0FAB75DC03FEC53E277007232971A350BA9B
-304B3A7B4716E10163BD7CBF72450798470880EF7B29445450D13AE9A791422CD0CEDFB39CD4
-6C1BABE0968E1971AEB7FA7928C1892726DBEE269DC53A555EBD451CC52E591E962474CB684E
-9937F44449A1EE9E3D8B93F24D9CC4FBE226E9FAC41461A08207F9FE9A75C015928F144DB071
-A268D413CAB561DA23DAA6050164AFFD9DA89F1E7449D3A18C2858095F9866A7C3C45680CA72
-D25D2DD7A0DF340CF9579D2E1A4B40E581AF74FCF18AA070BB02656BDBA1F8A8E81996E43906
-2AA6B6AF83AB056B7484B005C10AD72351054CA1DD21FF9241D02DAD97CBB9708B8605B185E7
-52BC036A662A8C62E26D6CDC540E282BD379891C4C0EBC66E4AA5A38F44DE31F1BAE1CACA711
-007E698D4150BBA0FF2119CFDB59FEDE5DFC8D7C6A9CDEC1259A2CABF17784FBB694F778371F
-6FBC34C855197B7F7802F1767CA4D20B1FC5BB7DA8E0F880834D3963CF321DD157441D38BC62
-E696C7A150590311D19380D77F39B05DCFAB089B8F3B613EDCBEA90CD8FEB7B367D92028055B
-BEC1D34D9E76C23AEE16A2E58AE7FB62672F971006E21863BEB87098CA1C58A7C8993299B605
-EE23EF6A3A75A9A54CE8136133AD569321F471A3B7E3B7277DA2120339CF1D07948BA3760E95
-F1CD8CF15283E8008AD78E5C391A953A2C9EC219010C6642FC23A1A20DC1A8E73F697AF075F0
-6CAF1E15E3A4DB743E1C37D2DE5569FC5BA700F657DC9957F684EC0B13F841C50EA9AE04718A
-93510166B1216A1BBE7C5842A7239C078E6F8AD0EC731216D9629F77A983F3E498ED41077E4A
-2833A961676744A8252D9293D7E8695A069529E44B657B61F0537A2813A0A94287D7983FFC02
-BD6A957C7B74D7C4D7019C9D1C474CAF66AAC0405DEC2DC968E85F1E9C1B59CA213C0F631DB2
-B2F203EA816511443285F7E83AA2AED0592675316B0AB1C295C6A1FE553CD2E9F16F7CC821F6
-272A8AD0091A775594BAAFDF10F704569CBBB3533196607F6594E5B1E4C77202EB599F3F330A
-B3DD8C26555AAA4612C75BEE3D2039570D307593394F5C1DC5307DAE998D1EC73A6FCEECA19B
-FC35F6C3FEFB45B9B462971F96B0E770A6A6BEBF12D3A3420477DD83E4575A2C4B9330336C6F
-AE6B474346495B68F927715F999D58E35C49DB5A5CADBC98591462C4A05A7DB4C202A89DDC0A
-409AB18B70DEF14AF44CDC704B7C9D5F169BF0DE52F685DAACB39C0DAE0FE6739CF5DEE51EFB
-2A51DA04A7E83D510CA9BD40649A0CEB8413DD8855B8DD39C7766084DDA3A0F357CF5CECE316
-D0BFDFB79B402E1F777667ED5D9D42A9BBA869E36CEB097FE7CD48CAC7562DC7408F13577497
-B16A6D5128EF89E5CDEC7C0E93EADBB1B699A489CE07CB51BA5C2F21B03D358BE882EB274DA5
-992C53237559EA7B59C87557F631EA641ED993CC06E3DB0AA291E57F380F509663F0795AF2A0
-95D988FC8D2FD4780B8F8D16E7946541761DC33108EE20A0DA2DFF7E83A7948597B2310CFD3A
-C1B1AF72B4666B2CE84559281A5DE56C7426F8E4D6FA694DD730FEF69D89FA29FD52204D5265
-DBD079680761B72BEABC9E016213E9E02CEEDE9CAF1647877D6B7A4DF31D94F1D77CFACFF1B2
-B252690152E89568FE5B72E99D0520840809DD69781A5584444C4475C65AAFAD4176F913D947
-413CB2CA27B778177A3768F0254E836645DEDCC11AE6AE3C5EA1A075DFFD32DF9C14C9E320E6
-302BA7D1D59D258E5C6B623107C7BE8AE47E168D141CCD44D5147A1CB42D319EAA1EF302D4F3
-5DBD9F3A418E39E2EED9451E8004881A79889D412A47CE60CE86BDDCDE638BE152BEDBBD56AE
-1D1D0A3DD3BEE23007696366DA89B8AB63B76B25089A3C3708E44D5E7937B815453EBEDE80DF
-F400FD94200734E0C4D2423A154177B85873557E1A4B6C224B06CCCA11D306AEBAE0BFB27AEC
-14FA047D4DDED23B7838166DC778075D90094ED5E92CD06F48468B18C30CD8B6B250FE605EC2
-AE857BAE553A4CB32388DA08084FA957AC23BD35208CFAF7964757B34218AEF5193B4D019569
-A19297A1D813A912C7D6BB5456235DE0451BD679807CF8A2333D94D086D7A91EE6FEBDC7BD14
-4122012EFDB631B1E97034400DC601F633F74955FAE6D6B49351DDFDAA070CACCAAE0181A73B
-501E0CBCFB249C41D2A28B16C455467099EEBCC1B9FA6C04EB0930913155EC1323B7A2AA8AE8
-F0F73C42898A164669EAEFC2015113E14F9793678D0F50CBD040B76883908534906B16B6E29C
-5F4776CB55A9955AA41B7E0342D258F28FF9BEBA84F5685F1BEAE5405E63EA3DC93C7F7D7683
-7C4861257AD2B899272D0D67D882FC9D4F9A5BB3A9B8B97959B45B081609DE1392DC54AA8953
-50142D444FA38EB15DB0964F0E7E000630B3502ACCE64FB6BBE18FD94EA04FA6DAEBFA91AA3E
-A470831BBDBD80C4E9049DBE100AE2121E2D12A5923875D07F37F2EE97604A1E15721B675E29
-0058C119F292E8BA6FBF47CD4195C3D249335AC7EEF34C419AD0896C76471D11259B5D5FBB23
-BA063FF1743C8A62A96293ACCC9EE498DC81FBE7B3DB8C713AF442F2845680B72670A907CD56
-51F5DF4C7E71755DF356D5C8C2FE28C7A980B84E4C4A0FCAA4D1BBFF42CB821EE05917AB5C2A
-C5D6A51D3E1EC9326F61F31E1EC511019196A108A47EE460CE0DF6A71F83E12A71A58A78330F
-87D47E5B8A2724D0C31D5B390CDE246436838F1586347FBF6931A5BE35D337DA071734794704
-E0563AC2E4727D0050215C4ED00B2D2EC80D38D6B7096E701EE2EB61B8579A47C60BB0902F3C
-39D59F5E6CB0743DFA3D193781250E5DCBEA329308B2FA89D95406FA860FC6754FFDB2A61A9A
-92D764998F74B35C183CD0E5435C8AB0AB15D83777D0F4C9744CCC040310332B6AE7F0411AF8
-15EAC3094E734CB73BE514B428003D104F6BE22C67FB7C6641F68D2CD80A071C5C85E639BEC6
-0A321AB9B7069F058994E6BCAB2AD918DAE026B5EEBF3C95CD286B155A01210D202AC89247E2
-996CBCAA5B0FC5F72DE1F2AFE326431B49E4D09E42958E2162CF67ABF2D2C335424EC4908D52
-D1A3534CAD58BC48E6549528E20DA15A13BE3F3F4656D97BA087D356AAE63C4FE4FFF8161ACE
-2A0C465671017431CB3FF45FAFDBA1A7D6CEC695E3EC2C91CAD76B0CDE6B8DEEEB81D0443DA1
-EBDD1475BB8DF19AB4BAF2FA6F91B5E753F5EE11F62C67D5DC78550991FB2D2BCA39E0EEC87A
-A1D2302D3B7FBC57297BA3285DFD875F00FA3BC92E74AEE09407BD8F16FBFA1AC0AB08E58F92
-3A10D5E4BBF72B034F9340BD734373F0CEB0DDD6CEE44A903994806252A4107AB65B18B04D59
-BC46ADDB5CBC56C2F196A458C8F1BADD3C24A003E42239F058D6B9404306F6B38505E827284B
-60DC3C22E878F92F6001DDCC3FE63DC61A04087FBF2B4E3F2339AA658147186D88793EE70D6F
-24FE1D72150546A34202FC655D0145AFC2EC4247F8DAC8C3B5B6BD930934BD435043057CE6A8
-6A5F6DE5B8E51D7DB3BA0B3449612951E35F475E83B659041DDC1DC942A437A1A2778662ECE4
-5722DD25D0A59787F7E58B0C96A652B7708A8C89DAA9B577B9492DC8C48D98AC27C4B73CCF9C
-55B430FBCD12B5C6EA7CF90BD28DB91954BEA34CFFA6683EC5A3B84F0413FBBE04E550D6005A
-5936816848671AC4E369C1D833F41607C041F85B9DD20098B41349AD5331606BDEFF5B5C936E
-C356254A6CEB340345BC9DB418E4093FEFDAC4693A609D4DBE0669D67D443DCDE4735A8918A8
-B8009D055C0746380B8A4195DB3371015F61B302C95B45E2B11F013F8522A44C49040748AA94
-37884F805358D1DE4036F01FBDA4B986114CF9842F1E5D6E7E88F978F480FCCA51E4E0C929C1
-E3227595FDA8EB1A5D6361A5028E1209CED07ACE7D33E42846571A784D2CA58D02058D20FFF1
-955D0E5D9497B0D9CD46D37DAAD7951B5C1F44F63FE8EC3D7209F5C7F4ED7806138EFD4F0746
-39EB7A5CC5419BFD245C95ABD4AD268D4BBE6721F74A1565E2B7FCE1A0A235E4A6802545E8E5
-25BCB645DFDD9B87BF0E7379676F0EAA0E4312F9FEA5D9EC809EDF9EB4A09C8F7110E3842FAE
-2B2EE8AB6287C3411FF5417144ADA4C05B9BE3FBE325243AE2C0623837D2648C2712DB61D180
-AAAA7E06B2952E288536A7366163454831ADD3C48D865B5455B97DEFF5CE69746BA056903A11
-8A2A088817AD86EA8E6A7D7C1B30FD957DAE33BA15EA45FF661AAE136FC48A7481C567FBE81E
-DD7094ED2C4D6E710E2A0B626CFF67E78D83522A6F84F4B1C3196BE60B1F1F1C1A55EE0B0ED4
-F9BAF40264960A149A01D8EB0346131BD6DC12D336F67421A844541B519D5FB85C0CA8E82B53
-AD008F3BA141B332FFC6B505EFEBBBA8DC78EB8FB94506A14C8B212F51C0CE2E1D978FE6DCE9
-488831BD96ED5F2A9DC1204402CA8BF9D2086E4E78183CF56A3B514046FFE9F1FBAE06A6F719
-8AC66C62182F636E9277BEFD5378F6853733E1823BB93D5BAD028647AEABEEB969F4CAE37597
-667C1F08091CF9F1704B3492F4834E93283CF29652706DC0C6AD689F8D3BF4FB7C3C2C01E417
-C122C8F7CD30AA58E8C01C90953F55439AE0DBED0800686F8D096EDE6A1934B5F34B5FC73D27
-89F4B4FD8C8DE720238EB84BC5CF71D6E9FD4E618848C9C4B33BA6BE7D8566DD3708B2C244AA
-86E256B8E5954734EDA64472C438A0436B293D245B268C5DAA08750AD0BE8D1B3F5A7B4881E7
-D74FD9E3E7FA5705050141FBC24D71844055960815A3C0949283DAC0C1CF8FEBD10EAC6F9F4C
-A3F93D54165B0F74A811C31C8CAFF4BFE3AD6C8AF02FFCB2576A04C25A9DEA57FC322836B2AD
-A20745D121C6E65A4839DC52063FBA8D67D42F09A5878C9339E06249BC4758859A994504A369
-A67A17D7F36D583256DD65FE76DC44981EF254D5CD33C9C79717C051F20EDDC7B2744148BC18
-FC2244F596340B69D9BE1316941D6E1F00B3EE56A0425438ECADEF4BA8A9452D2947C7CDDEDF
-5D1BC2C3AE280D77241C48072ADD5C4A4E8806FA581E39F45C95CB6C41D87759AB21155ECFAE
-D8AD387A70B27DAE6122895AB20793ED40912BEFD7C25265D8B52E9AB0308A5920A80C4B6C06
-38F7AEC070C3195D24C3D601546143F7E62573DA6C7203115A2613ACDD2C583B88A1C905D3BE
-019B392560425C82FDE2C57F18079E17A0C191C9A715338B94F72DB795C151C17B60B4813A15
-04DD4D91E7F1150EBD02D600C96322FA5D90E221562B7C5AFEF894F4A9DD36426964508E4155
-B2373221D5AFBF01CE958798CC04A4E6AA3EEEEB4C72A11224A645D7326FD4E704FA13663B3B
-93855501B9E12F3AD4F1820E1051285A9F435FD9C1572BC264C355C792A477A70FCE84813E23
-FAB7A325C03229A257B30A055ACD0620AEE4DA877057BB3FF16A78028FD293B8A5507858BC11
-9E83EA0D0CF842EBB9E4C48BB4480CD0EB448FD1925A88EF9F68AF1F868F9E477870A4A78DFC
-5B299AD9D3966C21C202716B7A7D830E7200CDDE60C5810D574AB778C9D59459A432C117C78B
-7A2E5F23C88308FDA04AB0A0FE01F2426BB7E508ECEF7E87603D51BFEB723ABBFEDD44EF9365
-1C1B8AB35A306B96D92D1736AC62CE8D2C2C60E67583DBCF37953A7ED2F22FED2B3338A5B4AB
-A0EB8282AE9BE5FB37780576D3BD2022C9D4EE4E8F5088816372A236A48534220B210BA346D7
-667A8E22843937D77B770B1BEFA195DE3A41DE77CEBB2FABC5549032344A9B8373984843AC29
-73BEC2CB18AD56D965CA2BC27081C1A88D5222CFEC5DB939B26BE8E78D6EA2DD75D58DF88EE0
-847E9A9135D39160233AC728E73E2EF3A6C89D60722D69A90EB0095700A5CB404F9014D51A65
-3F0605103DBE7068B70584531DF286F2BFA72CF6A8F002BDFEC2A411FE3837B11F2A01C0CCFF
-0D009299F9B53BE568584B919096398E1B88938F94D1DAF5965E485FCF5CF682BD85818F1B8B
-6D3B467B4EB10E8F2971D7E5E4EE3BD0B4B11F6EE7D720B1D69636ACCDA458C3C8EA663362A7
-6650F38BD3EDECA408FB81388840887B60ECB1BF5F67AF16EEA563DBD2AA4288615D9C014BE6
-C555FC5C2FD2EB877A4EE67B81859E96679330732A445C3C1A4FEF79318DCCFB5BC1BF5BBD53
-A51EBD88D9A424F1AE14B38200844296C03651E109446FD75DA1D35B9D7713E91419194EE1E6
-89844AC889D3B392243CFE12F129DB9C0FD0014848CD1A7109441B72E5C503FDC13A0541839C
-B7B8573443145D6AFA1D2B110BB3C9A181A82275641CCF5FBE9A4C2EFA5959662F2B74586C82
-6F7B39972928EF8E3AD74E7DF53E456A5B74DCFED01B39E4B4BA25D14B9C770A8DEA7145D338
-603B25E8ABB72E85A1FA7E28307F9C3FF1CF333AA3644A5B764DB438A358EB5CA7ECAFFEADC4
-53996CE17869F07F41D2E2A052FBE036239E548A1AEF946B0AD7A3DC7962E9DE93A55DF87CCC
-E22208A98DCA6B2FDDC856F63B13317791EFA4070328D10E8119E6015C0A46D868C568029D68
-9D40437BD2737F0B4D1042AE53973818F4A46B37119D18C1A54F43FC8F700B7534F806798169
-31556D5EB3E24F51F3979BF43E9D336BD90411FE10E1A1B9DDC8B58FD232F35C240B495CD27C
-BEC0F22560E5856E86EB13197CE095022003DAF7AFA837D86B96C4A3AF6C29D06AE252712B51
-53B1ACCDF58AA1D32E78BEB442E3904835BEBBEE01803A9191ADF2A6A579FBAE5EC828B17F3A
-C98BE2B9888AB8154455A3897913C30EA380226809F4076C72D81E8C02A59B6ACD1C97364CB0
-F4F5B01AF8F567F73840F93B30A923A002370F6FEFCEE17B62A7F1F352FC3725FE9F4F4F011D
-3075BD9297C94D81988C37DA6189F8D3AF00624922F1BD4D1B99D532488A3887E8F52153A95F
-56759C9EA10B4941185AF057F3809E84704D5EADDC50AA1929768265D97C8EFECE1ABC7B3D1A
-119DA997BD7BD6FBF08EAD3985566B3DE614A84A4FDB8D9077428AEE4AAF19092702499F2BE8
-7DC21918301841F266975D134FA9E8A9344EC6D2E23B0E1E81305FD0DFC0971A250DD385EBDB
-872110B585666CFA5621CB25DE4B332465CF1FCC2CC64AA3BE9D6372FAB5B31E1D402B12361D
-D7242D20A87F11B28F2B1F0EF8F24F8D23DCE63870297B8CCF0C0CE5B1890A035BC3FFD7F833
-031AB13636FF96E421D7823C8B6917FE7E594AFF9E1BBD16A11D991F033517D007156AEE9B89
-7147998F4FA943D2255FF70F556B12EB8D7E438471C012A7A9FB702FC31D2BED3C1006A30AD9
-71C708E5C971BEBC35D40C78FDE0EC2321FB3670265D106E24ACBBB449DED5F4DE5D170B1676
-DF800358E882936BC24B279123B89F8988EBCBB6B44D4F0E06E0C611F9EC17AE9A1BA48E8FEE
-F40BC359C40EFD821EA76BB8ECB02F4C596963F192A8E5AE7AE5917F7B0B457D0234F914B07C
-C61420E11D092068AEE4A44D2424D5A9DB5074B57659E0A7599AF8D3F4E2785430825C108691
-297C172659ED2869C0E23394E03DD550659018BC817D0409362AE35685B284D275DAC2626E29
-877670188B868665B29DA3D5E37217CB421EB2880D78F379172886F80109608B8A3FE464355F
-D86FAE057A8F6A68C2C3F884A7A62DFB74D20614A478B7D325CB89CD529772E27410CEF348DD
-E18239D13FCF5BB30B5AFD02E71D7CCA9F0FF5A2EB7B88488558201F51F02C6F3FD7683290CD
-A6035FFF4AA488537315649C9541186736AB7EF2727CB0793E5736CB0E6FB12213CAF845640A
-15083BE4AFC34F9A22FFB953FF36654DC266E24B1604950E41B2BAFD6A8D0EFF82F1A89B6C0F
-8AD52418A9A26DCA74E4BB4BCEF479735C17A657B63CDB9D4AFC434CD385A6ABA1D9ABC3F128
-9748BD29DEFE7E45ADA40BF65F700A029AF59F2D6CEA188C1E9887C382FBC986B26179D241F6
-E2881DB9C8F0C577735604671D53EBAE52FA47EEE9B74C390F186ECE2C962720F09FF261983F
-8EBB312F17B4C08AB2DD4D3AB3C36846EBA599C6EA8A8F64D0CDE5A4E599534232D4484C7A46
-FF7606E711C52E900EB668399BC7A30CA0F19A470078BA4D2858531A8A2AF2FB72172C242B25
-4911AA4D6DB4D491C7330CC922B22594578D02628670CBFF2D4D0B029F90D4A912728CD4FC68
-0DF3BAE8E905BC74D3FF36D1BC4F6A013F8ADE5DE35F0D22B89119E0356FD9F175094EA8C5A1
-3A5CD7CA4B9A1948FF82FB471C38031AB513036C08E8DA771220358EFC7FEB26B4FE3F99E244
-6F008DDDAE0DABE2732EEEC19A4917CF63D482D121708A416D7763B1ECB12666E936311EF376
-625B65B59482C6059AF535BBD80F939FB21452B72828419B635A7CFBA0BED1A2F1F639F57274
-4BE8E00B114203409742F65327537AD4B52E46D1285049D46E4D2E86B915C8B24E621732916B
-9A487EA1DBE547F83FCD6BD5B77132380D8566B3333C9C15549AF57975BE4D33E360530594B7
-7B07B10BBBC6F247A4EB65A8176C32F6751FBC26303655AF7DC43508F4525CBC18C73F6C7431
-F0170EA6B39658510F537A6992C90F968C6D3BCB3E0269FE2206A763E715B7C9B5F483176AFE
-2F545CD2B1AA662831146FEC8AB40475E6276D0AA4F3EDC8D8AA99076E7C55980D3729928052
-AE4A0D1E5249F0FFBBD28B54D208B496C4AAB0A5192832D9513E8856281B771D26921230C0D6
-5C0A84574BCC7531406B9E79F4F56627E692B5901D346A59EACB9328E96C6C02461D3B420BE4
-B3BB7DA35E2AB0975D6021C11DE28A0814D38BE87904329DDDA53FC320BFA72B136AAFCCD7A5
-AF2F64A7D364AF2AE27DAF5BE67EAB44F02BF350067A4191FA821322833EE156927F8CA078AE
-9D16B0454805BDCB57053C03A3A5FBC8AB47A425ACE5DEA6EA0E2A34B9A1707B651EA9FB3A39
-1C4DC0826A6ADDB97764ADB310287CDB83CC7A268FDC65A04418660A2E25A64FACAEFA106148
-50DF597A1172061DCB6F39E53C77ABFCF5B12D85636E03D3F150D3DCDF1DDD84F5B7165C4454
-E6F44B7D2E64E483DA6D3755F09DE3BA9EA41C38E37CA05894DB3FB1E9C1946E2D1B25FC183A
-D7BF236B8E410C326CD34850BB6FD5200C0481B4A9E62836F8093892B2BC67CA2EBA62BA0F15
-C5BB65CFD03A35F6629CD051A0CE3A78F7C5C16553C55C2F424E608512E5CE607521A567AB97
-8C8BE962E869C7FBBF582FB801BEA242CCDC8E8C84B2CAB30C9EC78220E5928F945AAD997595
-94A63B02C8CD99EA09D865F08275DE178123BF0D5B8BA2A3B1E5243BCF448EF3B882DF28F36B
-415987EA7C06B719034FDCD115BE92FEDD09526DA14BDBD78C23B9B3740ECF878825B193B802
-D0C36A059EF8D81FAA909E7EA951D0889A2D799A32720E93DE638040540BCE28F8E16F230968
-E8EEDC30EE042A7980D53E14888301C77A4EB7152D760D7BC9D13B1FC39BC592DCE2C1010E05
-0B6AECB896456D3C714353A1AA7D6F04172759E034534A9E9A07B7AD1D66F1EAFA7D212694AC
-385A4E20F82BDDD1B66AEB6866CC942D5448C15517B5D107F6E7058360417139F77B8A668578
-D7D61576CE8C44232F4BE41379DC28F3E7C180611D27124873BA90457D2A2F4BCA2A84ADCE08
-9855938AA62E3776959EF719D97C09468584C4B138A16747C63FC06290FCFC59A2DA0E73BBF3
-B444D90B6F7D4E050052EBB2577C2389311B6956E7D583BC378CAA3051BBBFA134F391EF55DD
-752332473020E8B7A1849DFB64803E961EC4E396053482393684556289EF802D3865167BBC84
-06A93F55E337F79050052D3E10ACA40E1FF68C475DB20A9A49776696ABAB21011DC3CB3E651C
-F279200226E8407E6A3FB8053B9C1096FF3E6A7105E76A09505E8CE7EB8481B9A9BDD98926A7
-7280CCE379A6EB0C4EE91B90E4EDC9A23921AB0CEF381D3E7F1A18ECD4CB46037C6296237A5B
-079E62614C08847E5BB18D1EDFBB25CE3C1C4947C709D09B7A130B63C9CF093A62A974D5C3D8
-C858F7D067042F76FDF0513D26463FA238D4D96A58FF898B31FCD014DA4989661E456D0DC620
-4D65AE05B89158DDEA46E51FA1665EAD01715845B575E89BFE34C629755A109BFDEA871945A8
-9E18D8CD452CBE93594D246FC0D31B50521BDD3BAF6D752949B907AF9505FDDBC937D2DB89BC
-64E503EB71D1043B66DB13F11EDA59E03FB64DE3AA9B2E7FE668607A02BB54B7FD5E84C65B20
-659D16841E09902A03A9C326D9D332C7622AF73787DC8B98909FBEE0AA73D43428D0674BE761
-72E94E57917363351ADF9DF2E8256ECEA99558691DE9EE2775624F7675210E6FBB4A8C648F33
-8AE65B478F4E94A2485B75187441ABF860D0EB31D657FFF520694F8E36E7F5052F2EC333CA7B
-318374F74EC8FEAE6CFA920AA04094EE42F2E111E889F469FF848BF0D13548669FAA9333931D
-8480E5159AB8F4110BEB98D24613DD4C1B0B3D078B799D00364027D24295242024EF4E323BF8
-32D54EE33CDE826945FEA9F7FCB8D5EABD8E7BA3A623CF1E1987CFE1A2152688C1B9C097F18F
-237DC2C3E0B10841CA7A26AF98D4A1300FC249347EDC2FE5158694E7884850A8DF7450273C84
-AB9A10ED07E68443483B01A1181E791B10248AE0B893DD2CF901820F710EC8DA35F18D08C04B
-F13841E77871B3A8AC3E75C5364A75CAFF0FC7B963169FC90A7D55F6661E7ACFA13D6D239885
-E782499325F50E24CF795F501BB91422E09E0B1A6FC28EF7BB7439A52DFA48B04D2584466AA8
-A5202915D70E93AC51B9B31454D2F2AFC421DC0F23FEBA2E54C11B58241B73D8D8D3728D6D9C
-AD925CF425398273D7724FFD3E2581AD147DEEFD5DF7A3BF404A767A56D51E334DF0CB49817D
-6042712F55C513DF1F40A052649B55BB50D09EACB2B49DC75C44EB6386D842CDF42037DA6959
-3C0D72B0954E55F4E94E30043F9A644A8F10406AB77173EDA4C260EE96498C0E3AF61AC5B223
-8FD73AB9F1685F13601D080267E514D666C8342AAD4C30F9E0FEF2255C2A4AE62802DC36C15F
-3C41F2BC1281D44E3D77860FD18EA7D67D335BC11EAAEF15CAE8D1CCD26741D738EBA92F5BD3
-C527B920CF883F7F4BD31711D7A467D4DDA775F811BB31AAE9311E4B0DAFE50A2EBF8301C2F2
-A50840C8E61411FF8FEEA7D3E948CC0B796184C30A500CF5A198694566812BA6CC2B02CAA64A
-038CD3112FAAE519EA9E6BA8CDA04592F5759F2962CA28743B2171B85A69DB0689154D69C5D9
-6649B621ECAF2B2B6E4477C8707C1C32C2AD2CA8AD0CE598265907D3BA86D613ECC85A93B3ED
-741F0BAB3D78E035EC3F0070D0C53D46E52C7FBB8DE776079E36D6512BA0F6863BFBD8DB3AF4
-3F0456B00416321DD7464992940B8B7AD3BFC002A73D199BCDAA01EF0B37CD871036923C25F8
-DFE5A756917875150ED2F442A6FA9D53651C715FEE87ABE09D097210A1BEEC3408B5CFF4365D
-53D196FB2FB5723A479ED6FA7C7509C038E2F248447FE6E6619B329AE5D948777ACA0D756339
-576960242F3EDC31CDBBDA235718A53E64FDB146EA3487C816992A733DFC1C43EBABC7A3D77E
-635BD4E964DA2F540FD38A0D843AA6D9E4BD30F16E17491C3C7AE1F6CC4717015948CBC7B7B7
-8D51DA7997A876A9AA18E3C4A09453B35B1961C8990857550F05583732192DB74F1A55BE0526
-BAB47F6875EB7F9274F09829A5E21BCAA61FBE29234E52E803A92262DCE826D2A7A04B497565
-5A39205DC80FB9C6C2845C5B3B98892AC2414C54D6DEC09C38CA648C51B24B6C22F5E73C5F9B
-60C495DDCCFA5F9A1153A5A2381A74D99D1C0392C3A668D6281A623FFB8CCF2336F68A5E1BD3
-9FD0FCA558DEE89FDF4CC7B65C324C1613C4A0777E0B99A24898298CF61BF898B1FE53E23794
-A433C612ED800A77CE6DF76C6AAFC406AEE44767FA1B109FF5A39117BE5A3B8BFE784FF34EA8
-EF65BA77940F3B7A662738CAA98E19714A96AC66F7E7AFA33C244BA06488B854CE55F5AA659F
-646C3A15F5C082CB11B979DBCB780605ADF4332B246C6DDFC3F7A8FDAFC78380EDC0BF3B7B4B
-8A843F78B49828811CCBBECD281B724EB6C146774A01513170165EE561C724B81DB0ECDA2404
-6FAEC3124184E30EB1C3080F983818D04F287B2E44F68C5B84F4B4BFB861063661D2A89FAECF
-49F30572136172CDE89F3585088FD3C12190CB7E2183B865AA60339892855B014B6F84BE699D
-C1FD0558D5AABD02C1282D5E74E3BE31CBB4BE35256101FAF84CD4883FCE88329F095EEE8D46
-CAF552537276E03BC886E2B45A6B080B7FF8B3EF1C3ADF7593B3E7123AB67F632B34AA129376
-1B85704241D79B9EF7544393153622959B3EF1DA3911EDD2A2C1EAFE61DFF8153D2A5F06CD26
-A70143A1F2C8571D6D35B20ACA92BA8EE5568F020B7867BFFB13B258A30D5C7BA5A73765AAD4
-1166B4F0D889D0E0B91B3F1D8EA44F425EE974DAC37064BC195A7D02DB8BE32ECC8B3718B599
-F3751DD224713C603F5DC2C78E802E491E0B0020BAF14A0396196219FFEE2C4DCC14E6E8D34D
-38F7830A9B4535610825E91D4E08DA8A31C6CAEB161F183F8BFDA5EE4A7CF1A78C32141DD2A1
-62379CA05C4D9D005779A4A175DE2CD14A0B49FEF14F9BEB98C8B967196135D8247931557872
-E0ABE65042048EC380BC19359F705B96C3C45BA4469E923B754E08973A8DA02343C972315498
-5CEA4524CCCC134699E8C9577E750C865F6CD7FED22EEEAC8E6B2B01A1B95F426FB051821632
-E02698F44B97E60E43C04E5A3B1A8DAB994C148DD438C0166A3AFAF9964A8CED14C34D07FBFF
-291183884E24B51BF050C394D591DA8E043FC358DD00221C3072DABDD75305A74BD92875BA8F
-E54F599EC04EDA89A20B8FEA682232E566CA147FF5BA08B64816DFADC69CDD39452BA27C9EE2
-1F8D974BBB2B8AEF87BEF53BB11D7DB870745FCD64647BBC60F538DEF1E2DAA92B672B23E150
-D206249DB7BD73273F58AFD7E6543E819BE241D7C363283E19A9E54C349764A1DF2CB4847D73
-80C35CF54C2ED12651D64CB6804DEE1C71176A008D6F6EF24CCB36C60746B9AFF719EBD587A2
-E66F3252A0B79B48B09F6D23E0DB2078B77949AB861B041AD4FB8B2CCA6462AF42EFE2E8CF7E
-178E59D3AE268135997FB721C703F7A9FAF9D4EEA0082C20F2EC900BC1DD30B3D129D6151A1E
-FAF52E9379E6437CB3F5DC75F03A3B1BF457F29B8F58292036084B4E58B296DCE28CEE150998
-3F399C24B7D065E74C51C8E54D07F6E8939D93B46170AC18E1E0AE2CFD954063007F00E6490C
-05C139864F37C4D07F67DCC22ABBA07D2B3965D97B1AC483E65AC14043F71156B3CD9717BE1D
-24DD1F642AF88361AADE33C89E1C3AA48FF9DA04C185A6F0E7E834B06100D5867C1133B8C462
-A60C4291B98E69E7B0419561AD1A42ECB01A8E2FC09644F84367EE7E51359BC07C2195C40E21
-0F66DD2E2CDB7C6AAE4AF98027EF3826E8BF93F10B1D5182BFC5E838EB95B7114A10B73FB4BD
-6296EEE0DCF7193166C4E903BDFC2467222183BCBD4DB6A772AB10C1C8F9E20B08D15736DB16
-4AEB29BCD29A1D59BE64C46338577460D04D043E406EAD24081198BA62D0EF6AD52246DB3CF4
-8CFB6E2B70F43F3D5B68464E8F922BC11DBA3E2951A5F98F1BC1FD6B22A3CF2D179B933AAA63
-2BE6EEBFD2BB965FC3ADDADBA49BD3603922E67B60DD1FE11E533FEA668B8AB3A73A58976975
-F9017893033E88BFE3543B3FFAF4D9575D1458BF77AF4A7CAEE22CAFAF8ABCCD31ADCA0116F3
-9D423EAE339EF19050428B5B835EF58D438A4A88695E4F2EAB2855ED05F408A2C8A574747C6D
-7D11F3FFA46AB02170FBA0C38DF2C4BB28542DF2E553D519F45CD54167EB2BC52B5BCC38E54B
-FB605B0A0D001DA0032EABCCF5C73F965CEB1C33CB820E00E50578B4C398C99AC5E71EDDA32A
-7606599B9CE8DA0B440B4B71BCCE74004F6E4C52894B4BA16E9E47E00B4176D09DD75CEE0350
-7BF43898FAA27D8BE506FBBB6CF9D7E225FEC783439F315221EF02A9CF54276A7CD1F6F8853E
-A51876A91EE00FA11135DB5C3F6285C12857D345DB8EEEF81EE2F61FC14FCBED119737ADA169
-7F584C0A5BF780913B5E5EDAA7356382EB77DED29B0A3B47B4D7A9B078EB32B82AC6824B84D6
-D89BB055421C29685783B37DF7370075A7CBEAE339EDE8B43C7EDCCC956AE0E04D903A8286BA
-0CF6DADDC9968C441230E2AD17174A8FCA2FF7791BEE9B88649686314F2C8136BE6CAA48E25B
-56D50DE62BE38633E7039B5F00713A7288FB1B5A5094A99FCD54183F2ADD51FCB9365813EDBB
-653AADC6AFBC475A695B12D28308049346549EC2602F42215A4F08E91864EE3AB5FF00A0FFB1
-DB66AF760077126CB747E5DD8E6158DD902A2A265262562E6B5B09847BCF4C182EE4E90BDF4A
-3DA1F725A203AF4EDA966E75718DA1D1C304941CF1E3347E36BAE78DD32843F92029516A51C3
-CAB2C8E4D899F24CC4F45E50E132501A8AC50D92BE48B4DDEA54B7E21545942337785B1AC5AD
-25C1A7C6ABD2A925F8593AE4AF9ADCEEF7555D6DC8ED79DA8B3DBFA39EAB384B0BB42A5A9F71
-84C31A98911AC23FF30E3ED0DD58B0348BE7FB0E3F465343F05FF871E502EA6382017FB79277
-56D5265C67ECBE9CB98B5C20D995F393E828273CE50D4C92F7CEE429B966D0A7BC4B57F9E5B6
-CCBA051E34BDB6E3F2932ABF50111BCD5FA0D32E349496B4F3827E591E8725F8809FABAC9AE3
-5B5B1BA79748C0C98295D61D4C83D040523A2BD4DE07002EE83BBEC2C6038C9F89977F06B610
-1D5D8E9EB3F0F7EA4D2CAAC71B8AA0180F8E0B675AFDCF608EF92355B9B31AD1719B90C527B7
-DD4F3E4ACBF3196DE1CD39ED0A06E8614A23C2B43B92DF172382ED02DAEED306AF67C7AD37F4
-6DC1D9209F857C1C0BAB085D3691DC707985980899E410F8F33A56B3E45F7909927793862D6B
-FCBC9DD9D40E3A4BD68728C54C5C63FAB65D9FF73817F1BB70DFD1F22241ABB7960AC01610EE
-358B67C90990410AE8583C1E9D8C9A827E8A2C2793597F1200E0BFC141C261D5C43FB596931F
-147A68B8421DBD6746BF32D58C044F2BF8082D717C1A196590AC9183BF07E7CA0747CDA6B2C9
-EA7B5244FE456F939EF8F82589024E4B23F81612795BAB3F074576BABA4C73F63A166E93DD37
-22F6142DE5ADB791982E52C6DC3AFA0FBF730D0504E857A5D9103CBF2A2DE364C443641FABF8
-2BFE180DC51AB2EAF10A93E15380C99B608E77568EA2E5BA20F8CD41D7F989CE46D1FB6AEF8C
-EEA02063948136064C240067661C1D1021B8660DD175EB9EB8CE9D634528E7333EAB1F0006AA
-65246F03D4DC9D8D86659F856E81BBFD6ECA496E67022C3445FC11295A37090EEA3FBACFF051
-4A6EF31050F30A24971F3EE2D1938CA86BD461845998652CAB1325EF273DFF53B0C44B5BA13A
-51EB92495EEB28BF467883B09944F87AB5C2BC7C516019CE49F0DA012205E81F9678D39D4362
-A63C153567EA9F287173CBBA03F98ED56C75FADA320F5F238B10998AB396A10D852E37A0CF50
-76CF1F0E99AD15C1FE29DD10AD143E3C46B7F317956D13CF75C45C8AC35F6FC174E7584D10F7
-903BF1801FE16ECE31111748EECCACD0FD054AC5AE8260A7287A9A48731420CA33531F351578
-06A591ED7644DD90B57311CABC9E47F5CF19B55F618009F095C1ED38B13229A50404AFED1B63
-137DF4833B8FA87968AFA64DFCA521257D523986FB6BDCE7DBC3CA1995281A5165B28545F638
-A24A7891BEAEA40CDAB87CC3F6460CA1B1875BDF47C0122E6A5417B86A04F51CD6933F5E0AAD
-D81F646FEEAB281F3D264E75593D03A9C03B183FA2E38484D831C3B1D3813864651B3374AAE3
-8E530C3850EDD6EE45F09AA3A319FFFC32788F82C0D9A7C71B0F43BAA1A77B31FCCA8C3292A6
-34817DB608761E96C0EC5F31CE1626BDEF2A912C9974942F808EF7D53E5E649725AF40EAB1F3
-12F929CB329FE013906E986BD0F956559B147153D3E0FA3FA057B9DC113818E0863D071EC31D
-A183A1B1C07A5091CDF22E53A6BA95E1D87981DCFD50E0F7DF2DB6F6E83E05BF41C7E4D17CDA
-BA4A886D882A0A9A76E9427E4C7675ACEB7F3D9293C0C18192506D906255BD8AAE0E8B80ABBB
-CAF6FBE733C7E389C52AF822C2C3CB82BD9B53D9BD926AE7E8C29EE4782EC841DAF2F40D3DF2
-38FB1D3C099E52071D742C6774FAFAA4CE4CD481FBE9AEB58AA2F3BFD9E80B0CA0C6CDEB2BC7
-85F586ED775025CA84E9EBA4ABA79FF34E93275D8D888557B0719C072E900C67996862FA3109
-1263D4ADABAC7B50078FDD845C98B16CCE135CD701FB4499E7FA1108553FC0D113B7A5EFC258
-5912EDA7D68DFAAF058E0134E0C276126535E4122A06289A1F2BCF7684DE89F4444E762D0F44
-B2E19FD4F74C9EEB536AD6CE38E24C8737A4810C5B2611EA03A73CB17480F5F871510A093A01
-A7489EBC57B20375CF8DB11D5D854BAD973599306FD9F843645336E35C8AD236E5990FC4457B
-36400B2D2DC8B3E4FB0148382FEBD864A3CC897C967DF0D8198487B0FF5FF25C62C9C43D54A2
-80E0F5FB77867A9F0FC23C1B6DA7E582345FB81BCBF9B512E0425128C50E6EF312A8E6C877A5
-8E5CBAA425C2AA586D86012CE8C2864447A8AB9F46AA7229990082C52E2BF84F6107C0A4C687
-8746243FC3CE6B09522B8BF06C4766E3E5A131F7EC96326694CA723FDA91973A49BBEACFDA2E
-3989E4CA48F74FA8AE73A280E24022A766A8188850D08C9B8F78743E3E11D897BA3ED9AE6E7B
-1059F5F6D36807069DE3A172205ACB8622177D41229F816AA50116E696DAEC917BFE2AA7B5E4
-F32AC5721061D928AAD96A8CDC5150321200A36EF61EE82D771A25DB81CC1E71A7A1B8171886
-1C9EB588C99CE789326F232B96EC72054DECA172564111B0143AD0836BD42173B4C4B04820AB
-FA65D2FA0B77D0B073227FF5BF3711DFA823CD1AEDE119DFAC288CDD99C7F6A80088FC9BCC69
-A9AD8775D8E58766D7083FA505C88E1829283ACF717E630EFC7A501D490D54594EA13E190588
-9B7344400A1327F0C75198CBB939AF89E93305A033B65CA255CC4CCB27377A85462E2A4B65C7
-EAA385233B6C4773D85DB846D5AB82F5C0B0A6BC9A3176974DA1A239D352B2E9B546128110D6
-D643ED714BFA82C7851DA0C66320FA2739A423A2CE503CF0DDBC3F0660B7B8B2E3B7155327F9
-59F77FFA144CFEA492E50B3F6D3278F33BF3EFE4C9F4901EDA676DD6D3B01ED3DD28DA965727
-3595DFA14280609886A498027A9F169F87A0D9A5C800D6F3233BD47FB18CE16A0380557ADCAD
-CAD9A8366A36C54E33B81B2F22C346EA4FD71727D4D8CFEADCD4083F057F3DC34BCC6D93D580
-C91AB78F0C5015B3137A242BD7EE24E7BC218472C48C92D16687AE450F6BDAFB2A6417AB0E20
-BDDBDDA07E8014F110551859FE25EFF5E84B30EA57234304225799769ED77AC5FFC65084720E
-F3B6ECA720E09DA6375B2DCF24CB495241674A1B6B23EEDBC3233E304D783218C6BCA464E448
-D6149948689776289CDDE7D915073948E0AAC04E4526CBD5915EEA88263630A679B7B7C794FC
-6717FE4885A77785BEE4731C9B2CC8BF04D580043113BB9A69233406374B06937195813B9D03
-419A6C75414047F672899FC68B1E62B12563C92B3DDFB3A8B0336EEE7556A1AF21AF371164DD
-EA91BADFFCF660E8261E51F8C0DF4BBA1E8D6CAA1011BF919E85CFBA2A5E69E4893D2EC98325
-99AB56484BD6CDABECF05EF40512DC42979C9357F5410E73D85B07CEA86674B380AECE9F3B36
-D072CCE3D9586D8C0FF6524E3210EFE54E179DD9801E06B45990E9F2C8F18BBB78F0BFFC2F0B
-9A0504F38B0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-cleartomark
-
-
-
-%%EndFont
-%%BeginFont: CMR9
-%!PS-AdobeFont-1.1: CMR9 1.0
-%%CreationDate: 1991 Aug 20 16:39:59
-
-% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
-
-11 dict begin
-/FontInfo 7 dict dup begin
-/version (1.0) readonly def
-/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
-/FullName (CMR9) readonly def
-/FamilyName (Computer Modern) readonly def
-/Weight (Medium) readonly def
-/ItalicAngle 0 def
-/isFixedPitch false def
-end readonly def
-/FontName /CMR9 def
-/PaintType 0 def
-/FontType 1 def
-/FontMatrix [0.001 0 0 0.001 0 0] readonly def
-/Encoding 256 array
-0 1 255 {1 index exch /.notdef put} for
-dup 161 /Gamma put
-dup 162 /Delta put
-dup 163 /Theta put
-dup 164 /Lambda put
-dup 165 /Xi put
-dup 166 /Pi put
-dup 167 /Sigma put
-dup 168 /Upsilon put
-dup 169 /Phi put
-dup 170 /Psi put
-dup 173 /Omega put
-dup 174 /ff put
-dup 175 /fi put
-dup 176 /fl put
-dup 177 /ffi put
-dup 178 /ffl put
-dup 179 /dotlessi put
-dup 180 /dotlessj put
-dup 181 /grave put
-dup 182 /acute put
-dup 183 /caron put
-dup 184 /breve put
-dup 185 /macron put
-dup 186 /ring put
-dup 187 /cedilla put
-dup 188 /germandbls put
-dup 189 /ae put
-dup 190 /oe put
-dup 191 /oslash put
-dup 192 /AE put
-dup 193 /OE put
-dup 194 /Oslash put
-dup 195 /suppress put
-dup 196 /dieresis put
-dup 0 /Gamma put
-dup 1 /Delta put
-dup 2 /Theta put
-dup 3 /Lambda put
-dup 4 /Xi put
-dup 5 /Pi put
-dup 6 /Sigma put
-dup 7 /Upsilon put
-dup 8 /Phi put
-dup 9 /Psi put
-dup 10 /Omega put
-dup 11 /ff put
-dup 12 /fi put
-dup 13 /fl put
-dup 14 /ffi put
-dup 15 /ffl put
-dup 16 /dotlessi put
-dup 17 /dotlessj put
-dup 18 /grave put
-dup 19 /acute put
-dup 20 /caron put
-dup 21 /breve put
-dup 22 /macron put
-dup 23 /ring put
-dup 24 /cedilla put
-dup 25 /germandbls put
-dup 26 /ae put
-dup 27 /oe put
-dup 28 /oslash put
-dup 29 /AE put
-dup 30 /OE put
-dup 31 /Oslash put
-dup 32 /suppress put
-dup 33 /exclam put
-dup 34 /quotedblright put
-dup 35 /numbersign put
-dup 36 /dollar put
-dup 37 /percent put
-dup 38 /ampersand put
-dup 39 /quoteright put
-dup 40 /parenleft put
-dup 41 /parenright put
-dup 42 /asterisk put
-dup 43 /plus put
-dup 44 /comma put
-dup 45 /hyphen put
-dup 46 /period put
-dup 47 /slash put
-dup 48 /zero put
-dup 49 /one put
-dup 50 /two put
-dup 51 /three put
-dup 52 /four put
-dup 53 /five put
-dup 54 /six put
-dup 55 /seven put
-dup 56 /eight put
-dup 57 /nine put
-dup 58 /colon put
-dup 59 /semicolon put
-dup 60 /exclamdown put
-dup 61 /equal put
-dup 62 /questiondown put
-dup 63 /question put
-dup 64 /at put
-dup 65 /A put
-dup 66 /B put
-dup 67 /C put
-dup 68 /D put
-dup 69 /E put
-dup 70 /F put
-dup 71 /G put
-dup 72 /H put
-dup 73 /I put
-dup 74 /J put
-dup 75 /K put
-dup 76 /L put
-dup 77 /M put
-dup 78 /N put
-dup 79 /O put
-dup 80 /P put
-dup 81 /Q put
-dup 82 /R put
-dup 83 /S put
-dup 84 /T put
-dup 85 /U put
-dup 86 /V put
-dup 87 /W put
-dup 88 /X put
-dup 89 /Y put
-dup 90 /Z put
-dup 91 /bracketleft put
-dup 92 /quotedblleft put
-dup 93 /bracketright put
-dup 94 /circumflex put
-dup 95 /dotaccent put
-dup 96 /quoteleft put
-dup 97 /a put
-dup 98 /b put
-dup 99 /c put
-dup 100 /d put
-dup 101 /e put
-dup 102 /f put
-dup 103 /g put
-dup 104 /h put
-dup 105 /i put
-dup 106 /j put
-dup 107 /k put
-dup 108 /l put
-dup 109 /m put
-dup 110 /n put
-dup 111 /o put
-dup 112 /p put
-dup 113 /q put
-dup 114 /r put
-dup 115 /s put
-dup 116 /t put
-dup 117 /u put
-dup 118 /v put
-dup 119 /w put
-dup 120 /x put
-dup 121 /y put
-dup 122 /z put
-dup 123 /endash put
-dup 124 /emdash put
-dup 125 /hungarumlaut put
-dup 126 /tilde put
-dup 127 /dieresis put
-dup 128 /suppress put
-dup 160 /space put
-readonly def
-/FontBBox{-39 -250 1036 750}readonly def
-/UniqueID 5000792 def
-currentdict end
-currentfile eexec
-
-9B9C1569015F2C1D2BF560F4C0D52257BACDD6500ABDA5ED9835F6A016CFC8F00B6C052ED76A
-87856B50F4D80DFAEB508C97F8281F3F88B17E4D3B90C0F65EC379791AACDC162A66CBBC5BE2
-F53AAD8DE72DD113B55A022FBFEE658CB95F5BB32BA0357B5E050FDDF264A07470BEF1C52119
-B6FBD5C77EBED964AC5A2BBEC9D8B3E48AE5BB003A63D545774B922B9D5FF6B0066ECE43645A
-131879B032137D6D823385FE55F3402D557FD3B448685BDD20EB05D5E7C2126132E33A59A717
-0609DCF4871A5D023C9EF57D3362D9F2D7A440BB69BF653364105F16F4D0F03582F9ACED3D05
-CC76489B16E3FA8A446094D30038B06ECCEDA269F2EAB9D19A99C7F939F9548F206C5A457A19
-270B2B82C43B091DFC5573468EAA3E7A4A32F8042891D85E4B180FCBCB3091D2800E54C87D84
-CE9CAD6869B5AABBBE47F40C68799893D22B765295E1E69E33AA048B7ED98BA480CECA91F3EB
-F8EF85FE9A3976909626B95AC5940D53F9B02215D84A44837BA25ED15CCE0D504F1D33506559
-4F3BC8245405407591CCCB11CFD4645DA60D960C0B93F187B0CF7B105543C0B70F89AF5D264B
-6C026E3AE646ACF145950202EC73282111E3E601CF2BCDE22CE3EDF6DB23516481420F26552F
-F4472D749811F27768150450D0D0EBE3C79F999E99B5C0F22EAEBB12D97782B1BD91B2A1F62A
-76412548AD53C0DD411D4A08C0F071C2C21863D9ADB75A4621803ECB84C2BB235B620B658984
-B2D8E0C4637E2811BC8F0D046C8935AFC70141E1B2D9C23BDB251D304B3378FAF8928BD09686
-AA0340FB0CB1DC48C996EF91530FF078666FAD227A3589F50B605267212D3A65EBC1019A8EEC
-9A0739A00279471A01E1505C17658C10030FAE32F274FDC8C8774B0D5406C384A1D17068DCB0
-C2575C562FDC5A2176609BAD9A2C255E426A4325A4FC3053C7B32F35F7BBC9AA50135F302233
-60FA2FBF019294B1D29C5E2AA489FE860FE1A8795EA9F2098A92C845BC50220471AB711BA6E1
-8C4AE211F10A172C60310C10FE9DF9EA1CFECAA5FEE0FC12063F8F48F26955F5809F8EF09DBF
-7515C7EE361CE74828581752D0169CBD3A850793875694C551826A537F8F736C280C09106A42
-A594173592AC22D8837B551670E96F78B71486427834F531C345F233ABE4153ACF5307CDC6BD
-FBF55788BCF0826812EA3D220AE12D403B7185C3DCC1D211F1FEE7BC12A6B005E4EB8DFF9EA4
-F388E39D6AC968CA08B6FDBD3221C01021D4559A396D424C14BA169501BE44275B5C72470C71
-8DBB92EE844F7DB9A5938CB340443B5FD69C7F0318056D55AD66025179B49D484507E6E1D38C
-9E3F20CA96B31EA1892C0E8B844938D9EA985DD202056BB54E7EB393147E7AB32B335EF1002E
-4099909E199154CBA70736830E49AAA18482203E1DF6C36FC9E661B0D4B4E9935A30127A1771
-ED815451983AFD8E6B9FDB0E7526D9566C2C9BCC5A4F8D8B69ABC257D12D97D2771FEB780A15
-BE513B3740F4252534F0BCAA9C755E94AE61AA697BF689107C4FC5841E9DAB101F3B4C00DFEB
-A79D50E74204835B872160AD3BD77E9CEAAE233C72DFC902BD16E457C7A53E665D60D6C87429
-D3CFDD52E7850E490A5F133082D17220A3146C19F0FB45F99AB9A94180FA8E8B19DADF23DBD3
-E2A099D533DC3DC0229415DEE606683EC67929805DF2288CFC116ECEE35BE81DFBED1A77044C
-BBDB7FF8924D50763F9DF29B0D67A7105C455F8A853F80E58D25AD263D860F2E31ED1789C198
-F2C9F71F8207C331E9A81CAD1759D93D65E6EC9F607F819C136C2642B192547CEB6AFA4B24CC
-B316A9B16426BF862316154579CDCEA5AA8AB67E9C68717920124953D48C429E368FE62695A6
-C755FE17963F4AD10EA11D1B043C73D03FECED825290C562041E819B81682CE89E3A8D712FE1
-A882B312A5786DC2AA373B8C733A1D29C641DA4666D939DD2C1ADA11B9358C292205AC7AB478
-7D14B38BB01AC9BA36540E3A2B9A1BE959BC5617EA85C2E67EE039578C98C94EA95FC8055EF0
-DC155B19ECB1E540FF78D579E629350A957DD0E41EA444B2CF13A2FE054E8952E97D9E2B1EF4
-9E63BF6AB23D3B1B208A2C030588B3B635BB798EBB1AF921F801B2711B301175CBA843307E7A
-3140C30BA881D13A533415D46E4DAC4BC37BEC8055A79C6578EE31DC696301C41EEC06A93A4B
-6F35B90D3B2D2B0CB23C7EC2BE8D09CE2C14425C1426458BB270284EC1C9389909CBCCEBEBEE
-A168B102B92D40FAAB8D73B1D610C234705CB6349ACB167578E62B9FD886776A51B59138A33A
-7B5A668427A8E86EF022BB8076C6BDAE83349F8C3B2FBC85DFD0E981926E25AC33983FE117DF
-14B94B1B6A3A3D9C8AA229D5BD71DE6DC3C01DD803CCD3EC679B64A9FF6E2B4B78CD3D7FA85E
-E7EB9E159F49CD814DF9C976F43A5974F66D7FC9BE938CA166FACD1F28F6564D3F7EAD59E5A6
-3E6AB443366F74E543BCD4A23073AA56CF26B12BE2498BF0AABF534290F45FCDB10AECAF7E0B
-8F124ACC5DCEBC60A1DD6F877C4F00F03DA11E2716DE7BDE52EC348EAEBDD3E2547051B80652
-69BDA67E6A946D735863FD50919FE961DED54F53EB1D5DB088B3F7B2D5EDAD30A731A85E19C7
-9B658F9D1E914E0189FAE214687B438B2858A0A660E2436ADD68ADFA8639B15694BBDBB74983
-8ADEEAD0A53447B6C9C46623EE15FC6A66B014E70E29999AC5C51699039C6504F026DB30AF28
-04DE56A155712CDE2F15B638F50C7BF4D3EDCF224C5825111A48CC34A8822C8315A17C1B5874
-B2C115C2FA893AD13043DE08D2DFD2F9521AD7C36BE22C61DA5D422E1E9A29B3C555307020E1
-29236C13257EA49F7494C5446A2BD8B9755566F68BE651E1E4BDCFD8136587B62DE3B79977D0
-9BA18CB403E84DD8194586EFBAFCE53697EE87AA0615C3CA76DD1023E36343C390A145A1306A
-D53D8932003D87263E5CF260AEA09EAC244FB41342ECB2F0C8F3D0E06DB67415034FAA24CD46
-883A5A8DF7128DD4A21C4CF6E76DA7696AB66BAB62D31CE6B563899CB066E2135F5DA74050A7
-7FD3808B550E477821B5B173F6585E3FB6CAD840F33606D65E2460E54FF603257735A6728F25
-31F69440280766BE322DE460EEBC2583CF0A8801D13FA4F636A1EC35FE36B02AE7E7C23E32BC
-0FC721A5F372B81EFBC8B58C3343CE67E34C1EFF8FBF229CE4F06321F1D580FC186505535989
-25D7AD09C9B9DD3447A2DC614E1A7A330E0DE134CC469B2EDB8E31F9663D2BE685815F962AE9
-B1CF51C8584062BEA53BDBDF52C5AFA07C0CAAC0C9F92F60CA7288DD7EBC53F61C0F670AD5CF
-C5E35488BA9A8305A0253628B4357993933ACA122D43D5FE9B76E881E86C7AC2B4DD1C96B527
-4BE5FD69289D9C37C95ABD38FAC01D7933458ACD06170BB1C87EF09BF45F75AFB40936611DC6
-7F3A809F5E9D82286A7B157AB8CC8109733AD1DCFBB95E6CCBCD8CDA7B8FC2478419C8AD0A5E
-51A274B11934F69C5FE4C9BA16A2CE0BCE57FB888368EC93AF6A1DEB09AD493BCFF8529A4D70
-422E3ACCF3E2F8B0EB8293333BEFAF1C0733E29F4BF1D4080B2F32778997B07941B59DF460F9
-B957D82EDD5596F9B110894123421FE8BBABA45C0D69A03A3432C02CA9768CF942AEEF860F40
-0A3E837FF346F0066C4C92C35BCC3E207227AC924E7F9E282DED77CF4CF703F89282571D3FA2
-A7218E75624E153AB91AC251D4B221A93BBD47DF57D70D8F95855CAC83EB021DE7BAFD06F550
-803A21063A06AA03D5458F697CABCA2E5AE12692C7D118C15B8BC9B299E8207A555CFC2B419F
-094198BD3F2F4F4EE78347415AC1559A224D8957F3B610A1E9FEC3DB3D4BEEAE4016C6950858
-8C9CC5C7976966A4D9AD20AF67EABD911A85F2DFFE448AAA8F6A0B2568E393BFC7CFE93FFED7
-861D8C455138DF1661CE4F8ED9EDB500CBE6DD745E43E209E81935CCC9F8CB75E49A3B2B1BF4
-3A59B993B361FE0F8C3D80561F071F93F04A1865A3145981F1937053EBED55DFEB7E5B4A1532
-C1004924D9214007D04B35C162A40CD23794F157FEE32C667DB378B42FA66C1365E95D770BAF
-BA64FD7659806E42BB3E1A33E0FD213B6B3CACDF3CA24881B4DDF8E615B616A32C27F00E9A5A
-4EF256BE36108F59CFDE91E473A9E403B675EE70AD5D165D2901EF28EF82D4FCFCD23FD51506
-41044BBB61B5B74B4EA4D0ABEB2546100AD4620B589586E5FA6B8C3C825FF0DC811045F6AC1B
-AC57EEAF487319917433495A956F4B0B4F2E113615CFC62DCC7D46ED424238187DFD734D090E
-96C308A573E62E96FDE5BA3C8F0D2DFEA845B48C34558241299111CBF1114492F790D2808CAF
-445FAC91499FF2F077B137C28FCE443ADCCCC1D5BEFBDEB8B64C452ACE58B96405C7244CAAE0
-58897DE7AC76549BDBF2D541DF175355DABDC292B4DF85B9E7127F794E1D328161690B0D11CF
-91DDE93C174288419871E8B7F0932BDAD96954FD1D955767D8CF3337C78FFBF8F7565E5FB755
-E252D86FA101A3A7FB119ECB3A893E5372FFA1DDF431781162B15ADC845D6A81021EFEE6B91A
-8E2267FDE73984FEBBA71AFE464F4A5FE5F908D06838D317F2230BFD733DAEA76A404ED4C55C
-59F3D37D10F8D086E5590352D92EFAF0D94A15D6FF0E9B9942FCFAD9EBCEEE42BD47D627BE8B
-DDD5A725E7893A62475DD24B4D9DB0738B5058A2A26A00D13006560C7F6D5C302CB6E6F9CBE3
-CF58FB6C6405CD6137B938E7F16B07D0B1A6524A3D34735AE75528E8994D2C944D90BE3270EE
-01D4C180989A1C6E0EF34691F0B168B33957DD1FE0E8F0B5D213CDE640D9490C8B68D8C7F88B
-D1B1369CA8A8911C1A8E3EB64A640FDC7BCAF45B3CF69360B1265E16302CE6EE2DE6F872CB21
-E4F6A2F232384B704F2F198F8603FBDA728B15A15C96F473E471F86663F25ED04CFEE32F3E1A
-B4E2EB241D0E347A245F164198B1355F15E237E4D00FD36279DAA3B8F826ADF11F2E241C4FF0
-8B7FF1F56E4E42312EB34C8F1DF5EA0F9C75BC377F8F3E842F09B1F0CE78FB0621DFBE87DBD6
-B052B7BC67D66D9242CE2714D54DC5DA474D7EA79DED5A510304A77A96C790D0FBEEF2243E4A
-3CFD2B393910C4039432ECF531099FDA34113DDF106E81DA407D901CA7CD99712B6AD589B1E0
-E36E949D0ED0ECCC5CA3E435B33FE26315E88B04A29D1BC179EDCAE2DFAA32980A32D70CD55C
-22AF4A15FAAC46AA19D0DE0A69A9D31F36388BDC7AC2572011DC394BBE8A602E0A0A870F6693
-DD80FA1FDDF9142533B8C5E75CDED67879560FDC4098C34C856C3D35FDA7A161BAF08CC4CE54
-C06C512B1DDF9955483CFB324E9B04BDDBAC1501A1C67491EF95D2548095D99959257582905E
-824E26071733C83F3E84F44F7FB1F4549D3245D750C5F34953B152ADB57EB3BAD5FE23821781
-FE2F256EA4EC2972243089AD1F7A001E1D9C2D680C1A2BEFDC23907EFDD645EF39C55A67A847
-73F0E9F5F0D64109D12BB9DFF604333BD7555D5E3EFA5A14192E19E38496E1D42A55009ED6FE
-53AD6979A1E56B5FF2C369686067097E6506B912CEF34F019FA6FA39AC99EB312B1B3AD02DE2
-35C40C65E0D505D062C7133D39C03EBE4B75CA6FAB5D02B5CFFBAF537EB5C4A00AE18FC50E89
-CBA3F7FA31FC6008C189665740A6048900BF94560F32AB5A34063822331D40B98524FDB23A25
-650882399223C8F327370073BDA351D165847D2F0B1EB6545AFDC31FDD29324563F9335DEB05
-7F4E1BC026F34FA9FDF0C94217250A08F5F8937C8C8C113FB126DA522B92FC8D70C5F04C6B17
-24E1B2978A5054AC9B43307F1C173BC6B545DC68BDFBD5BB5FDFB340870B2D147120817DFE3A
-8D382DAF3BAB1A53E85FE476B1306DBF540AF07C439A765CAC5DCA4AA6F4EB43B2F93C4985FB
-269E0BAA5794D3C02217BA8AD52D8121ECBF0C3E806F91206E0EF43FF82F561400DC2AAD277F
-C7665543C01C7E64AB886586A5BDA15C1AB6DD4E482475CC2E00DF1D355A49A1AC2E57CDF5BF
-A64235E1B23A480D0CE94E656AC92B1A48F2D2DE0462CA1C7A5B5B4FC97E73A262E9F17B3732
-38BC7C777DBC7B927ED471CEA8FC86E8547DEBF728216C750D460DA9FA857F92DE48A04C7C68
-05FFB805DE8F03293D803941BDD89C1EFFACCDD6A9EE0AC1EDFEDE144FE786F50861D49FDB2D
-089D24F98868FAC405264E4CAAB7B05696FD4D265764F8EA28B74E405AD12A4FAB428AF41756
-5238C72F25390CDDA953D478B6E564DE04A596B6210362D0B3F37404CD7DE65081CBDD5F4989
-436C4DFCDE6B238A4C20B3067F85DF8858A00CB8A4240F3CEDE461E82A44B6DA2AD6D494BE30
-10B46E7F0B6DE9702B55BCFEED0DD89A1FA02F772F52F8B833BBB43EE2671F7E78057D2CEDCF
-816C24EC3BE2DABD2E2E8A9BDDEE62478E432EE1BD71C7F79DFD10C1C184AD1499283D557ED2
-84CA2A2AA7876D90F0C9655100966BEBF404A28E19130F38657F107F48F705BD9688EC2F68FA
-DD38490FD9FFBE2113112EBE30598B9573BA39F1DDC2812A567352DCC3CDB54A677573407F69
-91E22153318B0507B35736307C1D1377621543A0987F93E99198CC9B022A99BC5F5CA8FC0435
-E723B2D84505E279BE7760FA38107631E82792DCCC71C3632904E220A16514BDB0080B84EB9C
-74405BDB19EAC2574C0A7986A36C72961A55579950929C73489F0A07CB0E2A69367F86B2E5DF
-88753AB8728DCB19FC80BB077805470C95A881EF657B842F00EF72D9FE3C754DA84AAA5E3799
-041AC567D789FA6DB50C1B88C7CF4B5E39CDDF19CFBAE9225B46B0AA3B83F9B4E8D215AADD47
-F2BE0794BBE4683FC3F5AE9301E7D16E3798E55B193A2AA1E0C2F61003CFD6E10BC7725693E2
-705B323DE8434A2509293B278D71922F4BD432C2481BFEC17873B74205C5DE79B2C57F518F5F
-D0F971A4EBEDE17E4A034C5EB5A33E2E0E03BD6F61BC7DE8EC3882A0FC885F020044D1AD7D7D
-7E0DB4C425BAE04FCA6953B2BACBD9FB2AB7C015B72EDC5B05BDCDE08455769EC2A76E4A557E
-CDE85DB30BAC84497B35BA2BC581CBB60112CE4DBB17D53A18EFEF227B049F3444CECCD0E788
-0E36AEF24920A2F8EBA2BFC07E91F07F186A80E390BB49E786DA672CFAD2DF5CB7617614D6E8
-65FA2BDFAB653CBBF6A617B1A3BDB61BB8750097E0CC6CC94EBD23438E0F8EFF3F12D4529523
-BC5FBFDCCB4D4F29271182B3A848A7B55A46262A9BA3BA0FBFC983D136B8EC45F4D29476779B
-93302C1DF35C73C87C1B340C661BC4984668326F7142227EAC1010273593C86198832C8D5594
-AD30E3C90672B49EC8C3E25C98C065EF239C4F0AB97413E8726E7C0766906F4E70034FA98314
-B9F3CEFC13E40F14A7883C4DC850EF4CC8D637613033D74CBBDB9837505E2894B9FB9B33E624
-659BBBC22826E155D580EA8B9EF3EB27C070F619ED7EE7EC0AED05F45E5D9C4C6B6D6418E019
-C81E49F18FA808C36E501CF9FCBE138E683F914B25AFAEF8E017A3E894271F01CC636FA29537
-79A4B783773DA5DC3C70EA5D0178AE18BC3DA6A4267507029D47A552C7091F18E4A7900ADA98
-73ED4D472FC0A9CBFC3B08A2AF391DB35093EB03A66A99499DD68A7D9D266EEA7CBF6268E81B
-2ABA51B49A9869891E0435ACC6F9CE402AA1DE2A496EE5C1200FA4E2CA0AE893409347A21866
-D32808C3B918F85CD4B412A3EE8668605860D74BF0F13CA4AA1D6DA673AF97CA5FE989E348A6
-756E003CBE748AD34E0341006939E03D47182EFC7C45171D4E127FA471A934FD2C788C9202F3
-0C7E2AF2C2D71F3A072A2A00A1A81AA55D8A546235ED9177492EEAA89647E6C3966FD7BFD46E
-BFAE9EFCB7DF8BF03B1A4597E686967055BEA5F3B05F48D3CF37C33CE1C17B911353560CFBBF
-0595C41CAAF024FEC83E4168239C22252F5992569B275B628C8CCA3229EA4AF64221CA6D62A8
-84D636E0854A1B5E0B0307B6FCB3C6E6B0F887E516E132F88E1830C0801E0197B9CB76CD9A83
-13BD3A7940D7C160DA3D9E67B5E89C9732C13F24CFEDCB7ECEEDF1BD323FB7BC5B3BFA8CCC5E
-081FE114A5414472DAA4F889BF94B387979A8DBE36998C4D61FF3CF84FD9D492E91CB09D668E
-DAC03C71CBF507E69A37C9C5884BE278EDBB604E7FA0A2B276CDA974571FDC0B645590723B6A
-571263FFB8A211E84B27502AC775A91856D368D3E776F0050AFD4F39FA9C321E405356D060CA
-26EA89021F48FBB183C9FFBA571D5E4326405C8B6EF0C0D7E8BEFB66B06E3329A4DE2DCB7957
-E41BF1345AE0A40538A9EC7C063360A81B3DBB7ECD87D28666E59B105A5EDF2822B64C34CC48
-DCCE38E4DC8D49E20025B3E3BAA8D798EC7ED6B1DAB67B87BB1CCC3C91390CE463CFC0FDF02A
-DBF77DDA6E3E3F812BAF56D822FD081B68BF59DB3DCEC4FC052E92A653C5E399F749DAEBFD18
-C9F20C0699E18E1BBE3E4275F2079E59587B4462B0D437B15CCABBD924AD25694782E0395D74
-6991DE195D5CFAFFAA3BAC74AC86CF67B66EAE6B9604E796306096CD3DAF489AB2711270F5DE
-B93097F45AC29DA4C7FB493DCF76A56278BCF2627F49ED20EB74852C00BECCF2FB1D69AD40DA
-5ACC2FFC20551178F4860776CA2EAE54779374AF1CA3485F314969403B33F4C61C78FD76275E
-D73A513B7E4FA4F2A9355C013E4A981D21FD319A54F139648244D6B05EFFF45D94394B0B5957
-E33BDFA11CFB0EABCA6D8C69079452DBF6ED2C934F01C4FF3290A773EC24907B81CF658FDD4A
-626F37AB3FD4B37F0CF439BFE461870029F5785B7E7CBAC4594DB8C51F8CF2EDEAA2E0190791
-A0AD7E35CD2CC66D5F5152EACFF983E4A12D321C6B20AE605DBD8A323A53F92B6F5B0FF37D40
-2D46A9BDEADCFB3530ABF0E0760C96658C914C4861016EDE63722F522AF61A888108DD9BBA4A
-AB5B5C6C8FE89FF9A4ABD054AECA521F88D96A86FF7EFFB5179EF3FA6549F473A71F84F11E1D
-C3AEC1719505D1E6B7DEAEB3CAC29972ACAB9B8770B53F5A6AC993C74921847D6A566BC96567
-4B3BF7DAC1718854A47CDA8F89574D566EC01917C43BBDA54D50419918A84A3118093F990312
-C37CB5D67A406142DCB578C810F27EE33034666D7DFBDAE9EC153B2889D38DE3D923778BD5B0
-D6A00FE38E7C3794D29231BB0118F54253B7F10D9F245E2C884AA5EE365D5689BBAFD0B2C2F4
-046C5703B5C01A2859D7E7716ED1DDFDFCFCCF00768061B060F48FE16496FEBDEF964227F0DC
-7311892CA3F2605B59F2BA1242A882BA4E6A986B61CC4146D3AD1B4CEB6533319C47E7745834
-FE9F46B9ECC7CDA3E283C7C5F744E55377DBE88763E48108A4A1E2DCCB57902E77EAA7C95E02
-C654950E0644E329D2B7875722CD3E573C5ACCD5CE0CD66BAA840D0315D693A8BB275FB821A1
-FAA9EC0EEF490578473B3705B327E41BE324B088F424E6BAEA36DC7CCBAD5D23B8C5E4A19A74
-2685477C8363BC6A84B2065A1E02AF63D6C33555D5E6A8BFDFE07F2E86ACF69609B7898790FD
-F53D8D5B9EBE7A33A5252EB71FFA81D8ADD80BC11004F725124BF3782DC99E76B702AC01E0A3
-2E1BA317DEC909ED48094801C6B52B58BFE0B9498CDC28B83472DB91D3CAB55714E934E08BDC
-5027D907EC675C94BE19185B62A8666A8777F139C58445FC70B18AB35D2E22C3B9824A38219E
-CBADE61729037510BC840DB0E3D62C8A8F33CB124F12F71DE5CA918E3F593586251FA72E8D36
-A992AE354423303EE1C2BF3D3D37E64EEC500B31ABE2188049EA48C93A1B99D9A82B2C813224
-68560F23EACED62565F3301835E8F96335592C227169FE2ADAA1F7D8D6474118F9E00FF26D63
-0F86DC9509E021A59DA152198979190F2BD180AF8E82ECDB204D9A2FA51B182A84E958929F18
-3DBD8BE6575D4C13943F00ACA12C2EA603B107D15DD7401CBD5FE46413BF16F99DC7618DC526
-EB041DFA841D964A6B5B438C1470EADB68FE4E97D4896395C893EF7177C2909DD9BD8AE8F25C
-1A71A6FB7F9EE515715F4361E0B01E6130EB3506E926475F4F1A8B7E162EDEE25286E7565448
-9CCE65A80C54B46C6EEC990A8F5006697FEC8CAC1DC79B6420501A66146EC648E49F5A1A9739
-25E1F4FE865EBDF350D841DA96C3FF7D9955A9F4C3B2991409A50185E20FE5FE331E569CA978
-6925306020833F9F304C9780E8A142617A0B4758AF5BF73F9DBEC6B0F3DCA36648CEE09C08B8
-E23147681BE470A753DB32E84089D2AF3BE64B9C242182B4AE759EDC6CBF43E5825DD2E86EDA
-6686D6C45D769BCE34EA05E61188E4491FD136CEA5AED10E228E44CFCC461B507D2DA03A10F8
-80341D2F0DBC12F8CEBF0A89F1587E249E6D2089C8DE18CFD1BD691EFA55C66C71076094679B
-46875C78397C0081075EB45DB6DDB519D685E98C44A1CED29E29EFDEFF2B76137F55E1B585EB
-84B8AF5AFAA54DA43001FFD88A84747777A7FCDE41025E6BE63F0CC91B97256903DE43D3BC5E
-9CC0572D9DE3D185579310E93D6775FB33EFD1838B2214EB2332C1AACB0EF116106867A359A4
-A69297B6735176FC07FCDF7DFAD6C9009F87C21B3181796941C2CDF8A2AC2EAD7A70143D8B69
-29E1E36AE6C45035762DC43EB810A2CC0DF7A826376E9C08D7BB38360E0FC2B60B6ADCF4D560
-99D954C5436B4F9C2D15F8429BA4792810FB13E7DC39BAFC1912E1BAAD45F6AB0327A5D93FDB
-9D947185FD936229E4C99280878D91E2018979B583FAFAA33D5D4730D57FA680922D28B12D7A
-F720A080207D4EB4A10DFBF3F77CA336D3A12326A5C3F710AEE63F2767C2074199C9E2CA8851
-7E46D019A8B4F880BB6D5745C168AA0492B7569DB707675D485067F313A2E1F56F7A2C93CF03
-3FBFEE635AD06C6CCABECB5CC777CF274E2A3A7EB96D0B28C928C59E77FB3BE301647EB377C8
-D2E0B57B647F2B7E7B01F44519A3C561F7CE84121093BDE4F9E799B0923608EA9CED23E03F5A
-4DD1425D2E6219AD8E1388C46027483931FBE864FCCE49934BE30DE22AAB04C6BB453D52C95B
-5683CA29C73281333A4E2424B3BAB34AD8AB8270327CEFDFCDC4E63B3BC7068692BD663ADCB0
-2C7BA09EACDEC25FD07D2E236E299EE0F9290E5B8B7F8862FDBD81710468818674C2BF94050D
-C77B479F6FBD83EB939BE411711746229018421BABBEC891DAB8C00BD15602A2236927BAE40C
-057B88F3834E2582F77DF3EE0F387F6DBB23282F887956D01F183D73DC0B074A81432EC813F4
-ECF07EC8C51620EB7B796C22481284FBDCF26505CC200BC1D73821A5A8E3975D9E76FE594528
-F2391911403CFD4993A2CDBF6011E304A858165C6BC439CAEC6723792A78F565BABC47D48B98
-FFF4CD726C9F35FE3FCE4B2955EDB611F0EDF6893D904B2855E468EFBA52BE7BFAC6867ED7CE
-7C226CBFDF2A4B8A545209294FD58C883FF1398D24800CD1823500B0FD98B99E6D24F11282AD
-35FE21D97826B3CBAE293C191D5107185A620CE3F30F04C540C9F39C1BF0F9F59FCCA1D06558
-B0A01B33308DD3B9EB11E881B9A8E317682DC53BD91D7A984A751AFBC2548DCD4CF62B12414A
-CEAE7FE31A084CF7C809AB5EDA0D7A65A4E1552DD3F620AF1BDFA235B908FD2A56A9D316C262
-3812C7967FB7E21330D3E3F4788D58F08DC1266107842337EDA9D915307DF3605168D3456D40
-048E36016250CF994AB7ED3343053B638E58C462B99EB88F1543FEB8E8DEEDCFD096D82A8DCF
-6F978ADBC77DC4DC1C5EBD1783A1BDE71D6F1658E1BDFF1E8FCCDEB0B9CEF930421519D60C98
-96906A79C2E5207D6629DBF4A14863F7A9B8CF3CD8447B92004F29092F1ACF42045240EFEE4A
-70D69F5858647C686CBD317881C53E59795B32757FC33BF15AEF5A830BBA62DB8C42D1153FA8
-DDF7ED7B9849A2080933EE4CEAA4327C2320CFFB858328E2A936A58D8A94D334B463A9BDC35A
-364A0DD0CAD4BEA0AD9B7366DD984AA3AE74C823BC32EB265A7B5D122834BB53DE3A37B414B0
-3320129D8F7B78E59D3459758F40D049FDAE7588828EA1218D828DE318648BA0E0A414DAF43E
-227CEC67E9E0DD502D9A7925D5A2281701DC636BF79362B724FD93FC64BE8BA63CAA4D750113
-25B3C0C20F8DECD8C325688493F3347A71309482A2613B5048E06FA6AD50C6D0C615390DD3F6
-F0F549C333D1654BDD3E372D3970FB9880D15A2D922E86002A98807F68491D29B1341E562B56
-5C220B70F0239538551C3CB40BD1B45061205A60E3345BBA21500E8B869533EB266C31BE4FF8
-4F42F0CC5D8B178B67BBE7CF357199A32F94C1E8EF8CD65515262B29464BCEE3A93491DC134F
-97BE64FB222F15356C11586C7DAC7B2529B6B6246A028AAE3BC08800567F633080963BE6CA8E
-CC8D7385670D9DA59CC3E33253630AB01D3D5FA4DFE77523D2CEC4A5DC8DC62D1515BD78E719
-2C5C1C2199DE7876E36BD6FCC3C27AE577A72247C807EFDC6635FD80081B8CC78498002EB78D
-FB2B4A13B3D1DB8E6315891915F445E046ADF168358F713B4358FDE076907DC9FD86A86E3E27
-1656E9C6D24BAC96C6D400F4A931750F9A4E753F1572E069B4D6F435D63EC1E8B223DE713CF9
-6A40F3745543B8A9789DADD2493FB1738F06F3F35E4C295AC740A44557DA6E74F26A23750FF9
-D51E053EDBCF2CC30D9B6DDBAF684AF9D0EF7A14C4EB101DD6F84C3D7047C32293407103A660
-DEDEFE02EC9869C2FF2C0ED1CFC46A73F879AA6BB207007C76FA776098139F84B4A0D28C3AD6
-D500517AB58BDF672C24AC6979967698B28193C40B10E6EA1A3705B6FD1C1C2FB3813B3287F0
-D46B922673275C63C1A7126A84605D0AB5701C88FEDC467328998A8564FCD953C31C7EA2B668
-AFFF9872F86DE95F3BB1D9DD6FA033AB0B8E903AF01DF2E6B060E3D5116379676EFF7CC5899A
-ECBFC89A0D7FF73E46174F1E452A79E36353595A22493FB7D98AFC04549860F0AEA927653A7D
-B1AC86542502FBC7357314CA668F66E7E44A33F60820FA1F53E4BDBC448FCEB4D685FDB06978
-8D870B203BB18C18C8970A3513232662E92B1E5B126390AEAF1177BBF1EDEB55A1DCA43B4572
-3C9E8789EF056EDCE4FBBF73BB37C90FADF5105312E67F467F0C0BD42FFD8E5643895A5BDDAC
-12557707FB82AFF8DC9813D0B931B3E7513829E0BE3FBE342A6A3786F2656FE68299968ABE35
-5CF4FA53DA08F6745130336F5AB56DC6549EEEF1F579F0060298737DEC827AB7ED5D627BF7B9
-D8EF165A429BEA10A1C2ACE4C668730912EDA9F75D17E215827042E1FEF5D72E643C1BCE5E61
-F8ACA43381FE6FF1E736654D11FBA39EC251D69CC3C5B7D9145ADDF5CCFA6DEAC18D4B9F43E5
-2F3B660A34610761E9937051AF51707581ED226DF239F8F7C735B4FE3D9B251EDA1C9BE2FE14
-D603BC91B22CABD7384D9DD8477F66983042D99C911B14DC0372C8350B3BB1BA4199BE157A08
-DDFACB5BD368BFAD813A05890B8BDBFEC778301DA6CDC7FC67B9300448505B7BCEA5C1F4A06D
-0D39C12DFB649412827813FEC978530C5CE020507B08F90FFFCAC19B3E7B7474ACC6F4BA357D
-974D0B5D5F4C5FDFA899F1895595A51FD2590F0DCF52CB5B5A01DCB94589900AACF23887A182
-4C95295BC28EB912FC8512CF1465B7D1469AB94C104C1C1A57F22300CB9A151B7E7431080FC9
-38868B5A7177569212C1EA0D185CDA5503D592DD781386F619661A4C1DF9D9082F44115BF318
-5419473A1A786138E344C49842AD159E143D671BE1F224E06E71B5E14385F7DB7EB083F8D3F4
-CB8B3FBAC22CA88025F8CA84BDF31FEB9D0A5A379B8637E4E60A5875860945AE85A04D731C6B
-C9F4EA2DFC84A4AED87C6A3FE4DA0578A82895D643D64A99FBE989CCC2ED431464F919B89916
-730868A3DFE6DB998131E15171E39AE509521D4029384D68D6629424FF243B8F3573B1656A89
-38F21E0FF51702A24DCA86CDB87FFABFD07E23245E6A0522B15B8B7F02B32AC81559471C6032
-CB63B5C2AD350F1E60713B798178FCC159EB02B5C033EE36F3D760C469AF012503AFD3E4BE8C
-85403DEEEA0F9862A1DA3BA9BB4BCC3CF4DDEE1D7EDAD994C790284ED6AD9DE0C6E6A73812AE
-1A9295782ED743777A549B6A99C55DE9CFD50D98B8ED95A54D9E9516A5976D8BC3AEE24FCA52
-52FE4D75EC8F2D9AA34BF988384A3A9F86B1C55B6B38519CCCFF89B7CD0165DDF9B5141C173A
-8539C8DADABDD4786452A37B2245C925A0CE6277C5A66B8EA55533009535DD1A456057D5EA3B
-64384D89BCA81F7467518138B5A24670CDDBCDFD5F392208917853858A8EE9A5B4F2AC060ECC
-9A0CE34EF5CA812E15E10328D5234D087E80AB369529DDADA470AEB75555C584C7F68C300CBA
-69CC30169760947DCBD0382E4F2C68BE14CE1C521440AC192F1016EAA233A23FF59BBE98DB9A
-6FD16AC25576AAFB7CD417033A9A43F42DA250D9AB53E96AC10849C36348ED4B59B9E758FF89
-7630E8AF3746394C15FC94E1FC0FEB30EFA61C962E2380EA0479C6F3521634B91718CA44027E
-69D235ACB2DC0F45349AF9A4F7BE58056F6B574AFE364128C0AFE227231D0AA97D45E01914EB
-2EDA018807B73BEEEC8BE2CD1414E898A47267F4C649B326D097152A984B0492C9F42059BFFE
-6189AE83031808223BE65067437D7A89445245F7C6F5F61D87B86D5F3B40F9261EC3E142F495
-F1B199F30A13BA2AC6770B7C73D3A333090E267D3B127103A680DA1A00B0C148F9D000CB3011
-0292FBDB1FAD386C862D5A3594A31DDD421C680A83FF41FA03E2B32CF3D0FE3086EF12DF80E7
-B5EACB2ED8CF1857A00DF3EEA3510F49D6F7540EC0BD08625250DF8A3B2F3A160C0A0035DFB3
-4B7E960550C5C040F0AD461A407AAA9AA7877A23CF99551FDAFF32B82FC87BDBD86A6CC9C63D
-F6EB8BFCFF2C1E8A3B57BE35C0AC602F7B4F3A8CB8CE62253AEC186B1BDC8320D5099180527E
-1CFC40F7398454B9CBFC5BA0B421F55DEDAB33CFCA11217608763D23685E1530D0FA1AABD091
-AC9A6867DDE46A733712CEF63B46E29FF70E9210224E0964FBC56A9AE0F600F720F81C7060E0
-16ACBC1BF97C16F26255BC4D82538BE080D37774642CD4D4C95B3BECC85E32D182BE7A1F450E
-F077500EC91B21C383E71F9B6690C0616457CD8768916022279A6BEF4589A85732149014F067
-746C3FB95DDCBA28C7F0DE2C81926DAE7967A387A06DC62A82F813230BF1F2854E214984001D
-0BD4A7F8E778CBBDD88D702863FC3C8593255627F352DD3327EE291A31F3DF409F108DEBA842
-3116CAD62EE04D8CA8CECC16B8802F698FC2939E3898771FD932799B54705A481C652314C8B6
-4BBCC883296CE238804F68346FCA2F20EEDEB4B9986B470B84CFC07984D5A1E896790B854B5F
-21BEF81262E1E32B949AA9EEE6BF20B9141C890A41081868E83E5780B54A456D1D5F21198126
-93DA3BCC8C88D6DE832717BCBCEF1ABF462E949A415839B98D91DEF040B43D7B1F174EC74506
-262F04F1E599CD6B01D8B5902C029C97487A1E05F84DF9D42FB47DE21F59F91B78BD16DD9FF2
-087D1A73C0C58D4C9C53F64A8E2A076D763B581AAC565E34B05368A456F52D64528EF327ACDB
-4AB6263C1561EDABAA3032EAD94A6FB956F24C2E18B6A8A47BADA39A20BA4BC035CB8408844A
-1CFB7E4107DA5FF66C307CCC01B93DF8D99FAA99DA5F3719FB26C2F214F18494BD92F9B6DB1A
-F2DD492330615C3BE70D6F3E2E77DC8348D16339CAFDBB96A90436D2ADC86A1864B5C2F4809D
-01717A052BD791C092D600B1E2F6A24A21E13138621461A3BBDB4EE06D880B477DEE907B0928
-A154AFC2716865D8F751C398BA65E70F325089882018E0CC5CF735ED15428F074C17AFA559FB
-3E94D43340136BEBAA8E47C773DE6C28D54B3ED9FDE6622924159135D49490EA90A7D91BB08C
-4345C0AE3FA4A1461422EAFF8E9C49329AADF6342B7D55FC513821DEB5E2DFBB36413E954E39
-F5A24B3C974839BEB97250F1566F7A7F2F9F603D2A5E43CEF4AE5B0BAED2B6304461E8C16A7E
-F458E6542E7F8933E6C253341874867965489527C80260E254411265EF559A5CD9E25F6CA0CF
-EF2DDD2DDDF07653B9A0FC6066903189D076F485EF03D52D14D676406AE2CBF46E818D964607
-759091C5B0EC784AAAB7BE495FD20876E11547DBD9CB02508BAFFE62D3DE7F5AB79A2C94318F
-E16574406E7933D4FB8FBEA0DCCFD1553BFF56B1E04BBDA2B43B71A77DEB8B445E74C574D7C5
-B530BB312EA0EE227242D4FCA0CD13FD056E30433DFABDCCF6274BC85B3968AEDD85E3ACA8CC
-017B8077C0991C4C83490F5C59F3F383C6C7BF0A7DB9C4F8D76CD146E3F80D4DB44ED37674B9
-0091E05114257103902A9966F75DDD14A0B7EF5711A429BE60E378A89C7E1B8AEFBA049AB1DE
-59D36A037AA41592DC56E731E911BD486FC06B17D66074FD7332452EE69FB9909D8D1B245E8B
-7C3A4061F3D2DDB055CB9C8BD2BD1BF13C2239E90D9581A323E8CE5865162AAE002F29A05989
-938C8F78C781EB5EE467854052B9B520E9B518FEEA9DE7B9D1D4D13FA85892105E383924EAF5
-C7634E06F9590E1661C0D12F3F867AFCA1A8E34D8F2D7FAFDF87D345D914E3065430E84FD70A
-DF71630C6C836BCDBAD5AA6A1E7C73F3DCCEB4E351C1B3CF348EDC47E8E1E27EA2CFC2B3E7C0
-BB3F7AFB557D158D820FB815B8FE1DBEBDABC179B0914064845C90D2710D506D17729D57F2CE
-A990A49176E3A3A128137DFA6AB814A9B4F544E290BCD8A13E92D4EEA56889B0D2B58C50A02A
-D9508365881471FA7FF4C028704100F45BA347F8F62B3B8C827C2B19002F712D68AF8BD7A167
-E6B9D53BE6F44496E5B00C8458C2A15B90BFEC6E5FA5ED5743F4E607B93EAB75FCAE34A6B163
-D919B80C7E62B1E119BFA65520E2D7B78DAE89E34E0A3D7069EFD1B7D69B7396E741C1E81015
-710F743D2707A8324FEBCF46DA37F1D4A26A62DDFBD03D93F13144296D728EEE912243F8E8F3
-83042C51E48CA1A5FB757E6673CD599BE052BE15E9809F4B1C1373F0781963B39763DC4B3473
-0936DEADE5501C429D265002A4A280A1B7FBECDD6FAD9518BC26D0886314E2C3DDFC5405FDAD
-8FA9B0FBD53863AD59BAF1B0BC6ED0ECB34B7EDE368A1825A896AC132D4A3B598989677BA608
-A26F6B35E46ACF78CC959F94787027C7509F8E8248B32A68F75A5E14E2118DDA7A3861C3944C
-88C198012E811C71D94A3FDF8DE7E842A1662088B4C014CA1C29BBE58A46A8CD7E82E277E8B9
-EA73F75E0BEBEFFFE0A2737F200CF05A697210847EDCB1CE17CAAE782554D57C2C5F206893B1
-49ED59028707595032EA8BCE284A9B409C69C2EE9ADF0CF13035A1AE3084B58557EFA596B273
-DD7582EAA8153DE775B991C27124F684617F45CE264F3749BF05DA720BEC48277BDE7320CE97
-B39317D2D88CDF1967D50496CCAAA99D341E448160DA81DB763F8D721324608E9E13E7B92318
-22C669F6F7472350FD411BB2874935232E47CF8BE653C8DEB6CF00AD9367CD80B674DC67709C
-8317CAF2D618170F99155E2BEE988A972FD99F28AE1990B5C8234721ED0B5B05BB714E2F84D7
-056CF0D2B0C46E1823163BDFD90642CE898A92016C9D05FFF063E7B4B489C6C452CC7467E64A
-59978C6F82DCB872D322ECEA7BF2DA6187D78F27E4A1A9FAF3626554BC5CE69DA195CA0837E3
-A44AFCE6F55C2F540D6B9551CEE083F75155E978CD5BD4782C24EF20DFDEA82E66527304EA74
-CDC79B59B7DBBD262A2B37841AB3726D24F9E2E12325A22EFB884127CD3028C8CF819B23CC90
-8878190F8436AC1934F3478EA966708280EEAE8E35A31D5AEA2AE696258DCF952E53307138A4
-5F54E38566784A07B07B77E170B04BE12772D72D31C58107360F403BC991689A984F5A2C9B93
-7EFE72A267F1547B7CC477FE414D754EF566495D9C13C90FE4C87963011B0C0BDFF6A7179FCF
-5DACA0E475A9BAD39E7325C238E0FD9185DF54465F976DD68B43B057399526DC8C7383FEF16B
-AE73E89CA2A5F55B6EFA49DB81D19712DE1E5BF9C7B4536AA3F01B1EFD1AB3E356C2C27C915B
-9B795780DA39875DE71F836EC0EA0DB3493C17817F07D34FFF52C4879DEBD510088A0A95D1BF
-B06ECEF3524488FAB2F01A5DCEDA2F1C17BD611DF5455CD09AD5712144CFF92037EA80CF0058
-6543E8B40FFDD9A9E34CDF450D4573A9B279DB0A3AA8F17CA50E0881F8C31F681D55D6641807
-A08844309A7E40DA8FB21B1B409B2093275865E0375D7D5F9A9641C2BFD25DA3B7247A8B4724
-4B1BEB930202A9622A7636296B10846CD78FB8CB04C7BE4BD3E74DFC7627165FAA80DB2AFFA8
-914804A91824243A59B76CED9648D12B47CE8950BF760D865136ADD4F83C7CEC9CD730416E71
-01404019522B641F0A4DD1CBF7F9166180FD9232D94248D2ECF149A1F6FAF7F7CAF0FF409079
-10DD32C559476F6B0DF68DCC97716F82BA617DD0CD2C9D147A621200CC26A7A768796283BD30
-65AB4F569933788C17BC564568019D169E5E43D0E46C6CCDB11B36D5FF56B54786EA1FF68A98
-E8DD16F41B2B224D906F09A4347C58278F0F151AE45D1144D05E2519DF446795ABD09BA61531
-F93553D6305F228E82C1F28A76DB47A8BECC7755717952FB5506E4F26AE8882F6AE8F17AE34F
-B80CF2B0D0D99DBA2AEB9221F82155E9F666C5F1709D17049A2E2AFC05CFA9779D10F7AAD03F
-E76F78A57B1C4D844262F5A889802360BB50C3949FF9AF199DEBF95BD5D0A4B13B8663442C92
-DF829C65654C4A85DA8E9CA7224DED598801EB8019E78994A5C70980B1BD09F3B255979194E6
-1DBA6513B5B6288048A9E93DAB48AF7070FBE494A1DDAA087B2278045AA9C5E747A5D49AFA72
-D0BF8BCE89F182905973FBA60C5F65C079E143B0D3DAC6E774F5B298A18CB11AF493BBE2D783
-51B4688BCEDCC6A897797CF30DB8B600A5D5DEB6294E0EEBCB34523E28BBCFECA8C3B467747D
-F85BEB5127DAF911375335C5D8B3F065B1B5D6D6FEE564D88B56897A635822C09C999BCF91AD
-E7143F09F61ACC8C1553B5BC2C9A2CE6F86AA616A89808846FD1A2986D6D82013F1A2B962EA8
-1A418EB77062B721C0A87927E26358378EFC5E75256842F692121D4C5DB323DD87DE40C507CA
-C183F305418C1C4967C2114A28B5B3DA494C01F66AD78438FBE9CB77059574FEF5409FBB77B9
-428355422F16E9DAD5AA5FD0181668A4C896D3A5AFF24C7DBBBB88DAF95FCAA95B270E519F2A
-8729D013A5775CAB79E656BC515BFCF9054154D0C12C2D056DD947D24F553F40DD7052F4287E
-3F0FB30E89190444CC52286959AC984345B11859DAFF089FBFA03D6111DF78DD4393324A5379
-71925FFD6697C14A5A1971613124B13EA20FF1086986DCF54BD6460AE956BA5C3565E43F8DF3
-5AF79C637FCD39C1361CE296612CFA01750180E7857BBDE7DB8F6EC757D71912328E41A69650
-271C5380931EA0B6B5227D9094679D69F04C2EB2B406338ED59ADE317E80198D3CD722B0458F
-0A9274087325380B7E0047196CDDC363D161A49E7EE0F262C3CD9CCF88D279953A6FB0C3CB67
-390EE89BCBEEE246D321E076AB59B8D3C25CC90DDAAE731D19B43FBAD97A7A97BC21DF4FE65F
-673753FE06D633D156E57D1E87433ED1660EE8C34FBD3B2A83EC5E88F98A6F45415EDB555880
-AFEFABF11EBCDDF9DC0F6045BC18B24534BED2DED3A5FD0769F16C521A7F1EDDB0D51776B482
-DE96C8FD119A03D63FD475D7CEF48C854D079AE0D10A96F51171437C565FEF88AE1831C7D56C
-5FB2B184A882200DDDCF0552EE7E809DE45FBAC8777BEAD89958DAAC04BF6142C21191EDE23C
-C7710950472EA107C0EADAF87CB88F7FF9D1ED061A7CE67DE329DE30613DE5A30EC6E46778BD
-BA41C95CB9285908705E47E9B52F3F570F17B0190D33AD6F8F3255A2C499E81AB148ED7A3745
-B0B8FD74133ED7270CD24B33FDBF49E6BEB88E1203537A40FC6D1D3C545C75C7E42663CE5C71
-8BECA455EFAEA50D18E57FFE96F76452B2FE6FD147AD36A88D6DBDE307AA7A0DA0DD21986A79
-2396784F604A84C06EFC40765FAAED39756917635C582A8736D875D7F33ED344E7893BBCAC47
-176555E4A11AD404D1318A778D924FA726A50F4E114C29E44E9AC842060310285BCB78ED2087
-876CC115C0F837A4EBA53F07ED6594A932EAFF40ED42D7B313B99BBC8B10283E8FF683BC93D2
-BE4D2F370F0D2B14E6EF4832FD47CE536FD40E8CCEA442C6956D2C4780423BCEB9680280F132
-FB93670919B643F49F05F205F90EA3091620E48699D55EF35EA51397422E16BBE79776C222FF
-DD5A2E5401435CE0A3690EBA317A8B3469CC19EE0CEEE20DDCA91CC41AEC64E4B483122B1BB0
-1B67333B7E1F39D808F44179CA3042A87DE4649EB1CFD5A5E8231D08221E2C79CCDAA4267789
-3C6417F47C0DB97D8ADAAF7266DCF9011AF7D71CCD8092FC994BD639D80A019714073F3C7EE7
-769056DD117FC021A40DD98DFB115A596CDC0EC9AA9C8F14EFD96B98EF64F1A9430D8DEC60A8
-B4F20E696B812DD9C0EAF3259B69097C81EE43F9546FEE7A556A26F0E1ED927455A1BA0CA65D
-1148A2974516E79B04EE02DAE02A211DCC88C5F8965BFAE5BE65A25494BBF7C23CEC8115B1CE
-2FAA0CBA14D99605A24F14BE129F80C806CD9C5273AFD406AE758B0FA5F81C0F2D7CBBC2B0CF
-D30F525E35D937D4423B2E0FB8C83930A38D1620F12068CF098E459E8E970EBB0C44B74305A4
-01DC449EB227B665445CA242D42B8F73667B2BA0C7BC356BD5C80FCADC2256CB2D0BCE609ADB
-60A22CFED22A5DA1F53A3B7E9535712031C472D34CF2F771520CDBCF9A3A5E878A47B909B794
-91A2D0F2A370957E4B1B3A516A3A4BF0F431BA695339E0EAC940FFC9D93729D07318D6C8EB49
-0CD8C923415A811C9F78B6E5C66601C552500012144B793BD0128E0CDF6360EA70FC83633CEA
-5E962E14AED5633D2ABDD5A1283EF893CFD1BCF6278095C3075A4E03F6738772437BC1F45994
-9F5A14BB014BC909E3551EF70B61430E7EB9C0DA2FB9274F399560B6BC1C754D8A770598062F
-409FB7D1165DDD2356D506E7D1D868847BC0D4F8482C99CD1D4D460B51F30BA5EB7CED25F56A
-001C0AFA884C8A65E241EB72DCEA5DADBEDC6D761B24B5789EE8404529CBADA729159B5B1158
-91AFC097F3A3061606AC8BC1E54649579F65F8B5081CA530F51CB230EE59C1E065EEE4F23B83
-D6CC89F0FD0420F8A76E224491100E9C420008C8E5CDB9F76FC5218AFBD4E8000D47EB6D42DB
-69C07A4609410B28BDE1CE27CA80AD1FDF6D18ABC7E15D8EAA112DB8444495DF49C309FB63A4
-6647F73A5FCDB9DBDDD2F6E5CB19826D7309C1DC3F2ADADF2E9526C6FA3B1520BBD7D1B27E4D
-BF70D32F827FA36FAA710F342373C3A1954F5EADE197609809555808C004C8791AD8B0307C33
-44C93D35B9709B2016E8859FEF495859161F17D49C56DBA5E32300C080B7DB9D40B597D9626B
-2A042BFDD249343DFFAA6D96B585D127D4A6304AED5A8CB39720380DC175C17066FD39CD20C8
-35DD3E87EB3AF936CE1A6F953351A49A8160B22B2BA09000ABBE2FF305FF77326D080EA35CE7
-D924BEEC365A1B0E37F194F3E796810A4039657BD35F21EDA3FF80089C2E3B2A4FF40A7C0685
-8E45FBB7426A1FFE0DBF2855DE7941F616850FC0F3A9273DC2D40DD60ED28C6A6F7C4E684A8F
-B224FA1A52F4F733529A7F3CD3B672A37C9D1F6C00DEAD801078EF12EDCB63CAE750B82ACF26
-DF9650CD502727EBED366E00A1E6EA351119CCC68DBD3E1AE37636E897E964A5EEE0F4E892AF
-88CE18FCAED77CBFECF9216AF5FFBFA35B0181025760CD6920A3732BAFF7ACCB611710F2F389
-CAB8C1A5AF75F82CE7C8DFAF55FB6FEA658D706BBFD8C2A979BF79DE4DBC69DF204769719440
-2274A8D6B61503BE576AC61D64995B9A4E5B2CE5B41C70219E0CD4A42715009B8AB11CC681CF
-501744A54F64C35248A0A5063CF725CED14BA757525E5BA7F959D0CC1918A21AE242293EA7C6
-57D0B1BF860BF3A994EF414D457EAB682AD8A4E0C859B6B0C7FB74E1AE202AC7011F35B1E104
-44734F4C04BFFFC35A36A2F856CA4BAA9B0A885DEB9068F2E635B68411BC397509EF8211128F
-CE9FB991CE287347F06080083E34E61D6266D9FA96DBD1DDB48040FC1B99BEFA11CEDDD479F4
-DBB987236B4C5C41C536F464CFA43B3C2751D2114DCA8C6B75987C326FF5F5859A20CF68FE61
-264D3D19C3EF7A88C66EE455F1B85F570BE50903F0171A46FEFA4DBC8F62EE9383B585DC0D53
-F6E249F469C631099098E3391E8B432E739BA60EDB0B28205F4F6F306A70D16C149BFCA8565D
-1141BC191D12B029D4A97DFD41A10B2A3BF0E7E3E56ACD867CEA04B85C5F7462262DBAF7BB8A
-413DA70DC3D5CAE2C3F3568FA52081447B8F3F7F2C608F9D89156C427136A060323A8779A18D
-CBFA6D21DC08116BB1AC5966900095176900CF3C35323F45CA4DF4BEC4CE8B9E5183847070F3
-FE8EC37EC98365AA6F3353E68E1B371CFCA937C5F3628D85EA35615239402340C3DA703F7A18
-1961AC2AE63CBE7B7A99322B1F73AF2334B569F9597A2426AA02A36CCE40415392F382B26449
-592723D2C7DEF0CDC39773062E13741536AAAD929A35F698A226342E624449AD2534973BC1E9
-B2F8E875B881AC7B061A8FD8893DD399F1AF00E30153B6285C0B591135FD1FEC0425CE1B2C86
-AC0378FA29D7FD7668A4C766CD922FAE0CCA3EF7481DF2AF98C4E204940B03C24CDC6732C548
-6B0FA2C312C041D0A97F51E47AA0DD96CC60F96FFE1F37EE54C835BF34C20EE115378CB582E3
-6ED4FB2DEDC45F06D0AFBA2098C71A914E8A2BCC30701765B9E1F4825715C6328677776A6EEC
-B1674BEBC1EBBE955CC6E0DF92BF67B4A62AFB52A90318ED69347CDA22339AA374C8C7D6DD69
-FE9A3460D7368295FAAE607CF34E7FF41FEFB7650B5CFB4F7B5DED38A73C005B6D21124F9CA9
-012286398270B0A1C18E7CEA3C5116E89EE7751CF42D4E4CBDDEF49D1DA91F0132CE636E4F96
-FA40EB20392338F5D2F1F3B15C44D23B62CDC36F804005A06DB79B70964B48D8B1D859664000
-611D5FADC7CE699BCF030EA53D1ADA7A87D2A6B64ED5A107165F80EC60051D1DC85666E50BAD
-862D961589E379A48D51B7B69336B5CC155C732B14BEF3A01003732D7BF05109D01A2202748E
-C305FD6B7DF421912EB377B99C49FF565C88EBF56C7782BC0D3DB5C634EC0508D9194B09B98C
-E0E43BB9A6EA2DC02BEF2CF3179AC90F7F45FF6AEEA1C6C7C8CCFAEA5A4E8487EB470D9DE00D
-055AA2AA0ABB5B22CE3D96B2D422352747E1BF64C39FF9E1F51D404CE4339BBF79D2EF1B2486
-01B02F863456EEFD02BA5A1C7F97F860D4294D8C85560DD2D29C2167BC2208F9F50259EDEFE5
-8CFDAB9722039D2493AFC7D71E714B886B2DB7E2A6820CD5642FF72FB47D20A8F45E4458880F
-FB30FCCB6ACA8048E5914F6256A439188E004A6DAD74DA2C6499B63A7852684DD3BFF34084F1
-DDE62B6EEE0A0F33C89B90E42E55A285CABDCFE6280929DFB9EAEB9B04AEF296C7A9381054BC
-2E6DA960D32E848C9C40E28670E255121A0C5B17CC9FEDEB5EE0126360C134440F9632903200
-7BCBD32AEFA3E271CE3A8947784CA6B20CB07ABF83CD0B29DF54D69E2C072F939DFB7A7E5B8C
-6C8354758CA5D33AB9D61C77EA562642ED8388F2F84E3733F13787AE7D1089AD08984DB6958E
-65DC5D02572816E696AB3FB9D95A9F64D7AA7A62B3D54D152394D6598431B8BE651435152BDB
-D687DF6330ABB0B451BB3D5198C3740A44435B048B19A9238BFC5805C162B61981D08C4FFCE4
-CAF37D11CBDF6769000093147A26E8F69DB0E18A96150A569097CD1DD8C9848355434B377D5F
-94A01A532002E480AA5E1027B696469F85577FF07B793F4B8F684C044C62CF42D4DFF1AC9E52
-8C05F6F36FDB87F2266D232B6717897E6C57DFF9D31CF3CF5F12A48296FFF2FE7D6F975A2463
-F4FB485171143755AF082FE102DB4D3A267B43953C4F09BB634D54047076056762C099061CAA
-D7B4F5B0BAAA7B1A40396AF29268C784A91F3674DB9EA724C3C1B88B110839825B82CB50B58D
-A5B282D589E0015497E1D887730DB88F6061AD0A0D92645892F0DC6FDB83027C0715EFCE2B97
-B117F11B7A1AFF37CBCADAE82D38B39984AB201CCF1B1AF4D98FCC527CDB75D57454B5103FD1
-1C9305091EFB79DE2B9B7521F3FEC034887D7060CF40C2D522096AAB5D3ADF56920FD7EBF767
-6924AA2F89F97639B42286151E76B04EBFAF1492BEE4BE5C9A75A42079503B7B67E7C81BF5CB
-DBD93E34EB1453EEBD4532BBD47513981577E5C3E051BF62E1753A5A457C78E78AE563A39A60
-069C7294D0C04D46F92CFC9B9620CCCD1E4294CBDADC7B97537333964DFA8CF1E17E01B0533E
-01F41EC9867A111C17556105071E0FA5F247DE49159055AC50CF4B5144E2AD6FF2104784C4D2
-73E948D2A264E4BEB7B47A87D5E3D415D3926734E8C3C05281BB872D1FDE1906E8A397000792
-6CFF3231254CD5DCE65DDEE11ECE87791727042B933C5AC119C848554A4045ADF6ECF563CF33
-A8F910412313BE5B4FC987A67D145686C0F5E11869FC8963B1590FCE32AB66031BD2757E226F
-378EA27FFBFABA11F916538D20E6F692F9EB9CEE02AC13188E9905EE62000B32D6516F32D90F
-F71DBC3C288F3FC76CDCA24AA0AA1D0A7D0668E6986AC61690388413BDCA81CC060BAAF071E6
-7301DE9DB664BC988CF0C4375FD51150F9E254AE93C50A1EEF2310BE1673FC22BFFD2DE463B3
-69B19E23EA76586C54FDA5C50DAC226E3F95D95558A8E832DF2BC4F3AAFE1F3C0C4AFEDF9FD4
-0F0BF05D747C82EAF7EB4A76F01587BE18086C4597BB93A8E3A3AB7DF298F8C16DFDB101881D
-1EEAD4B0EC16688B43ED77D78BF022366C3C3456462265E6A962414EFEBC26D17B67647EF49E
-74308D44304EE64556909765F068DA69E0D3DA44C0F7EBE4019E304039CA80AB806B47B1F49D
-512901895A3822BB0424AFC3157392E402DBE450056C3C6760CF29E257911DF9610DDD9520E4
-A7CF15276A81FF9A79786B22DCB2137B20E1F976E83E98CEC6EAAB2D9427BF8E006FD30757E6
-ACA232D64AB02399FE79A2884F145FC6351B5C186F97AE8BB3AA3DCDD34805EB1E240E393A36
-D1F85C8B396624CD69802DD267D5A9909987669DC66805C8D273B1D69197DA8FE60D005B1132
-3DFD5EBC1B400B9127D71DFBBAC8899E688C92AADD87B2C895813CC20E35B1F7D383E450E384
-9589387488C7ADA1FF323A3CEABEE249298D2A2DC34EE477169EDA7DDEA1FBD7C2D5CDD76667
-06EE7EFA6F42FA80F622F42257E9B88EDEB7BD64BD607FEC542ACE2F119AFE682D97EFD6817B
-545D2D424D6435C101EE6BD670276BA133DD7D16FED98A9DFF0EDB81A6127DF7F070751B4014
-17EEAB5FAB5D4A1A18611655228BD070C0F0DB908BD32F03E1614C95162034560A053DD106FB
-4C0690D0DF259947574C6D82721966397EA456B1E8889518A7525C23BB294E47F5838C1B5E98
-8400A58DB86DC86B35C6A3414BE3E2326892105B2CE25DC7F15B76571D0EC9D646974D53B8EF
-96EAAB63928C4A175A34881BEB0A801C14567EA395C5418E89C5F7BDFD62682CC5598E74608F
-47F3B9A43765FC435C66595772B04048EBEB1C39C18ADF1235426E341DC9FF747FC5554FAF22
-AD4962A609F0CE274C572F98ACE54BEA41930931DC944986485276FD38144A08CDF0DE4FAE02
-9213B596977405C220E2FDD21B7CBBBEA22B2CBF7E7964D58035B242E2C2DFDE43F21B020976
-D1AC86E2C365FAC822A38DE51E6ECFFFF686DD4A8589BE09AC9DC1B70D65DAF7EBF146ED7769
-B3B00AE122F4D322CF1B83FA9AE501FB15221E1B3449ABACC7B62D9CDF77AAB06FA4CF610311
-DA089012AEDDF34D37204DD1E7EE89FF8E50A9AEAF99DCE64DF6EA0CBCFFE8E16BDF5CB5F482
-3C50F6EABA621C6034EE2F2EB0F099950E4425B818F4052EEB014E2929D34D5333055012F3C7
-518BA32BE04B1E8676A14D95EB1AF4F045E78E6CC5F6B5B6AA976527325C29B7234770EE1F9C
-4609DE2CEAD3F39EC7DF027606771406C622377EC9E7C4AFD2B7BC78AE0BDB6F08D3C459D1FD
-8E97C020F39205BF16C907FC5234257189212521E6E57F2AF9AE9EFB11EB36F2A428F3BD3A76
-628FF7E86589BD3CFDA9E774E1655FCE1BA225DBD2506C7AF7078CEF60C209658B0435DD5AAC
-00BDA0D6B11680963D9C64BF1A01E1B884A65BA9CEB5D87B2492D44664879A5E9B046B8DC05A
-96E50AE5A71055D23F79E5AF82FEF4D124D7238DB91E1C8D2EF789811D45411B75785B4F3609
-91C7DC590D15EACCC950634BEC77ADA1B96ACA5E4A09849883104B6B584B2F345BF3A6573044
-4B7E755BB603A170424412A6F39B882BD2A6554CBB33DE4AE5AEBA7AE9B0E61DB5A7ABEB1553
-A4A1B92BB9E6AB4021C266E932E6F871C362E07883D4CA4767FE1081DD626348979533C1DB13
-B42B6DD7ABDA7CB7D84EC347279683E22600ACAB25DF533F564C97BB3EF6E39D935C59A0BD2B
-9CB53C0ACCD759BC4AD7C60BE4536A21D59F47130FB01D174B4BBE901E19FB56AFA5BBF8099E
-50ECE0075C340148F18BD6661EF75F5BBEDDCAB5621144DDA8F5696F8867A4CCD8336DADADEA
-B5C6CBCB6A114F0E44E43945F46C9143C3BB688E4CA8229E8621B2986CBF45646753B0CED554
-E09F8D961CC9BD0EE359A859B0B7313932F6E5013875FEA4B51268BB3D608CA579D999D1CBFD
-6AE8303BF21DC188109092A62B1F5E18ACAE78DF01F41E106EEE0113D253297C7798E58AB6DB
-A4AF1E93ED7968C0A73F5151FA53A5BCB039FCAF488B2E94F20E3D5794256F202FC5EAD8E526
-763C772245B590FC2F98B3C091B9844C332E781283632EF903CDF3D10006969E8BA9F189A1D4
-A3A676C0A459EC476894C3E28891A8D451B055DE02EC8943E3BB0AC46CA0114717C24C6B3917
-5689B22EBD8BA91F97EC1C6452711C57A4B8C883D2FD15D1B416C3D131A902887E567285E018
-D52E3F5CBD6FE3D7127F05FD400419858073E58ED1F47F24A8CFDAE157230C6C801AF3807890
-C9238E6682CE035436E6EB162D8B30E52D6E84E981D8A977CD59B0093E477283FEFF42448642
-ACC058CEC4050282EA8508441E064EBDA93A40EBFBC68EDA94DAFD3675172DFF79A52E352044
-D94D457E84702C33F1C53855D048E0D37BF96E9EE400D584EB9787180E3765A37533BDF3F814
-78BA837FC6D007AE96F23F4712F0D8C610EC1F8027A0D0A993BA0F55D91A92A6A2CAE17E2F58
-638CB84099348B02EE5347FA6F02851EB046095B77B5011E1A629834098593AC3BF9BAC56380
-98C30EA3F0C5526391A8736AFCCD1C3E5864B2CB555E7D979ECBD909DA06F5AB8419C134AB29
-2C296F4EADEA3118D7015DD6A3B9D7350229D8833E34A583AE3B8D084FD9CF948FCA765BA0CD
-47FB37FE46BFF0BB49E7D2A0580968D09D35B8A28FE54848CF0BAAD6AF6C25AB15A1882EB165
-1FC1DBCCF0CAB68A5DD18974C2C5A4AB08C52AAF2C89E2577A621054D8D25A340AD7B0D3167F
-31FC5046819C9728A338AC204C3B0F196CE64AC34435B4491E0E357F26EA356C9049499F2036
-28F5D3BD277A7B337AC5D89ECB7F26E41CC80DEDF47FBFF8DC2E716284F68E6750C0AA4A1933
-B38E114C7A3761030FC8A5B9E1E7FB753E287576AB81A3D6F54691B9945C11B556604BA9853E
-BD89536E56CE770BF86DD1B26D10D797DBB001FD0BE132E5B6DB3D6B9267587F30BC7EBB1B55
-66BC7C5609F9D1125B4AA2ACC850A6C39FB0ABA181AC370E8CD49C56D0338832AA92AED01A18
-B9D228FC350D0B2CE6E1543ECA4150C0ACE22B5AC3938A8A04101C3BD78777DFB6F7A570A205
-1C11EFC3CD6CEC68F70EE9F950AAB03DB2891F942ED550434DA413CD2EFFA9D2043976AA20AF
-8FDE9D2E2FD6F515BB36741CDA5FEAE7BA1E8AC777F68E6CE7DCC16EDFE02F57852E87FD8A19
-1D44D87E69F7B9878CBC734C227C6AF895F083263F4B5029A0013E5909AF9DA77EAC81F17A23
-284CE3BB3C74B26D9D08C5B7379CA8C1E8BE0C84EC0C69D07A645E613249369268FEE614BC90
-AFFBEA486BA07B993119329573575DB8788DA21E5831DA7ED91E18EE3157EC6C72DB051CAF89
-210460766DB6743D9C285FFFA2C2602759F88570367D5A2E3B97DD9B2BB977B3900CF9FE1B49
-0EEDAC8E3640E43331D6AE45355460BE7EE1F38B11925A9DFA3F943074A7347CF0E03F852714
-357115FA41A17531CA9FB49C3AF799EC0FEBD61B81631775290B1F38E205A78B317EF58364BE
-494DD1CAADA4D5F84980D1FAE9ACFF35E2E7D6EEA26F97DE826765DD5E1268A53F5D8AAA8FEB
-528E8A13E0E001BB04C033BC88B1B83E0AC2BA20109014C9FA9353E7F10DD7C1DF8B3FE48A3F
-9AC01006A895B0019E7F62F450960BAD05A70BC457C88F28759CD1DC5904237C1B865AB4F938
-6A2BEA70D5ECBAC9EFFFBA1680413C9B49B2E16485D0BAC750AE473B1B3F6A176DED4EB6A052
-9478F120294E591D34B90CD62579DC53D94ED0A98A5CBA48225E113C0F8BF5FA1113F695B757
-7229F5DEA917A5CB49B85D107CE29CBE83107465A98C6832FB5F08C6B82540123B675E69D442
-7905C93B701C3667DD78A500F405078A7A0B929BF5DD21AFE771ED05E73FBE5699D4D2C9337E
-ADA08B3FC5916EDE6AF94E7FF54A01BA67F72A3A130B61CA2E792D539FE4FB417CBC6E4911DE
-84388B3E58BCA8520EB12A0ADEAB59119CC9795B64FDBC6C759D314BF93E91B8F0ED753F2A7D
-1F50255EA5429A7C2C719E13734DC9F0AE706A14598F05609EFB370FF110B87D418BE62D6B1F
-3278C5072F7C2C609E92DE9CB15654CED6EB336650A9CA2EA1786865508AB6CD71AB57503C18
-7C0A68F532E3F05557C191D1029C4F6095ED7869A4258C893394D8636BD0DECD8F3851B42EF0
-EE6D6D7B23CD07698EFBF2A6159AD01F2FC5C69FA5ADB8F0CECF8CF1EC99AA919A57E6EB4E8A
-840A5761783C7DF31FEFF92651AA3975AF5A2003C307C6703089EC1CD63D8798CD4FB69CBCAE
-A810E5D70EE6F2DC9ED3CAE29039A86173132A0EC628814994191095E1F578582B79274C6050
-191EFA3BAFAF6F331B3109A04B5450EF7D1E8CBBD634374F200980BEDB6A34B19A2057523C86
-706023EED5563570F6F88FD9B8310981D7C73BADC9216AC45A4BFAC6E4800CD101CCF035D2FD
-D71A8D082BFEA07F3F9CF95F2E70FBCEDC89F27BA661EDC0DB4F2468EBA2E4290EF9B5670936
-EDC728E55963D74C444E03AA2344D80C24405EB031F5933CF414705C3329FCF23D93C1F1C6F0
-2AD7B4FC6E5CF38BA92ECE4EA1656327236FA26CAE52031F4F1FA98903DA4C361FC7DBD3B5CF
-37508E1F22C582DE7E3E9F5278D7424A9F824FBA19021AC81EF7F0D3C6D67D53A26E4930BED9
-CB72815F0884BF52D1DDC54E4AF35D9558345D498D63E7A7D6BC327FBF8F47C4997E2DC4C6EA
-4E1B3FB9587FE48978B67EF5C8E8D8CAB7627A982DE551DA90275AB90DEAC04879021D593430
-A30240CF389E10D256FA72A8656E8A8865750283CB245CE063C4EB16059F25AF21AA7F9412B1
-51A3404D965C2CF02CE88FBDD3A708FC7CFA73DBA0C664272F49F86915FDF58477AB8678D42C
-AAB245B637168C4FB33647C6A2CD7F96D54697C375D0422E27D121899462D8998B3624E203EE
-5D0AB08E0AF286D5F1352F0E4D5B58774F0324B5479621488FAF503415333C9B232119542A24
-54A72918BCBAE3628177414CE76975F786C68AB4FBE47B150539D9BD711C510A7BB87523437F
-CB820000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-cleartomark
-
-
-
-%%EndFont
-%%BeginFont: CMTI9
-%!PS-AdobeFont-1.1: CMTI9 1.0
-%%CreationDate: 1991 Aug 18 21:08:07
-
-% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
-
-11 dict begin
-/FontInfo 7 dict dup begin
-/version (1.0) readonly def
-/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
-/FullName (CMTI9) readonly def
-/FamilyName (Computer Modern) readonly def
-/Weight (Medium) readonly def
-/ItalicAngle -14.04 def
-/isFixedPitch false def
-end readonly def
-/FontName /CMTI9 def
-/PaintType 0 def
-/FontType 1 def
-/FontMatrix [0.001 0 0 0.001 0 0] readonly def
-/Encoding 256 array
-0 1 255 {1 index exch /.notdef put} for
-dup 161 /Gamma put
-dup 162 /Delta put
-dup 163 /Theta put
-dup 164 /Lambda put
-dup 165 /Xi put
-dup 166 /Pi put
-dup 167 /Sigma put
-dup 168 /Upsilon put
-dup 169 /Phi put
-dup 170 /Psi put
-dup 173 /Omega put
-dup 174 /ff put
-dup 175 /fi put
-dup 176 /fl put
-dup 177 /ffi put
-dup 178 /ffl put
-dup 179 /dotlessi put
-dup 180 /dotlessj put
-dup 181 /grave put
-dup 182 /acute put
-dup 183 /caron put
-dup 184 /breve put
-dup 185 /macron put
-dup 186 /ring put
-dup 187 /cedilla put
-dup 188 /germandbls put
-dup 189 /ae put
-dup 190 /oe put
-dup 191 /oslash put
-dup 192 /AE put
-dup 193 /OE put
-dup 194 /Oslash put
-dup 195 /suppress put
-dup 196 /dieresis put
-dup 0 /Gamma put
-dup 1 /Delta put
-dup 2 /Theta put
-dup 3 /Lambda put
-dup 4 /Xi put
-dup 5 /Pi put
-dup 6 /Sigma put
-dup 7 /Upsilon put
-dup 8 /Phi put
-dup 9 /Psi put
-dup 10 /Omega put
-dup 11 /ff put
-dup 12 /fi put
-dup 13 /fl put
-dup 14 /ffi put
-dup 15 /ffl put
-dup 16 /dotlessi put
-dup 17 /dotlessj put
-dup 18 /grave put
-dup 19 /acute put
-dup 20 /caron put
-dup 21 /breve put
-dup 22 /macron put
-dup 23 /ring put
-dup 24 /cedilla put
-dup 25 /germandbls put
-dup 26 /ae put
-dup 27 /oe put
-dup 28 /oslash put
-dup 29 /AE put
-dup 30 /OE put
-dup 31 /Oslash put
-dup 32 /suppress put
-dup 33 /exclam put
-dup 34 /quotedblright put
-dup 35 /numbersign put
-dup 36 /sterling put
-dup 37 /percent put
-dup 38 /ampersand put
-dup 39 /quoteright put
-dup 40 /parenleft put
-dup 41 /parenright put
-dup 42 /asterisk put
-dup 43 /plus put
-dup 44 /comma put
-dup 45 /hyphen put
-dup 46 /period put
-dup 47 /slash put
-dup 48 /zero put
-dup 49 /one put
-dup 50 /two put
-dup 51 /three put
-dup 52 /four put
-dup 53 /five put
-dup 54 /six put
-dup 55 /seven put
-dup 56 /eight put
-dup 57 /nine put
-dup 58 /colon put
-dup 59 /semicolon put
-dup 60 /exclamdown put
-dup 61 /equal put
-dup 62 /questiondown put
-dup 63 /question put
-dup 64 /at put
-dup 65 /A put
-dup 66 /B put
-dup 67 /C put
-dup 68 /D put
-dup 69 /E put
-dup 70 /F put
-dup 71 /G put
-dup 72 /H put
-dup 73 /I put
-dup 74 /J put
-dup 75 /K put
-dup 76 /L put
-dup 77 /M put
-dup 78 /N put
-dup 79 /O put
-dup 80 /P put
-dup 81 /Q put
-dup 82 /R put
-dup 83 /S put
-dup 84 /T put
-dup 85 /U put
-dup 86 /V put
-dup 87 /W put
-dup 88 /X put
-dup 89 /Y put
-dup 90 /Z put
-dup 91 /bracketleft put
-dup 92 /quotedblleft put
-dup 93 /bracketright put
-dup 94 /circumflex put
-dup 95 /dotaccent put
-dup 96 /quoteleft put
-dup 97 /a put
-dup 98 /b put
-dup 99 /c put
-dup 100 /d put
-dup 101 /e put
-dup 102 /f put
-dup 103 /g put
-dup 104 /h put
-dup 105 /i put
-dup 106 /j put
-dup 107 /k put
-dup 108 /l put
-dup 109 /m put
-dup 110 /n put
-dup 111 /o put
-dup 112 /p put
-dup 113 /q put
-dup 114 /r put
-dup 115 /s put
-dup 116 /t put
-dup 117 /u put
-dup 118 /v put
-dup 119 /w put
-dup 120 /x put
-dup 121 /y put
-dup 122 /z put
-dup 123 /endash put
-dup 124 /emdash put
-dup 125 /hungarumlaut put
-dup 126 /tilde put
-dup 127 /dieresis put
-dup 128 /suppress put
-dup 160 /space put
-readonly def
-/FontBBox{-35 -250 1148 750}readonly def
-/UniqueID 5000827 def
-currentdict end
-currentfile eexec
-
-9B9C1569015F2C1D2BF560F4C0D52257BACEE583A5C939393E012707B47E0C1FA47D284A1EDC
-9D01A497D772BCA8C543388E6DC0D1E2C4944740470E0914F65EFB0737B0851B2BA713A9A00B
-36D07DA6BCB52920B9B59EFE587734027A3C5E6566AAD332FE6FBCCED1417802822A3B81D618
-7875263D6BBDA04BBCF6E4870FEEAD60F104BB3C5A766610DD11AEA64A6B107B0B04439FA288
-8B8CC39232BB83F7695ABA81F0260CD5248B9E649CD800BCB325468C79FFACB9CD3036F2ED7B
-B65DCC883AFF8D6350D6E8AF1A9C1D2D7FCDDC4C554C7E350F2420586F43E2124611487CDB17
-6D453854B51CB8F9100BF241899BBEF87D57F1135E589813995C4D86CE91DB2B3757948F3651
-86C2A4F32D845F659A8E95028F9C0168851997135C67CE7654F2DE4F22F51311AA78F440E5FC
-2C58F6DF901FD071B2D53B2697BFCEBC43FE4508828B4756EBBE1F3E8512514162B3BB44493A
-65F11AB56DEAF152FD01BC8B642E6021AF01998EDC76B6A1929FBB9FA4FD1A430BFD4464B53F
-75075F66C726FC037FC23CF09A178FFC60908F77D8EA9DEA79C1E21FCB5A6A994E5A8E5CC6D6
-291891DF7695DB58A2DFCF3BC9A17888C29559CE8F37DD35C0AE35DEF3E8DAA6D8D574F457B0
-385EFD1611C048B9CE76393FBFA006FD23D742CFE70160355FE95302B129A824C7794A9DE47E
-463A7FDC42FD72B9DA3FABC3E65D90258AC9C1ECBD3654B103A488715F9E0E649B57146B786A
-A399C6FF39F362816306A1E0AC83A4C83410030D740E15165683E843F69579EC1824C36935ED
-4705962F494C401400303B590614E5DAA8DC53FDDA083FC854E535FCD575B38AA43D3D5C4A38
-D56BAFA926FB9E7FD69BE1AF2B502849500693DADD5F00B3185A8B18A0B64F289F33F5F1ECEB
-14D6945A22D0DEF259076CB9B11D899318852053D98C36CE0A0BE347EDAF1E4E777C303B4A05
-1E3DA7AE5140CE9DFA018A8669B0B551696E194DD54CA062EA1FF406CA877790DFD1DC18AC62
-288CFFF849D4A142CB1C1AC5FAA0AE05F039E3CDEC6EDE558FA3FAEBB0D3EEB981E561FDE889
-9D5D4EFFD8610A21B1EE990D09836D75CA7F8ADF56CE3802BB45FA9A0E9449FB97ED1933F368
-69237D6F9B8BA6E7DDD9B21BC2E53AF56E28C2DC8AE41F3869A71F23667AB4871F212C30E98C
-6F5BB62B784116C484EF32C00CC4E45E770B5AB4A1FB079F7F9798747DD659498E6BB4816654
-684334AE644249AF38B565DFFD2BB9558579E4BBA85E95FF6C45AD1F94FE8CBFA2002A2704E7
-9BE83E3473732731CEEBB44337631F3D30DFC6DE9E414DD85F61BDDBF86AA99746EACC23868B
-BEC979D09D338F70F3C14117902B39F9FBEA39FFCAF21AEF2ACA48493EE2EA88DED0AF6A76E8
-FDF4DB385D153B3694041477B5EA24A49302C41BF01177D4A08F66A5126F2EBF90BADD94D31A
-896E1F769E608B634C8F4B3FDE672A67D2F594B7E7CE371EEC19286D1068865FFD645F6F7D78
-02F050593D6A275CE4D764F48266EBD76175ECEBEDC6D19BD6F852BE3311D4A24259F7C5E035
-1D63EDDECF1C57D4F6A357C6993B72EE17447851A7290266BA4B9B5511509B03B1930DD3BC98
-5B9E46CD383D00E30D8743D72F8B5AB4440DDE8B5C192353CD4C353023A97B151D88D78885B4
-F7D5849B5264CBA905F75D08864A6FD161C60FD70D59326A1B6F142469AD48466BBDFAAA2233
-688606370943BC6AFEC630FD26D05D9B1F84AB8AB0722B75F9D795FDCCD510974153D9949958
-14A67025D2014684FF628E87A61233662E68367B42351624A3A0EBB10D3C59227EED60EABF55
-E465D8EC2AA5279D935043A2981B0C013A53A4EEBE8F9382B77CE54A9FA1190F3AD1D3A93247
-ED89072F531C2C9DF65707B1A3AC38AA673DF83BAEC1597FB1ECD1CB91519181636D6CDA7673
-D58E5AF0906D446F5AAD03AFCFD4DCC0F7B0646963FCAF7997DE9D17291B0D8ECA9E364A7B63
-33E0843C5E26F1B39308BFA9C90048C5C453DD3B812FA5A49424772CA517A08AB9D2D37E2B3D
-E8589C00CACD77E40E3BC3C0B4490771E4CA461ABD886D3ECD386BA460DAF14020E6F0CF370C
-337A71ADB9244F42835A9E9A617441CDABD3F80B9711A6BA7A9332B72E5B5137C69D4E38E338
-EE70EBF273A8A9AF61BA6638ACC8D745A250F15DA7C89AAFB99A567436CFFF74D3927767EAC8
-94BE95D155AD8FB63182D26D86CFD436D1BFE9C6CF718E69E269CC22AE72914B1FC0009B31DF
-DD899251A54D1CBEF74E915C4AE80ADC9843DFC68FB30EA0AA153E85487E2DF6B830EA41D3AB
-11917AD8B77C4C7833E6DFD6B73A0AD722926A59CCB7AA14933BB6D316012ABFFF8F8A41CBD9
-4B9E36D181F2F664BC38692BA5F4AE109FC2D733DEAE26C58BBA3DE216999712FA90D13171E1
-959F1E06479FCF07CFD3DC6297EE0247B36E6D5241FE96FE937852B06FDEAC036322C6B36FBC
-DC8A51472516612F57CFC158CBC7338A44357224E6B62A8DDF35E9E031D1A4BC7C0830ED91E8
-9B25FC2E697E9074935A7AABD4926DE4D9398E154F4A76461E5B7A41A605DF1B2D073913E821
-0EA35CED50FEEDE89F4993A095626046C5FC1D1DF7B534DDA36FEDD49123865269803C0721E8
-9D09569266A612E588636072C6181F36CA1C002CCFAD484C16F6F69DBFCFA679E8C21FCC76AB
-39B6D77D4F38005596751E7964ECE979562ECEC743D66D8805C8AAD448902D87E18D11B82938
-DF77A5B31BB95B7B7C05A9A8164309C2466A54990C0DE752707A9CC95D942549A0D87260856F
-EE5662E15F41099EFFEEB73C41AF209CA26281A924150A7787B916288B23CE7723D1B78827B5
-89E879D221DC86A4CB6FD9BD6511E2A1FC39A28F66150CEAEF6CF631ACDE9CCBF611D04EA4F1
-603D030CC0B4FEE43140C85A42F18CDDA76328276B44740059C9AC18B67E6E1CAC81AF336E16
-BD1FC6A0B5AE96BD4551FB9A206ED360A67AD21956A6636043A06D22256BC77414BBC2B8C5FA
-3461EA83E4F2BD938D1C6CE704BD6B3B35A1CC6E4D0D527CB195894C001DBD03934047FE5020
-EEA555C80D227865AF5D7AFD85DBA790D22AB902BD6C46D6361BC3FC397F0F27BE6575D2596F
-0924D55A4BF40A768AB17E0B8FD4E82D9726008DC777242E56BC59DCB323AD87C341C9D615C8
-AE467E6E76B641A3D4E7F2BA2A29F2A355E67823F72C01799D2F42FD79CB10559DE1A3C6D79A
-8F665AB77EABEFD535661A2E6BA72DD3BD208DC7E5D1D98C0AE1DF3839F50AB3B3F384136644
-DCEBDC2ABF0431C10CDFDC17757FC147D53701B8E585C77DD01B72C36A813F932A0CEBB23D0E
-E6647EA5A838A2548B8F62496D44A082CF451D300D314CA6C70FA36A5FC504A42B9864AEAA5C
-4FB465289BF4655D6A37F0B5866242B8E5347AC12D24EC6A1FACFA36CA12949312B750AC381B
-F875B929C7958BD9B742965EDA28D9AF6A3029D29A7123689FFDBCE3C8E958C8325F7B39EC22
-070D738A78ED780AE26DFB8350E0AC1560F258E6FC5685FD0988904354034A906F32A78FB228
-9F906E5FFE82479156CA2120955DF2F19964F60385B7D799E7EF651FECEBCD371A735069F001
-4BD86F9EC28E0E47D477E2BF55C891219163FB2CBB6C5768285578A1FDE9BB09E369857D3176
-26FF66A5EF98964937F000ABAF00B317DF452C1242F9604CDB22FB1FE0FB49F9A2F0E44B8802
-7CB7E4B0AE1E1AD4B3684ADA9A8881D3914D18C7B5EF15720103D8BEBE4D8AFAB2E63528D474
-73C86763EA08F49BBE887A9DF823C038E6A8B7F81615867E4C7CAF140277890CAB29174E8936
-B26D209B960DA18468489E29A5CAA9D209AC864B9BE5504C4220FA48843AEFF8F39A3718BCB8
-C48A97B650746E6A587390AAD283B224D93439D9905B2783810688FBD269310ECF14ED13E6B3
-452654AF5DB2EFCCFFB3BEEE66011267F7DA5904C11619EFFACAAB4088EF2600521FC2128EB2
-7F81C5DE003EF648CA431AE8B964BDCC8972A5F1B3AAC3C77AAF3C3949C480AEC3BAF390CCB3
-0F716A81BB6D950510A2473BD83A3744BF490922C6260B651745C74A38BC7F6DF57E27E89E6F
-FE161911E44357D4A2250E5156B1C186F07778C6D38C3CA9D81849B2322D96851150A8B9E147
-4928E7B409B212E563916BB2771F335FA34FF78694ABAA90A0DE53A2573A7E4DBE82AC5CEA05
-2EC4BDF8B812544468E4694D7E102EF57EDF9885596B7203DC97FEDB96FEB0EBEF164807370F
-0413A0144CF3DFAB845DDC64B800183FF4F8A16BE96F492CAB0D093BC3372872A45D938BE717
-47D7F93E009D593F9FD2878C99237B1C2099FB3B70872861AAFD509B99BE724D8D52A16453A7
-535453D882E708CA809B37677D494C423DC3A141AC2A91860AF6E8EE6C363B07EB985D81AC9E
-0D35E136A9A402937205BE2E5BF1E8294C645E55575FDB0BBAADF3EE2E9800FDAC4E0A97BB44
-5386C7E3B0245C8F9D30C70D7D585111049E5FA7DD4B5884352E1B3352798FA62BB611AC8B43
-95B438182B7C2CC5D9E9EAB9F7B013A96AC1A72254D1E476BC562FC0BC7A598FC5BB589B3273
-258FA04895B412BFDB2DEB3DE665A1CA61A44AF5BCC689DEFA611050852443BFC5DFE1FB36F3
-0587CBF40EA6D37D980E290C0E2048FFA483F8DCB5AE9906FAE04B159E0F273AE817CBF3CEC3
-5FFA8A5DF7BEE5FD411F154A69AA1519EB0E74CB84276CE4A155677EBA38EA5033877377E790
-007D49FCA973D1BE6A8CD9616A29122C9680A06C1F4B348B4424D9B992ACDFCBA2F9D2B656CF
-B5DCE9D8D248452AFEDA8E548CF1E6451354561F59A105A024D27645898E6E0025B0E44AFB3B
-F47470B5992A3438A52B28DA6F2FEEF2FB92E97DDDD38EE2B51E2AF841F7A95ECD0DE7FA9EB6
-801769351E7E74F664BCF58EAAA920F4092088AA43FFC1456A5306E2D194B34795FE84D5FF79
-A34819DDA9765458AAD738C86B3F2376C9FCAD07F3627B4BF2DB596B70410C30AB23BEBCC87B
-1C077CC63EBF27D034197C1BD167842B3B7DF8D35D4B8890C9103E65414585CB865D7E2F03C5
-6D6772AE4F99F32C27FCF1C9F29D3D1E037031FA27F192EF5E2EBB18AA961C374B73FF816F54
-C54ECD7CE8E2FC7F8D913C0522056D5A05B8297A4D0AE858E0152C9C6F6FAB3BCCDFA0ABA603
-F4C0E2F2617B1DA6C741F5E1B90A0CAD404D82B1BB9FDFC496DE7BD6D780BCF543B574079EFB
-56ED7511344DD511B278282917648DEC31CC8B2FC9E93A28CBE949733F21B660A1EB29C5DF6F
-86D3855BE9A47DBFF032D78C6737C62F6FE5AF899217D7925E73E13640CCF0213B7C4C8DF3A3
-1A471D7EA5EB0BD9C06ADF447C03B0AC9EDE1AE16C5348F2D199536F848B6A717B76711E2FD0
-204EE6FFFE795C48AC4BA99E5617B463E2E5D0B6F662DD82E72C8D8EC4F62A8BB0BECE8580C1
-7115EE312ABBB7B9B2EF00FE63F5E17B38915C8F98286B9BDB1527DB2DAC62BDD3D0A80EA943
-C6FCDCDA45341E7F8DD91544F9582F4C66BB67763C2D3CE5BA507251E0A05E1689224FA9C630
-0327EBFF49FDAF6E92581BF92F8E3949CE258E9B22A619D6CFB11760C7EEF3485C59132C7ED7
-F9AE094F5A96A82BF4390B29796D0B4E1C346547FA7CB909BA996968685AAB783B458C7CE05F
-BD91D1CDF830DE43EB2FCFED53F2438F8D7CE23B5CA59EA4411E8869FE65F09F06F91B3CED7A
-AAE2845AD7267C1421E9744B60ADFE02950242DCD39803F3CE4F3435069C8C15AB34CEBE731E
-84112488C436EC5619232E21538764DD41CBA8BF1D6F22936EFE7F9A2045B72E263AAA75ACE8
-224F17BE8AA7BDBC143068EB1B5CDF5C92A22B0E21B9EF4DDA5C3F320A1E385E17E82A9C19F1
-5C69EA50C70CDE0A71C4A55CE5210ECB75D155E0EC3CD1B8501588ACE12DCAA28F42EA5AFE76
-6C03CE329769A2F4609BC67F73734AA27CBC3BBE6EA66D9E73E7C8F886551364D19791F42444
-CDC6949FA1EA5010A4E0B27B4727B6D9434A09C2007534354F30FAE4A19A414E8A040D2AE36F
-A86894384097A127BB5DEDE8B41472119B138F92D3632E6E48EC483B288BC5AD6B499FF6780C
-D16F981A89195142E709955D3426B333AD994ACFD708848F4DE3FA2EAB945424CD939AB8D44D
-5ABB52DCB0DA48B41915443F34DF900474F271938E13DA8F4B42AE3EF5E2F7475FDA139D534B
-57C887076F1A31554A87FCE63B87E53ACFF6AC9FE8A5DFC97EF5B95F7FECC5E4E4F9353A5459
-F91F98210F3F908D4F89396EEF8181A01385F102FAD7BF1D74201C998F070B9AC4CE0067E076
-105B702DC3B8135FD90F174CBF04ECAAF1CAF7FABED59CBDB3789F7CC8772417B949918613EC
-4ACCAA991C62B1B88C3DD3F856291155DE4D4D56011D18AB7E3BDD03AC87ABBBED125C4CB92F
-316ACCADFB721FEC10C67F5ED94AA60CDCC3B43AAA8D2D1EA47EBBCF2E3D81CE5AB2DD2921CB
-57D2AAB32BA0A9CA2163A78D11F09AC520D7621FCB783EDE71C7E5B81B2E7AB2EE8C1D297BB6
-C1C6EC5086D9308F09A58F2658D67BB3954D243CCA14258B634C96B271BECD0814F43052C83E
-C369CC7725EDE6244AD7A15D62B3BE0FBB5DACFB868F5337943069CA65F49739A0D12628A2CC
-54D8363069009A06A38B42AA33E457E004F68E63815947A0B45C5E059F13A5810E77AD5CFBA7
-138A2E07F0DEE43A92D2169B432FC9E87ABADF87D46764C05DD4F2320CAB0738181CD1E6E69E
-54E222FB329139315A06FC9DC29609D8588CC91991E1177829F6CAF770A850D0FE57738240B7
-92E067BCE342A54E5435F5B0E75FD792FBC23CBC5DC5EDC7A94DAEF36E32A9E302BBEC085D44
-8699A804F20883EA593F7F731DCFF92CCA20184D671E65617C423F1A17F99D11990A00F9F06D
-D5C2DE82ED90B785D88D7BB2DAC628759464E32A89AC4F4F7B7DC92ACECFE21F09C0095EB6B4
-B4C406ED8E5F268C5962FCCF6F877DCB28CBF0F75235E0866EA325E5D0003A0B75BA3B0307E8
-153CDF86ED05090976C72DCCEC648EF4B27128EF0032D39E9827617D2BF22935CABE6F23409E
-D0FAB591E98D0C97630578EF352EF217D3CCA6AC5717887E1B5C1F2BB99617C5500B18173988
-071681CA79ED22E253D6003FFF1DAAE119BF7047A2FD083C09C1DB4E2A4FF067464C21E66945
-E1CBEE86834161CE26703A90109E2627EBD3BBC11C1E3F7F4A2FA63E9EC83993F6E0264A37B6
-62F8719941DA50EBB2D0F7DF3ED52554FD70934BBA3089D3590F468E3F98100A469D86925656
-FF0F235CDA4BFA93DC33D7A69271708E89F312F33DF172138B42D8A52081FBF9F22A60C0460C
-B8ED394B7B3FB35C9394C60E0209B7A0A937644E720B6ACBFB61A6238386F96BE67D289EDC67
-933F27018058ED00711A343F6BB1ABB3641EF3C734298B825DDC5F76D1AD0379E0CBB531DB66
-0032C4AE493ADD910D7287E91F8C3B6D18D2DFB74B981631D65BE35D6757465B87B5F6AA1067
-683EAA7C32986202201879BEC728A4FAF536E098E669A054ACD1FB6404C44F24A002E552FF2F
-BC95CA9417139EDD44806FDBF884D8FDE8FB3DC21928C1769C304F7DDDB41764142DBDD5985E
-AAA0194B13B5F3D0675BB4E00D4D18FAAFB525A5252C1196847EFBCFC234E2669F4BC862B5DB
-F51E84FB8209B0EC09F243E69299E72A5EE58B8F6D9521C290E8B7231319DE62DE567DEFE7CD
-CA7284B61A69CFAF548F7A7EAA6FA515274B1A981EFEFC6E3CA739DEDA5EBC2490AA9453229C
-90EA76F4BBD264FB992C84277F08F511680498162A4827873F19311F3CADE57D2EBBBB2E579C
-F9B385BAADC3AC9438AB79B0989EC5F156DBF8CFC4B811E50A8A85000310C4F4BA4784A94838
-67D1CD37D8CDC2AF2EA1C916A80F18485A3FDDA6E34D8DD8F29FD2E85637BB42738266747AA6
-A3712E1059639A77E5AB4DDF111E4179DCBA7BED727EA17FE5CC62F6860A0CD747D45433FBBE
-88F63B31CD6AF91E52498BEBC1CD323CC6C2777CA9C559BDF8BD8E784F701B8CA0311F19299F
-5D6E45916C828BFAC8411BCF55EC3D41F8E5D6BCD3A88A523F69EC154F7B3DBBBF71E7C6DC20
-E7F7064C053DF414DE37CA3CC193287CE0A647CD592AFF40915A7F89375008DC5C0C6B405172
-A1DBA3AFBFD18E0B3D4C56D82562E130DBAFD963138552B9E7E208FC544935A50862D829CD83
-CBAB9C4F1848EBD24CF08978852F7EA390837A9F44D0520B47E6F24176A24A8A40E28E14DD0D
-78413A806E72FA21364A3B7C22CC523B3782BA9E2B0159193499D98C787D3B6D7EF759E03C8E
-085E08408F8466FA18D18A8CFC81448AFCAF14FC29C3AF607B455EBB5BC5E6E6B0FF2CED1C7D
-020268647427823068679C27DE10178EB04D884D8E3704974DF8787D1D68DDE61BCA5C9120D5
-DDA138205C68F6E375A064640BE3D270A7BC7CACDD0D62504F5B71206FCC476019163004D4F4
-561F7E3476624E4FB58F80B0A914C398EACAF77EEA35260D3BDB99BF78774ED836C2F431180D
-7CAB49F6D2F38434CA1482D7C8DA87B7E2E7FCD278C1B920CB10FA7B7107B5D05931F7FA2E02
-D996A7DB09AD04BCD962259B69C14BD3B2196B891DAC6BDF0FCF606AABBA9D3004FBCBCBA897
-80176E520BE884F76CDB46DC434B89EAB35E61AD10E22F61FE472C13032C60D6F20CE7C93FAF
-3FD9F43F7D8018BCFC8FE14AF4D7C6AE41D5F9EBD8D6F4988562655026F2A6719B66C42A8513
-963EB6DA25E4F406D3D0362ACB541EE1B4CE24BCDE6FAEEFD4777B7C87E667D6187DF04AD629
-7A8BDEB71FB5CA107536B7F06447A5EEF869FF8BF8A874A8C9BE654834B52FED2BFA7CA83312
-F11F0E39A8155F143E13DAF2E7A3715C468AC68275F083BF3B7D482292ACCA372AE66896A9EF
-39B057D91D048458A584BA7F313A355FF0245AA4E1915E6818B00346AF6E218CEFEF3808889C
-696B22CD9708B88A663B07DAADC5E7F4C41C0E47B8477012136AFB7BB02E95AC8025D0B4E222
-73CE7AB65A9444697A99094E1D3B603DE74AA6805FB95FF8811F358531AE04A8A9B6D06B5AA7
-142637B5ACCF01629A1A92B398ACE6DB8E700629A138C9C276D42F8AD5FFE1DBC92FA3E39097
-2F049B652286F68B1C685B261E4EA4AC88697D0044634DD794C7625E3C1749EC918E8EF94A0B
-3DE8D6324C23DCE48A00BC3E9EE96BAF279F0F7D33CE15F5919DC8B3DB58101ED626ACF1651D
-6C80A515DF62E43FF50BCAD9CCF22ABF57E032909B7F64868B9E07F078D7C03A64B68D4025FD
-AB22A949AA44A3D405A3CEE2FC4DC06E79604AFDDE63D3DD8C1930BD51D64BB49C664519217F
-CF608766B53527EC09F8A752CE385169C80ABC5F49C6C4F4F9DE6905CECAA9496C90640BB79E
-53F38E5A7C28EF174BA72A7D6A0581DF4A6F4F698E30281599274321DB1A2F21F6B9E6FDFC0F
-1F2DB3C95B11D2DFA609E1FCAE8EC98EA3F6BCB99A6E0C5534F61914B7DB50E5128E76212897
-89C0C4F6CD9D492CD45694ECA87858D2CBFF5676200976C2AD5C36471C5CF0ED7879A1DED23D
-B5124B37670D72C8F6F56067D2E1BB6E28CFBC036679CD19B3925C7830DDF3E3C40419DC291F
-AC4507DF598B809204D013864C11A6D2E7FDFFE135712BCEE5622DF673CFF6685808183A5933
-4CE2D507FEC9774959E8E0B080FF900FDA5DF199A1D55C51C32D399ECBDEE8B657944E338639
-BF6FE8955D1F520142C98ADBFD5749B83ECEDEE0AF600D47602C57ED9DE7CCE55497F902C371
-44722DDC8DDC29B3134D38569705391C6A301C92C8908E8825DA0850C7EEBC3F788D677B6E7A
-D7A4936C0943B460AAC793ACD216017A54D57B1AD027D028F8B50B8909133A7FC9304FD9F6A3
-C4E6C94CD8B15BB1639E427D87412EFA8CAF5F9731F53A3B85BFB6D3F5A05CDBB7EDF277AC13
-A5BB274611E82A11D31A63B7925A0B4F55AB15C2BDB16F72F97B64AD59949D51D975EE6EB4B2
-EF2CFA76AE9A4F7EDE0923F2FEE5C6917E28C2BCBCEDA9575FEE1FCAB7B7D759DE0217B9A2C3
-FC5A7980E11F25AC4C332CB0AAC137EFFD4D0351EB7754450B80035348D9A6CAC40561FE61E7
-B1EE7B7B63DF27539A05224EF56F46AFFCE8EF643224B0A5220B7D5EFDFE7751BC2E96988016
-ADD8676D29730ACB42F2DAA802FD6B29BE91283F20CD5547EA6DC55D28F2CB40D79FD9EBCC45
-6ED5F38DAD0E8E2898B0D08EA064E98524CD0CA1522668E219FC813EF1F58B81488EBA06A1C2
-FBEDC646FCEC4FF277AEC71284470780E23B847694B8E89D2CCBB6AA0D000C645A8C330B2D77
-B944F003C6D82419E50F023FCDEF6A7C548CDE5F567D78073AF1873F10B30970C4582D9BBC90
-0994FAEFEF02D9D36D49F29F58F5B5481DBB9C23EF500203CE4F051242B4F63D38B22660ACC9
-2DB099269D4DF19A14D3B1A4CD4D9536334AE9AEFB7DF961C426211A90A13D343F26EEE3A36E
-9794B558CB0CEAE70E02B61B585287ED911FDE191BD8E03095C59E617DDBD15380CB2631B0F0
-DC82D6AA0EA17FA1195B6B04C1580508485CF63173CFEDF8E32E561F80790EDCB6FA6D7714A5
-B54652B4FDDDD14064319E852F731A71426FB0518238EE8C31A829E097D62ED92A807F4E973D
-48B3997194E391B6F22046A4A9B951DCC3AF13E1296F6B1212EDA41D293AB8B0D6D86E5CF1BF
-B9D783776F9EAF45CDAF4CB7D68C7738F6D2A10717D1421ED5E7483330CE03010643D10ABCA8
-900E49B8C4BD5AC086ABAE3BE9D423C719BCE515D68A1A41F762D33A634B09E7A2D20D130110
-95B67972B1571D7CEEC27198700DBBB7685B1DB353945A50104BF5B8A83AD4F1E625B7AE7076
-94E930EC8B9FF0394832B6FFEE1E64038312E74EC3FA1C1E65A19EE810DC0248B0F56A5DCA7F
-A3AFF085D756D493376D69F5396270765EA544617370C838F083D7853056207DC0E585F2BA18
-80C2B8C76847C7A08ABC417BBF5A0983397D88AF4DFB4AB2C7336544BAE4CE46CB4E9095D8ED
-4291696D2490CB3536F8B4E9DA10C96117F1C133B11712FE11A14C64CB315086C76FEDD37642
-5A57BB84DE760B1FF19FEDC86A27DCDFE6CFD8394F8FA7B1F892A968B87FC996FFAAAC0F040E
-AA431634E5D84D6D6676E2E1C021F8E0817137B63657696CA0D6DD875987D0F73CD9407C7F95
-BAD7473CDE136C2081F58A703D66D7EA685782A1500D1005381562EFDF553EC3195C914C9857
-73F91979B1055A13961C6709A509E49E0EEE8008C64F0152431F62B456C1E22B30B033EED0D9
-069B108D0C8B4630034CD2E79F6EAA81A0C136A826826012355C6AC712089997D46304AB1A8D
-9DEFE805E7EB5770F882AFD0D19B366F3684902B164FD4EBE4E6025BE0FD27673EB514B8704F
-904C16CBA495D7A1B490131218AC89EEC2D7E6A2687F7B10EEA2F1B3CE47353D29775DB69078
-195F5DCB642BE7DB099E4DE99BC64521DCA9D76C3AD98D28391CFCA1CC30479470722BD9B09F
-D14BD73DDC0A42388044F0EC62A417999058118B9303852E0DE516B6FB802C4C3791E4AF5C58
-ADE45696C589BC1AB119EE6C57F5DACE1B89254C0F7D46B5CEB3A4481C35EC1724EF92BAAA86
-17B70A9F63E229BF7B594208855CB538D1683C554FAC813C934438B7247A647AEAFDD4A69BA9
-2C5736D16D004FC47B4A545D42D8D4378B9CBFB9F68E1579AF7CFD6A6ADD893170BA2EA6B30C
-CA0238F9054EC1EE33BFAFCF0E477E1C59F41CA3E3B35C0BA3B589F6F35BB49820B443BBD171
-547479648D7EF8DD901AC6AF733FC29F93217B32FB2D1C75EB581FA815918730814772DF81A0
-681704F63E68D9EBCC17C29333D64F77DB910B3BE3DA14E0EECFC5ED4D37517CFD916C3205BD
-9F1A063DDA51BABC8B0088AF09BEDF9ABA780EE4A29DCDBB873C44E4591C086FAB826602F628
-09D30D2024070E4E71369C5E4CBF8832B08BB0BDD4E431681B38919327C02ED74FE67B81859E
-C6EFF7F81CDBDB6A2B6F0E4F4559D3018A0297ECC1843FA48942C8AE0119BF3E49A258B9C04E
-1358DE6724540FA503E46D3B701D5D45BD00E2332A823B6B7B6E2070DB9D124F3D8BFD5DA4E2
-44A6EB4CFC11B4F9A08298D09A63B9421533531CD8538DEA2E35F3F711BC04F047CFEFFF42E4
-1405754BF2F28EA21433A30C5538FB257735E52C3C57B812306783CC471FDAAF19C0EB8EC7A1
-CDB10C73C2C5B57A582262F0B5558347C513105B802852C1AEEE1DB066A693CAF032D9CC4CA8
-E43543C70525784288D22C32A8D867767509FF85B73B2ADE7A2B35C23B8E381A1E116AA0B47D
-F6A430A749BA07B02557B8B84DBF0F10AC52B40611EC2A5B258D8571B0002E07493302AF8F53
-9149A60DFA700B38A1498B04DD678B789E5BB0E0FE954044D938E8D52D37F4ADD2AD6652D4F4
-2ACC4D7CC0EF51DC89A3C9B7635F0C447F8383DDF380A97EBBF7E60671006317082EB0C23E70
-20719A09BEAAC4FA2F0AFD0ADD688D365BB4F4289581F1C5237565413C9824D8C082F3691581
-4370581C422B13D13F0117610E86D6B6C4835C764E8E71BDA6E8633121D3FF6060ADFC83FE83
-81D2BFEF2AF5E37B737F8271EAEA26292BA818BBA79F997DE8909DF51C3CB9009D7A5F7D285D
-86EE01A01555D46F6DC43A679A801700E8154512504D1F0454BA12E82A4BF4B1CED0D15B7521
-A98C2AF30AC8DB5E3E0361ED5A600D36AC976D3E159543D4452EE16CA6FFBD8ED26D32FCBC10
-0F6A414CC733C2E5C4206745BD2212959D583CCD01DCD247B68AAD03539755530A1A34A7D45E
-E494B5AED84208A04CF5937ACBE24DB3A4CB174F9B67B4E542CD447C1682F8862F20FD898757
-233FDC2F5E800952AFE2DF1232E5A9C43E63FBB494204E252FE75882A018E4FF9AB7F9F10527
-D3A40DE007593EAC4CE763222F7F16ACBC028212B862A320E102EC1DABC73C1296A370BB27E8
-F35641A95332A9420064806CF60E37EFB3289C504479F6FCD9EFB173496924ADB694038B6A35
-198191A1BB204B80D920B0D553605D8D86CEB8A171241ACB02E6553C6A6DBD59527670B40630
-6D10E7BC5057BF6F2F8D581015B5A51C8D8D0CEDE144424302962325B5C584B429C6F689234B
-768C66C708EAF88BB9D6524B0E2ADBED1FE894D95C030AEA716791822EDA9832180E6C085FE5
-AA04F8881E192131031A7D1732EC8417517199E1A3CB00E1651C2549B4C0C377C6C8F41AC034
-AB4C231AA4FECD90D13B64F918ED7058EDE31AAB1801FF20E2BB190B75BC47B8AF46A39C41E6
-F8314052C735283D882169FD0D6EB0B0EF1BDC70385CEC33F9B5F2304058D5F0E387E5C02961
-BED5A2CBA1AC834187B0C37A3DCB3E7A13CAC65E203F0797521E725B2D5963F8E81A34E21E0B
-7190281E8DF72A1EFC1B869072039291707F4309F9F5AD1BF2B16340256DDEC0B93499B59D14
-E33F288E30A647E0A49A74628301CDC2CE6EAA7C53312E70A1DDCBA45460F37AF60F249DE995
-361EF03EEBE5D4303BA034AC412F6458E0A87F3FBE446C8CD60895FE7B4E8C93C853B822D898
-318B57E75A5B75C2F02C51FD402D27B2630A45879A14FE9CD53697F2D83F089C36BDFA24B566
-2EED9AFB36A6763CB45409BBA8B807D122A3C0EFD4B4EA4ED7835CAC5B089605CE93847BF40E
-D1BC9D69FB9896AB694F86D8171CF1594B727073991AD72D1AFD9B9622CE23AABF45D2C3BDA4
-89D20C7DFF1EDC1784FDE32A8E0F06750813A0214CCA1361EA4C76CC14634A73995170B316E6
-30CD6E931421805E0380F643CAC33030D775FA7F5A9B7AA303B390DB2C70A1CBB3E682FFAABA
-011B7C7382A80C7B51807FBF3C31868A34625162E808CEBFEEDF2ABE72C15F56DB2E94246CE8
-DFAA41E4286BEA50996A61FFDCA1A08E86ADA3CA1E144F451ADC6ACF5CBDB4D8E47554B5BE2A
-4621D6A24CCE70259B95F5232733641638CE86D291876680ED1A62E07766AF3919EAA9332F4F
-2011582E638A8C69C1302FB06E9260FA718499F9090A5A193374708BB5B1ABC28A203BC0463D
-E8DCEE44B3AD28AC936DAAD48F2BD4CF8062069494407E380369FED53B32F97E9B097EFC082B
-DFAEFDD97BF87F841BFD2743AE658271CA7AF28135FE49EBDAAE6E1E716B8790DC974977EFE6
-F862CC3A14D3F90C68D5A997B0B345165177F9E7C22702DD96B9F850D7F17FF43425728CB4B2
-966767E53F2F890AF9ACF7292FB45484E7067A22CB5D41B3B0823FED70681AE72424EEC707FC
-948ED49EAACD5E26373D7A2FD82601C085068DE5692AC691A86600FE4C22848CF324F4FD1FFF
-2824973D2E2F863525C851CF66058CCB47D21BF3534CF1611707D5F7AF507EB9BB86242FCFA0
-566AF291D12164E9108EE0013FC97FE4D7AB61F26BE01D19EA3BC28A6DF0D68035C5B4762528
-B832A6A9C01F71B0864A53ED0D8E5A17373A79626BD762106950D0DF5F38424E1C98DE390314
-3FE94F9C0CE128E4848D6EB22A6EFA296F790EF50755BA5DC4E39309E25061895BD8A187DC13
-574A2DAE3C67C7DAC6B9053264C50267145FE126BE77337F07DA139B457608F6D16FE7795370
-E668488E3EFEC83932BA29D44E39A45F986C59590E5CE61B2BF95656D1F4E28F9CD205F651DF
-B00002EC4453C207DCF80C239ED99B8868E7A1AAAAB602019A68AAF1E9AC4AF9F3DBDA349DEE
-FA4AEC39EE940EB47F7D8231B9F9F878B6859B62A4C0AD635D9A9AEC317FA3D511F1BDB7870C
-7BAF7D7AF72D18D68B8F93206D46CD4C9C1C63DEC987A015B0C71F5CCB752A052371C4ECCCC6
-CA7263F4517442D117A35A90C57263BADE1AFDDA4207534AE0D8B1B09EEC8878EB6F15B4D5DD
-3A62E23AE42FB211E44D0036A212DAE643CD6E829ABC4F260D45CAEF62F9F187CE605D865D4B
-9115F95316EBB2555B0C2F2AD566F52380574E381E3AA33DECBC4496405006429AFF7775E569
-83A0096A2E703F8A96521DAD6817C1490C3F0BCDD14EA5C1582CBE7382BC01DB94A5DCC9D84C
-AF5984286A5BF4C04E8D9006C06174ABD8B657795A4975D17DFBB65DBCCD16F6693E9662C364
-498C6D1E386A401575893F7504B5D8582929821D9F1A3B3136D0A7D93BDBE3135F148A3612E9
-8AF0C3A9E7E67BF7CC1D9B691ED9DD6BC481DE792E539FB90C446A230F9C611F774CEB1DDA09
-35ECE28AE99DD350DA9924CD39A155EB72139DADAB1487F18AAC41EC676FA1E0B74B7B05955F
-BF527948A5335694C94E40B6027D065CB219A2597FFA0C0389969DCD0503904169F2E8BFC0B6
-BA8A74798A2202447AFE8609F110AA62E575A29A0045A6ABE25135BA7C2E30ADB4E911D01D49
-B4356AEA9375D63616D8A2569A9F3FBB638FB5A463B946DBA62DB314215F23A9452310C75406
-88F27D778DFCB2EA0D4EB4C9C4F16D181628C024B09EFC588FB80ECFB0A1B3CEDB15ABB4B5FF
-136287F84835E15C1C7E0BE7F52DEFCCF551A5C735EDAC2B90EB760608BA3F0F3AD8B7A9B49E
-4C904176856A6748D635314BC8D8E0D1B5F53B85685518EA0A0C0BF2D47BA9AED4BF77557A5A
-18AFF751BCCBEAFD135A68B3814D5E0A3A600191972BCF929C312CAFA0FF533FA2C27C015056
-014C3BA538151B37CE02A2737ED72A9EE174EFAD30A928760161B22AA52B23D9C144BDE9C9C1
-7ABE493E3BAA310544F73A3564F52FFBEAACE279DA583F5952301ECB026B775F82FE6D3E4878
-BC6D8C7CBDFA826BCD0AD206E5AA67E151AB76940EC97F606B7CB79EB8B508CE1076E9071AE7
-76B78332CE2A638C0846B9FAB68495921FDF452E43727C2D8881DC6828CB956AB190D41896E7
-3A8E829600C3F077242B710B6D55622C1A631C09E5E9EC67DE45D435449DF7C9B689EB122567
-6E3BCEF0B8DC364B2D30BF61F55936EDAAFD8DB4215A0193C4513D72291ADF499B815CD166BF
-95BDDD9A5DEC9D117D2A880E5886D50D897E7181A522D4DA0F172F5F818F3F87ACFFACF83ECE
-D87294F8E198949AFDB8A6DC0F2C19292ECD3CF1F6D2BFB872C725CAC1D0404866461206A182
-7C178A390A746267485481EC67E552BA3EEC5029DCC9402CDF792DA0D27C1041B5F48267A7FF
-846AEFE114A0B76804EF3A542CBBC9A999350458EC8B7909C59D690810EB7FF783592044585C
-9463ED6466C4A23F355FB8C7D755FF649260808FDFE9CE5FBF1CC1EA35A13906A67BE6154F7B
-6C5C5815DC05C091F0B421E4ED2E1FEAC4443FDC86F42EB9415B48D934278FE4125B3B10F7E1
-12B86A7B851416FB55799CEC4DAAFBF2AA88AA2BFB0439E5349564F7C96FF10BC14C3C9C24B7
-8C3DE8C097938659A429A121574E26F71AA477C5A8716CD7048F0F5058F041537FF0D71960EF
-7198F38FE43C1CCAD626C0E1F0E15567BCB25EDE013C6FF7BF35DE743B23868ADE038B9758DA
-E556FD1D044EB4FC38E21848F29C750E1C48626D2E423737DF630A0C7C7578B7FDFB726E33C1
-41BEA2DE60CEB9B6312743822E03A616E3E34138AFF5704A33AB450BDD09851E6E79DDA45395
-5A38C098CDB1BF1952389DFF0FB3F52235EBC1C0E6BB51266277EC7D5EE15207947E721CD6E7
-3990DC8A660AD8C129FA51C2C58EE32CA8CC94264A355510407E1F247D94E3DDAEA2943AA912
-C4B7928E2F666E10E7261EF9A498FF46F7F06C320D61A0AB7AD2EF3D709BFEC87A7051C38B49
-A21B20FE516E519B21A2B4955868E2E04F5350672E907DF3BA3685886F1D5FFBA96125E3FE13
-3CAD421B656DC3607B275C039804FED11794B8420047FF618ED49AB4ECFAE4EBC5C3194C66C3
-778889C1A223607A0BD47ACA10580192C06049B05E5D9C5A53CCA6ED6BF169D42487CCBF9694
-7FAA87E7B98909C6FB554134E3F584BB3950ADCFEE17461F30E37DD24B134E11FC895F787F2A
-D32B2D39E58020C26A4F06B1ECDC0D3FAA00E8E341FAE5626F6442216561BF08DF0C63154F6A
-B8B54FA5BFB12E687A28B9EB2863DC843078FFD3960A02FA3909F360BAADC54F57D54B248C8A
-337071133C1BA3AB1535FCD257BCC12AEC4291652BF588DA0B43D761C259E3E4ECDC3260F395
-A4B9A2702C0A220BCA3653C677CA6B37AFFBE9B53B743ED6272829D9385CDF2F514F2394142B
-C7CCE786B74557FABD4BC47132A613E6592E6CBD3E393A432E0CB14A50252ABF6E1E406FBB26
-BD33E00E78D4FA51BC87342131CB03527EF8F53B2145F4B15B32BC5DBB77CD545CF97EC363EF
-95C0875A82C42BF240DBDB139B35E8BE2098E5333C7B470558FC9E516474D79EDFFA3D07AFC3
-82EAB87C9AF3D0B7C33751C8F0DCDA9851FA7AA1AF0C60EE1433BAFA0FEBC9454E747DF4140E
-8CD510EA9085C5EC52FB119D07BFBD7D8A163B2B879AE1119235431A1645ADBCC8C15CD91EB0
-6CDACA7CFA2FC6B54F7F8DBDE3DDE4922DDB82AE0FBD8E1DE459907CAC4A1BDA982130CFE730
-F116367C036D1A5F88EBE905CCA29367F8E6096D665E666AF68108B66D970170B917E5A0C13F
-16282883E1903B88047950401748627CE87D439E0CCE6C7D48FFCEC6527720CCFE2FFE572308
-B8A1C0A76B9B4FBD5C27161A68C2F5F7A91A51EB5B54580B2D5AE60F2D2211A354C6B6B1EAE0
-41E50C2D5E1F0C8CEB5D065085A204BE6A26B962D5500C2378D81F867A73C032E31238E422B5
-C87D8E57057F60371BC2E0D1C703381A2A20314195A74616E8CD4EC28C37D6C63DC28F0279F2
-694237F52AA2E06253183E768D97A6B29A7383706DDF74EF98A498FF742B6BCDFEB9E6D98BE8
-A225D8508A532D502CBB6B8CB94D4C450BCF071CD6755C3A7C09E07A965FEB88C4F7A677957F
-CF597B848D0863FDF011B41ABDBED4C587F6614C5137EB27D43DC3BA5310FD34C284085B3B8A
-DDD392BF00401DCA951B9666711F301BC9CF53DF0CA4A24ED81CF68C3354436D146AE1D63F70
-85D17F21FEB91F2DFFA41BF0A7916A713464F849A2FF2F3399EC68F4705FB34BA8078DE9BBD6
-97FA568FD22121CFBEEA57B6F6F7B5BADC746266900C642251BD48468ABFD992C04177863471
-40F4E3194499DFCDABDA5F9783EDF6402B24B78ED4552118EB34054FBB30E05E33E7B948225E
-11720A8F828D60FE01E8CE444525C8186E7BB6940CA40F5AF6A075063430D951A70312B9AB0D
-4DFE9C821C35CF8A287D4E89CF9731D21E1D1FE6BBAE19C8E8DE2121733BCA7CF5696CAD5EDD
-C0319E99CD1051F9449B9EA790F91CAD501A217F56E40EE4CA5772C76A2221003CF79B465880
-B44750FC79622AA0B60CBA6648AA834416CFDE8CACB7EC3A8ACAE6DAB4B076691982E8DA2BAC
-1A245D3B5F11F1D6202238933BEC6169A4A9511D7AB5BCD8CF4C2B78EAB0F8C1000C170193BF
-C6BE59A1F910DD1D25C5FF9D60148A6692741190475DDBC0DFDBB132DE1FBFF498202B239A96
-D8843AE11074622A40DA5C84B26AEB55550E69AB20D8C4243F476F2C075BE63387B6E3C4C7B2
-3AFC7489FEBA4CDD28D9FC3CDC7C61423A835A21FC5D506DA317D8BC9B3789EE9CDB0D74CC99
-DCB2C00837EB0F4D15D00B7DD83033F3B514BE0409117A45811B63A8081EEA58B4E4F2DE021F
-949E0356525B8132B120F076F55E6F52AADAF7D80EB5CCF38B51DC37800650267DCF6F94BD1A
-3115703D9DFBC03EEB1BF4C034EA4E270E6136129E58C4708045F1BB543804F4D4181AEA55DC
-699A06707C9059426396FD7CDF4F2FB12A43F11671B134676E2F1ACCCF5F3350EE7BA3F022F6
-D80E47A4D6BABB77DAA0A82CFB54DC85FA0A7D006F14DACA397A0D912516B5A2FB5FF13F7A8D
-9F3BA394BC715EBEBD946A5AA9B5192A4E61560824D2765C0F5F8676A295FB7CAFBEB559386A
-F7532DA394A98BD0F1DC486E91E835667A984A6CF65A8A0589C9DE0F929C5D3F6A497AFA210F
-40DECCC6123707CAFF3F40B688DCEFA319F33D9D35F755C15BCE6144AD9D9381A6CBAA3AC6A9
-789152ECA73CE6274EF50F169E2D23562A4F486E90CA602BEFA7611F02CF6B5A6CDDB32693A0
-B909FB612A5734966B3CDCE9A2B0515DAB55519A85EFB95825733961ADDF651F4DB9175672A1
-27CC7CB3C638F9A3FCC47A6B91A7EBB6115DA6C88B1ED600108114AD22E3C8C2A54C34BB92D4
-22425C8267F9C410204128AD971EC7A9746936A4C6458B012EF750E13B21BE2C07662DD37EA0
-9D8C919E3B0602FF23307ADFED6E98C5020F6BEC501EF68FDBB37501E9C2BC88BF3C01A4605B
-F9E519696E5B5AF6B2E989692908C05371868B3455214B9D066580BDF6A95B0DC03961B4DF43
-6B18AAD0833FA61EC358F2EDDC697B7CD8A81A1DE0CB245FE0DECA918D5652E36545DC7771E8
-831CEA3902762AF3CD960F98A979A3D2356FFC0267152417BCA6DFD03E905D50EC7B6E5019B0
-B41F1AE6BB320E9818EC9668E7EEBA81DB8A2EAF308429E445F49B4E8994CD9A9C499BA8637D
-A184D022E3C097098A746A47BAB76CE92EBF1F3DD3F9A4136FC7F61DB8EAD92439AA128414AC
-D99388C376B92C4820BC1C5FE30A3BFED524BF27A055437225B6DC8EAFEBFC824C6D303095E4
-0FBF5961F866AA97B1DB8A26D6A1BE39B71D79F3866B63BCC625FD5D1F6B775E9AE489283315
-DE04078A7D48D65ACFFBA74E82FB88845E447927F72D2149B78F44DC99A71C431942BD04BF95
-2BBD928540580355427C895B7C2B8AF52ED13C53E51E7B8276722860AFC38966D6677621D0F0
-F0B24B90B1D8945E2A5C7E2BF50B8E829B13EB47136810EE72C178CDFED0DE186B59EC5AD847
-52D7D0573C4F3E495F521AC8C45A39170D5C7AECA917EC898C981A2768C6F415486E09EF62D1
-4056B86451B24E4BF95E0F63ACB22A6CEC7BB16387FC9A4712035DF5529C84DC7C47138F2A13
-9C7F33CD256FE57E3665B5DE1A39D257860C252A59119F61F1541E79895CE9ACEC996E782DAB
-E4E3C8A3F88C123939971C9B177143482B53323DC3F205F420A793CC0163657E8C9802889816
-E84C97863A0F08DBD76BE7A917EE0E364B58CE5FE0CC437A9EC117D85B57B8C5818FCF7E03A0
-75A6E65692FFE0E91B3E8F75586D9A1316319F1AF8FE21EE54C84D994AFD97260A6E26F745C2
-94916B76D9543FAA941C66C16D654EDC0C3D1ED2A150AF8E47B2445CE249F61D2F7CFB7DB890
-420AB85E10E9ECD46B12A1B85B18326F7816D8F365B852BC79859887334D4E774CEF5F3A04A0
-B99353DD21E1E77C700FFB719A9AC3A84E037607D521C1FADF08DFC3027F5B95D8B80FD05B50
-248D3B22A3C095DA383A9B565A7D05996B0F76B056F4C5BC16F99918A340DCCB30F6B1EDD805
-E05E4AAA90F93B45D345B3E608BA9C0003E66C59FB89A281F856769150D558156D85EB827252
-894FC36870B4C319715110062DC614769B04B23A98B2787BA37EACCA0FB8B225B6028412D8AF
-CC1A74838C0C5C4BBEC9874D21F461722F8B6F7F854D3D56495740FD03DF6FA76C2101F2F3CF
-1593AAF4052D91E496AE1AB3AC046183E8FE4322A61004A5CD98B81A0BA0C472002C1C386D86
-A36C073CF1FCD46E74465E0042F52E236702EA3D3A5038AA53C76727E364805DC1A2E42063C1
-A7B6535F1027A24981006C112F6CF3E0EFC396412702F7FCBF5BD40CCDB7CC3CB6D325CFAA63
-345ED5DC8FA0822BEAEE634A1A12ECC1696751A4C5C436573749200FD429BAA6F04BB33BA4F7
-8AB4CF706914438A84ACD81C95C1E69408CF8E50BC3C7DED28BF885945BF9925D03FDB8862E0
-D17A2C44EE9B8099DEB7BFDC5C69B6E3953731739C7CF610C6E62656E4DCDC1DAB172EEA538B
-A48A2D764192889D960AE257CB4D95410EBE89BA3A55F9317DB26DA0A6B66D7D4E47CFA28ED7
-7152A590138AB30F3D37CFB4EE6F5571827AF8CB5E7E746908ED2D7A36BDD7801152FA1AEDEF
-D25070C4D097AA35CD4F5378B9E07C58B1C66B751BE0863A137F46135916566DE53679BC5972
-053033391F94635BB50567F8E022E34C052BC8A19306CD1AE642041EC4176BCFD32892C0F185
-2E6A2CF70795174313C58BAEA6CC073302A883DDA079AA47BD7FB024AE435C0265BDC988F5D9
-023BF9F6BF8FA4688143320DF3CED108FDDE2786BFA1D9559001780E752165077F708EA620A5
-EEB3AB9B8C8D842D9CF34957C6F9A5CAB14C78FCA2B9D1377FFDF37B481AE211F1319E09FBDD
-BD7529ACC3A915CD0CBAAF35A8A767062199D017C57107ADFD6F1325FFEB5E5769A12AEFB37E
-4777F9888364DE7965EE28629E124E2AA3A392D4AC39BC7433813C914095ABFBBE5D1FD98F09
-5A8867AEA4882E7743D41FF09C8BB94F0E56577BF5B05AD10F931031A5BBEA91E0982D041607
-15F0140294A5B4CEC572006748DC0121174EA53D67D7BCD56E58678CAF284F459D1B682A37D2
-11D194C60A1A8820806F17BDEE4F8CFD134496BCD84081141C7565018099936786058CF6B0F5
-4B3FFCCC3F576634445B45D245949F290ED3097671F8FD91DF74C965967916BE407DD9131CCB
-081C26267E6079090A23D9DBBE3BA8AF4933D94AF54BAF0ED378EC507F4E0F8A3C7EC0EB9183
-DD9520EE08A632D7AF5A0B85B4813D5FC404C682DD20B06112212C133982F13946F113A84D5B
-842A6A4DDD0BF670867378E272448775E5B6DE072028ABFF9AD1218368D2B4C4B52CAD6D0479
-C45B045DD10F23F89D8E369E111539E4D1E8E9FC13068BE3A4161DBBD0D96A70408E4E9C1A47
-E9D975AC06C28F1679196B4E15D871A060303EC0F0B25E408110E8CE1D838C9A5A6A9ED8CDE5
-D17DE67CE3D6942B672385A8BCB9C6CE459D7CB579A77BCA6C94C5464C57D9A98DFECEC16752
-EA9DDD11E243E5E83E78D904BC9D2B1F6BE5D2E201D5D139DEE2A50C1B77DFE87695C37BBB6D
-3B758CE015CDD3B4E17F7257B3C6E970A7E449F1564F6559BB5CF04AAD718A7114E0F04616AA
-3AAE44061FFD3EA89CEDAECE11790F0B2DFBAB76DB02343907BC28E07D22E1E9B45307F02BB7
-AE035E504143F4AB67095C8670C0896F451404FFB92FB4A71F11DA3D351850ED39148EC27641
-0D2B61BCACDFA68044478FA2917912EFCE0802E2A80AF89AA0B8BA6BACC16B74F325A92135A7
-9F57E33E69460F242A75B19E37E496AC5ADE26A27622E9BE0B8279C83BA338ADBF8065991D77
-E9E02250B3E150301911E3247FD51DB0EF7CBC5CB9A8B64D8B293EE77571C4D7FB22B72A6CA3
-28C736AF5C16CD110B97C643BA93BFF696A45C5F9A385131294BC9CCE03FCC1924158C5F96E4
-01DD9F6ECF635B9B6E77952891A66131E0FBB6DC1D2EBD16665EAA09338AC5C248A2EEF77642
-267C8B1AF3A4CA9716BE19DD111C99BABDEAC07E58300FF4BC9A4B759E1A372D18197A49D41A
-222766CB33E067ED405A1DA746B6F8C41520B501A80F31125D177B57EBEA38F7A4D50112BD72
-6D40BC7E3483FD28ED53165F0D32BE518A1CA23584CFD86C867C56E16FFE98159C10484FF3C9
-47DD5C4003861EF1082F25573FCF1AB4178AE61F69608F936A2B23456EEF3CDEF7082A1A6753
-A332BBC06677BA18CBACAFBA5F322D7175C45BE3266F71B174FC6215D291A82027130A6CBD9B
-4CD0944488CF1DC8F75F8A60AC10CD7D55E8268A9C87A70BFB2BDB984AE4774AEA5A347A3E9A
-D9038B80167FA99F67D552996A8886889A2A24BE6CCA0FE226766CC4F542C37E8403A579BBA9
-56C5F33747F484F58CB83A4647D2A30EA0829AA93FB4A67F88372551723505C8E9F7DF5D4167
-375E17616DCAB153C53A1E385408CDC7666D997BE9E898DFB943C256DDB9B5800762D08E5A3D
-DF067A3D1DA0C7841503CB1B522B92656C1E42F7D78096DCB9445CC80E21328A8B24257EFD47
-88C42595DB5B14D695F7C29652A9A92CB05D07FF9ED1B416326D055F2F9F1FACFC762A204C2B
-B2A0A03E7AB51ED6E8BF3C2B7763F90F039F4D833B6892C25982ED6BE9363B0F992E94A3BCEB
-CE53E4936F626B6B366566E2F1972104341516D536E555997AAEF1E027BA1A53F05CB1DF4743
-B86E4C4E26EAC6BB8622A2FCCE5C0BF035219961B19939DABED6F80957D7FEE4A1755042516E
-7FAE6ECAD63F0D5748586F66E61443BBA19FC8CE731D6F37F02ABF972F24B1764CDCCA8F6A42
-49FF841BC7D5BF516649579B5805F4836991586165C9E558B34AE29E71B4571B2393CA0E7F5B
-D161D07D2D9CC28B25A7384DC1279353F0A6B925B9B007534260EF5833E8E36528E33E7CABF0
-8895A2A91A7025F7A8FC72A351DDFEC11E4E7F4B73C89028817EFD7431B8115CC630AE78204F
-56E9564108C90BD5A3A89AB2DF3DB2C37001A963CDC7E318E52F70B0F704C092E858DF90AC0E
-CEE91C392ADAF77578B1DDDC335031AAA09362D6217146E489529C187E5C5788311AAB5B4AD2
-38D97F2E3715CA8FC2E50F38AC51E3BDF83DCB5B5AB4AFC027BA8B8E9C92C77A33D2F1F64866
-F54F1B66D829083205724E1406EC18C52B0A2C4B804E085A4F07B2E46EEF384AC81374FAB9B6
-F36E7154BB75821E4A9E495EA12FBB9AFF3F440A72C412E34387E216C557C1AB906D1004005C
-B2FCBC93085BAFE8B446FAE863C733DC7B5BCC34247587B5DD14849FF0131A2F63084AFFF640
-8929D6BD24BDCA871947C2F99BD7BF21B04C2F060A2B39A1225D5F5EF341EB47A5961AE068E8
-E8547223BF73DA08E6464AB7406F65E790991F7F4A08F2978A1E44CDD4AA21588EDF7CB9000D
-A03D758BD7637C1B04F13985E48B772E91F727399AD051963F2D4B391C4E826D1DD69E818C71
-E25720A451DBBA2CBDF4DC3F0652E9814ACF1E931C76CEB7559DF7CE02A4E58E77DD70ABAC09
-E5A5CB6B2A4AABE6AD32ECC957B584CAB8B375BF06FBD444D7B74449FC0BB0A3BAD594AAB1BD
-A9521FDAE16803647BEBA4CA9700D74E84C65A97138764990378C15BE82A7229D8923DAC87B3
-FCECA93D77A9F0F6432BD04BBF7717BA9B876F2C17BEBC00862705E558DDB38DEDA0890E0DFF
-15EDFB62E0C5B33859F1EC62770572A83D85A13DAA7F1F1787D9FAC0D4E802614F239E9D359B
-8673297E47FB9FBE51F3A35C6EA5DEA81231A97F524DAEE990300F9EF89383AF069D2FB5111E
-A5230C83312B907C6B4F720613C7B282A6D0B0E1E183CDDFDB347E5AE595AA7A80FEB6584401
-29DCB350E903575620D22BDF17CAEDB5C8921BEC4837BD4A589F680FDCEB531A3673C9BA56F1
-85E3F3C2D3BDE1693914FE184AC789CB5FFFC2FFE5FC158968C331198E655531F4A916422D55
-C3615736FAD85CF6FF05D783DA9CB925E11AC080FAA89666482BD4723F8E23B1E9649CEEE076
-1242EC961D77E5CD891434D91D04AEA83755F45B27BD38AB0464FE5AC97B85D9A22C0C1B6B84
-38988B10C0EB4486C7AFB4B59AE126DB1C9573A17071AD968E86E606629F7F041A7E5E5A3E7F
-CEF3F57F6E6A73AFDC261E49B32F80FA869FC9B0ACD3185BE61E568E6D78D08DB445275FEA4D
-DFC075F8F6DE49873AD513E6D8560EE0FAE647BBD94D6E6C9C04E2B31A39AF036F9444B96258
-7BCEE897077681322DF0083F9144BC6398931C4E87DAA5CB5EFE36E143CD94B6FDE55EBFCAD9
-57DAEC732822D43276B082B385E3B11EF98F4FCC9F7C5A820B94BEAF8990A23A68040C79BF2B
-5C221E3C774948FAD9C987EC7F9D50ED82DD864D1809DAF5A8EB844CD8A62526C429F5140D4A
-B8175906DBCACD1F9C70045081DCD7AD78EA62DF3CE820755F7BA3D1D20169810E92EFC818B6
-8429B10D8EDA7FF9EF6CA08A9D6ABB43D022403CADC46CD40CF1B17331021F05DAB5CA96E781
-FB79F859ABFE6F727E0474E5499C915C20501CDDB9D84267D8F9B9F88DC3FA66B6D025CC959A
-260BB626DB92B4A090AFCD4F98667A4772B46C5F76792E3A3148C3B2B13926170EEAF8C288A1
-E7A1E020DBF818A96D00B9373EB0CEDB8B61D6ADCA7A47ACBC125C02E354D8845F2F0091E5C2
-09791048261F9D3ADD46007E26C82BC089572968C3681FFF655AA9FE7F5EB6D63A8A77EC1DFE
-F2E9772653ECA4D82C761DB56DB79F0C4C3752F93F7B39CAD7122FA9076148607293B41686D2
-365FF4FF063632FFEF3260F9F9F8931DE5180B59C8B92454F39EE6829EC9ADFD90AB7FF90A89
-4B46E639A3C2539AD2167771136FBC28103F6ECB5BFC75B785CB5DF9A32FFDEFD24A247C26C8
-15A96D57C29EAF2201797D362BE84FF64415963463C2617A8CA8D04030DD66B23C638C662584
-9A921ED3606C876DC3EFA086C73136F534787A28D4CBECFB05076F55D3BE57038EA209978153
-BFB3C298291936CB9C4D26489811E646E7C5D6C46D83C720C6CFE8D7ECE34E207D244DCB58A4
-2C3AA7C436B986A9491B9E842F7B0F686A71CF0830EB03A17917FCA87C79F62EEC363D3CF220
-F60013BF644E7822DC0B712D3F70FABC67B628D62652C4B46529AEAE03FA158B579098D96E2F
-F9139151F403EE7D774B18017CC85C56EFBFED8A8D375B0A66513E4F0C550913A8220BACF4BA
-BAE294ACB503D1DDC12367EB16A5C344DE1EA390183669E134012304121068B112E40942DEA4
-D86BDEF35AD1FB00BF6F0ABCB725364724857E81809E391C4A40271066DEADA959BE2DC9ED11
-660F039E9717D0198429A49A67E6AED2B25E1F8971DDDA970AE8191A0DE2D4531D8CA2BC5742
-C62521F00826A8D3C66A4AFB2018E49AA92E43A526E7A33539F340FBE47EE7DDEBC166D9751D
-7CB75A75E0373260B14208B89C569E505A5B65EBA5689F9E3E663124D2AD8990CF43E2CDF2B0
-89B284864F2B732CF172B1F4BBBB6F47D3B339DF3E7A70464B56C205B2D6C5F7F8842DA9BBC8
-D57C3A4D3E02CB913C92BA511DEE3327181047ED02BEA1ED373E9605265C7AD78FB9B20B3E0B
-5FF5FC35A4BBC94F14A4E516106FA3664FD332266490B4ABFBEDB4AD88A1897BB9F556B8A4A1
-F5BBE1791CF08E83038874183D4C2343CB4F4D8F36EFF8273A5DDBBD5DDBE09C0C5363A653F7
-BD114592F8D316DC75495CD78D2E69D95B190919AEC0D7161A665AC680A03C7BAD6EECB8024E
-6DA08317D475D3569245DA8CF53E217FB5F6E0EA6BEA848A87CB80C25ECF286F562AAAFE735B
-4BFD3C5E6F558BA5284F7B90D88494B5A13314503623E12DDE6749760982561A63AB10F5F5C2
-8B7D5F42893522C081A2B08A6767D30BAEDD136D223CBFB183001A08AD4226ED3DB2C462ED67
-7BF56285DB8887C93600431F443AC001BA1D9F173001CCE2D7C88031F218F0F1088B45DFB345
-7DEDE1CFE794DB66AE4200DEA1EF54016943B17540FEF6CBBB77BBD9465E22A116EDF17436D6
-E24523C0CF894042FF37007338C0B3B284D7657DC84F80AD5B876D8849FA10AF5E2573EC5165
-2B13516043AFFD9E0A5CB860C02AD2A43D4347DA310862195441D653046E7244C3F654099495
-1FD62FB5F8D85CD2C1F11F193347D076E7A5CB123D122DF2BAB55C734452C2B37287D775FB6F
-53EAEB5A798CDDCF0895628A15B80873E15E1AB1A95241D063A25A05C79FBA8907FB28B39F68
-47D4FAA3211E43FD93F146331277B33D85A65E92C951E43A65941C2DBCD0ED58ACF698A9DA9B
-F2EB7646FEC84F65FFBEE0111DFDA4C92B28AD109EDADD07F409327067696CDFC57F3D51DF7A
-BEB3B6036C77C627F9F4FC991A2DEA0F540ED7C94362F532593DC7B098CCAC4BBF1ABF701D1A
-34F59496BFC61A1FA2E72BD9F68639B8B92FAFD54FA651500F31C105022281E73683D9CF335A
-F8CE624A8CEEC267B512553E0AD25BE17070C46429EC7C580E1A6AC3D01F51109202F395A7E6
-8FA2BBCA0E2D1853C85BCB936BBB04FBC0613AAF0385ECFDB5B1D5721CEC321C24375EC5086F
-7F0F556C4AF461F0ABC6DF0B008D46828E867BD11B2393F4062F9AEA1ECBCA61EC6D913B11D0
-8485BDD2044C91CE5D57F6915CA1EA81B53C41CD9AE53D41F0E9DEC396D9E12CBE82DE131E4C
-CE354A8ACAEFF3D20F057394E98D088DC8496520DDFDB0587DB03768A4ED3FE2F35C45A700A3
-A8E68188D1ECA835EFBAAE650A904872B51DE70F6253F3A2F5E2ACBCCF541E7D53039536FDF2
-FFA99CE60EBE9DF25DCA97FFB8215B2B3037434A4A5F647428B9AC684D791D43480351CD21EF
-EFF6BEB1DFD868727ADEE7DE8AB58DFFCFAD37BD0207DA347E66B5A81959ACA8D08652A4DD8F
-3A43560B8A38A6B1653BDF64C3C48C851461E25E715D6E296CCBE06D33352CAAAF35365E9608
-0339AC078C158122A738B8635290C6C2D857B2DBFCF67F4D1C04884F3BE9CCDC53F08E43B8C3
-0EF44EAA6559B019BEF8BCE4640F6DE473F2E1BBB4C66FF3681A92A4AD7CB2F8CE600B73E687
-1955AF24CF5DCD62B0295F5A50AACFBFC41329A6EE712ABCC7799C46A7F80BCC47280912DD54
-F01836C7E63BEAA36D5E4EDF453D1B55AF254369836C02CB8F437C74BC7D4F940FA1D4012D77
-2883260007D9DABB546A4C7CB471F4F6C7A577F3527F078BC7AAE4BA915ABE6F0413EC6B9038
-C0FCBE6531D17830146E7988CFC78973C1B6AA3CCA8B0C7C39C2804C5049BE8BA799E19AF396
-1EE7EDF36A162F699582A686A9E0AAF5337240C32BD93F92F055D2D2365CBF5C551ED02FD91C
-817BE28E6A7A510D74E97D572AEBA8440BF425F1BB0728DD8022A96824050CAE594539CB8CD9
-D7FC41A93C1446222A1008E9C5419D3033BCF29C5F4195BAE02448D9D278F314FEE669CB0C2B
-6911277AF02ACAEE0232AC8A47D0D31392EE739EDAB934A73361C0DD745A05A6AE8F4B7EAEC6
-20CBCF882AED7F5CE89B0246C29233D7EB3AF04AB77F8F94C81160E76671D37F1E7784C03552
-EFAC86392BBD0A079CC2868959D24486172D3AA9CA52E9BA6D98BAFE6DB33D3D0C6934C25AC5
-00A0A898ECBB51E26A3C56AD78D5DCD8808E48F46A9E1C6B4CC5929EC5A5D03E154248D410BD
-F32538053F5DC0B19AF6ED119DB53218B08F35D650402EA96686A62773FFD3D41ACC18B4BFD5
-CCC849C68E1985CDCDC4DDB2A8528FEEFC6468472A546FEFEB77FB3AF2C54EE564179A4F0186
-74A746F8003A80F3FBA55038A4B0B404FA7704BAF41F69781E61805E110CF2CC00E82421A5AB
-32061EFDFB92272299691691F4D12EA614DFC240CE737C0B1E32EFEFB542923FC898FA79324E
-E30ABFA65ABF4CAFC7DC78F1CB3F849FA14D61F8789176C36458BE24389FF6C0CF627FFAE6EC
-51D04AEF621143B116561B954C18C182894495641CC19F26FFD0CFE8885787D214FB3C23DC2F
-C133508736FD65B8452696536C6590DF6A4846D4165ECCCAB150DF049685CBDA5FBCE14E92B7
-3D7AA0175D91CD4C8F5F7CB55C45386F5E8C68725E337B7B4D5D404A8E6BC42D3217A8E7E114
-88A47D49EA60CDD3986627DD42C5BCAB3E7665E12A299ED173716101506F9CA79863C97A1E06
-C9258F0130FAE02C19BD2F772BF4A5F984AC0180A1C8659D759107F3753ECF25B90DDB016EDF
-C9E6BBD65F9E074E347D450A8E7C7438581B5700A0CF55E642787BDCB5AB21D83BCAC8536284
-D202A844DC48CF64542286DD26FEAEED9BBC46A9F81769913FA8BDA9BCE984170F2B155D451B
-DC20C53727B80B030ED5080F9EE2E2248BF7127982A93D5B59C938287477FFECE565D442DE29
-4D6C3C2077C99C873A2723C29A0EEFC842DBE5FF64A3ECFB4CB88F5BA0665BE790FE93583618
-75DCC66C14865FE1B010D72D104077326A4F9EC5BA250277F82D539D3299B6185713B0F04E88
-10E77113B795280A4E95A4A6862A0EF25C760214B95BCEBED36315567F0C9B80D3ED97DD27E4
-C23F7AA05F82A000450F808D81EA19D5E04D2886118504C32BBF80143C7F18F6F652D5AD106E
-D03CDE35F8221BF8E2CBC79AEF65A7B1D2EB30F1DE0A8100207C7DACB8A717588F0B6FA20BA7
-38E9F02A879F3751C88870FC865BDD8D55D86FD6851F4385DF426C491CC38809DF42FF6607F9
-F15AF723619623F473971C352B7D31B1960CA7BA4FA299CE18217C46CCC59367DC950CE088F7
-B5568678ECA6CC5832E77E8A446FAAC03502EF3BF782B7B4C21D8E8A4B24CE69172D26346EFC
-97CD94B7D2C0A37687B2BADD1C329A60DB8CA2E6F4F5DC7372A454E1F90FE25F3F3C30AA1508
-B7E816117904E2AAEBB6FA7A576F4045CD50EFEE759AAD0B3DAB6F440F1DCD3D2DC1E311D15A
-654D3BA33F45C873F4441EE1EDA3940FA5FBD5AF39F5DCBB6BD390ED427DC674D61698287F7E
-3F9BE35C32102D634CAB14C4FDB924CA7A713A0C6FB3281988D239D2DA9952CDB54F28FF4162
-6D11E819F12A882879960E8AEAE2917E347C62FD3B63CD9FB18AD211B62562604DFAF420C32A
-32E9A0236A08B911691D84F61196C7396757B1579311D44E67389157A1463EEC89E5989833C3
-BA777ECD9E11D281A2A5F1360B83EAEA4B1987D7B3E231ACF714E40FCD467F827121F1526248
-E318A2CB3DB615356F2F45FB3D5774EFC4B6BAABEC4A5F870D11D58B1F1CE257A7EABE6E676A
-CA270187A312B4A6287FF3033CC204C9E1BB9A068D1F5B1D717FBF193C5541E879A77935D924
-899E79BF759511FEE38445EEDA05E132D10C23B2E6CE14E5B9F51BE4D21A462081320BC8DBB7
-7AF84E4E588C0E9B6BB37A57F785508950DB5289D1944EA2F1E73AE29EEF80164A463FA27699
-1C57AE565CFAD33FDC0277426263FBFF935D48E6FDFB7A8EA6B722D5026D26CD794E3CF4E8AA
-319502F516A07D9C74734D20A878492491C7C7FA82D328B6E771565B2AB788AF29A2ACF80687
-D67A205C4AE117F49C77F76FF8E807B47B68CA23C549B0F197E5263C846837FFDAD3BBF283AF
-FD0DCCC3D57ABA922CC88F6B128E303B9E607D97EE77BD96DDA0530488FD27D1838B65B298E0
-27825F5F02EA61E55BE2EFF2CF1B9B1935FEC3428F51E3859FC03B65A3A9A6652F35F635741F
-55E150F5C1702A0BF03334652FBA0A60BF61C7A6617FF85C250D6660C0800ADDB55842D46C1F
-B5F93746DA16065A51BEEBCC1AD2834D35E03BCDD9524EE280C364EC24E4FA62D72D3385DCEE
-DFACCBB96FBAF72D19D7DBD9E8FB41939CDC79F00E23881B24F664C943F0C68FED3CA1E35A6A
-E1F1B9EA39247C93050C9DC264A366491A578912466E0D0F37852E4E6548C03E858ED5A67C43
-C35B8FBD7E6FA4F079BB566340D064CB80373E93BC24D1C62063C967FBF74D8403486CCAA205
-DEB98E7C3EC41C40C868C62A6A6C3E105217AE8B67750D579C9DF1048FC34B0FA87A3058BA71
-4DCD610B790E7098326779ED24E2117DA9420AC5D22EE1CC7C3B93BA96AAA25620C1163F0A37
-187A9E0F1DC3C4C766112CFF022678B96953C8F2BBD2B76A3628FFFD9F26235AEC462A7474BF
-E000C5447E043388D0BBD480D5FB6A514DF1474DCBA302CAAF15921EA07735CC60F596137144
-F39D922F10AB14A51BCFA033E3EABC9ECDD80795EB4E422628C80BB9ABA641274954C4A5EB26
-A9D9E6BF231CF5571AB9DF564EBD6747DEEF7FC6D5F33636584364B61087F4EADBDBE945B962
-9710FE1869A2523CF46A7808E4FE7DA5C2DE5819E3C0B40654D665F4746A81D6216BC0FDB0D2
-326463961CA77066F6B3B84E17C3D1E1C54BACCEB784C40797FA7A69C501FC43B31A19D02B02
-4CA2FDA936B2022323A0E5CC15CC1FD848D4DB5B9F177860288F63BE416A3D2539B05BAD27EE
-C2DFDA432FFB22B8B903F5647B3A61825FA5DF791D018CB92DBFD0FA93DCA029EF2308CFB6A5
-630E1F6197B6BF195C5417C2F135D1224F52FDA9EFC990F275FAA3F4A1EE4550B1EAF4B5A432
-E1EDE3A0F0C9D1595A729D4FB008384BCA04E4C74D96644A9AC6BDF55FE9601DFF5CE27A54FC
-4E1EB00C01B70FE716BA0870210328AA90313E70310B8C65DDC0809683BE822B81F2F2F77D5E
-23C5689F5F5FDC8842E1BB19444A2CF00D0EDB5A8ACFD94D0CD02CD2DEFB0FBC81881DAC3FCC
-DD2C5EB025CD51AAA1286CB11BAC21D35F618542C8DF1872D24ADC2910FCB3658EF7029A1B84
-719149DF0C1B7E71AE3BBA86B6E828325474BD0DC33A84E0B2575A2C587AE093B20AD1192881
-7AEBBDB8C13D1633B0B7FEDBD2465460CCA3674F97E70192879E21C5E996A448674C80E7CDD2
-6B205AB4682507B0DDC14F972F6ABFEEB9B4A1446C394EFBB864C6AE60003D42A8AA8B64F44A
-8AEEE0AB638B417B5C86D2A55240801E97CA0F74D03DEEDB301C5B10F8AC7D4DB47AC9F3D0C8
-8895FF1031E643013B6973D8C55ECB2F32C1CD67C8CEFA5CC5D04E98BEDFB80533C0B4E238B1
-6E1D94F3BE5D349DC88D4D7424D5E43C8272099BA189DA7BE66D19AD19C62F93C5E2DBDD113B
-BEC4398AD074C79C6361F760601C73D1C4DF799171DDEEAF14E5B02C506D9F83E885B57DDD7D
-25322F53665E3E1F97ED8DCA484B1B2C868D9032AD24B15A7AD58BA8D8817DFED0D62FC7C46B
-5ED293E2B681BC4E011334D022D490C0FF5E4D75AE0B308AB89BF51DBDC1127CCABA567B5087
-E9345C31CF905A7BDDBAD42B69C0F63FAF87FD6FED01E10874D883FDAA4202506DDCC71E84FF
-78DA1BF22C3C7A48D10566F2D63478897EC03C71A1249744B1FA1A5C2518CFF163DDC51361AA
-9408BE1BCA3838ECEE3C73D3332AC2EDED9DB7B805F71A9837B905E0094D0699B1D7D20C4727
-3E0650F6A699091C3A951EF4A59E0B9E36F6A7CAB10F5395EE85A38783CA90061991A6A23BB0
-B7A6C5D1D81F6A9450BBEE18B4FA10772533140223A6B9C0D88007423ED79EA3D2C1DD133280
-1692FFDE210F7B2FC9AD91EA9C048F3419F85E2EB8954657BF7017653B83634E4463DDAF14CB
-BAC9CF75EA3E023B672DB2EFEBBC0AF0DDDDEE3DCFBBE5AF152680347A4E6E04CEDBBECB2348
-FE03F988DE5F9CB7E78D11B1DCA78EE541E2983B3D9ADF05D56F0917E4033DE2E5CEB984096A
-500C7C758D4D4E176077882398AC22ABBEBCE7FB4DEF1C9DCFA55519D46E867B47F1EACD7FF5
-0590E77EECB0F416BE78A6C0EFA1EB98B03C7A555721D23D8E380A7B94A9FE6E8FBF07DE7E85
-8F7E629C95CDB62F15B2A432549AB1BCCA39AEBDD338F25FEC35478A109158EE0508607A60D7
-0818CB336C5595DE9936387992A8B194AE94139FC0DDDBCEE22B5424610699AFCEEB01ECFB39
-AD7B94091787F5C0B1766A445191CC4D4F30ACCF331FA12720AE1327493423AFDE198B3374DB
-110A1BE58A4376BBAEAD5F2AD2C1462A90AD5A39D3514D498EC2901F48C4A35064DE76A1D85E
-1CE51B479A8621A6B123300D451C7411CA8C2C35E4F865D41E571AD11E2AB2F97F19E7BFD2D4
-35107A268624CB3793F4119FD7CB5402EC4C3DE13E25EF056522A827D9BB489BDDFFB398B889
-115FA6F3613EBCEA7A1114AB5EFDB2FE11F5202D3C5B5415C9C72D89195247B8DCBB4D53B9F7
-F4561A8C47D0EA488357C2DC48CFABAC32C22ACA4F8D63E7A7D6A8F0275C39606EE4AAAA3E34
-1EFFE7E3C0D08B06E606657317B24B63D8C562F0C94AD187E70F50A9BF2D7642077AA494D83A
-4352EF3D5B93A91CA167205ACB66D159E8588BCFDD80CEF101201581CBCB02D925431A50CD6E
-4719E7B393F7D427EB531DB3860C757A284C5D331C6BC3BF51AB061B91D04B33ED6BA7B41545
-1528C744E89B556E2464868557B67B21476094AA036B80CF04A63BA3B4265DD4EB704919AC32
-B490D5B3A840A5D6FB074CFEB4DC58D3FD52A8A816B9FFAE5145A141739AC63E5A6AC07998A8
-DA736E37CC3646EB0EF4BCC30286A7F3BD33544D00173344DF1BA5D9B630C907723FB22AF749
-3EFEBC84FC8142BA18B6B06A5918315668A180A57356E56DC65E05D47A4FADCBABDD602FA00F
-2544C2420D32EECBE6D59A6AB5DA5302C342062F06934F3524C5EDF956B6BE378FC9D2650612
-F1D9901A54C394D776C9C8DA4E682BAC417ED3FCD84D1A9CDCA68F97194C8BCCF13EB8BE8430
-8EDEC1CF7CFE4A08F18A0AA8C3BB75967890315C20AF9F9AFC93FBECEA25861531770D8E02E3
-5C5ED584DE63EB65F72EBA79BA03702535D0300215AC6153A99F6AFF983EFB822197B5FDFC3D
-437423672000D30F7E8522C81383E4BA6DD0C5E1E3201E60965FC8F4EA4650429C23A55AE693
-A3E918F93ADE7E9B27E922D9E54176C4B16CD7375B84EC1D73B5CC41D203ADCCA1B0099E0DEE
-DD3AE87AD1A62149265CE58CD0A6ABD40E4D34EA59DBEFA51924786419D1271858FBEE380F28
-958D9F1F84E13A941AFB132AA2746C4321984D68374952AFBCB87BADBEF37386CF7BC48CDC46
-FD6C8E3BF5AEC462E386BF341611718CDDCF759D49D2E011D5A27583061689AA60D32F89C7C5
-9B316DEBC337D06843FF0BD92002D21DB192AA92ED22C9DDBE37F30419F05F6E2C95B80FC068
-D6672FBC7D8057B2870D8D0D88FCDDDCF76C7B7A0D3198867EB5F54EB4F208B9DB303875232B
-78B9DD6D99ABEFD765EFB8DACA357579F605E259B84DEDA3B2C251778957C8B2A9A8D905C068
-3307A1B015B60A89A67EA3D5CCD1BEDB9B03D09501AE8BB57A36B29B7D9B970940FB114B5977
-AACA60C306EFC7916A94C7890320F7C8C541771083B8FFC6CF7E295DE9C91299365DB3B6F075
-FC7210673955DBBA33CABCBF4EC61B30A2A10002D175F7126F3A511099AE72B9D1EA59AC1AFC
-FAE9C72311924844E03001A454BA5211403FDBECE659F0A23B0F33C3B8290048936A8FACD25C
-7C28AE782016725C61845CE68FDE4D44396741A01022AE58F8C1CC046883F19E13D76BC3CAB1
-D12F645EB7CB2BE29C902A595EDD2F4B0549B4F38AA1918D1AC6A8BB0F42E28E87D8E2B0B0F8
-839FD014FCF09A2599DF733D7ED40EE5CD0D8F40B40856EA7839ACA91F1C7523537614D6E1C9
-E1D7ABC6F8F3F8026779AE2BC3202C6DFE5C854F5C7F5BC1702B970682A986DEF75B6BDFEBCB
-6D0FDB09E06930F7D305955D9E6A7F1B425C24BF93D29477B9744A68793D8305FD3E8624A4DC
-41CC16B6A3112177895799F65FB5F57BD6D385AB0CA4F7A7237041D6A475855547FFFCEB3390
-0CD2DEB24FDF60CE660E5B06EC41D94ED16465F0243AE6BE5E75176BE06262C7CADDF6AD2FFF
-15A0479877C54B05E7345FAE961FF54EC2F657ED45D395EB32D953215C355E5572C80D179CC9
-D35B8B5AAE9FE122EC2A033B5FAAC4D10859FAA612E3A95701A015AF859A04DD1A2F929B8532
-238C2011C02B154F3A73BA9AFDD861934E3121ED08DE073F6BD488CFA96B3088BD6A53E57F34
-C658A6533D0C84B26B1E7E57CB03787730E0187B78D5677912EC6B492B4058088659374A4746
-D383BF2772D8A1BEF93CB553FF653EF6D8B401CAC65147ACCC9EF5AE8B38309277F280C89E1B
-7723DB53D2440C59BBDB27B475C02FD6CDDC6982D18CC38490A360DB67755933D2561059CADF
-C6BC2250032305BA3F95823CE845F118552A987726256CADCC9623FE188254893BBC00F193F6
-876468916EDA8F05248B98C3DAA195D10FE381728B559F9D0F5EB168B868A98574EA59FC15DE
-000D63D8E39640B2808F3010B545E824BCF4F737235E5E98F9571396427F50ED9944D4177A97
-DC67F6AA00CB79324154FC8C1FDB60356740438D51FBF09D6B6B006D66D142D5FB048B2427A7
-DC1E586C2AB4569B6F253080B2701EB12AC3DAC3F85704D3416F4B993A09DF138AE9FB243DFD
-2C872DE34A1671C0212AA745645F8649D958994E5FDA7D36982509F26CD101712FAECC40B27D
-E95F1D332F0D12567FCC930AE14A2266FFB25C92C9A093D58608F7C446BD0F2ADCCB2087A957
-537B6D1BB1C3B6DBC002BDA32089BCED46BBB9B93432F9A955659D63587B746E20CCFC061190
-A6C6D3D12D99805921AD3DD097F121385E93E2A08C4F2F59AD668E86E4876126E7DA0469925D
-47E81D2FEDD8EFE590895E58F1D208D748C1F5315D620A9E942784A802CD6EEE4B0C18B67966
-025653FD3C7549C76BAFFA49381A333BE59783D9D3905E45478B068D7EAD6CA938E868777295
-114BED4159E6D420EE29C9E97E34A923979490AEB4DE8D75C53D5C961595903E1D94025EDBDC
-DDCA0FC65D9D3E1D1292F57F255F710B1E7EECD5AB7EB679845D5B2496B018AE62AA310608E2
-94BAD705781753FA1B17A7542AECD69DFE9310D8C9F9A8B9E08109A5E936CC5E1EE2ACB387B4
-ED0A529810B92FA9CD3A7875C2D60D33A709957C972694CD2674CDA23AD3E7166956A3594DC8
-8944BE232BD9EEF8D0ED9B970B2310170AFBFDE00006E93D8CDEC27AA821A83CEA57B55162A7
-9A1009CDC50DE9C98AA4DC9188921F84024931C609F00A8017BAFDDB151B61587C979D96A2C6
-E382BE29D7B192399B2457924BDFE99C2BBB111099AAD207DFEC210450B7851A3E562ED7862C
-5F3CC0290CA1767DC24B69F0A415844EFC3E02470E24E819FC394304645EA8A1FFFF25ADC107
-9EA1CE6D7AB2ECBA3CE38C46E8354B5ACA6183D61D8E001D00232C2FBA50282F15CFC26D7EEA
-F1153F44BB14AA200F9203AEAFF04FC352BFC0B43778778F729965DAC1571722E8B50FF28580
-BAD7E91CFA84436F1BDB49184B6D97BAC4E8BFE06F3BA6E2D23C5324B97223CDE027BBEC16E9
-5E7F30EC2CE730EF791322B27C442EBA301B32D0E8D53BBD092E37217D351B08EBCE5A017F0E
-7C095BF3A31792875FEC2EF001F53FD931995C7DB8C49ED33267D5DC1ECE471308CD01365891
-122E04C5DBBA7A8A576A02FE574F7B520BC1F455724B3A5E482E092B3E5A807DA0C0F7D4E7B9
-2017D3AB39FABEB5EFB31417EB18A0514F5C32B5D9BBDA218325D80F3A3BE8814B72B0E9A0BE
-81C9F2C083E8CD83292AF556D02B3BE7C385D10CB52E7EC9134C5C7EF535ACA4C83057AF593B
-351EBD11554C95AAA5772D2FAA9F88BF973657594FAA9EC9F58F749037C09EE63DB607E97872
-02154E31273DA02322491B54684B4F9DCBB86AC17850461730EAB7ABC53184ACBBDBB77ACDCD
-E696D094976C0032F0741674FA242CF039AA0D8286853CD67A6F781CDF7096FD9A2229BC181D
-A9A1BFDC1189623E7450B297986F72B22A3B205C41761891908648918F551C4E4F95C77F9311
-899986D9589F023F5364DDED65D4D614AEECBDF23953B460B79139D980C150A69307E217C903
-392AE8B4A7E5E6270AEF1704A5B843AEC274B98668C33CF82FE8A6E5FF804661DBAAF3A88542
-0531C5C1ADBCE818D4002557556D8A8793EA27DA4556B3D7B7EE9BA81FB83031E6A52DA89911
-86588199D0EE9951BF089D294FB3577A9581307EC44FC3D4EAFC6AD91232EF4D7595A496CC9B
-D2B85DD2DCFF52CCBA3F70C44885004D73DCEE4F43EB4F77E21C17EEF1233ED78E5CE7491C54
-AF1957066F46FC888E6E48B08AB6911DDD012FB6BEA72454EF9B1DA3BAF9B70CCFC2EB6F90CA
-0332104EC5E6952CA6077F2A035A55BE5FA521E428DE70918E268022ADB5110F1857C18B8F9F
-F825AFD94C7674041AB89BAE421FF7BD06A5AC660F2CD3BAB6CEF3F44A0721437CB45321FEDD
-1F42AA41D006C05061298B1FDD6616B114057082055158C64B6D4C676D4CF77B4D8D52A175D9
-742A637FD27FD024AE7B069810BFE471F77CD07B5E52A3E3A31EC6981BAD91F838EFC7DB4588
-DD2B064D71CCE33CFE68BEE4B8AD8C0AC862F407AA4E8F1F4B81561790BBB9ED868D42744D81
-AA34EDBBE7ADED15B4BB9464A1EC1A37C90F8E1C16BD1E6F46C68BD42C9F881F95DC84FB01FA
-3784B24007FC26D804A8FAC4F50CBD569D68561808A917BF1A50D2E2C0F3A18E19119776427F
-25912786FE931A697347D3E121739EAEA1D6C7F6311E90E3CE824694BE0AD61D3E453A88C7EA
-38007EE504D566AB003F7CA3768264E3DAA816A91B2BF0A807114B5D09630A3438A972067145
-B6D3DF98EF26D06F202A9ABAB507CC185B3382949BCA2D2F4219F2C17B6E6E0172485B6F94C6
-AD4A98CB068B961A19F31610FE5BDFBEDAFFECB188922CFFB80D4DA370B744D5A320C3379BFE
-7A62E6230C2E1C9579CF3160789465C45B21FEFF2E00301076793F1B6BF299A7BDB4A9B5EC51
-89BF2C7A1A59402976505441CBFBAD72AEC88BDDA1CF0B906E9EFEF834274830CC6B9DDD5EDB
-1BC08E4A608CD964B7470000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-cleartomark
-
-
-
-%%EndFont
-%%BeginFont: CMSS9
-%!PS-AdobeFont-1.1: CMSS9 1.0
-%%CreationDate: 1991 Aug 20 17:34:24
-
-% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
-
-11 dict begin
-/FontInfo 7 dict dup begin
-/version (1.0) readonly def
-/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
-/FullName (CMSS9) readonly def
-/FamilyName (Computer Modern) readonly def
-/Weight (Medium) readonly def
-/ItalicAngle 0 def
-/isFixedPitch false def
-end readonly def
-/FontName /CMSS9 def
-/PaintType 0 def
-/FontType 1 def
-/FontMatrix [0.001 0 0 0.001 0 0] readonly def
-/Encoding 256 array
-0 1 255 {1 index exch /.notdef put} for
-dup 161 /Gamma put
-dup 162 /Delta put
-dup 163 /Theta put
-dup 164 /Lambda put
-dup 165 /Xi put
-dup 166 /Pi put
-dup 167 /Sigma put
-dup 168 /Upsilon put
-dup 169 /Phi put
-dup 170 /Psi put
-dup 173 /Omega put
-dup 174 /ff put
-dup 175 /fi put
-dup 176 /fl put
-dup 177 /ffi put
-dup 178 /ffl put
-dup 179 /dotlessi put
-dup 180 /dotlessj put
-dup 181 /grave put
-dup 182 /acute put
-dup 183 /caron put
-dup 184 /breve put
-dup 185 /macron put
-dup 186 /ring put
-dup 187 /cedilla put
-dup 188 /germandbls put
-dup 189 /ae put
-dup 190 /oe put
-dup 191 /oslash put
-dup 192 /AE put
-dup 193 /OE put
-dup 194 /Oslash put
-dup 195 /suppress put
-dup 196 /dieresis put
-dup 0 /Gamma put
-dup 1 /Delta put
-dup 2 /Theta put
-dup 3 /Lambda put
-dup 4 /Xi put
-dup 5 /Pi put
-dup 6 /Sigma put
-dup 7 /Upsilon put
-dup 8 /Phi put
-dup 9 /Psi put
-dup 10 /Omega put
-dup 11 /ff put
-dup 12 /fi put
-dup 13 /fl put
-dup 14 /ffi put
-dup 15 /ffl put
-dup 16 /dotlessi put
-dup 17 /dotlessj put
-dup 18 /grave put
-dup 19 /acute put
-dup 20 /caron put
-dup 21 /breve put
-dup 22 /macron put
-dup 23 /ring put
-dup 24 /cedilla put
-dup 25 /germandbls put
-dup 26 /ae put
-dup 27 /oe put
-dup 28 /oslash put
-dup 29 /AE put
-dup 30 /OE put
-dup 31 /Oslash put
-dup 32 /suppress put
-dup 33 /exclam put
-dup 34 /quotedblright put
-dup 35 /numbersign put
-dup 36 /dollar put
-dup 37 /percent put
-dup 38 /ampersand put
-dup 39 /quoteright put
-dup 40 /parenleft put
-dup 41 /parenright put
-dup 42 /asterisk put
-dup 43 /plus put
-dup 44 /comma put
-dup 45 /hyphen put
-dup 46 /period put
-dup 47 /slash put
-dup 48 /zero put
-dup 49 /one put
-dup 50 /two put
-dup 51 /three put
-dup 52 /four put
-dup 53 /five put
-dup 54 /six put
-dup 55 /seven put
-dup 56 /eight put
-dup 57 /nine put
-dup 58 /colon put
-dup 59 /semicolon put
-dup 60 /exclamdown put
-dup 61 /equal put
-dup 62 /questiondown put
-dup 63 /question put
-dup 64 /at put
-dup 65 /A put
-dup 66 /B put
-dup 67 /C put
-dup 68 /D put
-dup 69 /E put
-dup 70 /F put
-dup 71 /G put
-dup 72 /H put
-dup 73 /I put
-dup 74 /J put
-dup 75 /K put
-dup 76 /L put
-dup 77 /M put
-dup 78 /N put
-dup 79 /O put
-dup 80 /P put
-dup 81 /Q put
-dup 82 /R put
-dup 83 /S put
-dup 84 /T put
-dup 85 /U put
-dup 86 /V put
-dup 87 /W put
-dup 88 /X put
-dup 89 /Y put
-dup 90 /Z put
-dup 91 /bracketleft put
-dup 92 /quotedblleft put
-dup 93 /bracketright put
-dup 94 /circumflex put
-dup 95 /dotaccent put
-dup 96 /quoteleft put
-dup 97 /a put
-dup 98 /b put
-dup 99 /c put
-dup 100 /d put
-dup 101 /e put
-dup 102 /f put
-dup 103 /g put
-dup 104 /h put
-dup 105 /i put
-dup 106 /j put
-dup 107 /k put
-dup 108 /l put
-dup 109 /m put
-dup 110 /n put
-dup 111 /o put
-dup 112 /p put
-dup 113 /q put
-dup 114 /r put
-dup 115 /s put
-dup 116 /t put
-dup 117 /u put
-dup 118 /v put
-dup 119 /w put
-dup 120 /x put
-dup 121 /y put
-dup 122 /z put
-dup 123 /endash put
-dup 124 /emdash put
-dup 125 /hungarumlaut put
-dup 126 /tilde put
-dup 127 /dieresis put
-dup 128 /suppress put
-dup 160 /space put
-readonly def
-/FontBBox{-63 -250 1027 760}readonly def
-/UniqueID 5000802 def
-currentdict end
-currentfile eexec
-
-9B9C1569015F2C1D2BF560F4C0D52257BACDD6500ABDA5ED9835F6A016CFC8F00B6C052ED76A
-87856B50F4D80DFAEB508C97F8281F3F88B17E4D3B90C0F65EC379791AACDC162A66CBBC5BE2
-F53AAD8DE72DD113B55A022FBFEE658CB95F5BB32BA0357B5E050FDDF264A07470BEF1C52119
-B6FBD5C77EBED964AC5A2BBEC9D8B3E48AE5BB003A63D545774B922B9D5FF6B0066ECE43645A
-131879B032137D6D823385FE55F3402D557FD3B448685BDD20EB05D43275929A6669ECA68789
-24767868F73576F5C7AD5205DB1A9C1A6A56E36D0437E383FC6E79BE0F2101B1AD492956E006
-C5934DD525B295FBBEA391D875577EAF43B92C5F5822650FC8FE8DEF32AF1FF2AAA741BBAA5C
-7188BB482EEA9FDF154A99C0F963A285AF7A9EF2135E96E03C69DB45FE42DD483C44A3108698
-9BEC1150DB7C28825A741B352734823EEF88BAF840E7B1D52481A9EF70D997D4F0F23BBEC3BD
-2D45A8AEBDA8858D5A540BBB6A171E034C8281DF7815CF71056426777CEBC1FC996718DBF2DB
-D4479D3CFD6CB606ABEDF431767AF2335DB3E51E4D72F918EE74E4D2302A453C6E795F7E32CB
-9469221D023775B65F222C5763D872740005B3892C46C56E3F574F673475C1EA2E5FACB31788
-90A907CDA254570000144F959AC399267BFD09F3E0A73F624ECB51FBD7A61050F5DFA46DD148
-8169DAE84C82F36935E031B424AEBFCE68F552592FC586E4BB59F15BCAF50DDFD352180855F5
-84599225227BFE8C3B719DC24BADB65F8283F00D4BD0E597797DDDBECA3439348036FF2A35C6
-CB0BBBBF811AFC9ECE5846B0FF18770C81EB888E15B5E7FB352CC4712C71136CBFEA437C74E0
-1EBF44B394CFFD09788EEBEEF1F2EA6EE4F5DD8D267B544828037C024D025A57F598F385E8C7
-E19B26FF09D29C9A310021BDB913DC84BE33B1DAD4A1A9FE12C164B3EDC21D0F2A096DEF8582
-9867405A16FB047372DEEE2AF1158ACBD7FE39F6CFFB23B0C52389171023D88B78E42A12F9EF
-C5DC63575A9A53BE7E1C06835014E8B7DBF60FF0DE6768AE11113E9F71DB809A152552F2DABD
-BFB9FA6290D8844F6DED074988D0135C859A0A499FB000080B4FE698078DB71DAC885743176D
-24987173456970477F9FDE14729097CE6BE8E9ED49E072947C2BFB87D55140C6B998029BAB15
-CCAF1A9D52F8488CF6526BAB9F78D33EBFC36D8D4A120D6DD58A85F8084B022A4A5AE3780C41
-2F543739D321BF4AA81F4771E704C8D751667FA86D10B5AE0D52DF1FC811F6946F99AE5F4701
-3D168DC128E6F4E2B7F5BEA5BF1DFE647C24382C0DD715A469807745878E68A630E1FBD8D0E1
-B508B3C1A7E946BE27104FECDA6AA03190E649B82A04D4DDE129C90A9993FE4104F9C6880972
-F64C3D880404DCA79049DD4652F9335F012856B0BB6D988CD2C70045CC711460BD54B4F1B045
-40B75685457E96A70E80CFA369A78FDC325C2465CDDE87F62479A7A15F63762D70945BA371C3
-18F4D5F6FE522C7387CB06DD961AC6CF46129E8FA9D73B318D0F17D46C5D9C8A1D3A770325A4
-238D3BFF0C31B301D1917C6D2D124BD25302B4E7532EF14D98475550344FB4EFF98B76342C7D
-CEAB6CA3ADAF8D741AFED62E78084C345F48A70DD71E224F627FEE3E06C1E5356BDD927409C6
-2744242A4FCEFBC7E05C9289DEC0485E535AC6BD08D3925758BC08AE59AB7328FDFF16EDF25F
-1E6B8ADE68C72729623135CB176A6BF91AF453B02E2F9A38D38F0701DA96D77F4E3E12714F5B
-2088146FEC99144C69EB529E7BBBBEDC42F0C61899B390E97B5EB811699D8F3978966B6E10EC
-BE2304C76BEBF359236B16B24584F6CE87AF26A1A01BEF4DD1FB1A03A5F11BD55F6F19E1D26C
-E2CF2CEE6B2506CDA40FA5B202D58A89583B333C00161F02F1C0029E4C6DF383F0F19CB29440
-25E7DF30171E94977DC0D10B7D66C6C71E7BE0A83F2B65AB9B7953222F2D85517A721B9E0049
-65158DA181EA49CB5CF40FE1C114931E9B9DDCB145CF4ABB5A87793970D5AB9E7A5B017D4CB2
-E814B3269E6E3536689800F7030AE72BC4E041BCBE8F6F976DFA24AEE5259CA01AEFC3A5A85C
-52E56E895071B1179CBEB28A615B2785433B6B092F079DF46573F3619E18F83B698CBFCBF58C
-E50570192005F41E8BDA96E77FDAF12A4C17D01731718013A030D0500752725148BC5B9AFC7F
-C3CCB56950D4633705F55943B737DEDF8D589268432F0669AAE0ACDCF4A779AB5B6FA05B7297
-CBCF0D29E24BEFD81A96816072584241BF223E62D6A0D9020016AC07385634789602AEF5D26A
-9D41A8444F139AF6B014AC33849859ECC22B128B806C88F38E8FC0ECBFCFA252F6D1B628366E
-4D8626FAE7F295C6003569BA6A741F1ADEB06CB7A36E60221219FF8D1787CE7B30BDAE28FF58
-B5B84086C6A1D95F4C081C813DEF5BDFFCCF9C1A9817FF712401F8B96336CD71481253AE94DF
-0517B46FF9F203AFE6E060A8AF4E4A161BCE19489ABC9D73B6D0136058F6C40758395A961364
-F71441452B915B138CBE4CE4765EAF057BD9D9BB23A210D7342B428F289B7320FD6B1C4A49AE
-2AFBBC41C59DD7EBC5BF7F1FAAA5011FA076EA1BFFEC62CD1FA5842F92EC54EAEC97052A742F
-61CBC48D1D65FF99070AD69F308F991D63BBC2EDC247A258A1EB483B3216D352EF33C97162C0
-1BA37E5F708D5C98E9F58DE87EAB7EF4A629A8657BD62A20478E3C5571218124AE40EC7A6D84
-A717255417324463A3164F167DE30249B11463B67F570659D81B0B2469E94B6C49ECDCC38934
-12DE84FAD23B55D03EE0CB087588219BD4369135869620E7640B9EBFE4A96FB1A7E605321269
-DB44A19EEEE712BE877FD8991BBA52E0B433478EA33AC5662AD7AE849E0D0E2BD4E7E393D5BF
-A69A78C3E7592B41A45EC7F668E134F0F0E9E986E5777DF1C55BB5749A95ECA1E339001A2ECF
-852F20CEC7159EEF6708074B69CC9564146F5D92EE9F1EE9BB4472D58173A18F9E288D8279AF
-550ED5D00294A1B79D236F3452943537B4AB25BBE03BEAE2E9DCF6C111C644DC77B9DEDC17EB
-68DE86EB92FB1AD21C08745683E86B10009706AEFCA2FA8FD1F28CB12C160E1A9D09E7048447
-2CE9D82F45E46899176A1A2AE42522962FFE950B60033AD2830F71B2AEBFEA6B2EE3A422C493
-83847DC7F20AB5F1BCA75D387C31257E46814A5BC879BF9C629835A8ECE7234EE86CFB8A8CBF
-34A3C16294728CB86A3F56B34881D7FB6B7E71B5826903CF184A585506F649B54BD8CD7B57F1
-2DCFDF5775CAD40A506A278ECBA8A9E83B9D9084BC2198102CBFDEC8E5DFFFC2327E70A7BBDA
-BF17AADF132722807E4E1686766AAF9D0FF18CA5289C3FB3CD1648E9B9997FFE8528BAB95206
-AFE577C39AA260D2595F565E7763857905683AFA4D9EF50B577E2801D1FECA616551A1FA52D0
-88FB4F301698C270A0FA25EFC3C336E9E822F71D0B40F953B1014EDCC988BAE04A39082D037B
-AC9233155FF81970840215EA15F3A5B1F8AC344BAC12722B4430B54A94C01BDAD758325C39A0
-5990B1B25552F8493A732047278214A9CD1760D9372667C9920BD72235798C11CE7D9A0057D4
-D32497563687186B23CBC059F5A41C81054DA9526A224FF59E32899012E91C66F0E9F758C5E3
-7B824B481E804A6DAA304E2C1B868874BEC2DE5DCB46F6DB56CC3658665A49A62DD30FBFE4E6
-7C5252FEBFBE47211FE2C52481F85CD9D5AC0A68C3A6718CA08941394ACAAFE1F20A783F8B22
-F1457B6985C30F07774CABBF4742CA93AAA9992652C9D927DD26D195CFB6F16510D613FB690F
-47F952EB943540AA1373C5683B79C939E263C34F1BFDD41B93BF241FD0ECA57C7753F8711790
-455FE7134F8D80983B8DD27311B42B2C02EAD3A2DBD5F277047C5890B34A7CAE0E6CD9EA9A7F
-30471B6451C9915FC10F9D596E0393043085CC34AF179A24515575307492468A6AAA6E4F11B6
-75AE2FC03FEE10335226F96A496BE2B15C22E86395F0968B9A2976C64339EC68C301A2767C64
-0DDA7BFAE7BE560BB0A54A45254A63E13DFE271532191581DE1C48CD2ECB411C7EE2D5E0C001
-7566B16A00508305E252466A1FE64DE9159C409049332DD2B416A8A185A8BC48D3338ED5CD74
-E6F375BDE9C2FDA512E1397EC313E10CBBB94733A8F5AB2186B6E42372550611396DD990AD25
-645C2F3011A0A60A64BD5B54796627D3F0CFD8F25F2097491281F38B7BB3021186D6D80B1783
-44791E5B95C8B3099E5797DCB5209959519B3579B09F5C893D5F4B2C3C03D0C7DA1A5E1CF556
-5A6C92CD6134F8C61184117D31940EAF23567E315A781A3194263DF61A70B67EF8862113D7BC
-9C452DB3BB0C1A187F13736D27F2C52CE5FB5731FF7E7B123DC523AF72B3DA78AEA44688A3FD
-244AC698D2F763F24EEEFAFD8E2F97EDD2A08BEAB39DF4849D294159DA5C6A760ACCBF5FF9C5
-3CD8AF4A69AAD26A918406D4E6970C9312EF8C24070B95B3481D12A4E2621236D7043092259A
-2E90D3DB0C0C25DF1F538DB2C995DE3902115A839420317ABF3D87284238AD36FBBA50B4489E
-A3913B012C0EADE8AF053FBD76F73A999BD3DF8BA9A6A6A9250F9286FDA2470B4894AE8E0714
-918E837F3EAFBF7DCA77085E4EC1C3FF884D266506265E614DD9893390EAA75B7B77FCB751A4
-89F847A3346FD410BCBFEE0FD2C3B7146A761ADD9B940AFCE2FF96521D445FA781684EC02281
-023F8551CC96BF3B7F83AF6C10CA9830131D9BFDE75E383CE7EBB186647BA8A919E409844193
-4DB990D74B8357B9584FE95A56025F2EB59E3CD096A89EDFC259907D3FF58FE33056AB60B7A8
-8454BC9342CC610488062EEB9D4A8472F8D8EAC7B73CDD32ABE590C4DB052BBBEA2B7427F15A
-5B21B81D8193580B845FAA30349F106C60E2D75078BF2E4A4C41556469E13ABFF7BFC71E0719
-BF724D251C47275BD5FD0C4B2AA46219893A5BE07AECAB2ABCA504AC82BF89409472C2CEF60A
-93316CEAAB7B7632FA167F409E4E9651648498137EBDD6F1ABA64C031F4414FF4E7713B244E9
-1D2102211B000DA124A4992E1A09F8F6A5B7C63482F6047BF164BAEF29217D1A7DDEA3B0BDAD
-B37C519E3B810E4F4B9BCAE5FF424575D6D424FC34A0392B8DE19D23370A4996790420FA38D7
-680A7395CFE3144BB6DF6C23B62D1BAAC0EDF6AF9DE50496D748B7EA064D4E72CFABDAD7C67A
-9C02C0D28B8E2B3BE8110B6615BB7C41F62381DE3C0C92CFC329A4E1D5809510BC4245A4E6C8
-E02EC34FCF07B0B61CEF527AA90293F5FFF3A1F9C909B188DCA9B9F09EB96A6D8D0BA4858DF2
-7B1973BDB243B30B975FE3EE901E4A0D2383D00E668BE1A48E5BC90B6FD6171DA3FFABBFF374
-FA40A6A2950272470F2E9D566C8CBBE298AF4BBD898A7E33154B276C44CE6751AF1ADD9601F0
-001B0E94B92EA69F705C9D5834AF1E3B5E10A41E4E993C772D8ABA0D9F1ABB3DDDC7C66829D3
-0C24CB2B21B8F6EB293BD5247BA7F3D234E232B4FC8134254AB028320C2D9A8F4178D21F4E18
-FA5C992466FAFFB251083FAB8599C3E386C39D81DFA2CED237A4A833051514D2490326DF3BFB
-D53C7C06E1479962693F24531A70B4BD31E6545A6676A86186F9208EFD441BAFFE522297F697
-D99F045DAF523BBEA8F5F661D88B0B0F8B14A584FDA2DE01787F62EB331BBA789043B498010F
-AAFD774ED20956D39FA7A144585B660EF56391F86AC06ADA2B4640AD1A4C95F18B65AF1A082D
-8A1D88530F2961AA98D1A0B9C54B7EB711A257C55F05EABE0F4E3B64DC6B5C96DB24B1F04124
-5E75A98EC2BA68A4DCC074FD8465EBDA063D6B743E5211E689CA6D7175195CF31B56C8DB9E6B
-76E746C5405FFA3FF469C6304D6D74D36C9A7EA349C718A69730452C5A74F9C7305C65FD8279
-EDDE0653FEFB3F0348FACCD1E0F898168FF2B6ACB98CBAB2770C160663FA856FB6219D37E0B6
-6629D5CB7F000EA0A2E4124D5D15C3A912AC92824E0FA77132BF12AA9A363D1C8D387FF5548C
-333CCADA7FA892DCC8939C07FB7410CC7222E0DA55B6CF27C9EE928BF530A3EDD37476EBBE88
-B1805201C6542F3464234F1B10FB010BF4244ACBAA32696105D8E5E312AA40E07C61C3AC4FA4
-00B9951988710E5C22656FBAFD7B31036301AEFBDAEE61CF538B56E8E748B5FE6BB0150948A8
-171849D5D271F8C8D50C590A61D9812449D23E14D572C1CCF163EDFDB66913356F2F45BEBEB2
-71F11CEBB141F1F9043C5AE8776A76B5852F6070E49A5C4B0FC33BB231634C04F2C0AF0FB4AA
-69D9F39D48AFC54FC26CD84F0E636285845C4C2DE175C0A2DE2CEA1B9164C7AC2E50602B29F3
-3C9D5244360C3BA48EE885DA871A52CB17F15B1720224E53A17136BC8268509147736BB58F0E
-A707648D49164FF3B82EDB183DABD30FBB3D6F8CE492AF3F3D8B65B4628D060B4CAA2C3FA3F2
-509EF60AF12801A90677A219D1142333FAE4FFD0CD79F1ED1E5A7499B3B574F08C0B7DBDF8B7
-20CB2F747EC93A67A96255B58E8A319F017C49BD1A15122AD24BD6CFAE8F0C939E4C145E730E
-9E844996F41DB173522FB8F920D669F3E915A69639DFD753562034713999E1C7B8BA5249011C
-449DE09FA83B3120F2A00C14292847C0B99C3CAD6072D9A6BF75ACF312A9D0AAA87A5CFF35CA
-17002BFB38A8949B5290DB993C73D0CDACF92F3F5CA877CF844A9743D901DE72CE656117D86D
-DD16D8FC9B938E753DFFB38C7E490FA559FBFA2CAA8A6115C3AAB9B6F886FA2CF147E06745E5
-350F9E63455BEDA6B510C93B146D40C2E3B43C51AA8B6F485709E9EFCF64B654D7DB18997326
-BD6DE4B2D4680A909739DB569A091D867A752BC7472565A48B6A3A05373206596324B6E84C80
-0B2365571BC648C411A379DBF68C6C3F66EF66FB3136D11AE7F8E40EF0222DF8C53D5002D99D
-6B2A68CB0D0B67FFA015383DAD9AE46785B4ADB5072878A64427B88FB24D7A859E50579940F6
-D00FF46825D78A6447EA4FB516A6D1B137B9DCA3A2E27E898375253534D7E299DEC079EB2586
-C10EA4A9F9F3360708032E5D7C18693477519BFE8EF9ED5AD5E64D49B3E0E8F4178088C37748
-DA11C428D2C30EA1DEF37F7257793E8FC277F90B77F3B663E10CEBFEAA88D4AC7AB1BDB6C2A5
-B1FA570D2DFA2C16B2F4EF4385685C46A68C931B31BF590C4A1CF181A9822BB75BE211DEB86C
-DEF5B8B4BBF6593506D9195F03A3DFD3DCCCE0FE14937902695C4F7EDCC7BAB8BAB96A01649F
-8F65B9BD008DF7FB0EC35D27A0CA8BC52FBD08EC15DCCFFC3848AB82BA1D99EA34163E2F88F1
-1255430288D8DA2C6B55CBC0C46477A26D41E9ECE4450F4BB390A529179F428A2BA700F2F087
-4D6692A8A0FC94895A43FBF1175833BECDAF0B79DBA59DB42742F9721B0F37212A3F89081292
-9BA50C0CC258FE260A5C92E46CE2966BC757248DD7B160C95375234308D78A7EB6CF26500740
-CD71F815AA02DE2E510E26CD1307787A81948FD8BC4777288D890F3C4E5534F61AB9C2379311
-9D8D74969E785EA499E8D3E5E6A74E735CF35C18259EF538BDF644888A4343E129702ED760A6
-F6FAB37B08E36D23F621E1C5F2761545D3DF0ED1D4E0CE1F90A831EC00847CBC84C6515FFCE7
-DB3E20DF8D417E7C3B5C8B236729460467B77EACDD7D4F08FC81F4F1D0D24ADA20DDCC462F44
-8ECC5CBEF512A4B01C0AFB44837C8ED9374FD3BCF494AE0C62E059DDD8CAFDB027073C5DFFBB
-1E19ABFBD050F4DB4B0FBE6BA83FAC2511F2258A191C10B4C6814AAADEA19F166C49CFF4E8C8
-813830E7E5239DF222D820754925FBC20A16323A1884B5CC7575DE4ED3C9859EAB568A64F56A
-C2A4EF869344C7F83A0AE49ACBF14D15CEAC5DAB4472350B3E6DBB64D0B9C78B4115D8ED4C96
-20C4BAA9B3C5BC9DEF182F17F14BC4B64DA17C74C82914B36AEAF4E7393576C9BD955C043396
-485BFAC78EF92F677E17EC467211AA4CA9649A2B8900F82456449D5587AF8DCC857B90E401FA
-B505A6855389148A3CCE288B31395DE818BFC13B6AD107B33B053E90F4960F8BB6D2E1FC783F
-0582C1E5E769FB9ED61E54D63E02FB9DF4F613C16E2B141FD4843BCA7541CBD5573D80049012
-6D01D18B2204B53CC5177A120E6B92F0BC6683AD3D29586A67AFFB2C6EC35EBF9281226D3FCE
-00D2DF2556AD01260D696E908B4D86567B057CD2CFEA6AEA6D62C343F964392F8D2060D7440B
-523F337840B4E3A13A73FF1947D47C8AD4CF24A43BCC7CE5809596F7DD5B635AE6A7111DA61A
-28699930A94F47AD119E7CB52966C1508C1B464014092D9AEAB28F6B639C45B93267F9EACE63
-B21A735C85A9FEC1F22C1D0C49F46CDA7DF3F3F56FEFD291A0DBC0144057B48016C6132F8BF6
-71695EC2577ED5652917AF615F81617DEE39A3AE5A7D504F7BBA0A0750248B6044730A5B3E6C
-5091A8558D45961E6986BB424ACE821D43B59EACFA9CD32F2D815231B8D5F14B6A44F27F87D1
-C87B8E463B33B0C65F5502750FB0EA03C0A92D9B2CDF96D851715A16452AD1BF6F1EA73224C0
-84F680C94925711C5EB9B1C7975295A27B45BDC99C1CE82324634D0DF0E22D41F2FCDA573FAC
-794919517154B60A8788369E80E8E5302951BAE4402AC3531DEC3C54B43E2469F8592B933730
-B2502AACCD64372E2363DA30A9086C507456005C05ED7009E100BEDDBC0AE7FB316F5F4E44C2
-CF3ABBDBC7863F7E4E350A25419FD1D9DA4FEA51CB9552425841F1FA53DB64CBFCF90C287758
-178DC5F09EEF863AB189ADD25A2E82A76233986D77610503DA88F95BEBA9F4D1EA1EFCAF856C
-546151DEC35E3EBA222B2254FF601EF21B10DD0A5BD4F28521C6D85F14AB6BF38BCEE4AAA0B4
-E8641761080CA22D22898B7259AF6C7A9E683F93A40EAB4F2B55399B42233BF426E7C365DD9E
-26E9D9F8BF411BB832F601AA95F652E017EF7B35F862EDE879B18C08078BC7213B2C53107B90
-1C335EC116C8D143C069D09163F2D0B1F245E058F9E3578C0A5DA6369B6E96FAE9DCD5B0E851
-8E1813CDF03CA784EDD06E2E2B66E46364481A2972B95D7B5D266C6A74103FE8DB80D9161C88
-D278C29307A3675903ADA6CD0940A7407A2E9E296384F4D68F564175DF8B116D0EEBBFC31121
-75EF8089697C234E753ED158BB228331566DCA24BF3B2EE44C96F6C0029050227F9B21592BA0
-B8F7919A9D12741E15AA5BE0D967E4AF204BC62B84D5F0B4309C72BFD7EA4FC4DB7F8542ACD5
-96EF93DDFBB8B573705A2E5F2A1F29FC3C832FD11736DD73CF7566BF34072B245691699D79E1
-4D221DD4E696721F0C278883C558544474074C2539F9DA28A13D91EF77490DAD7D25F302F8E6
-546B9CB20515E78FBD5A37F58BD0441DEB91C9C02DA246A45AF0FE09037DEC07E768D406AF1F
-C3C29A7D22854096368E0489784D4D03A4F68023BC98D46FEE7A3E906CB4299E6B0A56F7B982
-D266ABF5F0FB87EEB5CE257ACA46AF56D7BE40D9E380BA079D8B4BB3D8A832943C4FEC71EB96
-46E041BAB1032EA0C09936BBE6BF11E183FED01D2D5E248B74D4F2BF0D25AD4D6BA078542E0F
-16DFB9CB0B4EDCA3A93596C549B88B54489FC19B3C65EC70D46C3FFC73210935A30A67B4CA36
-E3464BE036E79D7C341D3FF41B14B08FE883595E4F8DC365FBE01CD97146CDC69B2AFDBD7B29
-5864EC49F39E11CFEE53CF06300F33CDACB435465A75C8B72F10853C456ABCBF3FAA5F5D67C8
-714263758F05B3E9122AC115DB7CB4EA3138515254A1E7FB7081892CD4FE64BF93D23BEAF6EB
-599BD4C184501CABD8F9045E3998AD013D39BFB2442D835BF6002DAB27701477A6CC703031A8
-1101DABC754429270C4231CD7EAC9356B2ED7475CE26F3E7BDE80C7967414779984FDB54C2A6
-B3C55791AE9E04900CE9CF2B52713C853F0CF9B3DE8AEE6E1F7123F1C77B6A999D21BB6CECB0
-F9EAB93872229AB5B9C0EDD3BE1BF86136F36DA70636E680E729665EE9E2FEB5AB8F42974E17
-F09C4703FFF019BD78AE0B9359B595481734CFA2473DA1EE92BD4277888606036DFA122C4C00
-CEA775AFB73CEC2FD339167B3083B908E7EA98B9CAD205381D0AF0E2E88F6A41B01A6E935F6E
-E6A4B18A52A5A3F2925BD4ECD0AC3C5C34522FBC28DA1412C2C2127C80588CF9CAD09A049F51
-21529AFC1C952179942F74AA38304AE3000568D88543A1A2CC058140D1DB7D90EF2586ADB1BA
-952865A3388B6E8CF63CC32057C57DC7AA47380E4D8210679274D0B9C93E627D49DEAB09AC8B
-FF64124F250BFB88A331D3FC77470EF6068357F3F06E273595968E6BB6C92624B71688544D3E
-FB45A8EA0DC377B38137E1D9FC253A9FED235AED46B1DE133A2E0F90636411CEDEAE9913BBA5
-21FB173FF5B11B6D9AA5515D6E917FF5E8449EBA29B5560326F3D4738C8D3C30BB81EDD8A1CE
-87658C64B5266764329CB37E3026067A09BB0A0D8C2D69F6466AF62FA32C093C6A93C186AE9D
-8DBE2F2E377C24D21E80261933818A5BFECC339ADDC054299D73FC6D426E2583ACB8EF84F7F2
-77032426F541ABF434DFC52469F50538C73151035C4667FFDE8AE5AB356EE88B442F0A8831BA
-1CF2416E4E4A53DC6260A1DDC19AAD6F752069FABCF4579267A441153DB8F4FF6B3E45B261E2
-228AB6A3E8200FD271D0B4FFDC6A94C2ED3D019F4D546F8F5FFA46AB001F8FF23B7A2EB51531
-31AEA3E20B03A888A2DF51311ED903B872741D512877CF0DCE075772FA699F77DF8443E6CE66
-C6FE9774EEE9FBFB249A958CEA0CAA3C5AFEDE741F711FD4489A6B8B0C1C27B777C61BA32CE4
-A3520C9D74F872AC19C82C3638D5BAD7E3D575FBC0DDB508401257ABA66693CA82D67A66D83C
-A6D8AAF405B2F2E300B957D1CA92BE7B9EC4E76805F17C85C16680AA132D98B161066CA10DF6
-485EC84980A0E1A06CFC658B76F2C6AB4CDE375EB5C9B14788E5B350ED85895FAFC33BD4F50E
-36BA1994B58B880D11B19571AE72C332D26735E839FC99230FEB19BAB254FE67FE8B5FC909C2
-A54EFB5AC4810CE2062534ABCE2EE488AD716A4CDFE914E1D58DB608BFC92DDB22AD287E3127
-4499F6108A9FEC8DB98F13BD11F11C8D26E162FEB779DDC6D0FE1EADFA083B821FF57CA249DB
-1F347EBFBA279DE907B891E94F94FA8D36C6DF76875925EBE5338C3BCF42D72BB2A3EADEB091
-274F7BEA68EF82AC2C9672766D62C92375074A64E5E11D8A2EE4AD5CB42D910D8AAAB15ED74F
-EC3BD509FB91996EE9C3DD2A4464CE521F983788DC80A68770059239B05577DF775F879F94C4
-215B6B61E13627E11953A5446B4F308A302C150C2537FA4C389C703DDD2F7FA95DD3C0FBB03F
-9B2E47EDCCB07DCC9C6D81B9E9B3D31C3605AF90201087BE6CEB43FE562C6421336CF84C7F4A
-D93B698FF301CD68FEA61E75B7CF05AAA4D6C8AF98B36489AE94B0C7198F13F616E0366BDD61
-EBF4C58EA891CADF19E5AB442FC8A913D214BC6895348C7F74CFCCA17927FE1C9D82A8B87DD8
-268D12C6F164E7D076522971EB4A694E888DEEBC6E2257F69FE8B940562B17DD70F99D6397ED
-5F32351A5784E894F0ECC157CC20CF0B27C2DDAEDD49CEF40ABD4D2C058E98E31B648DE59887
-D87F2F4F1C772BB1F6BEAE51B8FA75F922FA69CE077D5A2985C36AC7250262F908348B732B5B
-92F92F24508D8F233C31E579B9816C7675147B6AA811996AE3C8C3828634127065F805FB6FD6
-F74BC0EFADEEC9EBA1E24B3CBF9202132717C2F3F09267B54E14D9CBAABE0DBD94C502A8FBF2
-61CDB130AA678A3C09B1D3384D681C525AEFCC8AE0AE8881259DEE1CF22F65E95B351831EDC5
-16A82FB0AC9E2968570ADF1D4B5406AC2E02BC9D9AD7F7F0FE3501C4668CA8CECE97E10CC4B3
-47927E9AA1B05CACA0D53E7CCD5DF47A699284EAC9BDB3F7E16EBCAC0E86B98DB7371B91B3B0
-CDD857EC9B09EDAB1DA2BF9B2E36C7AA8F3D3DFE65FD44B088ECF9A5DB2766366247AD1FB6C9
-2A6303B563E452B4FBAB1F682E74B7237C900DC5E521CC49F26C78F34967BF2F25164DB9652D
-AA7550549677CA45DF3FA57F18E3A0871C5E1F18715ED005E804FA4948C842BF04B951470A17
-5BCED84E6AE871B6E95DD98BE13AF4976FC066E16D3601CF91EADBD49A8EE2B97ACA7A4C9857
-8D84169F8DEC5A924BC75C08D141065E5BF06D7C2BF594E3558DC7828EE0125AA96068B9854A
-4964A1028C2F0FC93A95308CABCB1A9FFDB01DE59FEC5FB0388CC3776A40AE05E72E690BB1AD
-4C7DD6508D976979FFAC9C352339AF76A51E2615BEF6208BDCF5DD11CA07609B867AAF38F4C9
-0FBEC4018F6B13F33AF5B8B029C28BC4584C9F644F42FA5CF8B56719E09F4DF0F220D6C27292
-28A3738D16CF7C8DBDD56E3B29A00136619FD18FFA7A283F89CCC094C9D0921394D5778A5B9E
-45C5C5F26E10B19A8D2A9EE58B25A5A0B294A3B15A382AEFA89E7B79D3A5FB49D71A4BB0587E
-1138466171F37ABA47DFE4C3993BCABA52A2C19E0BFAB41C99C246C2228F655F8C30C5F041CB
-3429845BA00AB05C571AB6AF069ACDC33F00AAC92D5A73A4501887FED7E888C41EAA072A53B5
-5B35011FDF9AD810E8BB2276BB514498633AD9F1F668E13E577A1965A7B6A90C3150C0C310EF
-FFBFE82EC6319F16002B6E60A11B85AFB0FDE596C0B115BB66BAE0A33637900519A28CF830BA
-9DC6CB323623C1C3E40BE508862A43AD528F95F5BC683F1EFD337FF7C5CC492283A48EC74A4D
-5265E0908D14097F7B0E1BCD15767B60AE9852C08A9AAC69BFC41A68303EBE376AAA381EDEA3
-150C88A1D62162C3B7220193F7201BAEFE131A4E8034EBC64681DBF08A43C03899E0DA420D1C
-527559629E42725A6D0325D712B5924AF6F7FD371261D1783A00F50FC8232BFD93AE4FE3FE35
-6B62CFF1AD255CF1F65A0008874E1B4A6BBB4E8EE1FA63B85A651B4A7A802EC241310A962561
-A2E5A516E2D951800E15F53925D1037328521377B097348C072FEF9178D2E415C35107BA1DC1
-78D47C6799C1AEE518A2A4C52D2C7B24A1BA422D7FD64C33D9D53D12E839CAA55EE3CD2FEA91
-136DA0D7EDEB29251BE13142C0C5B473834743EC2327879AFD4D803D0433C9E0CD9C3CF3FE90
-0A8DB2250FE1791CB407C8E8C98F8D36FC6CDB55FE130BCA55D9FB219C5016D6981653E289B9
-610D246CE3A4C598C5481892062FD2EC64A537524CBEFB9BB1B5414E63F536A5CBB7EAFBFD82
-76EC25FF291DFD94369927751602D8E46D9A7235A1DE9466F1B885BD225DD0CB4A658EFC5729
-9295768309CE0080BE53AFE0756086AEABD59B508746B212DE9DAC70760D1D0519CFA21C4EEB
-01205D861745FD8708370B5C70DEC45A2BDE0342A7A7C643DA376A453C68C6355222B57EAAF7
-B9FFCF5E87AC8D95C0F0BBC74DB99E5ED335E1C51BFA7D1207898808E68A46E5B60E9BE16561
-7BD8C0F0B452D996C0F696058ADDF96A87CA83922244EF94BD582AB2F3D882084DFD5425BDEB
-5BDAE99778B2A1E703DA1DA3E77ECC2C067CB7EF941DC2084EF6B9174C060A1A3A271235D4CC
-BEF85DB35201451324BF12BDB3D91E2202061B54EF8BF0F8D7807BB30402AE73162FA6C40681
-BE99B543A70D5F0EF22D8989DBE6CBFE9E59FF5067D0F0E4280B7FB216B3DD3B7075973316BC
-940E403539B03F4C51F564C4016FBCC691818DFCE5CC2F3B6BEBA38D7F684257A2B0F484118A
-6ED7E19B73699FECE064C0E38C6B4154D5DA4B069EADFFD0B8B0FE811C76639A21CFBE289C70
-C74ACFE70DC916521C27A65F36A0B46F3DAF5F954C268FF64787DB31C98F760710AE3F929CE5
-92C8697D2E023E27D5A17AE819A0DAA657087408C3F6FDA4CFADC09D0C2D06DFEF232A23E518
-B6B67335161064B2AF0DA5CCF27C1B4BE2C92C480C130DD13D385A4EBE8DEC27A0ACC5936350
-3E4E82B86373193980AB295D7DCA9EAB0C712EC52B08992E94A3B892F06EA28AB8D63AF56922
-59B0DAF190895E559D549B3846C8C44EA4F99D90AA4AFF6096F603E2AF4D7FC49B478CFF690E
-FDABEE77711FFF1FCBCCACD0A6D7C454352429CFBA93BF52A985CF207D49E6C253BDE2E3BE08
-4F2E7B618BA8043CD54A879C1C614D27C07119883F9FF201897D8A62C7552910F03506237272
-A33987FBB80083FB16CB61B33CE2563C4E15A492D7942B23A78990A9D2D2BB9956B3169961AF
-E0404191726746D40EF3510DCA5319B1C0212B95C673138EBEC2E264666A378EF2349EACDE8B
-DD6B437BFBDCFE76829FC5E944DA80DF0754098526779058AB11C3FBA1F72693B4CD1F6F2AD3
-0ACE059E1B6743FE367AE0F207A2B2F1C8DD1D772073E1B5497C1AE5F1C3C7ABBAE1C679773C
-D99FD1EBAC6438D4EA8658232ECB2A4BDBDAD93733379EB5815844BBB6E23B4832A43AB19737
-A22D296ADEACE54768E81474256D3A470D7EC0FF00326EB1644F7EC35CC1E331A0283E7548BE
-7DE2B81E88FD5E78F8E512306FC97FDAE74DEFC39227958AE8EB11469AD9CD581D8A3575DE30
-8D8FAD75CE8CB0BAE718D7C668D48E2FFC59D383753800DBC1DDF9E8AF2BE7B0169DAE3BB797
-7BD5459F5918CE9E44267EF0F19064CD6C94C65FECE4D2BB228BDA6909C9F36F374728AD2542
-25DD16E2D02BDF03FA7208D6134BE5E4A312D3AD53894AFFCD072D2D6D9BB2653DA472F72873
-4EDE0A8C3D4B45E87CFE584521732CCF4F389D7A172912C292C65D4E0ABC96AFE293733875C6
-ADEFF17556D770E364A1298174BE007259D7600165ACAE541C0771F50E3663E371115B707967
-1B0A57F4D098D1FE937E1BC73FB334EF568A5A4D192EBD9EA036703EADBBE0D9214C0A37731B
-0FFB39C474FD48396FD9C0F5605DB6DEB4B962FEC907C5B724465FE4E9CA3E317205FB6BFCF1
-CC54C7C738CFBC09C3368C28A53258CAB99A5A7C5E992BF28BA2D78556EC6BB75B826B7786B2
-355E5F3EACF3B2A678B1F751E7749779042AE4A1C2ED51E03F317C8CCFD4E2F7A281CE6871FA
-BEB43FE6BAC804E497888EEB9021952FD04F36F1EDD2519BC52E18656505843150CF78111911
-8EA53E1DE5E054C1501704B5784E6F216117BC96D1B9EFD53545D4E6B6665080E6A912A4D4D6
-238E6A7BA9DF20E2A713253C5FFF84AC8FFB7B1838BEE8727A46DFB6CEEA64352E2C0917BF55
-4781D85054F638D37765873E7390158799DEEF6BB66CD40B557BE7956799B6B53E48C22A9E1E
-02CD726BDCE3C8CD7CE170CC673E7AD37F851FFB6B08635E3682E6648FFA8B9622A8DAA9E1F5
-F0156FC382FE422361703A5A039A1BB0406354DB736822F4E7CE7E19026F19FB968523EB3959
-ED97F6627A6AC3D69896AE1433D402DF465D0402DDD921B1D30AFCFC52EC8CED4C6501680BB1
-C8F4ACC42834E878C38A04791DD8DE570B816DE8A0A851BB3F92B484A2F90AE09046B687333C
-A942CA0FBA02E13DF446F99E152938A9073692DA83E18377882335B1A39CAC32FBB1A3B28E48
-1D548A380A7B4791841E6C3C651865E526CA70FCC247FF4CFE6202003B92EEFFBEFCFE0C6C55
-9F68830558CD2849BD733F2BD1679BDB649A5EDA744EC6DD97A5390023679D017245B6E0538C
-2821EF56F603BA383092527147CF55671A8A536606F6A3A6E55F7EA4FDEEFAC0DAA95C868128
-4A6E63A1F8A2B379DBB3EC32C88C8CBEA8630D70A7E84B900B5C465E9C71021DD647664125B0
-C76E460B7757ACD774A0EEA0FB78FB21D14FBE8B03C9FA28B208FEB7A8C059CF33D3DFB09CDC
-A24A639752F52571A46A66A41EF56BF2382A3D1CFFA1CD5174F67C79343A7AE2A6CE6747E166
-E9A627835B316D306ABEFFBAF0EC87BC9D8480E7D2B450B3FEAC0909C93F15FF26651806E4D8
-EC13B34B5C521C783FF316721E0ACA320B0F2C78DDE464D9525659C75BCFAB85A543CB416842
-DFA6C7F55D652262572D32F586B5E7603BE592979016EFE48CCD6B7D04A35F4CD0D6000AA95C
-34D1290304D052F9E438A5798E4D41DEFC23C13D74425595EB151A6BEA7064BE35467B99522A
-8E6282C9006CF7398D4A03A7AB8E7123D0A516925586D90D6A7DE5BD2641868243F57C491822
-3B18B1F6BEAE5187DFF426204B83A192A76127568623180FDE339ADEFC71CB9B7BD8B82F8643
-835BBD366A33CAFDC421F84E77F5CFF147C4002086AD05C623487B32E4E39A951DF6AD35F755
-892B92FB16F538755D9DB459AECD16F4CF7E6A53FE624C3F8ED48B77E44DB327F03DCE872B5D
-46121F7D9248BA9F71B11D822305C10F6B7669063D8E6DAADB7292DD4F9C7CBBB636C2EB726A
-EED624E55F6115C982BD26B23C5B29A8B58A1322C7153E828E71BC02E815E95A8091DE551489
-1DCF97E34C209B2A94262A081B4480F1E01DA307CC5A109B8EA3C99C20C63125BD6087584680
-1B7CDB0F30598E9F51A0B2660FCA868937CF99DE600D2841A70DA00839CF2D053A005CA07DC8
-4B0432D5CADDB390804F1B27B0D208A111D91C747F159D08468ABE30F125F2E9B4203BEDF59F
-4CBC915EB42C5EBCA5C718984089743A2B6911D2A2C2F3AE1EC6F88BFB223E07BC47459DB6BA
-CF2F15EA730FED2CDDFF86D84ACCE18CDBE7F1BB3F2E87CA07235985477FD9E76F68D1328086
-E043739F48D51FB898E976333DED533179E53976100944FCC39FD8D581662F26FCCC8FD1C827
-8F03A5D1BECE1519EA85FFAE0942A0713CA0C80967B2B18190A5F6BA71BAA1588FAE98CF324A
-3CC91F9AF6DC9376040340DCD2C76258048D0601609217B0B2A95318373249A0B9C3099E544D
-5CF173DC7CE3B17D52A04CDE07DC40BB5DDD3F9248904FC3D356B3CB0D878ACD77F9B976B742
-AFC97E2FECB823BB0481B234D9AD084BA974C4556E760A15CB361903C3BE6F478B881E91CA97
-48D164A1198C75EF6DF4D0FFC6CEF2E3C8BE8F310A412DF7A430C8651D320CC81FE13F7AAC6D
-EC0D61800A11EF73E4887317D6C574FB8F87B1B8BEF453EF660519E28A927B13DA4FFAEC0671
-4D1267F4AAF61255CE49BC59127268E48B05B464683DCBA5C8C3BF7F49A9D817B9BFA70D28F7
-127A4CD6E2904526726445E538DE5D0B80E48683E12D44054D3D92529093C273BCD69FC2516D
-C870D3B08A5D65975598A42BA30DFA5B655B64A1D1077C485740E8D07B2FE4C4DB9F7E6C0A37
-B25F32CA62913BAFA3D53E8F98BD53B90F613D8109E7B3062ED10211B08CD90E553701C29CE0
-67F8B5D59F307D75A0C8DA8E5F5D75359F5A0F757CDD580F4DC0BE5B5666085005F9F34A2F29
-767161900B7E867D23986F466F040CFD277546812C48D7CF55CECDD82E78F0956679CF23514A
-C2F9C3BC526B62A07C8C01776F93F6BCEB1A81BC375EC9579D1497DD967DEFD9F07566B8BE75
-94A316C090FCF462F21836632EE2D9A8D8AC5CDC0F808E988C6FBAEE5C2A1AA41C2F03034137
-B71B9580FE39355229A349D0E47F78A87B0DECA520FFA9C9CC5E297C9527C5CDAE118E717232
-0CE99D309996BC717C4B3BFE4ED92D1DAE8DE138CF7C7297F8985B2A057B0B1907D8B4AAA99E
-54D0C2165C79F7CF9B9E7624D80A8800C305615D77A3C7B2810884DA28BD9D8A26232996E170
-41916966ADA269249F30A32F3FA08DAECBAD630C34FDB12245E0639586A5EDA168958E8CFBFE
-C8959845ACA43409ADACDF512FA81A6C507B24C82B0B868B9AB37739D630357EDBB00B405072
-379A8A8498E9DDFE0F844817A3A336811C164BE516F9602E4FF5B8D3F34B674ECF9AA38CD50B
-B2B5FF57F0954388BCE6D6C0CCA48A8763394196F319E0087F81E5FD27692C801D87E2324C56
-332DBBC57085CBB4B576CCA7316D4D64A3E5DF65D7BF049D722BE450BB0FB81027CF9ED24879
-A13F532B03D902ACC7192BD16A6BE76BCAD51ED0A0EA34BFCE34438A0F77D33E22C826EF89CA
-6FF54F6AC257EDC5FE4A752D35193DA165A4A1599D2C90D9953071FCBF0A85C759498F9CEBE4
-07E4B3BF3ED239020472499D488484D48189481FD5DACB221EBFF00E4931248E25C202DE7074
-A575C79524D1370C689B29E041DCB2EAFBA2DC6C32C19231F5757A0B52B779FAB67F8689D33E
-9F764536BFD254D4C7464DB92A0BCAF1EE4E11183DC064004025DB8A49F1F248D94998895761
-1704EFACBF19B38ECAD87EEE4E455EE4B1D8C75AC6AB34D6552AC8F36426095DE38EE3EA2E03
-42219494BA2165179EBD09179C90C818302872FDD9F854B8BC10B28E69A9CF3EFF4C7B576FFB
-7E9E062BA244EF2B9A05A68478B06B50255D8A4B4E8218FF969AA6EDC6BD75C259D7935BE918
-66EE7FEA303034214FDBE44ADFF04FBDAB371D901256E7B9C7FE0A6F32606B8817063EC2CFFE
-8FC64BE24E37D47EE3249FA7C3D952C812C76608C5775F34A3BD98C559EF718BE9F9DD8ACB68
-AEDE64D4330D4F4708B8DD5C12E96C6FB8CDE4A462315E54B136E156A7C5147F14A564452132
-F585F03FAD1CB018E7E46F5529EB951E8167DFC02DA30C94D7B87E0403E0728995DD9B257820
-FB2D62252AEF4BBDF069F09B2D4A7A73B477789E88FEAE92BD2D0EBE70D6DB6F319EB30284D5
-070A2739CE08BD4D4EA71B63E7A32C34E52180ED1FF7866AAA64F374C4185D83E083A2C21774
-BA72D685C23DFE2B99AEB32C5621D4EE7A3BA6E3DC4A88742B26706C3E44BB102C230C2A95D0
-4384850B90CED4B9C613DA3D40AE962E98C005AE0FEB08A817B1E5B7D5B62B8429925FA3ADEF
-2FCD932AFBC326BDB321393648B985FFDFF93F4BA945D12CA630715014D741DD8EF53E2F9A34
-A4AB5048A42FEF44E0AE73821D9EF0B8CB35C3B2B30ED4D0FF1039E17257525F32F5D89CCC5D
-50DB59938A5A58D05A0CDC1BA2375E4D325645684745172B72A67CE3CC45BACD849EF3712E19
-4EA386992C5C36EE9D7A94A35A1F40860FB36DA33C85FC902521B0D936F7101ACAD9C65AEA76
-B89797CA3C39A59F152500F8575B01CE7F5653844B55498C039C14BC9D1EE480F44CADF07AF7
-D60BADAF8DCE92C9FFCC533F5B8AD99FD9AAA5B6A9CFEC605237B73906EC378F24CCF936C5D8
-048F5F3D63A1DEB6E0337DD99DF45C6FE772569B86356954911E71438BBAFEDECAB2C84CB414
-2AF5F7754D8810AEB3DE8B5A1F1956F87CD9AC613368424ED8D7A4C6ECE6AFE44682BCBCB29C
-4DCC8F1DC4B26F684BA5B2731792E55F1AB5DE1F7FE17E5059673EEE6704F75A192C393989D6
-1BFFB8BF765C4630737E01249E5775C7EF5DE483634FFB046BA2A538137BD9C3C69FB4E58CB8
-13E40E77CB58494984759382160B29431A99CC68DAE094AEAD53B0844EFCB1F6A35657CABD0A
-7C2000A536BE7790BA2BFA6759FD377FE2593DDE5A3BBD4F5B51941EF56F7E9A6C9D16FD29E7
-27FB0BB7976DCA2C32D8CCF8D9055CB7D4639A6510753D6E2E73E7A6CAC9A45AABBB5D5E0C6B
-810D5EC0294CC293D2E6AAF480B53AC9FE1583072017FCF6BCEF0249A3C3199471CC9A8F56CF
-1799862F5B97CFF62F84D6EE36B72949ECB484B0FB244EA946E91CDDF0FD8AE9FD5DB5007058
-487B66D146DE51C6227139FF2A40D808F4CACE4904BCD045F615C652C45F23DDA6D91CCF0B1C
-24523A217D05883B29539FB7F24D0C33FD922D3AADE5380D4A8ADFD1F4E5D15AF089E106DD91
-789E05298591F8479EC535318830AE46333CFAF75ABD01100F391F00F0D206622F2A608F914C
-1616FA11249BC934F23815E5F68A56E5D9E1A3921E3CE53345ADA25640F3EDA18539C65139A7
-3936D840E15D11803DF7A051E7255356900A27CB82FE1C5A0B6E9D7377CFFB7A7E34EA59BDCC
-59A563B4EEA2146B23BF9FA808A6A98E9C7334469702AAC294ACCB322546F1DEDC1200961D20
-A940E0DDB8A693F4EEA9085547C78F033A31E7FE9BE57EFADFD1CD5653B25604A3C24799F267
-D1C88415D27B07B04AF7E3E2E87975C36F8928FF0968E7E018E97B1EBE94DE49099C486303F8
-B0B4CAF401147B586052C99F6623BBA8139179492DE079F72B36C3B9424ED9B6A7000E65B5C8
-9B645A47F3BBB39D638F79F84BE4986AADBA00BDB7568049BE7E4DF1B15FA2F0E187E1235255
-F3F485758661379649DFFE0D465AEDB0C6763BCEBB243E79B01B8F5607669E55968C06E7229E
-63F2766D78A1B4AE181C360F28E892BE5006CD3D7128E4E7AF4B580AB1B11BF3A0D6DAEA8837
-68E5A2F8F5FE34A9B27AF6759C53053B898426673C3F3BF2CA60043B9CBFA6C3AF89BE98E258
-4F4CD1F07E607FA3C49AADFAB3B18B9DFEAEA78580233F32AED84261DFD3D69336B2B7DCC252
-64A38A1A3924C37161E432545D18C4323F2A3FC68DAD2F9611D0E7F1FE8628284DE7FD791B6B
-3AD12B6B3C34091F04175D3CC555327D073889F053CF0908372089DE3FF735DE350D4C20C577
-40E28E8DC6A52D7B8607A578D12BDE52B9622823E4117587D890F81DCDA8897400BB66388BCC
-2CB91D7A95FFAB2688EF0E09751D115F488189380D3BD9E4A326B282568556278CAF1C4F29D3
-1B77227573E9AC087C3349809E261DB0070A985EF6B6C283E529CF9A126C4E76D5876CAF0C6E
-13165C41CA3A915D4F1DA89A4E0B859F67834DDDB325B3E715206FCF2FAD9CC77957732A6B9E
-3BCE1A595684EEA956EE5B524AB10DFBF0EDC74BA24E2C57CEA6C6B821A716A73E19A4A7DC15
-B8D5B3E5F95EA5E23D5A342590893019836E557A9854C88963FBA367D1EEA5FD195A69912D11
-9D57AB4888913DB6FE4F0E58F88541ACE3B14F98C600F4ECE70C75E5C1EBA14E43DEAAB0B22A
-A70912119BF598BEB1B913658903076523BD04FEC0BAC8AB697C13510BA8662C018F85C2ABBF
-FC48F1660C1E84B5304FCA3938D8BBDEFE9BC70A4A702DBB0DDC10E51D955BADD33EBAAF7DFB
-2AD087E95A2E91F69E44F8192740C0CA4D10B94C3A6B3FE755E78165BB1D9B584C2061E18298
-87F8E086315A0A1589DF8CD8567A2EFF66877B079FFDE64F169086FD33F2A93492A3C507CE9A
-5BF05A4C10FA5DD06E721F46ADF3936BE8F73CCE6A0593710A3129569E02A27E62F088E8B662
-1B097C372C365218CE4D232878D89B4FD21B02CC5761676215C1935301CF69C84E28805E9896
-5DB98052ED612BD7C8958CCAD1CFF832669F0B9C09F076C6BEB773A794A32636D349D04C0768
-079331B5600DD9017C1F9B4C91DEDD49208839ABAE6970C9DCB139F0B018B815886EA12F35E4
-1B83E4D04047FE9C3B083CD952A811E771086811C02591A4DB4DAFB2399BCCD4C426EBCB7109
-71B4F581034CD0D5530C6B77EB2A1E533F514A98C95C1F81198AEB39EE2107FB9817375DEC84
-609C6797EBE41B857E76D4649EF3068D47A43EB6CC5E7E81C594BFFCFC71AE3E70666919EF08
-F4E417AB491B8D96056891595CA507D275B977B25F253B3754212C5188F9D1D0A5F5C27F64E1
-E63A2E3484658A9537169A793A7A968B7E3F87D3D697707BC84A06120D55207CDD1C12DD2673
-0E559E46D066D26B6DC69F3D8A7E9FC9B939FEEE772636ACDAE86D0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-cleartomark
-
-
-
-%%EndFont
-%%BeginFont: CMTT9
-%!PS-AdobeFont-1.1: CMTT9 1.0
-%%CreationDate: 1991 Aug 20 16:46:24
-
-% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
-
-11 dict begin
-/FontInfo 7 dict dup begin
-/version (1.0) readonly def
-/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
-/FullName (CMTT9) readonly def
-/FamilyName (Computer Modern) readonly def
-/Weight (Medium) readonly def
-/ItalicAngle 0 def
-/isFixedPitch true def
-end readonly def
-/FontName /CMTT9 def
-/PaintType 0 def
-/FontType 1 def
-/FontMatrix [0.001 0 0 0.001 0 0] readonly def
-/Encoding 256 array
-0 1 255 {1 index exch /.notdef put} for
-dup 161 /Gamma put
-dup 162 /Delta put
-dup 163 /Theta put
-dup 164 /Lambda put
-dup 165 /Xi put
-dup 166 /Pi put
-dup 167 /Sigma put
-dup 168 /Upsilon put
-dup 169 /Phi put
-dup 170 /Psi put
-dup 173 /Omega put
-dup 174 /arrowup put
-dup 175 /arrowdown put
-dup 176 /quotesingle put
-dup 177 /exclamdown put
-dup 178 /questiondown put
-dup 179 /dotlessi put
-dup 180 /dotlessj put
-dup 181 /grave put
-dup 182 /acute put
-dup 183 /caron put
-dup 184 /breve put
-dup 185 /macron put
-dup 186 /ring put
-dup 187 /cedilla put
-dup 188 /germandbls put
-dup 189 /ae put
-dup 190 /oe put
-dup 191 /oslash put
-dup 192 /AE put
-dup 193 /OE put
-dup 194 /Oslash put
-dup 195 /visiblespace put
-dup 196 /dieresis put
-dup 0 /Gamma put
-dup 1 /Delta put
-dup 2 /Theta put
-dup 3 /Lambda put
-dup 4 /Xi put
-dup 5 /Pi put
-dup 6 /Sigma put
-dup 7 /Upsilon put
-dup 8 /Phi put
-dup 9 /Psi put
-dup 10 /Omega put
-dup 11 /arrowup put
-dup 12 /arrowdown put
-dup 13 /quotesingle put
-dup 14 /exclamdown put
-dup 15 /questiondown put
-dup 16 /dotlessi put
-dup 17 /dotlessj put
-dup 18 /grave put
-dup 19 /acute put
-dup 20 /caron put
-dup 21 /breve put
-dup 22 /macron put
-dup 23 /ring put
-dup 24 /cedilla put
-dup 25 /germandbls put
-dup 26 /ae put
-dup 27 /oe put
-dup 28 /oslash put
-dup 29 /AE put
-dup 30 /OE put
-dup 31 /Oslash put
-dup 32 /visiblespace put
-dup 33 /exclam put
-dup 34 /quotedbl put
-dup 35 /numbersign put
-dup 36 /dollar put
-dup 37 /percent put
-dup 38 /ampersand put
-dup 39 /quoteright put
-dup 40 /parenleft put
-dup 41 /parenright put
-dup 42 /asterisk put
-dup 43 /plus put
-dup 44 /comma put
-dup 45 /hyphen put
-dup 46 /period put
-dup 47 /slash put
-dup 48 /zero put
-dup 49 /one put
-dup 50 /two put
-dup 51 /three put
-dup 52 /four put
-dup 53 /five put
-dup 54 /six put
-dup 55 /seven put
-dup 56 /eight put
-dup 57 /nine put
-dup 58 /colon put
-dup 59 /semicolon put
-dup 60 /less put
-dup 61 /equal put
-dup 62 /greater put
-dup 63 /question put
-dup 64 /at put
-dup 65 /A put
-dup 66 /B put
-dup 67 /C put
-dup 68 /D put
-dup 69 /E put
-dup 70 /F put
-dup 71 /G put
-dup 72 /H put
-dup 73 /I put
-dup 74 /J put
-dup 75 /K put
-dup 76 /L put
-dup 77 /M put
-dup 78 /N put
-dup 79 /O put
-dup 80 /P put
-dup 81 /Q put
-dup 82 /R put
-dup 83 /S put
-dup 84 /T put
-dup 85 /U put
-dup 86 /V put
-dup 87 /W put
-dup 88 /X put
-dup 89 /Y put
-dup 90 /Z put
-dup 91 /bracketleft put
-dup 92 /backslash put
-dup 93 /bracketright put
-dup 94 /asciicircum put
-dup 95 /underscore put
-dup 96 /quoteleft put
-dup 97 /a put
-dup 98 /b put
-dup 99 /c put
-dup 100 /d put
-dup 101 /e put
-dup 102 /f put
-dup 103 /g put
-dup 104 /h put
-dup 105 /i put
-dup 106 /j put
-dup 107 /k put
-dup 108 /l put
-dup 109 /m put
-dup 110 /n put
-dup 111 /o put
-dup 112 /p put
-dup 113 /q put
-dup 114 /r put
-dup 115 /s put
-dup 116 /t put
-dup 117 /u put
-dup 118 /v put
-dup 119 /w put
-dup 120 /x put
-dup 121 /y put
-dup 122 /z put
-dup 123 /braceleft put
-dup 124 /bar put
-dup 125 /braceright put
-dup 126 /asciitilde put
-dup 127 /dieresis put
-dup 128 /visiblespace put
-dup 160 /space put
-readonly def
-/FontBBox{-6 -233 542 698}readonly def
-/UniqueID 5000831 def
-currentdict end
-currentfile eexec
-
-9B9C1569015F2C1D2BF560F4C0D52257BACDD6500ABDA5ED9835F6A016CFC8F00B6C052ED76A
-87856B50F4D80DFAEB508C97F8281F3F88B17E4D3B90C0F65EC379791AACDC162A66CBBC5BE2
-F53AAD8DE72DD113B55A022FBFEE658CB95F5BB32BA0357B5E050FDDF264A07470BEF1C52119
-B6FBD5C77EBED964AC5A2BBEC9D8B3E48AE5BB003A63D545774B922B9D5FF6B0066ECE43645A
-131879B032137D6D823385FE55F3402D557FD3B4486BE465959B1188F76DF7824C135A7FB382
-C4E11B0DDFE856B6F34552CA48C24B57DD8448FAC257C4D93FB122E1218FD36B99B2D79DADB8
-A66613EF11039BE77816A7A2D6781FE0D69CECC36232AE87A172AFE095532A8C7FF40BDD4A5F
-CFB6CCFD8F73C572231734E0609C6743EE11010F6579820B367C49069AC1D45760C0CE8DC1E6
-7E2BEB72BE6D287910547E1B4BB52465B75947D224998223275A2C340A1C14C04328CC2935C5
-3EE2723F1DE6357A291E14863F6C1411477B3EA074052B30A4C1EF03AB324D678307E1CB02C7
-AAFFA90B73975378B9DEBCD3F6AAB626A361BE8068C60EDD0B2BE4A7BBF4EFA21CC23607BF7E
-2556C17B5923D7B0FC036D5FEA743C55F2BBF72E8A61995BD63D7104A085B8DD5A83D9077376
-A4AE8F0C048606BB9E84C640C52D0A95B35D48CFAFD6EBE3D6EB0ED88C1212818CA93E406147
-40A77C446B792D02945642014A5D55BA970EF97DCDD5B839F621A052D1085A1E668161D3AC02
-27C7867324EDE883589B39E5D152793009CB5A08EB6719A44DD01B6147E3436BDBB28215CE3D
-29A3B1DFC72BD942F8CB3376B4E12CFF53A793D935346CEB2AFC3428CD9313C671EC1911252C
-19DE341556999AAF7417AA5085DAA35EE71353FDC2C7D83EC2DEDAABB1B89B99E4CB5282159C
-BFB45A286F8D6BF3ABDEDCCBDCEBAF6994B17F1BE07D10F00CBA47F2FBF66C0F608F28260BB5
-BC3CBBBDB7BCBBED9C322336A2F2E4F954A05D4879D3D976E201848ADC85B763C354E8017EA4
-1A9998C520D720772255C69EE591BC3650E52286E6F7089AD632DE39FF9B40F725F8E6E02B74
-8929DD79B8FC1B2E7B761909AAF09F1F3F0AC59488B0D281D4A95C25CE830A5223D618A634F4
-71F737BB211E4B04D6B18D867CCFD1694E1586595AC2D3C5D7EFFCD4773DBDB8424C817414F0
-0B2653598571C3F5D9712671A0E8B555FCC8F9AFC5809CC7EE6131E662E3D3066E183D0B6711
-A18806B426429B44521AD3F2E33FA53438F676B958752A88A2416118602CABEF91C101C94A0A
-5DB4D11B913DBAC9C62AB4B754D3A0B43408AC835B5A5A28BCFABFD6CD9961FC83A21FDB2FC4
-8BF4D06F623ABFCDD7F72E7B2BEBB7995C67FD497FF0008940A6D6CB04A8C65C7697D91B817E
-D4C57962F7A49B2C878F09B012D9AA9B871E7A9400C962982EBCC9011918B1608506CC061ADC
-D18A1CF71E088DBD1AD5D9B7459FF3C6014DB70839555EA342E2F0FB91308981FD02A7143549
-2DF066176634326152BD0C93E4D83C8F9687B5FB6C349A6342AE87F7C15E711172939230DB11
-75B419575780C0A9824F8F4F9ECD9F2C2721BCBAF73E937BC57A80B0C4D7BB6049109F21965A
-6D9EA90B3072074E1C9FB3FA31E81754010705568382F8BBAFDC25E4AF9B1458A8D8AF2E8357
-A9FDB226AD41469A18944F78DE58E996C3888B75315D8F08C991C04C6300B4F111F5FF14FFF1
-2DD6C9FCA55B96142E2FBF4941087FA600ACFAA02A2137720766A431070DA574E0F60E0F3009
-D004671C05C37B1B45583DC371DD2FE9E6108AAEBE6A8CEF22B0AF23F68B44FC3E0CAFDA3AE8
-0188F3AA9DEE25B099A2C0C4CB327AD3275A7DC004790DEF9E2BB835642B81254A1B8CE1F59C
-F7BFC693E5CD4706AEA03463575E0F44AD48691C206B9800060EB2F80A050A332008C9D05832
-7DD3E42F4D046D8BC37244599046B30D65B9D059B55827EDB2FDDB79F0DC693743226EC0B995
-A7AE8D5E9BB55FD830753AC63EF8EAC1B18414F060F4613F868ED7E8D96F69E9BC82DA3DB64F
-939BFD11B4B64A5B2E068796C3E1CE3926B2D6F442408820050D585EC449E2AB037AF3257DB4
-3710F2BA8BC72DEFCAEFC767F92DAF7D91BFAB2AEF7FC7E98C8265CC5604F6D5F63BCBE954C2
-ED8A58D0E312DFDCD46B7CC55981DAB0AE803D65C3BDA792EC50ECD030D41ED9C19C5515CB07
-7EA57787A78A143F714B020B5E51790E3FCB235970C0B930CC0CCE33EC3851F86E0F8897E4DD
-BCE32D78876F1BF668971E0E172A3BAD17F200F9AC224ADC85198F538ACDD6BDA89D5337BD49
-1A19BA4344CEE1C0A2603FE9ABEAFBFB96E0402F1290EEEAF3C59D9FDC7242EC59F006146BFC
-1687A5081562629ABCA526344DFB24A4067339DC7D781F392B9135957D61E7C1B97E29A4DE64
-B7891054A6C33452995B9DF8C063803B20A5A85A8ABBF35B55BB4CA57E91916F619DD920C24D
-7E61D85B549BADBA2BDDFD0790EED49AF54A6F5F62DFF642F05F718EA14C92647A4D20EE0489
-2F15BC318266FDE7DCE53847E6F64C2F48E6A0CDE461DE6AB0A114B3A72ACC43A18458A3C9B1
-3AA722E999DB610EB0F1DDFCD098314EACD40104FE9869325A9B847B74DC1A6B08F55241F8EF
-BC6575766B3A0E3977F0BE297E20C6DDB6ECDF051C0C0DE64DEA23B02FC649A86CA14E0197CE
-B2CC6B9CA3529DBA5ED27A7713EB43E133C60D2A12BF5AB59A2032335856DE22893F2C03BFCC
-046A22CFD609EB0CB3284AF5F0BF9FAE44920C2BC3E3B667A40AF42CAA78DF97286DD34F5CFB
-4BCB44026073A80030017223C4A7F54B49E87478F6B72523E6869BFFBF243698CE53DFC56A2B
-12907580DBC8F29844F7B7780FB6517EC2692D9E6F7F0794E3F5D1D98C0CFB3395AB9D30102D
-95496F3A07D65DE1F99F986887AD42A5F8EA442FD964FE2BE2887DE6A249CA455102091247DC
-D69C6A3763C057BC6AF8E7C2C6140DDC5AC8545820400A6EB63CDDD1077969EFA7791F8C9BE2
-0C35415DCA882E2A177A51EEBA30B74A4A76674CFE4B0FB190F3ADB4EA053CB2134FAD92B66C
-81FD0CFEBF3EFC858B1B1343B32FBBD5C9FDE90B343D49666F7DB0577C85C0471B18A3680349
-0D0594E602EE96BC3400413D97A0364C1E4057AFFEE0880A97D0B07D4C72F2E0DD6D59811541
-BEBEAB66DED21D76411722DFF15EFB5B90BE6774622E29C7C3B409C76159C23CBB0A7D4CC542
-B7E52B6C20932CD32D6A8941A231AB9DDBAD60DAC9494F9F92CB2A04BE0753742A035974C7B2
-CC9C3C58A7986DBDD7490B4F74EC880F20E94DF05BCBCB29AFB8CA3E72B5146B3C3C19BC336F
-07366B99247C3DD400B2897334786069C5B5928DA1867953653629FD9016EEE4CABFC071C3F2
-FCD0A810880B26F5CFEFB5E23193D983D1BC6D954B292167A9729DF33590CBDCA442C7BEEA9B
-97DF822D87DA8DF90636BB48005ADCF821C7ED488D0678601E2D2906391FD595553F5F4093CE
-87CFAF9BFD25DFD3DA97A0DBEDBC884CDFD62F8A7DDC4C8D74EC44C6EA4C264FA9FB7F83FC29
-2FC6A85D8E5540D05B61610166438F643CB4D2C5C0CED49CE885E53678F0207579162CD53A2F
-07831EE409261B7F1DED1600C730DDDDF838365A683DA6A6F33E7A49016CBB9D94532191F666
-05CA66E3BF54590AD680EBFE92A91BF600BD4DFC19B1B895BDCE26EE4C8F632EC828F60E8FBC
-39B9B0CA053A00585100B96469BC6B3D372F84BF41417FB632D797DB504F449D8DE720531CAA
-84104F76F1FFC2B4550EAFAC8CFDC48E709D377972B9AA380D4C9457C84AE40CFAB5174E33C4
-8CCDF3509DEEE1422F7AD69CF3D90185C720B4CB0B82B5AC24019ACB847FAAB343BE85273526
-9F98C12D8CB5D78A9A21996B72A71E83EF16EA611DCC4DB16B16C0960A99EF50141DFD364AD1
-09F58E8FD2414978CA4C51A6E030DFC3B02CFAABC0DBA4012EE2449D7DA7B2970017ADCD8736
-1F43A7C84E19C4A04D610F57120A69ED4EC38D1B2314A66B048A07B268590DB0271F95AF6560
-A1CC21066F31B25E96BCD6EF687DC1D2D97FA44FDF5B1E3CDD84BEE32E032A27310A0D15CCE8
-BEB01015EB3E0BECB7810288324E25CD0528BEAE974357E02808DBA3FDA1DD5103B8A5C94F6D
-9DCF8FA92189D01DAAAE273EFEE45AB6EAC3A9780915272B8F7C3075AB3074FD1526590FC446
-FFEB68D49EF85420ED72681896D377CD1B228F3A54928AE764D1C2FE62648EDCEE93DDA16187
-B710B70D980C50116455019FFEC47D6D9B268FD9F6ABCF089B7DFD444B62219FC5808E524D18
-5A6EA9D8662679D82B28943E457FAD8A8BB804BAA0A26ED7C4604745E01BD94FCDA7FD47F593
-89CBC8E92D81028113DC6F4129C21A6D4C06F782400487E50F615502C1A0DA3BE5B5A168FD12
-935DC67ED163886FCA29A6987EF8D92405315AFF0470DDDD1483D198004ACF11833E123081FC
-BD5B022EDD1D8A2D77EE26E36729B40EE54B91ED653216806FD1DE1BAF894A5FDDAB7CD6646D
-6EDA39A1669712A59DF05F35EAF0C0E951717EF05E05D47BC0619872FA285BE4B0930BDC5993
-CD32F6269745563EA2DF9EE4CCE7714F9D0DB8CC65AD139F8854B3DAC7BC2ACD047F287C2C40
-48852CA95A480C82C9071A22BB14E57D47B3AFC6D2A2BC96DB1D9DDF663E94C7CB0A06686B5B
-06832690524EF5122776265FDC5852FD4DD935FEE5A03F17E05E71BA218E181150784BD37DB1
-7A6ED6BFA0730DE1FD118EF42459CE5B2305C0A2CA87CAE5CB5A0FB79BF68077A7599BCCB759
-E20E3D911B36A717937B094AFA50C3FBDBFEF20D9DE18284B3958731F082D2112F45780E082B
-7259407103A3826F1848EB7AE0840FB90659BBCB57124ABE3B3D60012D43699EC1675A9A95C3
-7E9009EDAE906D332FD771EB666926E18A6B772CBB532BF51482D92D1A2E0BC28FBD7246EAB4
-2BD41527F2ABD4E1358F0CE1B09A1E374F9DE128FD8A499C0067CF3F09A5314898873FA04F02
-B6CAB409B8E6EE6AF895313819ADEFD4906974BCA826D933BF01BB3498417BCCF211CCF67AC0
-41016C776676B6D550480C078897ADB8FB0CCDBABF089C74871B48CE64CDBDB0AB941905D6A1
-0203BFEA7830C0AAD6C32CA155844EE6529396E96FDA312E4C72E0407919D28C5BBE52050BD9
-2E5C4AA61DDE9D500B243D4ED8F30FB0ABF32238E4AB33529E8E9766D93AC8B62148C004ACF5
-03CB12A7723434D9B52BE24421C7DF41188289386C6F8DAB25434AC3AA3E176C297BB067F066
-830FF1C10740772590753065CEAA59FF2561A6683CFDC73AA2DBFA299A25A48DB7A02211F586
-249DFFBE444936DDC4A4D8B55AC00B6400A70031A96D6545CA37AC1863A9DCBC665C07A40293
-7A16B1D1B65C08F1812613119292CCBDA5D628BA1D240630AF2FD59145A7F7DCDE5C82159BC7
-EEF6A8492B0E3E69DBA6E45C7CF47CF7DCC96B19F57B5C9672203EED3EE9001C6282411E1BE2
-157C5E38E8A747FAC83D72DAEA38C96DCDC614B819F6E84B353B3C4AE9C08FDC940005F14206
-AD379A666D9812D032435C25A97631EA4D9234963F4183BBED6ED2B65C2C8DF6E43C4AAE5F52
-9DD716D28FAA91268FA7A366A6E5B1875BC54746BCCE65F1CC98A10300D5F000E5F7C67FAF88
-48E68ACAB634B5D90B8E4C69E1C2317139779388C3D13990EA4E40A3471DF1158EA94C1CA90B
-83E70F5435D8A9AF64E5669F6BCFB437C12579F74B14D9CBAAB148169A652AA60066A43433E3
-AF6200E44C20B588BD126BD5565BD38038FB332A8605811C8E83EA43EBFC12F9F1F3389904D5
-28B7832DF876305EE931C2934CB9C6A03E66A4F042BDE91AFDBD8882890AEAB62B8E878D4283
-65B48DE385EA6D52F020E6588D3D42AC7384FFC99B01E6218ABA693F5FFB00A22587A42D9672
-83850D058130862EC67E143E3233B4D145BE1F96CCE0C50FABAEA414370893EEF80C0A2ACAE4
-F61599E691E94684C30250257CBD657754378A83AC6329D50BB7E8C0DC46CDE54E13D094452A
-0FD0FD3C7C7129DD6DD9F7D15B8B4C0882A24C33D5DE5B80884E2CB81CA458576CAE48514A48
-7A94D747F2AAF4F0AC018DD254EFEBE96145A8D3C92F606094D2A4C73BE183B44108F515EFDE
-8BFD968CE20747E2EAA0BEF32A6862CD1D6666C2D276134CBA9AE56DB71B9C030C87964F0B12
-BCDF85FE8CE5BA1AD178EC4BF77FC769E90A0D46A88A9AF8C10022DC7D15E4242B9E28B9E65C
-330A6349A51FB2BC14CCD4593AFFAF36890AAEF531585F6DE6060A8A1477BDD5F6D1E70B0F2F
-2CDFF62368AA89177347CD20BAE4A96FAE26F9B4524F9AFB7109D57C10D2F576DDD9C1FC3AAF
-DAACC1CC566F9BEC007A682F604279280D9799D94F6F23538C63369CCA480C9ADDCAF94E8E3A
-CF47E4DA69039ECE2DA0F61A12A6D007F7680425D0829E67121366E8B4AB5F911E3A04EEFA86
-E366462F7A90325D58416ABF8FE464116A33435358C525038BD3ED1032164F6BA3B68A66745F
-73758DD06AA5693C34205F5B33E2681AF1EDC50493764FFC0FB42224959209DA11620C0AFEF4
-4C7C328DFF7D2712D1E03BF07034637B4459D8D2EE228D01E6823C9C477CD234DCB6C929EE49
-7128C73B5CDE048270440A9827B8282FF0B20C28A6635B1058B55C6F5C8809C6B9E3AC4A29AC
-41ADD9AAAB2AFCE31F15B89F8210E03AA78BA8E63E4E3864C2CB0A8B577E5EA6B3967E0BD22E
-C08BC53DAE6ED470BD049B7CB53FC213B8E2B606DD507111B221DE855CDF73573641E0F8CC5C
-CFE1FABEC1AC991CF868155DC346A3631BB469477A6154392F13E2394CC0C61557BA30BA0600
-980893256B6A5CE6C123A9BB639EF6C6A9BC782EB9222C986ECFA9B5C159675567BA5D69AF3A
-476BC03CADB9C01EC2D79861B5F3E170B07D6E7991B27014FA33EA552C8462851E7DF7FD9F3C
-76DCDE16E08B3FB44F0845B1CB354F26238E89EB656406077153A6BE189D2A6FA9C11031C1DD
-16513FAE9290CAF9DF420BE75392D36A9D1E9B5D05B57763AE8270E9EDB4FEB6134D0B1B0BFE
-1270A4402797BB76566D027F31E33D1FA703AF8F2608F76F99E02636127D836648AB9BAD0E61
-EAA4FDEAAC409195E044032C2856AA780605946E59103DAD1D1525B7B60BCDCAFE71C78FF00E
-83A77B8271AF411716A1509BC9E0ED98C19DD08C457D3F608C220DD544EF4D4B6DD31D8E6CD4
-BAC1E30FF9FA95945A19FDA3F61B199E2A96055FAF195C66C0D36B300009BF5025E3CAD01AC6
-2FDFFEA9E10FC5E6BC7AA1400A9B7D9E6734CA2C1BC4DF7D82142DFC09A56370FBD6F225D5DD
-06C449598578BEF70513235AA24951323896B07619683DB052288C816FA08A752740261AD1EA
-09860CCCCE951D2C606A2B8175C49F5F4DFCD17A207EE4350784A751038AC6446EF0DDE8BDDC
-01D3BCEE3747A46A6BB2320A5003345597BE631C03280F5A699AAE3F3FF772F808EBB7C66107
-40549DCFBBB26153031FAB1452AA82DAB8E37435D6C143E954E4DB00285FD48681EFEFA513FB
-CCBE3C74B6885B6693957808AB234531C3D08B1C77B3719898BDCD93E763F57935EE2A74E682
-2B11E298B46FA218B91BEA0847DD5163EB4B8355E1A59B4D93CC6E13D9E9C443FC5159839931
-403F9B371AAD48C3EBFA13F6869CD6A761C4A025DCEAA2378BA1141FB3116CC34B676F57B800
-2114D0E601C905C0B341A1A027968788D0DA1B7538EA1CD3FF889998B4C33BCF6708DE5BB6C8
-DF46ACFFEF462EAFB048EAB6119F8ED70D5CBFB7DAB71AA6F2C9DD82F17524581AA73CBF6A22
-5CA70781A2EB82B2209450382A55685DA2A087523264349A55FFE19F6E52BF52ED8E392C5FD3
-B96769EF3C1CD2AE7FF15D95B72D9055D4B1680EB5DE98A399ACC23D8A0A8A59774E73519209
-E0AB48CBAE1FD3B4B5057DF9DBD199A59F4A6B33CA2DE253B55A86DFF752EA5B69D5F0EFBBA0
-A5C53899CF772B1CEA68C8AF11503371DB48872910677CA0D870E1B42D6AB2B4792BFE97FB8A
-0B60CE830DAF6643C7307EA70643644D2F8D0408461FBE68FDE7379F09F8EC87FAFF2F6034AB
-6BB2AD1D13F5714301C878C84253F85779BD1CABB29E07CC0D9C543D50B76C416DC6E90D985D
-01E75A4197EFD2451336584779B17F1A8B96BA920CFDCF58BAA576CFFA1AFCC32B506640DCB2
-E0AD1CDFF9D1907AAE407DA643934F1960F8008B9767342CA9A75B5628A0A3AD9DF5B6F44E08
-6E3B2269AF255F5A4792851D403CF5901DC56E047670F6313233C3D11FE76FFA7E03C1DCCD5F
-A9EC5627BF58D777AB50D413FB0C72DF4B71BB352F81D8045F30C1B13247FA86116167AD6AC6
-D75F6650C34BD5F3EBD4DC864F414FBF34E2E3CF1D4A921B7CE7006AFA0F435A530EBC612E65
-506C3547BD53E20508D39A930931D68F74FC5EA4CE9D5397CF925E6B445AF0E38CA451EAC209
-D7E3C5FDCDB48CCBDACC57A062814CE2FFBD3A5CDE97D01465E790F69D79DBE79C87ADA92230
-650349B98EFD94C6CC80EDD22B559A3228CED6CB5AD5E79507D23F3C54471D0C0B1A7D0D8759
-8038310C232667A05EE8138198191A82815D852E967A9534F360DAFD0790FCD471CCE0AEF5D8
-07D7E14D4AF93C6F7213CA1E5D864DD2BF9BE6879A3212AF6259106B3AFB611EFFAB7D87B03F
-49E0F67B8BFFF7E9A9E30F1222D92FFE5F6870BA8A2D7856038623D5F280DABC661ECA98800E
-758F8CC46BE45E4DD795C05127264B226236F26245FBF409DFF34BD05579086B9002EDC817FA
-0A7B82CEAD24991718311D58D6F29F971C743FCD54429F0CB7DDF91F436D51A52BE1D99C99FA
-FBB93E18DAE89070A9EDCB07674CEF1214C0FE6A4DF86B63399D82C50989A03541D23279B8C8
-3EE0AB2A42D60945D7B1B1AFADFBBB9DA8384C4CEFCA6F6BB8F6AF3303E416B545939C3F6348
-B8D055B9E31C8B450D5A0F292057349D639B0DD79AB8428E872AE2301EB48DAAE5AE32D63518
-8536E446C22B9080C34C8781AAD7A5BC12738526FE7F057494960678EEDD6FA007FB8ED34A0F
-61BEAF4168B770E8C21DD19D7C3C56A96ED431F766E7A6AD8C8AE6447B9902FD10A87197CD6C
-3F771580B4634B807EBE94A169AA42081323E0DCEEDA4AD8B5C2BA36E82194D9CE4746214D69
-84BBE12142BF9EA6DDF0CED377234CB504B40B203A8424E94CD0D9C7EBD7AF6E7781222F6CAA
-3EF9CD4D547B9E8840A766B07A196F92B16CF470BA28AE3A705F1AEF3F7A4FB78BC7E5315812
-F34E832A2625DAC422D278EDE610FB689C42F7B0B2E9BB1C6531BFB591FEFC35A17469A961F9
-66A8065063FA2B2F57850270AF7AE77D6F2E2C62F7C9F046EBD581EB0EB709DE3C50B789F4B4
-5B8B28BE16422D01AD6451D6F8329C38201A9AB27A162AE428E061C460CD7F4524C4A988D420
-114340BCDEDA7F4D6E474DF4AB1E7F6A7EC187C195EA7ED1F0D6B78AFA3E9BFF7F452A86FDA9
-6C2CB0F67A37BE74FCA44C985AC3AC6223AC391EE2E3E47F42C9594A5138B505EF6D89D2E38E
-6DFD6C81984D85D60068C3A887724876699F0D9967352B5058AEE865D210E11ACDABA4F64006
-F8AF65A2988D97E1B21B09CCF382C86E8E28191A30BA425F64970058F5ED779EA9EEA82B5CC8
-6D05534152DACC07537E7EC725B57C4325FC9E4401AFA76E254B194C97907DBDAC3924138D78
-EE81AE2AD9D75260DB65F48DC236DA93475354F57BE726B03F9FE8207FFDCEE6F5504B9141A0
-4B2CBE7EAEC208EEFDE56D00E35D79916E8F05D4D5FE6FA083A76E744059416DE1516873E641
-0695CD5171589C3631443DCBDDB55603C871367412C7170857E8BAE6677EB8680D573E691979
-1992C929C163E8603467CA3FAC801E4F23637BAAE85BC118E3A0235FA797F0FD64CC5D15149C
-1B1FDB4A50E7D59C777CE67F7C10DA9432C73417E5077C18953DED59B7964E8C1A1EB07CC880
-43B095B8504EE177D1F752428D44337C506B1305BA105DEFE2F5F8960C357175DF864274A2DA
-1CC1026E1BCD5D8CFC2AB5BCC63D1ED6994505BF969D3CBC8F97F0683182760B5B8972851A53
-A133643146A803F474FE4118121F507F2C9A90D815496ADF3C9C2EAF333544711B612211F964
-844FE6EB3ABADF202D5F8425DA350922C9F79C1C07AFC03640DE8E691CA7CA74D3CDC56AC0AA
-7C82F99A057FDF9B1E0561BA8D0710D2B9EAC03B1CFB7B141D39693B7054C6815A0C335FFE5E
-36DF9A6AE13E7313C6C7C2331FCA2D6A6F2208F72A0E62C74C91054F9B778957DD83C487E924
-29D63A0FC567290B6B49749D3FFDE6173B5D5B9937EC597FF138547921EDFD0ECBB4055FFAC9
-3A0C3EF17C8898AC54605DF8327D0E77E8AFD22CB8F73AAC3EC94F1AF338B3714A08C2132924
-D19BD496BA4F1F76F410917F4429D5D376310DD5E12339D823292FEE8425310CF40D9A41844F
-5EE2DCC2EDBD53544DAB4A120AB999B516D8B90CB30E1D8CCEE638CF1CED2DF52AD14372DA53
-C9F82F876B77BD8E8BFD3189211657194D83F0E57034319AE16492E17B0766AD4EB3E6771F12
-D8E8134A418A85888DA76A9727BC5B5B9A59B996DA8F135AA7759DABE5E3AEED9340169ECE3A
-9731A6822438D8E8D44196B73FF8BCCA83FB3B1E0F66064ED52C8EED2E86F57B2CE5D4AF398B
-F2664D15F3F411BF5E9F6470DBF2BC6895862819570971BF56897573634856B5D19058DE78F4
-693A339225D1815F62F95204A9E42A49E834BD1B97F20C1332287248A98CD10690414D324E77
-E2AFC5482402EE516F444B7DC47818A64F3F826F4F0EE067CE703D91F52B0C6A5F2A04864627
-15F8562814D6EFE11493F110104FD734D0EB99E02603137C4C0D992CA67B9A7323E946DC8AD4
-16EC380038933E3DDC4EE24215EDEA8125913ECA6EC727715C49B62320868E9407AEF30B07A2
-AC23671A18F3784E6D98A8C1C191C68DD7598A65880220B8BA5B425010EC489022BE4D2645C1
-7494BFD8F9F45FC44C9067075C790F90DEC194655C517A6DCF1FD3122CFC276CECBBC5FC6039
-FDF134F56EE78B5A9ED678C7272FE8FD1A0E8B1C7DA8546F132B9552A581C4FAB5237E254036
-92BD33D7BAD9DC56FBAB7C69CB79062527ADE0C8D8DA93E5C0CF8E4A9ABD743B65FDAE84ABAF
-B0674650285796C620AE873558F53C850B0E82140216B4BA3BD41F6613519935F201FD0BB4CB
-FA8D5223D1A26E734444920978B42620F21CB5A4F24FE225D141B5A56FB25029CB9976ECEAE4
-8647662E7C98796899C3E1723DC998FB42F6F1B13D536F82F908EC3CA7FBDF24D4A0B954A821
-71CC936DFA24706A5D4FB4118D87BAD0BBE0AE559C951E22A33181F31B92E64984D52744A4B6
-E358189A88360B9521475CA082704F444791CBFABEEADD5A6708B864599843A181FEF2905498
-30BB0B4793371A04AB1A9DE37FC3FABB9921C2E2924495C6FEC58CAA3FF4564E9EC89078E84B
-C2066D27270A1E17A92E64AE34134FC6BA7E4F70E8743218EFC83FB76E76AAED306D6A5AA873
-0FED52EED087B46C4583445DFEE680BEE70374C232FE1FBA1C16E2A4FBCE3795B030787D42FB
-F97712B887DBB1685474C68CD30A6AED27392CF8EDE33C61C5CEAECC44DA2B98944D3CF40EB5
-F24240B12E0567FF6F62302C5E1D611CACD6A249CD39573F7C09DE6AD8D01C1E69EEAE713372
-5EDA78387FD25BB4E96FE593928A6915F821EDC87C1CC2B2C288C3CBF49FECFDDB20F47EBAB3
-CA3818D6D61BFC825F164DD73AEF81C6CFC01EA5077B0564000C131F4D45E66D8A7594E95D8D
-07D647C8D4D31843E10C259684E5DBB3660AC7B8D0119924480756B4355E1EC82E74509AE4B0
-2C9B8F5AC8694B878A9175E7FA476420189C487AB4C9016404F13666ADDEF60F5B9211343453
-5EBE0D172942EAF0CF53AF4FD7707D3BDDD04508A8D4DB1362B67F9BBDBB80F942608D44DD69
-88DD8C4BC013E035A27DEAC29397E1CB89DDAE52F61B52AF29EB39C15DAC92F3EA57EE7EA66E
-B8BE4622DC4CA2A7FD49BCD5196A125A74DCFC7615F6759E4BF2FAB28D7E82002347B89AE44D
-4623A1979C1A885F700B157774ECA8B5138C427F7FBE595D9B7907CDD24E765659AC9A3D8009
-1E16FDFCFBDA8A6C9C9656A0F75F4D391D4E503C00B7A4227B2D89437F97D0FD2334CF8613C3
-19FC8AE7B9F31B9265A4AB8C3B40806762227A7D025FCF7B209AF7C54455DAB03D708B2FB3C7
-C4845C0D539A6E3F055F19FFB857365FE39959BDB251FE4D36420F6695DAB7764EE1D9B0820E
-17BC3C918AD2E52F839BE0015E6C596EEC813A6CB06CF1A57A00EE4785159D493BEB152CCD9B
-4EC444998DD44630FE056C5E96E8FE9A7C44E810CA661EEF841A4F540D4F71EBA5660404D39D
-A8A03AA50E4C6293F44413E90B4BC05CA73C0D5BB2169A4EF75E1225C6555B9FB66F9AEA4A55
-0DD39B1DA07754C8F4430067325F1CEEFC48A524E521A4320E0884BDF5965DBF3EF29D6B4F00
-815C26CA5D937E64A24A7D6FD126A46EE733BD1AAC872103DCB28CA020D3B50EAF57AAFF6266
-40F8DD8A859A84EFC1504147C444526762DF179457A6FD46F7F02014F761D24438CFCB0BE296
-69E6025AA73236A5CCFACCEFF05F45FAB87CF96EACEDCA999E9D49E1E6F2D3491B5BAF9BDCE8
-5DCC9A49251AEC2589B56037E5032E9674DE1E92C97085DD5E3CE2ECB83DEFF5E4D37CCBC092
-CBA6AE8521122C21E23A1A13FC3A3CF970E0182851B7B729C37C154496250688D7870E9A0902
-CE7A2B9B22C77E58C057B2ED5455962828BB3034E5CC1EB74EC8F68C75D9EA50739DB54782A8
-D55CCA7C47B05A75E4A8A51E54C383716C0E6D59ACFC0A66CA5056A413801EB193F39DF2E7B9
-775CC59F005863B9BD98BD69B357337AFFD8BBD9FEC8C0D4518AA6F201E1820852A59908C10C
-D48F29EB8ECF19D2168CF7BB39E60E707A41A3D1B84A4704097D689DE1F111F5CFC33D8093C1
-F0F811DBB7908433E1B4354AC6A0DA559464E8F2CA94FF5092DEC40E949F6C92B8F311D3835C
-623B734B3D1F9D54F2E933A7FB664FA0DA414F608BBF614C65E644723B2A723661963E7700EC
-318115637D017AE4DE8EAE0CE6AF903487D0B3831511AE19CEFCB8B9CED35C65F94DF9806BDD
-D82AF0E521FB82A5B4CED998E02AA316A04CAE953D731DB0854D9A4E1CBD272D74AB47572D68
-05067C60E4E960C240CC2BE2AF7AF9E1304C0CA522DF528A18923F97C5CE7B4409FDD09EBD6F
-BB800D11A1C6882B8965DB70E019970E1F6E541BE5219A156C58126FBB21C4726056C1C33561
-32B85FBADB5EC981F8D9334FCBECE6785EBEE2174BE962766EF60B7C02F249882D05E8CFE782
-237F4D9B746323EFAC8BFF9C1CC1B375C620CC708FAB00BD2EBCEBA4E5C0C3B3A4011A72FC47
-4F974CECF0DE5DB304E5A06C04F3D049724DA571B64697C1A80E2059E6BB6A21ECD81A544F7E
-83C6AC240F22F25F815C20C0AF08EBA624467DD4B5645E2E1F4DEB2C760CE6CEFB0217E000D5
-A7895B4A62ED06B86F2A0082F1AAAC04AD55A5102401E897C6C50D616845CF92F25A4147118A
-464D4B94003C7BFBBF3AF22AE709E360D5DE1831994B483350F1BFF514ED4782A9C6982F7501
-63DE2A6B91F67713F49DDAC470C19E9FD4A1250030CC67868662C0015432AD7AC88366B70E0C
-6023F518500BB68C4D84080EE5B692AF1C78C35D48F4F3392B5E4FEB298DF0B869A28B8799E3
-0CCA3C45400620E9AF3D61B5F8074541BE64914343D177B8988AE05ED460AF0180D04B8E78C4
-67E83C0ED5D2C7F8613439A9808F28BAC8F3ADBC52093E4EC591692233562CD5A25070014F62
-C54F1F1A6083C275D7DD8A776E86C207DDB702F122CC92F73AED70DDC65D627525FEF041B597
-03B43840D82AC30DA54B5BF852E1A9CD5FBFA4CB3CAB838BC2D4C0B8ECD37C641885258AAE53
-0BBF0C2B279DAD677F152A1ECCBE6C89A55684483127C3D03AB7F4D8A31FC04535259B242C46
-EFE89BB7D09161ADDA5AF99D63E6B023D3E01BF576D0F4821CB5DE75CD1212DD292172006FB9
-DFB70539B1C9A1DAB6D5AFA53A133FC22F337F5DE642B50EF883F50E5E75037C2104D36E1C31
-5251E246EAD10D073AEBDF2643FD275B6569F1F0EBC42DAC6E322747C3BF4312585C97171F03
-026B81FF362E10FF09A9471224453AA4B32887553668BD484B6356FFD5223FBD8841B584F3A6
-983DBDBE0B4EB9A4915FED5639BFEB0EF3506F6D5F28C9BBC24F1B2943FFDEDBBE4FC4EB4A5E
-0F3E4B86D57EFC43582E6AF464D5808200E9B75657D53F2405132E33AA53D708D5FE901C483D
-395F3013DE8035C445B3E6F3FF15866A0A6A3E6BAFF3503E8B4DE6A37600F24F38204FC7DE40
-77199E4B77E6C9370A6C9355668E70A65F1548CDD938E52CD0E774BE108CD6ADB0894B86E954
-2FBA918E606C88299C92B250E2D63CB0623590BA7C715A68ACF004FFAE4CA76866EC20E14588
-872B8E80F2FF582531601762B7E4F677021CCCDCF94E277439FB143C5D23F360A8ABCEE65224
-FA153538AD8DCF7599752C8353258FFC0AFAB471B2046C5C99BFAC2A639000B23A6497C0D78F
-BE39308D619108E17132FD52C82AB56FD72827CFA98C9F4C27249C111552D863687717974624
-AE692A23EB6BB6AA6AB2904D066CAB913774997E38DEC623488D3EDA398552852A6E49EBF291
-BF4E10E59072CFB85A805BC7BFD9A1A6DA4DB243AD53FE158ED1EA3DD8A8B132773630C6417A
-0D6BF88D63FBD35C4EE55AFE23D45032939A052E42E5D2CFC4669521A45043A37DD06F61C145
-C93DE5405D6AB6FD632A1A2D74A45FC944C2AFE7D66C7F43AAE63CCCFA328DBE0ED033EAD7C9
-4B4B4DB0EDDB167A845A2C918FCDB3509C2F966E45E111EAEEE77E4361E86671DCDEC06E6358
-87DE61485270E33672C0820B62A53F97E54032E1D3782BD14C7466EEE397A31B58ED534917AF
-A762A63F4B2A1E6E3E46FAEF8071B99E3A749A5AFA03F45E285E6F69BF69B6601B5B92CBCE3E
-82230130E220BDBBA449FB5DF7D67066447DF60748EB5FD9A4D177B4DEED8443A55AE6D63CB4
-D960E3C7488B900A95210F1D8B667E8356E8F4BECE485A1401ED4BDAF19021CE99E09E359AEE
-DD22DE12FB4EE8933EA28967704D31639E045585813DF66B97FA4A740FF3725356254FFC53F6
-6375EE08EA4181DC70EB27433BCE126F98163576A7A75CC27EB8A9A9F8B11B806E21CC27805F
-D892417724CEFB8E4B34524D9CDB6F900C4A08FD6FB0656041DC4CFC907954C77D42F130E2F0
-79631EB8D67A289E74C4DBEA0600355024A1AE888E907D73C610A6157AFB983CC041AE3CA368
-1F518E7574E9F0C6A5C89686E2479EDD4C098EE928C74F8AF2129E476A60F069CB8F1F2F04D2
-33B47C74A5849958A4AAB9FD031008E9ECFD5034887F36CE4899EAFFD38729E4021F9BA94A93
-04E186726C24CE65A4AF720CA086CCF6C776221DAE558C9F3A6A70ACDB6B34B0946DB062967B
-C2FAEBEC9FA64FBEBBAA90401982EE34333ACFD2F7287D9FC0243A08B369C394273B39C31EC2
-DB1DCD299248371FEF0428C8C84C4C064D77BF5663A8F0B067E784BD3BA9FE3AC720D9AC048D
-70EA93BDA9A9FC7090AC36BE749A514198C6A9A953650C73F7C40B496767F2FFA303134549C8
-C0B4CA24C40868D22126A8A2FD6DDCA8535792644F109D9A53628BA787E76B6A5AB7E4B2BF90
-1F7547388839D9E7FB2E8114EBA3357BBE95F10BD70C16693032234DC8CABF420D14019F3C3D
-BA872D5C5DE8EBCAF343E18345A2B34B9D299E6E0FC49DC07A6500DFC27F7E78724B01B56523
-87F59ECD335856B76694CA43350FB40E3CB11D4A1334E86057A88F0372622DF11B4C6AE2F02F
-4E6475B4351265D1730BD0CAF5D6891764FB5E5909DCA5A2ED1AF0484E3677CFB38EBC5993EC
-87B8B4E588FF9C2B85C61F8704E1DBE89AD116C55163B88DE73FE118F8B6737CCEDFA28A37C5
-CF58BD59CD6C7ECF9306D38706DB046BA0AB5542BA5391808EAB5B2211A3824A018266BBEA06
-074703E70093CC624824A387B1240023607EDFAC35FADBCF2F238DFE1A2BF494A7BC16D16381
-ED52303AB267FE5BC9C9C71848E8C0AFC603531193E8C6467ACF2E96FB22E56C502394CC8721
-581E305295002C7316B251A9668945761CFA5F1CAF4A6CBE8F74B69D65A2307A4195F6D264BF
-D4C2DA1174006EAEE1CC0F66D50507A887834D9710C40DE47D2B12F62827585F49A76E266F9E
-4A787F6012AB3D03488F0EF975CE81882C42D8E14E2B23696B0D7D220DC130F660E758E70FB7
-0C9F6D77437B10DD6245324FC6BA21DC0E9C4FE5BC2344EEC300C6B9B53A68005A38493A7829
-EF51F89EF4E049539E237687763E951D64741C0714F60850C8822070E8AC4670AB6A8BD9A597
-733DB0C113801698E89F74B5892050170B1220F66D123396D40AD0E32EA8DC3D7ED58E26C776
-1E997BFFC7468D884A0518A201D132D535067138EBD3185E84505B8B6606882270FE2B9CD394
-66D6BE7583B3F9B1701EBD49129ECB6D8DBFFC80382A95DFECD9B35D4F25D35BD39CB264082A
-0A12B570E2394DC231CE55ECD62122B7E8AF1AFCEC346B0AE92AAE17EFB5B29B234187C79DCB
-1BF735CEDDCDE0EDF8F1E514A2FC236F37D9D41CA8E0AE329F4D79E1EAB3F345EC8624153E6D
-E6F5C5F0DBF8FA8B152E3C10412714466F2404472A78FC7B984B3AA8F6F655F13C3F1FDB10FE
-BA215E5F3561FC3C79D07BA8547005665DB79F64C015552AE573BA32AD62FEB08DBDB5304B69
-869CDA9086F7D180D6179CF2B6ECD220A5C20C4E3B6892BA80B3FAB8E03C05963F1C176DD30B
-510189AAE745AC5C03042CF95AE3885B5407C323B4349A3CB445EBEF3CF298725E1D0FD3F240
-FDFB4E38C917BF10F34FB5E5BA2D88F72BA1C08FCAE86774B0535B31DD7C7D47B066C4716214
-671E653A36759D0B19AFFC93F64B41B83DA2BBB8087B07125CB115410A8EBD9D712B6B7A54AE
-25F79B7C6C70C7FD9542FC0CAC783F3B8F8E9AC4E089E6F4A6ECB54F312199CEE85833C5E077
-5DF387123383C3770056473B166AAAD7E3615590B6EB9979D837C3558EED8645354AAB3BA242
-71D1B18B87514847C59AAC104FB483C47F3CEAC41ADE74F9CC5D86E1509ACF12BF29C6632D96
-05BAE3AE921C6EE6219BEA5C54BAA03FEAC7E4BE6A2C2570DAE022C8172DEF353FAD3BF44F3D
-38DE838BA4E0F785973FE639321022C83EECFEE606BF47567DF300E3CCCB041766285912B011
-249F26CF35049E368B855C960DBF5A84DBF4CCF68A65CBFF5A46BEB6A0E36335D886534A295E
-A6D6573DECC1556D3F805030A7853FE4F6E896A722CC5C4B968302C7CFEF896F38EA2A1B5A1C
-1C73D462974AC0730E6829DF882D4FAAA72DDE7FDB838AB8ED392257B2A1B5FA781FD572990A
-4493DF8130BACF4730D9FF732DEB84A31CD76FCBB154E383A55A3C803B07057A92BEF5D56559
-A876E760C6A08D2BF31F4374A9796F9FD71D323AC9036775B346B99E6AD2948885461CEFB703
-B801EDAE4CEDD9C62BD1AD15C45A6EA4FCE3A500C8B9FEF22A8D8D79402D6538F9604E11BAF6
-1F779323B783BB1BAE7E12679890835D150CB3196A70FBECDF3AD7279AA24189D8F75C82AF58
-A9B0770BB155FAD8B400656CB89DEC49B3A14E5470040FA18E6D86D071525F1AF2A779E4B274
-1EAE1D4055DE3B520E143A660A7E9ACBDC354F4589C59E601B8AD3628998DE791CCADFC45486
-995C69FD315E09A33DE11A09551185C1D843A11AA015C2C43281FFC38306A3A07917996E344A
-3F230465D2C42F743506DDFDB0B1CA2BE16705F4299C77DCDBBB78646FA07EDD0A411082251A
-B49CA270CA0E9C6066E475C5967A4FAD7264638C63DF88351371A7A63D0E5B3064747E5E9D52
-B99179E13FBBDC9461A54342DA3FB98FEAAC7A1DC10B6E708EF04E7A02D84C72EEB0A3465FB6
-B1C76C2216D218E64CD54CC3F690CA6BD98C6913A0BA9B5496C425138FA04797CBC1566F381B
-CB46BF2C68AE286EE2A27F8B9EF7D232CC89485114A5B2FEDA1D0F8AA97425E29A7CD47600E9
-82EC75AD3C58F4F28E908FC0F0955560F8502A4497DF6D4F7A6A1343EFD6974793E2DBF11B1B
-4630DF9FD380F0CF01BB69ED2C2B13010AEC2173CCFB4682552B2800E91074A31A83CC3870A5
-5A70CD79E915BD08413324E14B88C9EE9B34BBD92717BEA8F6DEFC8E6DCF1868921BBB6AD163
-057D42EFACDA4048F5DC536C7B6F0E5BDC3DC46CCBAE5B538690E679E7CBAFBE78C27B8DB99C
-672232AD764FEC2D7218E1FAA216CD1EEA7BAFA5AFD6220CF4D1EFF3F013D082B93565355032
-0D6663C959359E4E6B9B439C868CFA0F01429B5E976A1FD2876E3E722B1AF363466361057EDF
-0CC2A93BF4A459C0730DFE9E7C7B19A1DF3DAE8C7871C78180A1828578883212E0085A3BE918
-DB0966D965F3199108A68B3A307433E82341B50F5509A9B87D4AAE3F536FADE30418BA1AD8D7
-9E1237B4F84067345EF75269EEDA1F024030EF28CF9D79F81D8C5F9F5BCA572737D60933158B
-CD61655EA21DD1018962FB58A52891857FD440EE871E6054ECCF8CD1546827852F1843F8EA41
-C701AAC221E97F4C74DCF6C8600338486FDCA002FCA6A68B740D288F6335CB3687E9BAFB5D6B
-B7F3740DD9E1A6AD6825A4BB522489766C84AF2A1E7DCC8CA0E2669B9348CF47C77A33896737
-57CDBEDA57A079EC6DCC516B51B67C9F2DF4835AA127CACCBB0B8E42DBBD4B635DDEE5CC4EAC
-1FEAFB5BA3E104E495DF8A264C27B2358B286F81CEC983E9F3CEDDFB8B5910D81C128F532398
-2FB87588DAB1CF1646AAC3A4BA66799A18E0B833751625531932BF9096C3479DB8696049D6F3
-86B4768155148F2EFC928DB22BCF53E59F01663D50DD7C13885925C409F27D6E1D6F1B229A1B
-AEB68B423F74B99739D5912BED355CB695CD7E15285B647C53EDB6AFA48F00795700358EFB2F
-0E01F9E4C5456FF11D56D3EAB0578287CDE175F095251AB337343F4C2A6B40EFF271467B5250
-35F99F69FA9ABC2CBC78BBA342BD1730BB5C3CC07F2B8692FE89328243B54ED897A5BA9772BE
-64C9234EFEAEE62AC5155CD65A66CF188C50FFA53B5F3B78D49FC9D35C677C6A827EFF15A10B
-1BA2DD7B1F90F310220333010F36E9D93A5DACB253666A18D2346DF08C9337CE0FE806A745E2
-898B4C4890281AEAFE22FA42CA6607AEE24719A416B0BD943E416F937BA8988C6B520E6EAE78
-DE0E309BE1C2A719641517E4F3D5482F090055900BBA8FF70675B214126EBE3CAECAD24666A4
-7770325903AA8271F73ACC39F5DE03D250F2F33C6259CBD823F6F5D49F1C34E361C987160749
-13F8B0FF2521CD08914319D171CFD1E4966562083B418EC7D9647C4DE55E4E03BEFC23006E6B
-5F13CBC4D1AEC758A676B5031DC080713163873D3DAE457BD75ADE007DA158CA2D08D104C39A
-D9B316FA542C1420B6AACBAD4D493E1C772E4B1E2407BFEC694C1C9639A008883F73A4E67827
-E27BA8181E791B10248AE0E2FB581E1F6C4E7E32C64ED03ABA8E0A242C345F3CB9BF07772BFE
-E7B2E6B81AD533D03F528D18E31914BFD244D31677A2A2E1DEBEAE80314E2FB38E76623CB1BA
-725DB0CE13BEA97CF447FA3938EBA6ED2003451C6A89D948BD1EA907E3A94FF65760D53A20C4
-134A4868F81F97472F08F4EE4F4EE59F0642768C517997A8494AB19B20C65CE2F473FD9B3734
-4969AB121734534D6CC1C6BF4D5720CD01D4D889938E71BDC3B8074062849C8A1145FB69CC1D
-BACAA7FAC653644FF5F484F705793B33F73C321FBAA03557E85D3B9B7923A84B370F6F4FC577
-6ACB0C1423C7F3FAD934C29608D9BCD4899C6DF59C2398679578D26D5237379E41E967DDB67D
-798EE6EF00F59B50D8F0DE4288DAE073824275C20D609981074A583C43DA7A1370DF9009BF8C
-058D60508AC2B4DD97165063100B5A622D168C7ED7CF690769573B7EAF6D8F99DDEDB0962B97
-B9C40D38B15717C83538FB95E243E1AA1165D2D3040420737382296C3EB234EA045006C25818
-D743D175EF2A80CC7E0E2BFC77AB2713011E4FF718F5166C101B279598D7CAA9A45A9FAA135E
-69B8B73EB512A6ED54EAF33E324F5536500D5290F5477DD3385187E29A3BE985B22F7909DF43
-848372D2E013ACD50C6DC1FC4B39F5DE72DC293C565CE65712BA69AC33C23B230757C40DDC5E
-517385E211C5E1DB2370A32E4CA069001AA0BBA1D2EBAD24331AEA8692CCDE85AB67EB80187E
-F452AAF1F741C62DF21F715F51310FE22AC05F63C4E43C879F6CCD874648FF6A0D2A18E241D2
-0562DE8AF2D795CE2F285CD4BA64ABDBCBB00BA796C812C61597B2D8CFDD6BDD8ADC7E90850B
-5F60E14A2793DB4393DDB9BCAD01D623D2D9C07CC893AAE088B42EC99493E0DAC521752B42F0
-E4BC6ABBDCFC92C674167E1A6DBDBBDC5D382D1932791C25A31E114B48F8FD110D1507603A56
-9018201869834B631CA6F91178D995D55E9FF41D4CFBC08D7655209B25E74AC28EDB08DA4049
-A6400FC8F0852A6E9D1DAEE159F507A56FE59BEB177B7A8967A8B043EAD2F2C0C1117044F6EF
-5406B70D3CAB63C7C97ED5DB5A4E81C85622E7AC8255F292A4B1557EF5CAD30D5B7576E18DA9
-43AB5E37CD446C92EA057B137E47FA377E1638B32FE24660E9B27D2165583FC4FF7EF86CD53E
-FEC4C70AD65EA222AD2BD824F02754C58796DC92A40472BAFFB40A922E5302610FDA9A6DD657
-A3B3F24E2A85B0C3C84394406893F297589B942AE91C0CF13B4B25FFA2EA8A4C642A8D1C7E5B
-C6CE6DC6F89B76797568403070B1BF1A70DCBD895A83389CF931A2D1E95DD24A013CE9991FFB
-90202661C150EA7E92243DCA1A3412C7AEBA5B94681594D8DB66328DFCDA7D0705377AA311DB
-242351C6308D41068B2B0307250635DE75F9E0E13B9BD7C33F544A2044F393EA8D3F7209660E
-F6C9C24A62935686063A0AEA796F9591B9C7F9E57196E90E29881D5B4C094F6533785E48CB71
-1A54A839566386FAE57D9B6AD584508DE26CD1CD1DF3E34D41FF74793C2F05E3BD7EB18BD477
-6C0C49B215EEB498D45D21BB60E36F85A899AAAD74B45AC097C3EDEC6C36D97B38D4E769A502
-6CA365EE306B5DC1EE932F9AF12EBA52495A6B7B737129EBF6B1EC165867D545C87161B99EB9
-FD7613092465BD80BEAE681397D56E3329C63F1BA8F48D4396412D750BF3E816EBA24C189489
-7B39D4DF8806B6B059D9F0A20533EF4C19AB8DC8A0AB5F49446689E44233E4A76F1C2FBE27C3
-15C078872399803E0E58629E3993D9B62FFC083A6C957489131610C9F9A130606BAF011921A1
-5D44DBB0AEC19A2BD022F2B713C41BA0C7C7F772D24ADE8C8DDAD89053F6416305D80E2554F5
-DC9A7E21A1C5C4FE68F7E66B500D8D2E4FB2BDA0CC3A65CE609DF98B81B264484290807A7FB1
-73C4D5A01CB1198EE883E8D8862AAE210C12E84330165B83992245D5E3BA046C3FB742C55EA1
-D72A086B36D9052E981B2D7800C94B4EFA2F9C9555A1876A3AC1B49BE72FFDDCB6DD704A5BF0
-904C4A2B10BC93E81F26CD4C8893954D150299FC47959C0597F9A828E86CC2E4F987DC140B03
-C1063407C924B5F32799CC29B231E0523F6DBDC920D67CD823E5A37C44904515BF348D75C969
-1BFF90132AFAF15121EFFB54943F01A6550F06C6900C7876BA533BADC79F499278791C677418
-7C917570EBCCC804DC42930804BA24AC79B84E59476E5FAF06F155AF18353DEBD21958FEC90D
-4D7A5C97511DF5BFFEE9C77CB5A3DE52A4C4AB6AA405E3E4847900D8B774486A0DCB2AF23166
-11B246545FC590329A62EDFB76CA767F7676FF21A9D171C77F5CD310AF7AB6FFB10B74581F9B
-10A6CEF863F86427DF0ACCB2BF3EFA7BAC5B0104FB2750FC111060E5E53B2244DFADCB548C57
-3FF3771BDCE4FE8F3593B2A47C05C13FA4782A75D1C2A8D0AF80AC1F5A4AB174EA3F18060202
-C599915BEDB9DAA2373AFD71099D565E976F59642A930465A7A45CD42E11F657ED7B5DBB49F7
-FA58FB739070484823526652FEF5340FABAABE7679C3B951552EA2518240422C6EABA838EBC0
-D42F7BE627D7B9CF20D13688FAA0A80ACC5F293C00F580A3779F331A52DDE22D5401C78FA42B
-DFA0B9D4979671FE7CCB642972F047764CF5CD0E180CCAC8154F0021ED854D76CD9364A77D14
-F2C7AB778CF6459F47ECC750D9615CF348FB94C4A2259EFD2E5D415749513DAE442F840B5046
-79B5AFA05F381DEEB9885E21E8DF0EDBCA9DAD5EE47428CE2E347AAD48EE6111CBA5A3F77802
-7E8B62540E1483577235D6F02445C2C9E3364C986B51FBC03B4302F48DF7716C225DB04E1823
-590E5E64F2D776ACF6549768554DD376E617AAD05C2A764DC3BEAE8CEBC12976F1DAC2A41F5B
-32B48E412444F3616DADE10AC585E1D8B7936C4A293CE678CEAC47D5CD1395A05E4ABEA987E1
-3A0D1069AD62450135170ADD22E02570F58242CAFEC1E47089116511049AB559DF9E1406CBBD
-E0911FE9D3C1658F58DFF363D4E5C5A57FD862C347E7DA9B109AAEB6797919A358DD4C1F23B1
-04F8233A7DE7A256AEE9C9C149C6B447F417D55CDC029252B6034C3AE9FDB45C35DAB333B430
-890D1A2D3F32B6951EFC53DD1FC9E521A5A891BDA1F19B373E6A37B9DE1B1E503F78F27DE920
-911E16FB07833C734770F17227F7C76D7067DA1BAA3887EF8F16D1643156AFD52C4910977784
-06997633A63C58FC63EE43BC83C99BD80ED298887B91F3388685F7B13C0027547662CC0F64BF
-2AAED6A4362D6C6671724451AA003D87AD6B9F8ADB9636F214BD5F013C4C0ED7EB685C17CBBA
-CB44D912457FFD0F3726CFB8500487ACDD7E580346548BF165A3E8DC2C2385E105656A1E1953
-5263A09E9BD3B7C0AAE2D921F3321D9DAFE7B08814DB726F7B56C59804862BA095E4C5D3E5DB
-454ABBBDA041698F951E8808A8A5CA86CECFDEC00A2D3462861D871AE19E8F223C755B0C0F3B
-8710217E747010502F360401F98EA2F46E34DCA6BEF7EEDA347C0351B17555FF0C61CD568135
-022A50589239770213A5324AF5AD0F68D85B5A87D5C600A966BD16C7F2C4D80600935F14DBBE
-FEBE33105CE41716C0E62A1CD18080E5407B0E55D575BFB41033E62D7AEBF9E61415D3F32C02
-E56F018BF5D9A8B8E47A0929EFD0E4710D52DED51BCF60B8F5DFC9FE7774F0B3C4C0D9E797A1
-22B7A899B5FDE3F2E6C636E3BE4EE21ED25DA5DBC2890ABDE56CF5FE7C47A7F04E86B204067A
-E4FD00C32BEEF0939C128DA0F8601902AEED5EFE6997B313B668BD3594ED719175807F03BB72
-16F722884631435E05598B91AF40580139CA927CCC50746DA1D6D68AC3567E9A3A8DF77A60D6
-03DFFDA9761C3466DF50C8FD7C298AC713AE7674EEFA2537065A3B6A67523E5C4C4267BEAF87
-1B634CD26392ABD8538A54C4CBCEBD0B7748B31DC43A05CA835730D0E5EAB7AE8A861256671D
-FCE6A8B2D37A4B7BD7132179163EB4CEBA50464B439AC81FA224C8F7201C345E185D53D14609
-728BE0EA91B584E78BC2E7B88FB454025FD92F0423E6C821A2B783CCF3FC8D7F2A7155B61096
-9785B346C5E4447513AD7650CEED23C2F991A8C3DC215953E222AF8456734C56764E3FA539E1
-A502394367E0F10CA1EBAEFE867681C6E7C5BFF5433F68B9735E6428A9166EF62CA43BF02036
-DAB15672A3805C9EB03E9EBCB6FC16F943FBD67FD7F5F1D25F64CB417E038569040A4250A15B
-C0933C6C86B6EA0D59B85BDBF0EDC9FC3BF184B440AE5E5136B762A23C743DB47201437510AB
-1EF118990945F9752390B73F90B9CCE7A281CFEFC1C643CEF2D67398271F7236B1CC124AE22E
-AACD04B922BF405DD4ACB82A7BD28C4DC9B2AFC7B152777798917A1BA5C92C14AA34321E94C5
-C3307F8FC1AF824C52BDAF7F51E95D8CC79B6759D6073D68AF725410BEB923DA287CD54C202B
-91C738FDF5A9F41A4C5F0A8DECCF48F5B579025FBB83D8B039AD73C82EDBD6065CE626F5CF8A
-838ACDBA15A05A4BA26984C1D7D906A30CF6FF28F443E967FCED75E386F14DBBBD58527163D9
-050F1F290D0F9E43873EDFB865EFD791D1C5BF64582B59B8C1FFA03BB8CC90CA43D466E67333
-A7AD14E3BF24F786D05395A04E32C9B38D5ACA70E10A3AB5373D59675DED7D69B4463C43E071
-59EA1A028A5AE1125C637E430C0489D5E15C8C303762165F66409AF63C3A2C91506C56FC3641
-193A79A630AF6ECAAD5D83739E70909727EBB4F16A4BF721D6FCDFFA5D954F7848DA06E28FEE
-E1D283708165C220A63B4A864D9DC5B6DE0354BA9353CD2B494261F7B257B0A754B3CDD52E9A
-EC7C566F507DC292B2471190D3EBB72A4E1BD8A42591186FEC62E549CEB7BF65D7134B36CEBE
-73272E8C5F8D72CF2A9058E05AD3C9765937F570B802BA20D0D02AB6E1B6BD9C2B396591173B
-4C10976293041155FA4EBD82C51D844BA8F80DF7CF5FD65E3600694145CB35A450DDE4A8E878
-56044CA5D60337ACEF015930AAF0EB5CDF0B1E1A68BCF30BC0E7EA5D5C2736FC3CD262CAA6A9
-0E9CE2E04EE096305327E1C8FD3AF080DB83362C48A141FC130479EA1DA05EDB076109F2387B
-F05FD252CACC81FA1903DB9992ED1F64E2FE510C642CF4CFFDD4EE7B10C089B5C86E5BB69DEA
-2319C2932FF45D29DAFE3C82CE7E40674BA0B3D7989CAE0A727DC2A235D87A59DD3B4D35095D
-FE3C155C6A81E101E6FC517657BB369AF76F6E1F336CE401711D6AC2A8F59EF8639E3EA31CF9
-0D658551B204D7FA73620C0F58C9C6F31BC04AAA58346EE5EAF4A909E5718206FA5DFA33224B
-864E096092FA5F1CADBEF77CC148839D34BCB5FA37534E9535A89677973CC3074D1AAB7DE161
-902209CF7BDC36A72D819B3F56EA89DA7311EB7E9127FBD76A2E4622A13A8A75D086B233E4D7
-AE468EFB3E891B98029C051CB3F8CF448559121E36CB382CC227349D4E23DFB4B48850E43703
-1C43BAB97A3DF6BA3745036817D9E2A279A72F79FFBA64CB7BEDC67F003088E4FF9C67B955F2
-B27755B4BEF3EE2004952A07757BBF35D867213C46EEFD6F3E7C3788B2FF6395569C6DB9B754
-51427FF8C981F93CE4465B8752BA981245C02D0114AEFB57B7281ADE51976D058020C20316F4
-03545D1223F759470E0F13615344F39EB233C4C3B7037D802F14AA61AED76F095928A791497D
-A74AF37D9E6BF3F9279100ABFD656C69D3F213501D581FB5431208691E7D1C93AE306642391C
-361ABDCCD0E2C8C2B8866FA235E37E306EF9E3A7F8D157D08ECEA6234A118749C1F8F86EDAE1
-B65CF9925D7661B86DD88CB1284CAD6D5BF8288B75A7FC1B4D368C67D4268268617017A7FEAA
-2BD49F2121AEE3A5B86E1C3960AF8AE73876478F2FD6E77D9897DC2208ABCD87EA3461A32717
-CD25B7CE6B7D67E233D2B6203B779017CD300D07463E856BF938B9BF72FAF468426DD43186AA
-9DA5DEBB10E12D1F35F89B464EBDA702FFAA4673D99CCA3ABA68741FB9C2EF7AD75A262168FC
-807B6C7C9BB8DDE5397DC7C03AC3FB540DC24C55271675696CB79456ECFC51D623CF61527471
-48ECDA89B0AEEF1388F71038A8CA246931D47DA52D3759B66491F5F8845C2B295BDB417512A7
-30E9A1D9C25E0645E5031ACF12E66D25DE534D3C0CCAD4FCA5C1E12177CE38447A8B427941E5
-42650A8E5F4430A46B3F795718EF2CFA98E9AFB21A67FB9FC1C2F4A87DCC11BF15C9D1420AB8
-37AEC4A0F7F17916055B0E829711C1499538AFAA83DDFD7016B89951DBCF1A8867A10480DACB
-7FBAEFE55A8E2A54405804378F11CC8085F7DD8BC2325268FE003507EC92B0D9A3409B0F0E01
-BB0A05C4658C8B31F7E14D5FC2620BB81AE267E85641F246C7F8682B910A08939EB3AAEDBB71
-7164658976C32A53A2F93B60BC4731A1095F9314F31CF644D6E84A5F6F2D0A33F2E647257EF9
-D5914E315403FCDD5A7592E11DC091296A5E395D89DA8B85E18ECA098CCABC775E64DA9F30A5
-33113A5122E2E2BA128B0BFFC9BEC9EEAF669E88AAE1F659784F2907EF646A58F6644FF150F0
-EB3671758C81C882425249EC0FCF7583B160507EDBDD5A362923A741E4FE92F868821C006354
-25548AEEFF3C7C3094A651F9927CB7D0CF9C895DDEF6EF8BECF94C29816A680323DBA8BF395E
-A2371445E74A09F90FCA22432FB4FF9006AC2789421AEE701D54617ADAB9CF351C2DF459626A
-FF9DD0D5D5D9DDDD2CE5A42A5FF736C5881A86B94B8B20100168DCEA4466F5DDBB73352B0C81
-C69179938DCC461696702A3FB77DD6D7513535635E20788DC2A018281FB7231D95C9B9A94BEC
-E5C6AE242126D38FC6F69D155857C1555031925AA842771478C43F7E2205388949A231B210F9
-696D14D18CD37AEFB0DE2D2AFDF53C84781D5740FC9C969145FCB45787E3E26EDA743C91B2A7
-E10F248FF1B38558A4F1447798FFA9A41D5E5C51F506560B3A33F632D856CAD2BDFD470F330E
-6697ADF166B35B56C6BF3BE3D9B7896E714D84AE2E0FF67ABE9ABDDEF5CC0556BCBBA38F6DC8
-BF0D40C6F4970C3EEAECE1FD5A4C7461049805A2055C07DD1F200A34D6EEC6DE966624F05BA5
-714EF2404AFA3DA5F86F80EAAADDDF370DF0223E94B61550130C45E951ACACDEDEA426731FA0
-FE64AC5065719E3D691F57B9B6CAB4F0387BDBC3641765D01CC4203DAD4EDDC2E1EA3B17907D
-A754E2DECD4BC1D14615AE2C35FD8BF10BCEE295D8D3E2D4940821FAF5877D1E8AECEE7C0803
-01D7121913D27AA371B88FFF45DBB25BA77D2710892E425B9F4266F5B0A076F5ABCCAB3C0155
-95F04440EE4DDFB38823038C993DBAA31E5AD5448DD4972A2001C707372135CAFEED9B5A456F
-48E638099DEEB70EC360E2DB2B39D7A77328C6B3E84E318117FFF84BF85C6D195EA19811965C
-FD9D735B9F211BA3694F3535C4AA235DA0D68C20B37CD60454BAAAECAD3104FABEA8707CC478
-1D7D733A0D51A7E51D8876AB30D741BBB0F09D9A9656A4B00464D9FB215F8A890C94126FECE6
-CC5639FD3800D98D1012BA7B898C71FCD011F32E2E2967187A4B56E8429BC3884E18F96075D2
-16AE15C8FDEC5AC747F3588A89068CD72664D7BF721EA7432C919C4BC40FBF926EBDE67C2E19
-133F967394241852D1B5C7604581CD8AE1AD3F1C52898320B336D30000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-cleartomark
-
-
-
-%%EndFont
-%%BeginFont: CMMI9
-%!PS-AdobeFont-1.1: CMMI9 1.100
-%%CreationDate: 1996 Jul 23 07:53:55
-
-% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
-
-11 dict begin
-/FontInfo 7 dict dup begin
-/version (1.100) readonly def
-/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
-/FullName (CMMI9) readonly def
-/FamilyName (Computer Modern) readonly def
-/Weight (Medium) readonly def
-/ItalicAngle -14.04 def
-/isFixedPitch false def
-end readonly def
-/FontName /CMMI9 def
-/PaintType 0 def
-/FontType 1 def
-/FontMatrix [0.001 0 0 0.001 0 0] readonly def
-/Encoding 256 array
-0 1 255 {1 index exch /.notdef put} for
-dup 161 /Gamma put
-dup 162 /Delta put
-dup 163 /Theta put
-dup 164 /Lambda put
-dup 165 /Xi put
-dup 166 /Pi put
-dup 167 /Sigma put
-dup 168 /Upsilon put
-dup 169 /Phi put
-dup 170 /Psi put
-dup 173 /Omega put
-dup 174 /alpha put
-dup 175 /beta put
-dup 176 /gamma put
-dup 177 /delta put
-dup 178 /epsilon1 put
-dup 179 /zeta put
-dup 180 /eta put
-dup 181 /theta put
-dup 182 /iota put
-dup 183 /kappa put
-dup 184 /lambda put
-dup 185 /mu put
-dup 186 /nu put
-dup 187 /xi put
-dup 188 /pi put
-dup 189 /rho put
-dup 190 /sigma put
-dup 191 /tau put
-dup 192 /upsilon put
-dup 193 /phi put
-dup 194 /chi put
-dup 195 /psi put
-dup 196 /tie put
-dup 0 /Gamma put
-dup 1 /Delta put
-dup 2 /Theta put
-dup 3 /Lambda put
-dup 4 /Xi put
-dup 5 /Pi put
-dup 6 /Sigma put
-dup 7 /Upsilon put
-dup 8 /Phi put
-dup 9 /Psi put
-dup 10 /Omega put
-dup 11 /alpha put
-dup 12 /beta put
-dup 13 /gamma put
-dup 14 /delta put
-dup 15 /epsilon1 put
-dup 16 /zeta put
-dup 17 /eta put
-dup 18 /theta put
-dup 19 /iota put
-dup 20 /kappa put
-dup 21 /lambda put
-dup 22 /mu put
-dup 23 /nu put
-dup 24 /xi put
-dup 25 /pi put
-dup 26 /rho put
-dup 27 /sigma put
-dup 28 /tau put
-dup 29 /upsilon put
-dup 30 /phi put
-dup 31 /chi put
-dup 32 /psi put
-dup 33 /omega put
-dup 34 /epsilon put
-dup 35 /theta1 put
-dup 36 /pi1 put
-dup 37 /rho1 put
-dup 38 /sigma1 put
-dup 39 /phi1 put
-dup 40 /arrowlefttophalf put
-dup 41 /arrowleftbothalf put
-dup 42 /arrowrighttophalf put
-dup 43 /arrowrightbothalf put
-dup 44 /arrowhookleft put
-dup 45 /arrowhookright put
-dup 46 /triangleright put
-dup 47 /triangleleft put
-dup 48 /zerooldstyle put
-dup 49 /oneoldstyle put
-dup 50 /twooldstyle put
-dup 51 /threeoldstyle put
-dup 52 /fouroldstyle put
-dup 53 /fiveoldstyle put
-dup 54 /sixoldstyle put
-dup 55 /sevenoldstyle put
-dup 56 /eightoldstyle put
-dup 57 /nineoldstyle put
-dup 58 /period put
-dup 59 /comma put
-dup 60 /less put
-dup 61 /slash put
-dup 62 /greater put
-dup 63 /star put
-dup 64 /partialdiff put
-dup 65 /A put
-dup 66 /B put
-dup 67 /C put
-dup 68 /D put
-dup 69 /E put
-dup 70 /F put
-dup 71 /G put
-dup 72 /H put
-dup 73 /I put
-dup 74 /J put
-dup 75 /K put
-dup 76 /L put
-dup 77 /M put
-dup 78 /N put
-dup 79 /O put
-dup 80 /P put
-dup 81 /Q put
-dup 82 /R put
-dup 83 /S put
-dup 84 /T put
-dup 85 /U put
-dup 86 /V put
-dup 87 /W put
-dup 88 /X put
-dup 89 /Y put
-dup 90 /Z put
-dup 91 /flat put
-dup 92 /natural put
-dup 93 /sharp put
-dup 94 /slurbelow put
-dup 95 /slurabove put
-dup 96 /lscript put
-dup 97 /a put
-dup 98 /b put
-dup 99 /c put
-dup 100 /d put
-dup 101 /e put
-dup 102 /f put
-dup 103 /g put
-dup 104 /h put
-dup 105 /i put
-dup 106 /j put
-dup 107 /k put
-dup 108 /l put
-dup 109 /m put
-dup 110 /n put
-dup 111 /o put
-dup 112 /p put
-dup 113 /q put
-dup 114 /r put
-dup 115 /s put
-dup 116 /t put
-dup 117 /u put
-dup 118 /v put
-dup 119 /w put
-dup 120 /x put
-dup 121 /y put
-dup 122 /z put
-dup 123 /dotlessi put
-dup 124 /dotlessj put
-dup 125 /weierstrass put
-dup 126 /vector put
-dup 127 /tie put
-dup 128 /psi put
-dup 160 /space put
-readonly def
-/FontBBox{-29 -250 1075 750}readonly def
-/UniqueID 5087384 def
-currentdict end
-currentfile eexec
-
-80347982AB3942D930E069A70D0D48311D725E830D1C76FBA12E12486E989C9874C2B527F092
-5722787027F44470D484262C360CDFDDDF3657533A57BB16F73048BFBBFCB73A650484015441
-FDC837ADD94AC8FBD2022E3EC8F115D4B4BB7B7F15388F22CC6198EFE768BD9FCEB3446EE4A8
-DC27D6CD152485384EF5F59381FFDA43F2D20C8FB08AA27AB2015B774DB10DACFDCD33E60F17
-8C461553146AB427BDD7DA12534BA078AD3D780414930DA4F8D58ABEFD45DB119B10EB409DD8
-97923C6E705479464A4B33AE3D31BFE98EFE259F07F7950237BBAEE4F7B64FFEA83A757FA717
-D50C3298392891BFD60E34A056B0D6021DA3FD9B8B01BF78B0B23DC4EA3B0605150D20B27FF9
-EA5F2524661019D982A2E47CD7C21EE5CA9ED9227821F8D07119397DE26838C11D7B35BCEDC4
-3E011626E300F8249ED846D7B5CCBD8902550CDA17DDF113DF658FB13F816268101376686378
-4EFAC128E01EF997E1A430312AFAC8F3D948EDECD7D5090AB5C864B91D7F6E80256F0E1FC99F
-1102B74C61CC456E2E1BEDF0E627D5F70F738F963A13666405C51EBF23CA077E97570A2F1118
-536B92E585CF1DBD5FBF36A0A5D197538C05304E93A15F3C931168D91D42483EF428B1A3F1B7
-E548FD1D23B80F5523C8B8BF450B02804FB689F915C86693F7DFB1C43F681627C7F027021AA3
-A8179148F1819CD94B8890A9AC63DC1E8A68F254B64F6F173C3D02E3EBC037DA05074017C8B1
-13E94F4E2D26B467EF9DA07EAFD7EA7C86E025AF52C5CEEB835951E67E02A38DAF1E9E698BD2
-1996C1071E84FF7FBE77F6DDF9D06FA92B16ECEBCBE0370B5CD9D450D0F2906A105CADF93ED0
-0ABC7FCF2C996E7C15E452006E35390D99B9CFC59925F1444C20D9C3E1FC291032CF9FA9D580
-EAC1BC956334DA391B870A757C3BEE6AA3E005B20988D85AD41A64402A9726E3051FE2207B75
-21A437931A774CD81BD14357BC48C26AAE9DC596E2AE932D2CDF44BE37B4C1E1295B617F8EF6
-3CABC67C8A938B2A220EEEA74FEDFB346776D18855ECB95CDC8B9DB6ABB77CFF943A70AA730E
-2C658B257FABC2FEE2EE3E493EF0414CB723B3506C47754012F5F7237A02093880174A792EE5
-FA8B752CE148C3C8CBC461A2A99CAE1170DBFA4CD9A0705A56226841D063066645E020FB6881
-FE939F56D3FB1D699C40E84959124895B9F3B0325520BF6085607A732E748E4E15F24CEE1C62
-6AAEBBC0928C3A1FAED760492751CE835DD1BD691BB340583EB9B1E1B748D5BC20A81679FB33
-9DF3E67C71643E9A48787193064D35FDAC1EF3B622053B386EDE969C1DEF9F543F006A0FA0B9
-A0E6401472559206F7CBB7F18BA80643CE2C061220797AF1A7B71A343BB98606FBA19CDE56BA
-E4ABD7B13ABBBA07B51D859CB2476F848C9C4A80CF7495B2E639A2CA107D0FF8E49681BF789E
-F580A791BB977848EB65906B7D3CB9B357B3169E05885B0A106D3594019E86121AADF10DA23E
-6BED3501D54C1EE2A1EC6D9F1352810E0B8DD8B1F315C8BCC399303A4FB8C52ACBD003E3823B
-FA9D77EBBC4C89A4AB5D17BF8CB0F3468AA2907C363204499CBE4A5FBF026A71874D1DB8B074
-2C50220907AE017F9305DE741EDBC2070D1F55B160262A864D9F423957B031BB1CF33D797E85
-11660D988C9E9F13BAD0B3A75DF70F06F2830CDC404EF0D31A9D271AA8720BC5CDF7BAC10555
-5C2BCFDB7112A924D46198A72B2186936F8F46FBB51086F8E217EDF73CED642B220645BC2143
-305DD1FD72201E440A93F1FBFFF84202C4CF49F6C474355686F5EE03233E82D45A57C824EB46
-3F949E52E9B5AA6FD08BE515E63F2698FE64DDFA4D4E6F7365AC8F9BF1190F1CF08E39D64A32
-D056EC24584267F3B3F3791F0635C09BF772C0491D571043337A458FBE782269F35E45D593BF
-34EC6912440885EE5E3B3A90B4244967A4AEC638E79A6A1B4FB49FB38889E119E28955D0F8C4
-303CB14F25FFCD1E1B7D07764C8A87140C32DF72B9DABE354A12835899BAAEC995A6F62A9414
-A86AC906DB64C9E75BBAB7C69FC870DE60A4CC338174C894D98F1A8FD3E538F60136C503D787
-DE13C6498807FF70DFF893B5ABCFE38C7ADD38FFD9AFBEC5A88FEC05DFFC34EE8066B60DC102
-05020224D15908554D18D34460FECD5AF4C7B090219D75EE74BAC0F29A40F3D5354C069C391E
-6D7074026B36B52229860D6F2758884DA1BED0D70473742963748C2D0DF5C6C8B5DCB30F88E6
-B1783211837E29E6E6B108C96ED5C6712E91333E7DF957CFD111CA9AD2753DC5BFC68BFB6D37
-2AD1CDD9B9948677DF53EB2C3AB9A3B4B6DDDA6E2A90FE4A6E22450A9DF9B08A4DD20DE27D37
-032C90A0772F7B602D7F5934E3B23B7FB827FEDF38E2BEA2F04B51091B30580DC1C6B2ECEC5A
-63661AB969DA19FBF6123A399053F680D687CE94582AB38189FE50338157DE2F1C37CE0D821B
-05B521CDD37F7442CC7BE8AF1D011948C78EBCF7511A70312B1FDBD153EADDDD2BB9FC57B07B
-80680B92C69B32C3BFB99994E40E9F67BBD5F0431FC31746AB7C838BB58CEDD6D09E1D68D872
-CC01149095EC1028B3FB2B0A06A0591F822AB95583118E2F4BC32C870609E190D28B1F4B467A
-680CBA25759B4BECB61C5BFBAD82D27AB53265F2049D3BA8CE82508D0CD2F2137BA26A51D651
-61AF43E1F468070A821D9FCA52BA8B5C97266FAD6A7263EF54AAE62B565F0BB6644B47EB5C5E
-58639AE1D907DFEA567A630A06CD80365E6F4FBB78AE3030454433DA19093EA926E8C26429E3
-89CE695B01B7CA27DDB57E2C7CBC8F6E81703CFA475ED555C923F111F8F7DB9E96F3D7A17F00
-D4E5CB0355120286C8AC5FDD0B235BFFF37AED227E1BADC58471147E480409AE6201D66B5153
-588C0300789886B3DB52E03B3BB51DABEA1B67CBF2448D7051C3581ABBE1CB6D4233099B1FA4
-4F03328667838B6249A68F16D765CF8D6FA83A86D24B896D6672B51B299550DD306D6762C4B0
-768352358C7653D8641F10DE410A518DEADF7D788B65D4663E2215B9A814B55D3512EA223833
-6C9026ED0ADFB6F9B07E0F79E5933967073DABA142D9C988694761E120370DFF71E0CBC126CB
-51B4BB97189C7A415E67979DF20323F97E3270E098CE7501F83B05F40FC5F5AFE5A78CF05E99
-172479F98BC4CD3B184F73E01F971A7EA5CB61F3206E4E38DD971E320A691F586C7EA399373D
-0E92CD5090E798B08E5FC71747B9BB6FC10DC5234C720320AFD1C83FC0BCC99C07ED99D5DA32
-AEB68D13977810424255DFE60A29ED55D3C034A6D33C84504ED2510C7B3C4BAE0C22E7C3F489
-0918AE7CA0DF3876FE93853C82AE8CF04A455D9F54EFFD965CA948523CD541BDB5BA14F824C3
-6743B47BBB0CB596940312CBDBE9734BCA8BF186BCF85E785138550CE03E5554C246D6C916BB
-D62A3A13FED45536DA1F33FCFC5C0BBD233F652C8D95CB76A929ACFB744FAD60A1AFDEE91C71
-8B513A5D882DE757CCC554EA5C2987F574B734ED086EF8C4326A2224D9036C40DFDE09C01E28
-B51CB8F7BF1FF6EE74D415DBAF2C8DFF2D4FD01728B4938B83419FE1D3AD34A6283C2CFB901D
-E96A7AF6504BBBACF74B9C4E88338DCCA8DD7C3CCAF92086CC3D57AD0C5BFFFB8E3493B08DD3
-2312F94D56FE894C7FD3D2FB5C17B741F80E02FEA98BAE3717BF31928747ED162B74EEB7D0DC
-4A4393644DF9144F878A46D8DB9C84F415669BADBE80F7081C34CEA23F35B2C9F7B6283F9275
-84017338962680388CB6AFD258DDA685C5B30F53CED93785831F35D432E27212C8B058746D0D
-2C7D22026C7817ED1FFF510A0F33A1306BF7E45B4A7B46A9FFF672ABD69A42476081E912F338
-3358E275652C56A33ECBA633F775748655B5742AFF8AD12BE746E7EFF80382DCF96CD5734F7B
-15C5177A9279D6C3CD1440D6219071F41E2BE7D0B18565C83B6B37ECC9EB9677DBDE0F0E49C8
-8D245E39CB3B09B659549F2E8BFD1AFF39F8825F98308CD7BCC3EF382AFF37F849D0EE58D952
-E06F6AE55526A100CFEC9C59C90D78E885F650164FDA8E8F630D2F4C15EDCE829DB22BC9E23D
-8E41E8F97C3154CBD5C3649C67AA889B9E016B6198A7DF68F60A2C08799A36AB2CCAFD8EC406
-38E71FB509CBE17A67C4C82ECA9902E4C72537B50961A6B10246442AF4F7E73C4C47189F25B8
-0DA2B00760274C1A8ED651ADF77D7ED3CC7DA0E9EFA9A70692E4F8BF1B1BA7A92E840D46CDEB
-1A2F0030BF2F2E29100725338C61BD9FAD588B81BC735949D613EB5CA1101DDE4D7991BDE01B
-40B7F512C8ECC5C153794633C5346156B16F334E158F68C9D86BC97E1C968856AEAA7083E47F
-4B4C890BCBD776522DEEEAFFC35949452217F8183E5670EF43A858F0EB30E4018CDD6A314667
-BD871F907694A1A8370B5F9A859FD165E9C4477FC9C5474339F6F2C93901E8DD82028FEB1A53
-3ECE23354EA9E3969C32780B98ACA5F09DD9463A2B7B1B328439E308625BE4431A53AE74F763
-BBEE179D218770C77F8CCC87ECA55608CAAC0110993D8DEC82F28851FE0426A43936E8C516BA
-9322022ECAE9A29430DCA01CE2D0019D8D5BCF56A5D3606C417DC559D030BD7C83D5BB473D15
-5938F79416574D4E265FB247AC145CF90A2B22871A86A0D573BB595596D9D0EF34EEDB8E0383
-C49D75219D0C5D898A28BF72D2E7224B303086FF7739264F52FA56575D7FC888574AC083365F
-75821786C91A15C515D03351B78910FC66390E4A6F466206BAF6817F1D95745DC26E5F7C9C69
-BE7A029B9CF8039F9844881872289F9893AE57158B88525B40872368E036C52B0B2F2F41A54A
-72B5BB5B10607783AB9EC8768B659771E73CCDB86E7AD6180426D3BEBD7AF292D72524ED16EB
-993F210A57BE699D3A83524BF16C07C83825EF720951D7835D7601A98B3587FF97DF4C0560D4
-185225195EF4C13D0D678E1B8DB98D0D4F91434C4F2A9DBEFB742BE4A47440909B31092D1ACE
-ACA3B0FC2EA4537C73509243C7DF0C18C704781EC2501AA1D574E36BBA0C4F64F1CC9227BF4E
-9719DD0D519DC5FCDEFDDAA6E344C5602078FDF6A37A5DFFF1C8654BF5A149522136D19F24AE
-D50B73BC442F2D70F1BC67F416353E19CD2BA4DE87B070BFAFEC50DAE947C84166CFD8894B6D
-ACDFDEDF5B550485E9E0530B903F9430BBF77D585ECD580B91FA6B2BB9FCAD50D73DA0628723
-0993AA5DCF929C7910C910AACF0C27847C0396CB0F4F214E1D39BBBCD9ED9DFF06D52E8A5287
-1412228D387888899FBC871E8F10E888C22E073CA6AFD7EC67C15433F77458AB54B4380D0C0F
-1B79AA392E2E00D88FA427D10C82D8C02A08C7B6110A8B74277BCC1909A3CD682137909BC769
-70B993624E93D9FE671F4EE439E08897881F0F53250B020FC3D90D04CEBAF7B132C1175CE407
-9B338435DA92751A0575B5CF5C1F7E24C882B4E9C6E802C6714203EABB1A5F9263A3721D39DC
-CBB683E2CA2A177CC60B42A857AA7726DEE42C21343814C1AB20E5FB512612CA727623B2F527
-5610384E6DAFCD4BCD62CD904D0DBD9C0921AF0911C0F1F8FA378876025DB9ED5E637FB1B15C
-187364652E54CD9034B3D13A8D249EB32752DC5EB9E642807E8FB53377C5BD1D703CFBD0D839
-41161DA750CE4F127CB127D0BAB979AF8675CC78BB03377B89EEFA48A542DCBE9FA5FDE6D34B
-7DC0D0404EFAD99E657734DC314A44EED83E56A308E90F4ABA9534A2129CE3ECDB0759CE9D11
-A652AA16F4D3A93094DEBB334A74FF1EA8E6599A5BF2AAD29B449CE39C3383AE8BD766C90890
-B11953113C58F4E2CDCB23A508E8083FBD5B71DA34F6CBE724944481F3D0137E48D2C6251C9D
-2217DE57FC804A60C61F74CF9069DE57BE29B8908DCD7548DB124E2CAA02B66A7FA3ADE37456
-5ABD151BB52FF77949F024DA1B57E3EE7ED5E769D142FCA015BD2E35C619068E7DD217F29795
-A470F3402216025833167CF382F4F6B0380DB59E7C6AB570E4EBA7498722C73AF3868D150D04
-20B3384FDE55ACD42C2DD317071ED8F034E577C7B3AB3945BCFAD2223FF8C8B699EC49004D6D
-5184CAAEE662262B27CFB509A7851F675A1FE0C0B216693D004576B36295031A1BABADCF4C9F
-1764096672F390C2579ECF451048E700436937673FD4924976FD7235B553B4D9CD40875E27D8
-6DDA7487DC8A4FEAAEA08895243EE981A203A971F5166AFDD8B0ED4F9DB19F22A4DC98CF0847
-B47B0279091670E91A12E3B54AD7FC4B755DD963E784E915741C0CAFF91BCD2E49A1156B688B
-D959F5FD2074F12DBFD15A38F3E39E5FA97C8F7382589FE5A6C7760867320BFE877681AED666
-A247C7296EB6187266F7770A8C0A838B006E7EB66D293D7C79985533381918175EAA47E3C3FB
-3DC0DE23CD6B4BEB3982122953F97E7864DB09B0B9ECFCD4E568D49978E34B5CECE02C588776
-BA04367357A2BDE614F50E9A993D6B9447A48AD2FF369AC56096E904C8B6C9259C2307DC5A2F
-FFDC6B91E0D03CFC8A772F9ACEF6FF7E3421F21D118C3715104624C666C37D9434FA246E55A5
-E02EC9E6D755080AB5262674016F10F3E46CAADEE10DD6F86FA71E9F430F0B55E08F55BDB6A1
-787A9E7F82281E88DB5C63797B16E671269D95E9668E3A03609129ED4785816A0667F2743AA6
-0064273615BDC0C23C93F9E693714B0DA0A703DDD7D597CC15343B60F1B9259089A23BB9F0EB
-BE1BAD970F534A2011DE3CDF778C9BFF72F89838730957339063245A05D97E4CBAD5012478C2
-0CD520882C7EFDF205C275C4289130A8C86339521BD589801A59768B246093A9292E944C47FD
-B35B99957BEE3B70CD9E451867455B1A1B4652D145F7FFBB76B3CD6B130C234E89FB31B46EA7
-D9F2094AFDCE1BAAA45F38F7D4F7874AAC3BD920456E680025EE3F2490BB2B20897307771493
-7FD3FB33619E0A5389095486331ED2AE83FCABA315763B1A76FBB76D5857174B2CE341EAB048
-82DB3A3AB3DB63E017AF5DC6DABDBA6B2857C5AA32BDDECB27BCEDD53FE6FB9F74F941FD5035
-BEA495E8317D3FD8F501088724AD202CBC4F39B6F5B2F920C333CB4F52DEDF07D539361CA96B
-CD33BFCEA9704B65205B8B9CE9496AD8D2C235D49673E51806081C8BE70883B1F14F07FA40BF
-F62C2D7F4C9D502B3A936D0970DC6D61147B576654EF5126D7B6A7668D9670CF58D8B6253D16
-25E09F588178EAFF26F4CB123CCEA0FCDAAEB121EFA30AB4FB0CDEFD30D96465D9B6FA76C3ED
-28D140C855164EFF974C912D52B032301193F2C4312EC52B6DD21B206A1EDDE78692F885D714
-279461BDB7DEA004F365098A70A1DBF047A320542BA0871562F193D85DDEA51862FBD3A53328
-2F14CA6AEFB0FE438FB8979DF85EEB7873E6BCFCEA1D593EAB8D427B955D22117F6DF875FFEF
-1A7611EE091A56C9C2158F6C4BD9B72FF00392396B81F039F635929175B8A113119961E534B7
-AE7694BA76583FB11B773C4E87626AAE45A528F625EAFB92A1F8901CFC8871CD091962F75087
-1D20074B72E906A6DE6798CD250F2E61ABEDEC3932225D26A04D6376526F7A9FCD5878317628
-CE98F399BF8E9BD8EF5F10C836C9176CD4378386EEC58C5399A26294807904EA947C389941D4
-78D3B5A5712B9CF3983FBC1F5553785EAE01DF59C8BD022B35DD57D4B5017AE12B16A059A890
-FBC927D637665581B1772C778DB391D9F36B69F93AB3834210C10649151484182B24493D4A0E
-58D96FBE9A2EAB3CCF5750B61CC336BDCD8A0F509F11EA38581143D7154FCE222F4DC1895DF4
-2E95570FF23151367B17D2118794EEE865E7905A81991496C604ED4AA40126B53868BCD36E44
-60AEFB8CAE3A10B6751C97ABD657665D39B2A26DF751EE6024C3AB68443E3EF38724DB7E3179
-1B71C74D7502D9D88864075BF84C22DBCA774BC7D41CAAB6484EBC6E72F2AA90D096E90F8E36
-F8F1CC5F3BC8F3DAA2B203F54671159C4C734B0FD630B4A57FCF84C902D70A193CFE4A973B78
-4EA774CABCC4D93EA0BEE60AC8FC8F5D19487479284D9E095CE37FAED476519C45DD5E345169
-728EFE63196FC6814E038CFF4CAFC1F55179C3F39E0377A53F3797A5F4FC8AA28C469BB2CBD3
-065DC5DC0FAECF09A5EAC62C1E5A5376E7503E05454F02C1CF5D25732D19105B3DA2ABC4ECCE
-3484A82E69DA565F0EE5770E17539EE77EB54DEC5342B9EC8C695E66C4F0826B5D9F1751B058
-FD89A099652E15278C813D784397A56CB892D0D89282B8C32CB8B40711D0F966C00312FD00A6
-51BE42DD9490168AE709A07ED038964ADFDDF971775F67F3F4A65B6C1BF569049E3880D5978B
-56CA987D2FF853DCFD8DC33AD0C5D58579D09F431AECDBF339A636DDE2BFA9EBFE1C7D36F802
-FC33F32E1A8A78E26FA38822567686F976176312AB34674213DE4CB86BCC0004F945136C677D
-0C01B9D3D0D8C6AF036A38F52FFF1AFB2CAED092844BA23F1F617FA5B9EE4B4849D2BF46A495
-1FB913350BD62C54B9E9199EB378E5963F845A2A884278B91942CA1284AE683071218A1D0F52
-C4578BDF9986686F907B248CB12C3D635B8F9E084E2196C2BDC750867E76E518BA10B4F75AC7
-67D1EA322F4FAC137A99F0314A2ADCA931F5FABC6ADD4AF8BD5D94F1FDB30E4BAF1F9B103CD9
-C4A82789426E7AF4F17E8786952F31BFB5FB70C33E5F37DA714EA075BF30CF0A17E6313CF4CF
-9A17EE43264B7947354096827E61587A23EA1683B7DF5D1FC499EC6F4B9F0F217B46D488E98E
-29576A4808AE615AD8C0F1D8CCCBEBFCAD15D07BCFD1CFD1ACB8F66D6386329FF41414D4F1AF
-268664C0B81308DE6441B8C19A8939065E233FE6DD54A08D61564EAB4A3834A96D8251B25D06
-BCD666EB2675B57E9CA474D4254DF75ED8F5EFE0228301F69B43E5DFFCA737FC80E7E1BBE117
-00A92C094FD099C7B7D958C2C714912CFC6D49A77A0484F50A178417F35BEC164ABF80859B7F
-F8E0C66AF7E4D73DDCC5F71627AB717F62981FF8D5E72F81AD686781E2B800CDB5424AB210EF
-933F247E87117F867EE035DC0BB514B3D689BAAEB2CB6713FF2F91EB164FACF4A81DE864796A
-D289EE2CA3F08C298B791E36FCA10DD4D5E039A5A4F812CBC3B57E13DC85286A92066BB33A4D
-3C5B86708E291CC72C53EDFD2490E91866234F273118F811C2B09AA3BBDAD1EC5C56F090B1D9
-32E6D1355B28B57B44E6657FC637902B9B672E0393D35B72043E31613939CBE2C70A18A409F5
-B42CF96820F235E667A46A9DBDE68360D5258D6DC7A862A493A4CB6A5F8985ED5A255AD61A55
-D09909FECE7AEC60418A61F4187CD06CAE398D1D8B2201AFD33C732AF91FEA9355ACDB87DEEF
-6718E7F8DE8AC941BA112D85AB8C56146CC8FA28CC9E9BA455C30B3E04AB115523CCAAF1B69A
-AF889D46AE0E3762CF3C9F9800B45CD435F1B6D4E5263447008FA5E0871AE94E29F3D6DF7BBB
-CD8AF3253B8A185B4BB275D7EF5F73CC047CA8891660AAE38412E64D5F561FEB14D6084E4F70
-FEA9F8300F16C29059F73CCA235A62A5A371046A7D31DC1AFDC68C52FC8CD5414E8E1CC1A950
-8F0746B81A7B0133380AAFA22897E36D8236C0C72B5039D60FE1085F3C258F5B9EA49AF4E8A5
-C2641871315A01F79F874449376AFDE2A5C5556214C6B3621F076A126FF346BEC23D4954A3BF
-975155D45B8AD5BB9FE5192F5209F26F1BE23DF5F81B53EF6B6B46ECD059E652365DF778A9EF
-8EEC53618B2FC491BECAC1DD0A48148E67678F13EE0C2DE9D6AD34E840D94AD189444D84E365
-ED09485F9CF0F90126ECD1858437B89C2CD64E2DC38008547C763530129E3EF3BB3C530DDD3B
-1AC7EA497C7EE28BBEED64F82DC94CCEF1678A87D80F2A26EBA06E94FB891A70F1D7B7389262
-EC472546B270B812586BCD2ED22B4454A779E0170D2C1B04B1C7727AA9DE8CDFF4B7F94847FE
-96F0A139CDB6F48D5CB67A5D704DDD7390C5775605047970A028A7E16D393FDA41C160F56D67
-AEAE51B7FA3C1A819AE167EE2834F45D0664EC790BB3E8B17035118735A3EA6722C87A847322
-09811C7A60AB5EADC11D2509BDBDD8E9321EE73578A7CC7A82FD0F23C594571D0382E7BA5BCB
-8764200C18C7046932AEB95377EA45C7FF01D484644AF472F1C08DB9EC14F9E1E779DA337307
-4A0428833954FE9D2D45C6D9220F57576B96DAA4E973538E1E0D77317521C8D69D46B12AD6FC
-13A7DAC082FAB9C19EE241FF06D5B0EFD2005169653B51F73FAD0694F0148D677BB09E194613
-F0F71DFF61B81F8C412A1271B156D2C1FF4E5757CFD70EC8A5D755A9CD85F42B3084B7987893
-E7EF085C8A81EA9A00376020BCA88038F7670C5CB7451801B1209CF7F544A68A6DB363E9B3CB
-6A035F9685BB6FF7243B12D3294C217D8E4AE70C7F7D37D3A6288AA71D7D942789FA8EC2EFBC
-A8025DAF891B01C8033BC11FD4404292536B45AFAA134EFF23EEDC0DD016AE49B482AACAC0F5
-0E9AAE05E78D3C06229916E03FF1D43E6E9377AA2FEF7FE8C864DC7A9E07128126C4FFC47027
-CBB9AF9831783640A1673C0B4EB40EE97231C1896A596A0A524C3BE6786F7A438935BD8800B5
-90991953F2F1CEAF54F78E59EC88810E233FC3EBBA633610FE2309A710572358820042AC63EB
-278227F5848AD9AB3D40B20D0FF9A9B04EC03D8D28900D48396F1ABB786A76E0E809F7EBCBDC
-3105A33F59D7FF296038A524BB6B2260A222DA5F48B0A3E535BB9A7A8B37551C521BD7944395
-F21ED3402C79E3CCC6FAB5F51E3823EC72402D5D837CC6C20AEC5E222820EF6B459FA6071234
-9859B6E1122B52310F1A80EEEA6CABC47B8F35C56CF461F526E502290201C73E4BE66823159E
-25CE6E8A70BC2AA68198907D94A635BDF9EE0B537005E35048B8304EA274FEE5B88BFC6C9622
-B311875965274F83C89E2CC9C8BE4D2DA3FD358CCEAFEA013FD5A150160CB351864593C5024C
-B12B08F3EE1F7F1F8828371697859845AB33AB6BD3E6F6F298D1B444582A3D1392A95A2C24C1
-F8E743962CEC54D09DF8D3324173FE192771B6B5B54FF7C2887A719BCAE772664511BEF610FD
-15AEC48D5BB0E959B414A5C2BA4EFD335694037EDE03F19ABE4E67171D2051883DDD19756D7B
-9519004C5814E055C4B94F02BDF8D564770ABF2F569A5893B374C8175447E319851677733D64
-BE4C48BD3B3A2CC4EF9468A7065E3F84E8EDD061C5DF676EA8455567EAB84416564421BFAAD6
-A3A8048C4464CE3AFCB31296E9500B638B34BE54F63EDE677C32F9BD7DEF58C29E50D4600768
-E6F86DF06548388B728169520D50C5570183E008DDBE2B758977FA6C5B4B10DBC6D1A58C8426
-B901792BE1F2ABA6A2F64DA7954BEF6CB04562E6553FCB46FA6BEA9086D0D2D7CB89CC47FADA
-70BB7130C89CEDBF4754E971D12B65743DFE7CE7301400FE331B87B9A231D0B3ACA0F63221B7
-9C59FFB89E2F486B3CBD8CD93115233F4ACE35644BB9BCA9349B5FE29F6ED76FD9256C66B090
-E147AA3D32225C155690A4FD9AC84FCC31C638ADD9A472115D49E89C5386F36B95A37AABC593
-716A4318849476DA94F4737C2CC32FE94E06B5C827A8F6EDF872B2510EC5FD2E2FA5AC99E5EA
-12776BBEF13C86D418E8DD2C9E6073AB2427DFFA2E7821674FAD1F061A0ECEE8F96EAE9FE8C6
-67E2E87E394026CF743D1F7F3F2F871F6BA01EF256DA9D56ADE58118A1357418E861C4CE1606
-DA7A0BCFBAD58C966C82BA2E67E804626B08412FCF52F38F14ED1A8686A4A1AC7480EC72987F
-EC346048C965E59B2A45A1422CF8F576D93DA877AA733D93EECBC16C89DBDB34FA6CD05F735F
-9DF62B022581738640E2343C857D64200A2DD293A2DC646270A4854F25618D85978FF7E1D172
-BE6B35B57D40F76A8E62CAEF2C712C12A87845371879AD55CB8F1B71615E4C7B05F7FD0D81B7
-D05442192BC08374BAFD865B29BE14B967DC391701A6F895ACAEB2AE806AF97184377D9ACDB6
-B00A912C4EBC2830275F009B7E615C966DFA0C5AA2063D713B4B147231048CEA61C516779014
-AE32994A7B42A2CC5A86C1FDB77B4B7402121FD8752341FF4ED6D068D290B7BB6C4F5B11CB94
-5525DFF9C613B124B440ACFD97AB66F13DFF4E31E8A9B21C2602E5708EC97E9A5ABC9DF3F38E
-F493BDAB5C242D2F6C3E07F492929CF84E605B6E5E03ACC113629B4181F5FAC4A4E547A70CD2
-C486CEF19E3722709E183538C6B191ACD844C5B8DAF13F95814286875D07915DE27438E16663
-6901E19997D5ED64F7E5B1944A8428E74FF3840E08A1643D3C2810AE7E00B1BE32D165730E73
-118F0E29A2134EE856E472736483D5037CDA618D3A4CAE8C9B81177FDFA9E841F0CB381A0A31
-9558F8B15AE49892CC0A0C5DAF3E31A1A24663D3E1BE4859DA044ABE199632D41B4A7E61ECB8
-10CFB2214A9420AFAADB4C4334A90FBD09B877796205DA842E5FA0BE4DC5BD0D98D2140F90D0
-C0906745697C36D973D09318A9050008EC1624D7A6B8741A39F0FDA24EC4A92C918DC8039B52
-BE9D68757E150FB58DAEF9E0928557DA9064BF5D1C9E513BA4520876D34D286F97EB58457D59
-B877F4BE68E3C6F56AE6D702B4031D89A0AE598A04FEF232A2B1578869E731275BA04024E5EF
-F764E8E393904A72C7C982AB8E2EACD1C2FA4D8563E3F981EF49CF4598C5C481ECC1EF953F7B
-DBD2C11D989AB7E38E46614DD428A9A99E97D15C283E3EBA9CF389EDB63DA593232572250138
-13DC5FA6BA6C05D4D3130CFD56C25842CEBF75746DA31A1ACDE3A7C0ABC157D0762952813831
-E7EDF4B5A1DD415D0DFFBF5BBB05D5B39C412E82F0B9A6033C13BBE8902321BE12803343B852
-BEF2C3E9DB8DB30CCA0238F8C9FB012D5E56FCFBE78109AF1977B9431426829F6D058AA52C3E
-985DFEFF87FB2AB0E34A9D1C39FB35162557199404584798B837E2DC32BC99AEAEFF6C2B0D56
-A0AD82591C309FD5697F157369DB9EC3D4BC935DADC3AAC5D1B042F39D7E2A819D5ED8BAF717
-FA859780BD343B6C53B1DD12CF49EC709AE2215559760FA710651EA9FE72B30391126C01E756
-16BC45A7C14F6BAE90C69095E45D547E8D2F898472F80D010C4D42047A7C53FD1A5DDF5F0859
-5D6FBAC3101A5693956AFE71AFD53FEBFD9394C1C7FC21434979347BA67464F103ABD4072940
-03B5470400E05B14A42D662D8C0654898DB59D8811B3F4F8602CA3506AB451D530BD76FA4800
-9D3370643FBD45641068618AA82DDC95575A533768F24DA4E3D322504A5A1DD8B42A3B289BEB
-52E6FBF6A8AF118F56DD5A3D9B4F32AC0581BFA108FF1CB445905DC4595E14FA22540CC2949C
-A19281AA0AB8696F3EA07D6FA69FC3806BF581F421DC626C80FA6829EAD87105F2E7F0917A20
-E5010634647618F975A5BBB6160637FB6851690182946EF173581990A27A3043115EE6585BF8
-396D10DD92E5DB79C414567156A3579B8DD847F3006A7843D9F900B827891AFB282F75B2B2A8
-0AF39343D2C75C999B7ED84445FFF5667B038ED5613904F1942029B57862DD3969105A3CCF22
-452982EAE1071F35B5660008520C9DB31D6291431772075B358C32458AD6C9A7AE2007BC24E5
-E870B6C97F284433E79D7317353CCC058AB78431CEE525176D990A496B306E66760ED881E76E
-204632CA184A467F7DB8111F69D2EFBDE3C298E6DAE3E80D0439E72983B9819B5B89A63382C2
-26581BC11D6283E6AFF1CE57C827D43D310409572F98725D9325F31C2D212EA1F504615F5F22
-173A12BE11160970A4D6A875D8F296203091D578E66D0703F7002AA1468C41F6B06192724915
-98CAF9A94C450308298746903103BE00532C243D6C5A65CDACF81AFE72F6EA3A3455188BECD2
-CFC1009E17730077D86134D5876491F8A934807F5A7C2E00E29C74DB7EE5957F36E409478E61
-9793289751AB85F054B4757B01C461B45905ACFA3DECF21DB6A5174849AC58ED65843D8E6592
-9591AF287C2F436ECB2A684488A586375584CB71A7E804E8541A609AA2D839D4FC029F7A127C
-40D0253502B2B2B7B2187FD84AB2E97AC1C4CF79A1315537DF2677B332CCD3DDCB8F3254262D
-985D2466CD92C6A6B60DFAA01B34BA5B4453150D6D77AAB01F228D313DB535B454B6243A949D
-B45263510B10D320E1FA0E6F9E16965B3E00FB14D8D4C0F26FC58D4C79521DAF016008D2404F
-BADA783C9743D9955460998A3D523731739ADB13F5383141D4EE06EA06CE58F59D35D6BF383E
-BCB489ADF46C0554F254FD47031CBCF7DA2764E5D9DB749DCBCBF3EDDDE2C2B5EE25F7519778
-24ADEB1B2F0E13DC38E04ACD498BA2A90B6F6534EBB643DB47CBE536FB5D8C2523F9993BB5C5
-CEBA248485AB36243A4B7808832CD9A872EEAE77388F5558246F562DC986964E9348FD045E98
-38A17C9B18D2804BB73FEE8754752AC06C5A735396A8558E4F817FEAA83785E002798DA4638B
-ADFCFCD6593FDE22E128A048805238D5F20A0CDDD571ACF1AD7C6C923A6227F0A4E8664C1810
-E043622833BFFA0326717EC4CC04AEBFD7175E00F0AB417464AF967F9100310600A7ABE7D3C2
-083C118A4A633FA1BC3248C4076337CA2ACB4E08A66EF0D71ACB26D7575B53D688934460DDF4
-A1FBC5FAF8636B9655D820C901DF786CA72FB8D241E4EB1804E7367E621E00158B0C4E505C8D
-F0B56BC7DE5A002A19A070776DEF9F85A4E9BBD7D41135DEED3F880C86E3937927272768B821
-59D78D165BBD9C00F9D4FB7870633F83A58D060FF42D727BD29AA7DA6979607A8DF61F95F462
-BE8AF94805C4515D8D8A837B1722DF4C551306AD8E7374834A5771C6478E3340EB7A6B08DD61
-42359B5FB3D88902F0CA1079FCFFDDFE07FE95B533E27703DDCDEFCCC017E52EAFAA39840D0A
-A1D9915721F9CEB67BF787D433D88C9523DC46C262A525E901F1D230F8BAC74ECFA6CA4032F4
-19D08B542AD2BEC53DFB2666663C939A67BF946C588C357272F224F01C90E0FA7E2E0665C30C
-A404199335F8BADAD46CC9692C382482B157CD21E93272CF161246BF3D4AE21FB1CAE43F92DC
-7C9C4E648354B02BF0A93CA94042113FD758EEB2C7F57BE7BB67070E05DC4265E22C072163E5
-606E1D2EB54CE8E269E2FF07ECBC0B4CE6F0F19A4083976B3CC613F2B17CFE112FB35B17599D
-6C3965E34D9CA01A414AADFDD36D83F09612CB708D535189B6CAD0F56A2A6E7F811CDA7332C0
-17818D3851855929C4F35278560CBA999BA88E527E0020FED2E52AA1455BF3E4063DC6F7AF67
-0285D052B2DBE5CF4320C17469582982F74309F0BD9217BC750265151B5863F580FF97A3FC67
-9E8A1828DAA7DAFA4868016DDD3907600C59F3E8F68A2D82509104318BB548F22529D29A5C49
-2EDA6DBE8BD515C09E66F3DD5BF56E86132F1FC2849AD057325E5FF17FCFC2EA2DBFC8509817
-E3CD20EAB73691CF7A98A43AC2DDDA679DAD814BBDA21A015F3511ABBED23D767DF054D586CE
-8502EDBA8634AAA8F93BDCF4E629361BD7CB3285519A04E32515284149042454FA630C68E9EE
-4F0C0288EFACBAA180AEEC1798EB617A94DD515196EB6A8E9BB0AF66D53B25A000F8D6AB6818
-6B71445EF0CE9DEB8D71A2756A85F3D64F916DD65EDE488B03E646EFF6265F3E354B66886D03
-DEB3F2BD51F9DC3E9ADC418F6FE1B7CEB6F73B4F3D6B7175485855AFA03D60BFAB6A9C7E0DB5
-7D9E16DBC9E909C7896B178B1696729FD4287178BE74566CF15B3324148032B62EC53C206408
-7F61F28BAD07F026EC3748C6E1E3F9BE6FED8D450FB2157C257DEE8A0245F02B56119EF20530
-10C95788BD7C54160C4F8F881EAD72F8EAAC6ED5CF55CEF444C1B8DA4174A4F5ABF45CBF4ADA
-417DB717AC0B8FFCF7D98499CAE43BC97FE73DD324D61CB2C2B5655F9530EAA766D40C6C4A05
-E5B5FB04F34E94C9E4674FA64AAB727F6B886FA60A734D43201AB6F83D14ABA4B54B9C0D3AD9
-3C5EEAE88899D908865836DA3D212C32EC69377E2CA10DEA84B9B3EEB710A111A07149538EC4
-9BA53D3819C79BC22C985BBF2E2503691F8A82E3E8258AFA8B6327E2EFEF4764BF75BDB6A48D
-A91C1960EDF589384E0E270CE9471EFFD29464A231044C9EB77A266B4BFB186CE117424DADF0
-674B377382CAAF8FB20B639CA0B338A971446A5EE6564A34F74B3B7549CC35AAA36E9C3E831E
-5ED3AD8317AC28DFDE3C60F6193504E2005D7B94BF8A032CA8B6D9B07151094C3508EFC8A3D5
-5A219847A6DDACD9828A80636F6BD1019A20DA784C28D9E72D5E1052D80BB4E9209D839EBAD5
-9AE3BA0735411E48BE1E342B11883D5F53B58373E33BAD3243D1D0D298AC8E5219F7334CDAAE
-49326A6365C7304ACE7E02A75236B7416D2A659B755D221C42BCA81C37140368FE72BB33F362
-8F7BF8AEE504FDB92C57B868D5F847FD206D40E40FC5D19C67E8025C28F20514488DF16524A4
-3F24E297BEC579B2E15C35AA1C181641D637EC350093E5B7A50CB97FD0DAE7E28F734332E716
-1999541E32906E34402399ABC8B93C0E25D2696B50946D9900A82B2DAC3151F674DB06BA0250
-23D143A8BB6FEFDF383C0D20F6DCC329F60FC08C0DF78F8BCA6FB1945D17911EF078CE0D7256
-2BEFC44601F37561E4CE3C074B55F6A33BFFCFEB6C6477E5E9F190D212246707BE911DB4EF8F
-D022A130B4DCF061D883C3093B3B2B99F864ED02854324C496F177BA5AC2560A92E195FDE3F3
-80AF15A74E0A67FE676C3E28B3A59C7832DF21FE0F45E27BEA787B8202E95F403017AEC41A20
-3960F560479FB3B0DA14691EE02CC5E13BC10EE27BFB551693C32265B714C3100053F8A3D169
-79B692F2541F1EB692B0B6ACCA51FC368296F37EBBC05019E19DF37CD5746D357EA2FF4FC61B
-D5857BE56E1054BC368D1850C72E24278354C0BDCD87FA832CA4E0C20CF8BF8FADBEEEA12468
-07A3E11599E1135D8E0D884325F6B24FAD298BFDDCDA45F5A0DBA71426C4DE5546DE6D637346
-C9555D65BE5EAE5B05BE5912077A92F4AC037B35868EF0C5CD7BE586396F1E44DD5BE2EEA249
-9903E612DC73765AA9C796C8AEF0BF7E2D9956D3252029C1025FAC46A186A843B6A272EA095B
-D962D0FB33353CECED7A652A5B6A8E402EC749B3F1E6255E9500FB20D750DFC41DEA740676F8
-C42BCD153C3D97A890604522E3C47CB17B529DD120A930C45238E6F38B8BC9CE6E348F988FDA
-4277EA45C7FF01C4FE228A00526DDB6F206D0ADF7C8051F8912AF533487CCC97C3317E589BC4
-73D4D02BEECAA220115F8931CE45270A8E244792BFC0B4AC4D71DBE7F08BCC42A8D73CAD9DA5
-D530A3EAEC454D4CCA9FE54A2E7F2E570B09FA76CCACF67E021063E4331FD83440018AC9DC35
-F274D2486A73003C09F502BC471D451766E68C386285910CFA0347A316B4C05945CFE2949D0E
-626BD7B4B29EEE6E7A0B028551667B4E180EAAC72F11B6E6BB400DC8720795699ADEA2053DD7
-E15842498E3B512D9869079388542C8DC0B6E578FBF85829429F7A84EE05FE594FB41070109F
-B9CE01E8E2DDEFB57003C1E336063C561EB8642A2F8B5AFE4A43EBD47D394ABCB44D1FD1C835
-2D29FB3176DFE8C0B608EA45270893808283B62EE2D0E835C3FC56AF504274CCFA1566C2FB7D
-2A20FE8FFA74F2276B0690F6A7DD21567F09D8D4175C9F3E541E3641F6995253F9D2AD53ADA8
-87DE8CEC433B0EC5052CB87D130E01456299337F4D0090C124DCE321CFE408508794B67E3A87
-223550A70C0707B41A8DB412CFD7B693B871EBC7A399B0153AAEF34028EE5F12F7E1EFBECF18
-A6910C42C1260BEAD0A2675D87983F1236AFCD27DD5AF073A5E74EDF5F797D4347DE1663A3DB
-160D0BDA26F0FBFD973BF5C03E2C49D81EE0604E52132F5EB78F34F18416D1A677FEFEA61F18
-DB6142A035EA4864905695B270A0B80B78BF90836E5E2FC4984E069FA7D6B8EA85818240CD0B
-2A402D149F8E856E7EB89D3EA4E4FAB850FBE8943D83C29BE612B35138C742F1FF69E46161A0
-660EABB6591914552A5DC9D1314993686E3DE928396BF372BB817D599898C63CD9E33DF40860
-D7973405A61A49C9F647603B0291C03F7241B6E74A1FED6884955BEB16F579DA1B599E44E6DF
-C6C170CF7F25785FDD7B53ADE2EA6F2E48A6216CEAF3D6D6964B884DF8EBA3673758E1320DB8
-7DA1D41DEF1C22E66FAC1841A7B5DEF30D4D0E5C14E2CC1621EE454D123110D0578142485CDE
-5A98DD7F6E0625ED5E242D3FB151A1ADB20319DF8423BA8EC238BFFC39955EE44D73B5758E15
-0439EF5FBE890CB6FBA600EE0927BB4A5C2285E856ACE4FBE09BA916D0AA55E9C0DBE44E5CE5
-4B2424511958E3F135246E72891DA5CBB049D6B5854C2B0301178C71F2DC2B0B498C7E92606C
-F6D64F96B434B962597F521B36DF6362C1B7270FCDE66DB58A7B3F951589ED0FAD667C02C1CE
-CC562E40FD62E35C4367D29F6902EBA961D497EEBA1D1E7305BCCBC1D905ED04739D9868D390
-C7385F7C7541A3162737BCDA9A94B28F4E96A907201705EE427C4D8E8B0EA6FABAEDB556BE77
-E6FB2864E4CECB67C53817637413648EC2018C569CF1D642374A3D62253F937AB67BB72A7F52
-E4E4245A22D8AD38580D0A4AE0A37F745E385A1C3E292D7493FC367D64B400B00AF7CCA868AB
-A9FC7D3AAFE671E77D52EC0CE1ABEB5D987A0A57A186BE8EBE77E418F56D1887979C7074B893
-7DBEE2FCD620B7E6C67CFFF43669D1F901266C0B4E5B98B7E9F76443DCADF482A3CB44BDAFD3
-B43102AD2FAAED8D6EBD5BC7AE575062FA164DFFF38B3DBE1FF4B00C03E9484CA7F84E247C56
-4DC8547A9ACB746B1A8BCC60D37B47CED4E0D7ADECC7321EC7BC54D2882D30D66AB800E17E0A
-E60602FF103B22BFA1E3D898B71C19FA70D3E10491D4F773C28FAEC12320EFF572CA73888B7A
-9B333563FEA596A6BCDF3D16E45F8C92FAFC3681CE147B723B838535C4F0F7F285C9EE21F735
-9EC64250CE6065B831D25FE0F4AA0C36780476E440731B80DBE049798283DB9254FFDE764701
-7F40B1D6AAA28590EE61EEBB59E6C2E509A35B2767FDFAA6F7A6C19C1830CAA58758F85CEAEA
-6A862FB793A09E6619BBC7E24962D1EDB9BD9332D44063B6A4EC4C24363E907A69BD6177573C
-B792DC69BBFFF0456C1CCBBE1D7FEA2CD747D1347994A4B686DBC47FA0EB089B5B2A0BB10FD4
-F5CE4BF2A519E5B82149866A9D718CC78F29C39CD1D762DFE80BBBD0C30B96A969AA38CCD02C
-DECE00484DAF75200194E122C83E69AC216D53FEF27527A5C9204ADA1D14B43DC18E44C6B587
-FC97AEB6AFC837D5A9F5B6793480113D77A7D1E7829337ED1797E727A22540495F1CBC6AF639
-0558A30CD5042E790025BE8B1C102014283F724A18DD5CBE18DAB70962E7D05263EAF043FF95
-0C5E6500B12960F872A4E62BB67F05F0F3C38F09923055998A5F86E4884EEB278166375722AA
-F59C8307C2ACFB3F31A157F91C6CAA0ABB325414D5F2DABC2658245B1DC4A1971E4AF4F32299
-C42F17A35CB8F4B5546067E34BE4F89C9B33C88194E0F247C5E2A1D8F581B1304866CD39325A
-F0A8F303454B0D5B97A335DA2AE9C20237B82FBD9CAFBCF33D496F593B96996FA68820B8FD12
-1504A1243AB81119047DDE851AFD79878271B61F03CAF4E714469B62946A47BB4198AD2C623B
-59739677397282C4AFE943961F11BDD76B9728BD3E5F83710A9D01DAA204D0F1BBEB53D052FA
-11B60AE4AAFD906F1E43810B052121F9423AB8C27502D2DFA0FCCB0F09D5CEB27D81F193A7B9
-7243C810CD2FF4D1D198B307CBBE01AC6BA21EE5D2DB2EFB6EFDE6DEF5395A5C8D244EA8F14B
-44AFA24A927467550BAC61E102E86F9A3B8DEC638927B148AD54690F90CDC77D98158556B639
-7658A7532AF9CE0CC56CCA247002E0C3444E57A70498C630758DBD987A93CAF1C00B07F47323
-0375F55001E4D863F9501A9D595105F2A661E192B1AF34F95764EFFEAC32CF03882D0F25AEA2
-167AB48B4719CCDF60EE61794589CA62A149D9957EC56F26A248AD5F569BE83779BE51C3EA73
-BFCA8B08965DFDD7336B75274A2B0329BE494FAF127D600B1E03B846EFA35579ED5F5245EA4E
-4CC08492713DA928B987795CCEE7B4CBECC88C96809050843D52C129F905C1826570684364FD
-4E3681BD1420911661B2580D904FE9D7DD830AABB6B548F88C53F187FA9376D316ACEE3B2D93
-D8785085DBD1239013CA45A33C871E4D5314830D0A7CD54260DA840A2F896591B990C3F4CA70
-BD024643666341397F1C8751625219A337D4E8B26B31D4947B831712AE517B7C77225870FD4C
-E3E4E4F3DCD76BC0778913D903593F9C8C874DB46CB6BF7616D8337205B8D2C175D1A009E0A5
-4CD6B8ECAF18C385EF7E277BB62A52DF61F0C6326E670F8B6299D015AD040DE69D4976B6916B
-5380D474BABC6D6D41D783E3216860EC1961253FE6DCC060E16F889D403FCF56549179B41D5D
-7C38595B660BA89873CC73167CD67850C1F8F222CBC2E3838A0980913D38E2B2DA10B796DA68
-3EF7EB24EC9AE2524578E2DF2C53A31CC5E5B3003535D912FF5BA7AE196E3DDAED71BB4E32B2
-A64A965EC800ED6111E9BB2E282E2D20420C4897A2428D336AD198AB54739C8D3059346FEBBB
-29A78C42EF99A5F8AB91CA63289AE9DEEA0327F8F595E8B32260B1DE01412123AA6DFF518124
-783D6CAB940EA9F4282CB1481B23102C9A01D48A1ED7597CFDA96EE32958FE89FFCA3911550C
-70B558ECD8341FF5B5925CD9D7D99AAAC6E7816ADB1D84F45B3D3C2BE59DAB71C5F48BADED99
-C6AC02DC291B16E89B96C4B3C57D8A4484FA2E31C8E224420433100ABA98098AFC5C82F679B2
-266E9515FE2DEFD14A97CF98DA94FE13A321A5AC6C0EFEE5B4196B2A4B053BE872C29909A0E2
-F54FAC01E90C3AC25289CAC8AB899FC7967E5825247E7AFACDC3E8784BF7A16654AF86CC60B2
-CF7632B8DFF2479B63A79A0491E1497A6BD52E6DE49EED4A6BF7E98C5C85358CFBF613197C73
-2EEBD9679F82825B636516AA140FE474836FFE9AE3375B27586FE4FB95DAF8C2C89F5EBF104F
-1487895ECBA0FED871109C93EEF10970F2F1A2F8217F8AAB013F36C6FE04A88F104C1733D4FA
-BBCEB0F1A0895BACE2F48A3DF177F1CCB85AFD46CDCF14E1770564D176CFBE7FD0D13BC95D05
-5C61C7EACA2123325CD5FFD6C70C89DEBD2ED490FF3D84B5F17046B858F8A575F7D75C21551B
-52FEE3D4A4053353E8603B6E05ED71114A37177BE6C677AD4ECD45EF0F3C4F97C82486B71467
-18BD0DBAE4BD5A003BE0370352268AC6E8EDB0A411F87BE664DC2767CF68895356BECE4FEDC2
-AD987D1513DB869EE3E740A60B8D05DBD6C5CEA86CB82F6BF54316DF8FB8D08F660D3E004B68
-3CC4886DAF3C4BAF2D67A75FCC71EFDFA8EC1D37E58BD157756DFB7FD9114590A0288AF3B0A0
-4DDA6DB4E2B96F22D1D0A1EF70EFC781CE901175B2DD32F13AF07C92EF8DB5F3B2DEC7545680
-68494F1D398BA5BFE55F3000784FD266EC7633323C8818E9EF32B102369939ACBF01C18F8B34
-E93EA50B8BF343C70FD0933C54623A7C53F30292BCC852BB16C9880EAEA5C3F1DEEABE76A24B
-D1F3F2B885B942C34E2445F02A3B7159A9B913F339A729C6C241DE2FDF022BE7B667023B2730
-0D4F19E984CEC0ED02DD2488D7D942FAA57379A00EE641AEA94DC8119EEDEACECB198E24CFAA
-DD5731206F63497351C8685D70CB2E8F9B767C7DA0818E015C0096EEBA49FC55CBEFB60B481C
-DB3A14746E49DED240BCAF1C7C6AA4968C9F5DBB2278F16D6C3DAD37017BC25FDF201E5F5800
-C96E97AC0FB187DCFC1759BFED3FD1547C851D81CED1CC44D453444A4F4B0E2F64E526CE7F32
-774CE5D8F8FB00C50452B51782205987D9C4B8005F7478F7E175E0119991441C9AC9A42DD416
-72DAFB5C1998F7F031FEC4189229FFB15BC60A12D006D59D19AF8939456CF18220E319D3967D
-B26127D14DAD8042ABE32B54F1AB2ABAFE5BFD8F72E2A6750B516E3B8498B0C4A494C0D1E265
-467088FC7502FF4CD521FE40F884C1F72EB2054329F5820459D023C94DFFAB04DB5B90A9F16E
-194A7A274957768F4B151CC13D7C59B734E3DB5A59A42BB4E3786FE7DC965523DBBF18CF616C
-8D10437B93333282C9A702A2948C17193025C398593AA6942A3D6E71DAEBC1CE10A0FBAA6A0F
-6D4F8CA5326699B448EB747D39648941589D44D323D65D9BFA34A7B8F7235C010AB6B878ABAE
-D9231EAEEFB19A5345F6A52F4FFC475D09AE4F24D8D3E9158464FE20C2788DC18E8559F36AA6
-8FA3FD92F46CEB008A95843600E080949BCCAA96DD3C4A53D0FE4DADEFB2FF906DDC108E6524
-93A8685D6BE5ADF7FEA3246993DC57AAB9B576EB5CB8367E097828015687CA2D66F790E678F8
-B57AD6D3596FEADE392B4A0A8E904F8AE822F143D6C3D2907F1DDC8F5F5E61C588DF28C4850C
-5FCE8C4431CE9CA56F8ECC0712D8A7DDA36B484853E05D99C441DCAA9BB9D3D25DEF3F2EF19E
-49DBF515505B701C893DD1F42C8E59525E64C8976B2DE8F8FD07F041FD1F57691EA15589BE20
-CEB3EA5D0389470CFB963E115C93E2DD0304B5D847DB7D2AE5E0D1EC453DA44B8449D5A5325C
-E08926498AC2D11120F42EB117B2279642403A87BFE312C4BD6EBFA86CF74FB4F31A475854B9
-CE2E08A23F97FA6CD574870AB771D804E4A6FEA77B280CD92513CBAD6B05005118DDB1B8C36E
-83EB4E681DEC0A2E2FE3A11720A820B032F264E48056BA61999ACB83D3E8984726937BDEAF10
-CEA7DEB61201105FCAA84A149098731382B21BE1A95526860F70C721A5B355671915D69370D1
-5CB1AEE1B5BAF19D0EF62018F6C04257E6B2C4C1A9FD4D5E7054B475B209C3584F196A1EEEB3
-87C37CBDEEACBE2D1E80BBCC7124B86AED17BDAE18C6DFF7F7CC6A91AD7793F9E4CEF592C57D
-0CEA1B7A43EE0054B3F7F295F6FEDD32B6DB2D9E7B4B9A49BBE7CC8E3C1DDD2C716BDACB07A5
-86B71A4A4AB369AD5A19E98DF6E1B798C2EC2C516DA76BBA20367AD75914CE12B1C7AECFD804
-3689E5356BD4C3F50F06C2FD180414837C7BD2575EFEA4CC955081EEC436C91F298FE78EEB42
-E52079C43C0687890604BFF5D06635D5028AAE5091C35138F7BEC5DECD70CD2B59DF3F3DBCE0
-C200CE3D40B80DC6246208A31C02F60465B595097EBB1E53E59425F9FAA9AFD00C95689333D2
-7328B795E57191D254F38B714F1CAADFF8B239E8018A8826AF03C5EDFC16EA8101CE439C8D90
-4CF0FA14CF7D5CEE3E57E6DA63ED52121E21ACFEF1EFEC861B1B6303B71DB630081856D7CAB8
-24448FBAFD677305205FB1255537C66176303461C986DCEDE252B5F6A56E8A5589CAC03E85A8
-DFA5B35C048301B85A732E21D0C7C4B212421ECAF7D252A5072BFC65B42ED71A3420B6D9F5E6
-DC4DF85FAA4F5339714F7087595BAFF8346BB1843EC8523B073028E2CF8AA56502DE7C898973
-71B945F0AEEFFD62D0EC9BEBD8B9334924EE13C38D23B75148475AD497852ED141B506853C04
-3E221B8AB56F2FE0EBBD28554AD2B489421D1B977990FE47330DF198EC894E3C5CF213C7A8AA
-D7039FD2F01C26DB63CD0B2299BFDB05B0CE8588604E11A718775F130E9591ABB1954B1B451E
-ABB37A56BB1E213DAEFC35F699E41E998F8BBA2147B15914D5CB0621E70FF87906361F8D97E3
-ABC32DF379459DCB777C59E36BD6F634FFEE165411867369E5FE5420A86670904C191F1B7248
-3FEEDA79D33331FC7710350DD20E9549C0EFCF1D85E1D71A553274F582AE6B724E5E4381C347
-0109189D4CF08F2B9719E9BA016C0556146E5624CC70161AC90BF7D0360718FFF93863EE42B1
-54947D6AA63A9D205EE925B30F48A8BDC389DC392FE66349B8DCEE2E4E17916CD4662EA076F8
-FC7CEC96B602CE92F54893B8FD90710183F1D61CD11389D50754B0CD200B0CDD619E6807CC90
-2408BEBC63795A5848B6FE436C949F5AC0CBF614396B89A5396161BD5BE367F8EE647400DF68
-EA5A4655E03FC9C5541E10FCA31C3DF0E32FEBA99A36FBC1317D781AAD97E57CD9B0FEDC471E
-D62D602412463D309320E7BC150AE97924BD8E258838D74362DF3EA0F6209ACA57720577246E
-575696CC74A18F653BE1ABDB940CFC63A30AB40355ECCC219E8542FC324B012EB8DCF2167BC7
-463038BA839EB3E473FFE36D91A84D00BC8BE76F04A8BB0A14E3110CD873B66BAB9E6B0C1BEA
-1199A86E6D481CA1C363D4AF9A97E8A4626F5859B72E411F2F3E1124D3D0A823D52140239692
-979D0AC75EA91BE6B063187FE75385C3EE5097543FE19287D2A1DB002A4DA51552CD2BA0EBE0
-DE3B9E19E02842F366256074B4D1A29395F53BB26B19C3F026D8D90B4FE861CBC00789EC9DA5
-6A6DE665BB408ED13C514A4D826B44732163DED1F6FF667552F804FB1E40313F7E42CBFBF788
-F7B36D61A50FBA8F693D8A5B8CCA497BAA922CA5E29641FCF276C86E1C663777182A7AE7D0EB
-AFAA5EE3A129FAA057CEE40063B3983EE3BC87BC6FEA42975ECCCFCBE031F95E8CFB0FF02E72
-A98BCAD0DEE16671AAAC4462D6D2CCA9C439634CE81893BDD4DC8AD145F6180B7206578ED5CD
-ED2E660242667062C2E45E92385562B0B15ABA872B166E8DB8BFDCAA7454309754BBD9DD4DB6
-97EB33000C5EB23614754CF7272DF4D864D79D4F452AF252AC295F3AB4617E68BCFFD94A685F
-32F3256DF146F2EF16932863399F3380D30BEBB3D9DA2B9B993510A0C2A17D29D78DE8A02E4A
-93066CC5A6232A8A4733DC05E6F89D1471B4EFAADB0ACB035F5BEBB90A9A6A98959896AD6266
-82BCA5423339E246BEF8E7FCCFCC57AB8F9752C34D33E9F3D6D3A86E95576C5F1EADAA5D7579
-0080DE24FE890621D0771834C6288C7C71567ECE1FB89AA70FC3545E7DEB8CADF4FE57C4BA62
-F6E724D492873657424D12C910D09A3ED9AC344AACD232FE2DA7A79F6FBB6861F4AB94772D9F
-D4D945DBCD70EFF8E50B19506CB75780994DBB082B165D97F7F280DBFC3BE56259D82A495F78
-E515867A39BFD44662151C279B7BE9F36E6263A7461BE0078F81D27AE4001D7675D8DC2581DF
-B319FF772EAAB2BAA831D525A2E7C8A3B907AE61FA5D77B1D8A0CAC8F0D7455357C879B609B1
-DF50CF81BBAC75D3ED523A297660881BCC924DDF827E62A372489714303748592FC12E0CD88A
-81FE23A8FEE97FC50DF784F117B0E39BA7493C67F145BC606882FF183CA65BAF1AFD4CE4EDD9
-F6DAB92C6C892F4D6FFC675FEE470DDB83DA0754C5E1369390983D8E88904A95E02454DB5C3F
-CD08D15E8252C048D34CACED05A1BA12C00502F69A4B4F06A57FF8ABD06C574AF56AF8167D1A
-D94288D641A9DFC473446BF7B0F38E9236C83871CCC17E3FCD2CA34D9EB2F46998A768DEE5AB
-9B806B230050BB70DC57DE3E958DD0F256E21D54F3EBB0504778E31CF0FCF4CEFB3164B9A575
-7EB8EC5995EB07FD46380BFC66A3F172F7D0D2DF36AFBCC024A6D67BCDA7FD72FC6AC6953F40
-358EA5EA285766EC613253E47AA163E04E8EDD5B6618242049C5224ED0CEB84014C9B89FDE31
-5F852511C5EBD33B7F01394072AAF515DEC74E28747B509145175555414DB9CA11F6A66A9CDD
-C45B04BC4A43F57CBA3394FB62FB96DFE6B811A415BE3EA5AD7B30C1409E9BB5FAB41D367D66
-7ED10B11746A111AAB1CE2E6E938E6434DA7F962403496D390B8201340FF5BEFF94E1FA3F932
-881B82F1268E11CD5B5CE6E98E309042BDCC54B994CF78372F0E3320EE745F885DEE220B5358
-E4861AD42A67BB4BBE78A65515E8909B2CAA011CC01B6071C0A1AD0FF9520788CC836F9119E7
-5000314383CDCAAF2F62729B37891CF44771A9C2A64D797129F971EA87E815CC4AFD1BD433B5
-4A706EE696884AD7573F4A5AEDCC66840697C459D30DC1E979DEF88DE4290BDD507E2DD282F8
-4FBC0877E6206D70E652A3C45D5A65D195B206D2922394CFC1CDB8221F63EF9E2515CD0B6578
-2E34637E31F19AA89F235200B437C815E3D2A9F719B5624A15D67AE87D662DEFDB8CF3FD0314
-26372335FDDF18FA9D62B446C46B2076D8276588651E08D886F2C0ADB1F5AABE481DCCBEE18B
-C56FDBB92563E67F93A4721C2515AE140050787ECB4AE3D8BFEA40DA6B160484E3D23801D211
-3E29A22662AA3650C58FF55DD480283A93A77C2EED89733F9F210668209E3014F33EDBE43822
-D28E9EEDAE84DE4D382F582F980D2EC9014D3EB69BC9E03BC58E694E2DA67960D699F6C1FCA7
-333ABA6A305BAD54A610C68534DEE6AB31BEFAE497840AD236B2DA5ADE640F79940AF98ADA85
-913071EB097A2231D545C30DED8290047F8EBAEB6E0B8CC47787505ACA6ABA1515EB674C8D11
-976158586F9E9AA2F7AD9F50919889406347DDB06961151DE231A3E0FD703952521A5E1DE9C7
-33775DBD19CB772D36C1D3BCF222D8A9C278AB7DB4BD13BC36CC2D9AA2C88DAC4752CF879F87
-3BFF4B5E0272DAEE2DA0B3849753020CABF3AE732A389F5F1E7AB6511CFBE515834938EB3649
-4592D2BEC58169CCBC6BFA8FB7BDFE8DFDD8ACF14267F24C9C72571C628EF2E60C1395488252
-F2945FB88AD8FBD3075683AECF40136AD2FB5EFA9561AAB86E2A352872C02DB2A70A5451E378
-D3B243093E21999D37FE1FBC881FA2ADEFC8BF98472DB06C71731A78E67B7AA874D85F1C5B65
-EC6EF5AF8947240741AC14B4F7F4A166A385AF0207DA87203C898C8F63F9BCBAC8CA335FC956
-FF530FD4D2EC87D46A099C554DA08F8ADFDE352F12B8154A59F8823653796822E2ABAC6B3FBD
-DCEFAE2E84265C21526A3EE066A67068AF675513C3DB717738A28C2BCD608E111EE07E89A932
-92AB85B6AEC8418B653542CE64DA2336AFA8B3A0A4979B2287DEFEBBCC52C17C709A643B8C87
-182E49D90D2E4265F02737E473DA03E600C7DD90C3F2B0B9FA7E334534C42080A740D1F6DB83
-83753EC0A03DD1EB7ACAD779ACCC0EFD2FBA6D03F70DBB208B2C38D67D002A85B33856791C8C
-46D42D886F626151F012B7BCFA593D935DADE59958199FD7237FBFE5BF82A51952F43F3836BD
-C6B66D58FDA265837F94028E1024A2F9804C284568D00B4DA80F50C3611D111B57576ABAA54C
-0F91CD05265CC4B0EC84F49CE7413803A7255F17AEB87EF85B7C9A34E756A779EC4229D8D9EF
-5BF15760F2A9F954982843759AC554BACB8776F9BE0823AD9554B1798BC3D9D9B36FDA068EC6
-7D5B5CB950E0E7B777CB711CFC11E6C2EDC6E2B6060360C023EF3B0FFAD5A40948B8459A9A4E
-AE8A75BFF733E370999919766A88CAA1A27316E846E7EBCB71735C99BCAE66AA056E18F59852
-542DC8EDF5742633EC7554783FAC1A8C64197D65187D0DC55CA95D051DF0A8EF293AE2F9FE44
-9AEBF55E2194C3B4B058EB6C71CD4B7F4FB795806DCE597F721E847263E675DDB86505932EB3
-CDC4E7976284627C9D74171DB94D5FA907F555E6D9F65CA12A5DA4F5381514DD9EE2C68183C1
-74CB97809359E7C48BA17F5A5B2F9429C01296DB7861977C7B9B311E4D7C73669CCEB270F0EC
-492E2E790EF4BC96FBB052864C72F1D2F990BB4DF7B802BC160FCA174C5A485B328F4B44F1F4
-3C83F9329C23C14EB3BE842BA49BCF715C04B66F365CCF97F2AF508EC6C61BD87C70F902C1CF
-646206D462C6B251093C662B89BCB77917AD865E7EDC4E86B7C735FDC64C3A00254C5178DB84
-91F3D07D77EAB33FBBE793F8597BB549E44A7483FE0D7DB8B842C66E6F3DB99C886119EBD403
-1EBCA107541D4532E01998BE7BEE0C2D2063F4FCE02FF4A4EC1B8433C2D823EDE99DEABF7DEC
-F590EB8C35F6F485EFD380061191A0D7AD91E5186046FF7935AAB05DCD937D2C702152494BC4
-C832D3A987D81061F7A17B344AA48C29E4CA6BF366E21A69E4422654400E56CC012543D12829
-D03140265F4C9FDE078B8A444BFDE0D3197656569AD8B57ECDB7D2667B204A1BA3111F500CD0
-F70860CDF5363F63B1626B67F5DF72C43B4C3DC7D42507600C8C09657E2B521628ABE4622636
-8D784513217AEAF52B74F214282062574D5EE1B65253C0B6945B3719A776320874E81021C9EB
-61D405441360B20A8DF84236912BC3C870320F1C896E31205518BBD50E6F36B77856590A7E93
-23ACA65A0C1F6AB3970C306DCCA42EAB9D01E7899F738568379E2331277B3F9A7F36D14435C9
-E50133B28522995788C1E9FB2FA2C4EEFD0DB10E2D77D8FAA6798911DDC6C99C06EE346BEB67
-7028FC72DA4A16245ADB46A396E9AA90652800C0ED97886D58F4F1CD8AAF7442FCAE1221ED12
-0DCC312E10C914A3344F6C240A3C5355637D66B683B5CC8951CA11A3EF53FD1BB7AF3B13F7BE
-E55515D9DD2E1408CFB2DEDA74ED49A48853E066B17CAA34A84D4EAB5630D86298D979F50E1B
-648C45BF834FFBE926C2042F9ED5F751EF7477B065B38093CD3C542C9EB4D5E757200A16E219
-5D37771FB944113422651B6E1FB24373BAA1478324E250D085E1D0E555564D5B6A1CAA209D48
-33671CEB4EBD896BD07E8A697B7D1E1209AC76FFB0D41BF12FB107FA19E95D90B39475368A8A
-5AE436F41CDCFCE07D4071C698A22CEB1D9E57B0C0D1022615EE9AE6CC1FD3DDC3656E1D5BE2
-2A4A7B6152D53642A5675D187319E8D7AF9548E5B9599F7641F9B4ACAEC158E992CB88D9BB22
-072C2F3A3CA95B9A96228F6BCB5FFCD647978ACC0E67ECDBC58D079C7550866AEBA0260C5F8A
-AB45902B4726BCF58493A23010288114CA35C0F52A89B1B4743CB649A86502A838953DEC2982
-683D9DE00A51C9575CEC2B92691C23A2D12E973A20CFFB787CC759D1394A475C5FE5DA04E624
-2820888E6AFD8B4D1C9F4B105C8DCFC88B032182E7CFD6629707B19A55CD5B7F9AA0991EA99D
-E9F2984CEAFA45A0D45EBB192E384AF80987C002CA8EA3E2719A5D5201584889A3C6CDC897E4
-C1D6D5F482329FAC3DCA58037BB27AEA8A3750907FE3821FAB79CE83D33DF3D4C07B690ADCCD
-7364F0C12323494EA74B931FE6CD814D84881F8ED78ADF11F1085EE3F6CB379D8D865F1CDA31
-0308A3274FE10EFB3EFFD711D6002F9D546BCB35A186B0A149C32D239DCF7608489E512131EE
-833B7E8BB0E96DCE41C67456698BD0DB6335DCE9B83E74E2B4AABDCC14DA977A3602A4B2061A
-DCE4FA1613459E82774E48B318980962944351389D41D74BEE89CE3856B5FA5DE0D6E870F4F8
-ADD50CC9E128AFFAAC1096A42700CE530061A9C473862F35B2ECCA9E2B680707075B415FEA99
-3AFD56FF3E68A944C55E193BAF876FE5DDB98DD81C64BBFE578F9EB540F5F6E7E8B1376C4694
-FFC3C08442C59C3AEFAF8E464EB94E1A45E689D7F71BDE01D4B371A25C98E64BC72BA2028434
-A2A9417B255183751C143FB909F194D2F5B2C20B81662F026A22FE737F2C32C3515C037BC084
-B6EFADF2FF68943302A4EF610F511BEE0CFDA3F962940A7263C175611A6688A005797A3C5446
-F9B41E3F021EF1967CE5E3F1BE0BA746C4CEFEA68753791BCABAE5177664525038917BAC6A98
-CA417ACBDF4B5AB8C98802C0176238794A694A2940CA374C493A7F33B20F53D1C5F397B2DBB3
-60CA0C5F32EE4B6851392F40AD1AD6EC0AFC2451A62FC6DBFDCB8C2AA60F2BDE83164DF3D052
-71AADDB35B2CD53FB4E255B90188C57B51CA4D77C2F5307FF8D2FC8AD359ED8D8C4F74A338F7
-FE9919FC8F106C7C4E92D42075B458E29EF6944E2A4C9F8C6225CB32758F8E465AB2C903798C
-24D3C869585D3CB7A23D2B133FD832752F864B60205FEBDA77131D1395DC6C4FC65720BC654F
-9BC0FA978D5D93886F489A7C1323E569CF42A5B9D52FCA94186E3650F71CBC66DB02DCFB86C9
-84DB629470D90A5AABD6CA8F1228EB650C0CFB81912D145EDB3479CAC0624867AC8857A6264C
-1B93AAA554D85A35F7EB0210AA7C167202DB3929E22E36455869510EF9CF4DF1D09D2B539E6B
-04C700490C79F28FF2A2586481ACEADD14A5889EAF6BF235070FE9AD2FAC9012A6091C25FFED
-579D6AFF38318C92AF4D275373A00081C309BA26258F2A3F83C7617558D9693D1D029F6FC843
-683046F3A7F5B1ADEA1EE32AB7D2F22817AE06089E583EBD7DED487D45B71EA6B17D7E401F7F
-CEE17FC13DA623691D2B4A1830DFB1803CED74A0993FCDF596B0699644F7A18229AC3CE82031
-3E92C71C2EA836BA426A1C73417EFB118B976C0F477C164AEC368716E2E37AC213C7DC99097B
-A2FDA04C76DB4C871B07C41824B3B337449A2D7D35883A3964B888C720BFE87000710FBF6BE6
-09C89CEC3569685E6C55C7DC455D1F044902F305F2FE0F92F993086010DC1D3BA1CE5F0AC416
-FE71468D18745A189D2FA50D4327FC97018E0FD91848D21B1615309B54DA59199818A9556E97
-944784AC4DAEF5483A8FEBF6430AF3A48CF40DB1406DA989D5E297DF6E96A6F106F9EA7F5D1F
-42F776B80E4044615A12B4B8900BBC80D4DE37A373045356E5AC6BA1AAA1EAE1D34BABAA2043
-AACAEAF9B63AB2094F1D6C72A8EB7BCAFA0639D1BFBF007AE7725E63BAC4F5A6E3FD833F574F
-435E3EDC48E805FED1E077B8B581B6656125FB1C2205CF756F97EB5476B77A2417090716189E
-8CEC268029A2E3207233392F15D0269A3270666CD4753CD8FD6A74C7FA4EBF5ABBBE85C6EEBD
-FE90D2A8F4D31EBE5815EE687E077ECD6D7CD00EA57B4A8505C0E6A301A27591F342493D0A98
-F86A97D346D4F524F89EE1ED02775DDED2AFA93C9685C77D2864DB042F6CAFFAF2F449A83035
-2F524E30F9A8FFB7898AE3AF142962835B3D070E99A8D86538E9748A8CFCC51673FCC6047326
-5D2321DFE3AB3131C3940078918BD43E427D32BA0FE20E36949C12A53E73E6A4A6C4C885766B
-32295CE425ECBBA9109916BA2344C388A2C7F82F91477838C8BFC82BD677A3B84EE0E9D3D336
-61D9B4388FB7EE6FB0317DBCFEF4D13770B683B6946A81B255190AC4704996880E92D6C015E2
-5B0E74553A7C6B542AD709C4CFF3A7CD17517DF938095BC00626D452093114997F2709DA184B
-3692D7BDBA7CAB8F26EB9B14AF957B068E94C2C5C3733D7B80089E04824094A20BC715D40367
-1F4D8939B3128D05550084551661E019431508516D892ACBF7F0354E8677107CA9214E429FBA
-5ABEC4FDCF73E3B1003D043617C8FD26704652E75DAA971E24CBCE7EF65015BE3522A4D30C9E
-8AB8A7C8DEE91212DC92BF2CA9264C1F847C7E6D914D19096F032740FCF2082D432924650169
-C3D37AC682A18A751DC8145026CA9DA8C7701666A8B0E653F00357F30FBA189FC19CE74F87C8
-DCBE37990E6A6844E8CB1374BF2EFDBA3A1FE8D251825C7BC2C44B45E5CB1A154448AE5F66B7
-900DDEB3914897754128AA474A3C9B9BDE43FC4F3DC7196F471E734FC0912380DFDA9E09104C
-7C575CA24099FC0D8950F75ABE97EB80F10B6EE80245B45E54FD54187C96BA00CBC2F05C77BD
-8D3E61C1A60381AEE6757E83256793A9EDF9CE22ED8B35C96DF6B6461D3CC4C1381CDD8256B0
-ED1E7FD346483CC7078CA76393FF5DB218C8127BD7281A03E87407C65D46D97ABC93135F5664
-2BD99C11FA05127EDC6025DCCAEEFBBD8210318B18F4B6646430849ACB10F6F442F9EB9448AD
-66EF9C959D3A917C4B25269C66007076D5B4E9013FFC355D3D78329E4700590A5F4AFA85B6B4
-928ED4412CF0B9CF6C7C7B3061D5F95271B731C25001C173004C738F6AD6DBD8B4319A2DA207
-A8C346906A675FA53636A2BB48D8F137D86654EA08F6790E9A45EA435EEFCDA0E434CF1FA7F0
-3E0669D46A9B29188D57DC02868D7ABF014D1E3C24613323305E62C85C46D89BCEC1FC5C93AB
-7AD485C46AFBF9171BDF531A219BFC0F51EF38864805BF9B258F2D002F93607B3C8CF297AEB7
-05A048C7A1A89CC780ACEE5756C119CD1E85049245D9533BC0B802BFE20369A21438E8723BC2
-D630B0F4CD44E34CC160EEDA7286AF17516BFB5B2CE47A5B2C648D7543FC9D8AFEF72ECF0CC1
-9F8C5C53BE73A5C1E07C1F5E392B39469B212FF02ED9A505A0CB9EFA8DCCC9DD5AC6742F0BAC
-F7DF4F719ABAE0CAA568494BD269B0F4AEDF86F9276BE9C13CB3B884DFDF0C7DC6FA19ACB468
-8E4F819E26544377F052E72B71479513AB7DCE2B8E9FC3902D49F7DFFBC471237C57E06D2205
-6FD09D9307191E2F2AEE34FBCB8E7D1365A1074355F0F8E8BE57100F5A9A712AA8745D296118
-77EAB4C049BB34FC5AA0712409F3FE052F4859A1B809A4B325AC0BFD669C0731F8BC48355AD0
-B4747762B4933DF59B71237393ABB6D8A54C4C724CF966F0140FBB2E208C13CA909AEB4A13A9
-93081B8CA58A9A48F679B86F1AE6539639BD9E71FC4D8FC7B51B5ACEA07DB2B47B7594CA795A
-FC18D5295E1F3FD74EB1D572656F0231EEC3E59D2BC6977826A378B9B5DE482E3F1EECDF1E5D
-84AD00F05A4DA3BD3708ACF218654FE04831171097CB9514C6CACAF04D1C460272688962541A
-5A30A244DEFA1C33112BED9D46D4CED8A1AA8A132A99537EE95CD24DA03572230B0DDB4D74C2
-E15B1F47D8EAF889CCEF971DDAF8FCF94DD8F7E49DD9D769C03CA305F9B8E0E7CE1AA057F40A
-8A8DEEE1FA62F4B6002674F39790CC4817E69587C3DF93C08D3E1BDD2789E6AA3B9EBD51FE18
-31C20AD29D47186C1D6459F681D9DA97F6B413914ABF7BFF36ED5CD5D5D8FB1ABBE422C23DED
-EA7E0998F3464AD0C418F8DE42E012B2B4BFF55BAA443335D58C127B626FED1214D2D988AC21
-0A6BE95913889A99742EB9C6AB9B28D6B361DB263BEE5E2C240E4F9EA20D9D18DE516E7AB64C
-812C3F05EAC8BAEC75A44AA829AF8746DBAA8983DEEB4AEE9D11663E90DD8DCB6EE5DA790A55
-ED27DCA6B9D1B759CC084998AEF4F740096FAE46D21F8090D0EFF5D17ECAD07EB382388AC174
-D140ACD94DDDF5BAC69B0C37ABA82D840418DD9A2947D4836EEC1C12AE2A47716D018AF88F4A
-426B4A5F1804ADC160528C9E830F5468389D73FCC29012162C24F0370073FD2415A62A6C51CA
-D6A6F439285AB190A66528A27BEAEF8724D256FE7A92D983174D8A45A9925F3B439AE8AF6C39
-7CA52CF8699D0C95E867997B330AEFF4BB1A023BEE6AA3C32E5063F0ABF9A7693BDDDFCB66C6
-7EC3D064F672161B8D324DFD507F5437FE84A0308A58195B4EB4D91DC9D48C11BD4001F3169D
-BBBAE233151369A81F2598ECF13280516330D8D336CA86E4D9F124F743C7A2863592B2F4408A
-A3C178E465A9C95FB5F83E4979DE63C638681D33195AC791150C1CBE665107655653441D4311
-FDAF3D9235E3955434531B85CE83D58B72CBD2D07A9807D2BECACAC208728DFD039B38AF2760
-9B339AAD2C6533722F8EC35D9670A9C0EA93964A14F276F304D0C6124F94BD0E4AAB1726D034
-81B50336391A567D0491514580C0AA86F9207295AE0F670AF26267C708CE554F235525ACF6C3
-5AE3236305340E7074B138B223DCAA191AF31BFC506FD189C18233A31612C3E3069D9ACC7235
-60104FE2C2703454AD989DC74F3B9A16D329F69C2C2CDC90E668AC68B3171B9535F06B9BD8F3
-230673E4B00166652AC8EAB91602B04F1CC6D42BB2FEB69B4C53025886E767A45D0DD659893E
-3FE48C46AB822C05000ED150B1C5092ED8905DE9E4253A5E689EE206C92DB84ED8593A48172A
-BDE598728ED0160474731D3A1282A0C3AD58FEF57BACEEB2807F09AB5668A3A7942D715AA7B5
-BA1B53A793116769F555C4CF5FA0A88AD8637C49E9A7E304AE4B991DBA10E0F0897D6304B14A
-9BC880184F2467381FA57847EB96EAE0E65FA455E6FBCD3CA359009C7E8497061FC10555F814
-54346A30A956782EEDCFDD1CC88F2442281234B76084EC3666E3DDC215157A95533780841A0A
-F0D880FC3FD5AB0E13F85E1A88D55A61757F5C9F5F8AE5B1C67B532A2CAEE939703407D2AEF8
-D0A8C131C538089C4F470DF596533702DEC500D4BC64933B15C006CB68C0235E49BEBF7D75D7
-FFFF6E467E24EEA086CD6AD34AF5289072933ED565D15ABB4E7C23877B875E1797D1CC7334FC
-E5DADF069B55F5BBF26DD3DFB845E3764A431322DB42E08F5A8E1A43FE270658E615FC9E5336
-3D0C2E178FE3E982F56F301AC83DF7E0BB16B1B4EA9AFB7F4BEC59D84974CB19D6E6C26D6069
-651E4B17F6292AE66E5551DCDE4631297F81E3AB9A311C7327A6901BD556B8EC6EB282D39BD5
-FED9C922A88BD52A4A3AB8BBA64656CD65C694B1B5EFAEEB8AD5C3895AA59FF26223B3265097
-C49EF63D7C77FE55A96C4A4CE511A969FF80791396710BC0AAF55948A6E917336397FC5B83A1
-69FA84783C0F7B5BB7CA050A32A54D8FEE274D20B5AC5EE7879ED4C974380E3DADB6FA505974
-E6631ACEDDD4D020DF902C4FFD8152A60C37E81F285D4049DB78A032886291870E17E29D4E76
-034436B6D5C79F0EE9D292B839A2FC0051129EDEE78EBFB35654B4C80D7DF9A2CC674DF333C5
-34E284C37DE21D387C9F3FECFE4F435178DB778C412C9E010531A91075BA8566C7B2B4EAE73C
-3317013C71157349892996B025B24CA1AA1C55E4758F6834224BEF066FF7B997F7EADF4FEA91
-6BC5F1F761BA3267D9D03843B3019275CA1842D5BAEBF99FCF5237C8F69771F7BAE960B46ABA
-7C156269DAF7580970DC950FBFA799DDB0BF95509CCA47B2421996F0A5AF654DBA059B2786A3
-10200D050D09D81FB09BB460CD4D078E3648A74C654BD4CD79AC6303F89D9AA07C42B0BCE338
-CE1A849F9755E4C50EAE7FF00D0C0DA4DC4F7D7752B022B5208A93A5C0AADD8B35F1F2A263F3
-58845D80883F716881BED236B1280EED7B72BFC9439DC0A62FC243B704A89400F100EBEE7C96
-F4769313126F27C0F914D4263803F09C5142A13F44F8590B3382B922D6CF4FAAFD48613F5F3C
-8119066D1177E604FB3591CB2396E68CF400C6B3800210853EC49E8D0A789BD6D29FB1D22A0F
-A707EA9BCC604E602C1518FDA79C116A2186938E49D2BF63FD9FC10D66B7463EABCD2AFCF91F
-609EFB5AD6693132B33442B0EE0206AE2CCA6C913C7E1ED5CFE40269560B7397EFE0390D0E8A
-7CCD386C498929A4A6880396B434BEFBB6F149055ED45087EEAA787FBD73EB17E8A43F0BF705
-EDBEA87B56ED960A56F0E08BB09246E58C08367F5DDA25A6689E4F565D54C84362B612432FBC
-612993F7F74A829F8FF835B3648EEE823339B67A54FA7602CE4641E948F89E1BB5AB9CB75682
-D3D2C823F2608EBDCFD39E8DC1BC28E6AF5D57881A47D7D60F956DCD28A96E52C4A545C4CD87
-DAA393AF0673809716A5D1D84F0FCCD4805A040CB8E879026FC66649EF36D6FB7CD7156D0D38
-FB1F15DBE16A258915B2865AE1EA75C8BB90FD22EC98111FDF80A13E20A7970E9B2A74EA83B1
-CEBF50D7A8ED34AA1153FA7459AFAF2C0B8A7F85433F573E6979AA33693544151B62E4BB64FF
-1A4883B8AAD77123BA7B160E71D1A544F71FAAECD439092BB70926C506DFD6C0100B5FD8B75B
-47A93C7A068D45098E254157EA8515314F6473F4AEC2FDD165A91D614C5A2679631C53FFD428
-5B89149326DA683E5E693C5B2B3DCF300A891BF779A4FB293009DBC236A91A1E17BD530973F1
-D6E2A3CE9CB27DC3A906FC32243750E2DDE538941E545E9165338DBC0BFD892B4A727803B704
-411E755C1ED22A65990C931230A010956276F345A0694420771C34CB97FD07C6E3ED51C9E808
-ED87D74BCE8439077B5E08AA5AFBE733C402543C760AB4E5501DEE7F723CFFA43B8764684312
-8E169ADCA16526DAF66228290F2C5BCCCB4C9CCA1BBFA4E00696D7CCE5AAA91A80A7212FAC7C
-EBEDE91DEBD44EC9781102BAAC812EB646E78F654F2C7901B6414ED830372882CAEFDC267879
-AD9E08C63F5DA4194A7A8D8521228D11F4BD0A48AACA3280FD9E24CECE77FF8450FA712D8514
-4DBC636028BA8F73A592633B6BF5D5D29692CC37CDFB882FDDF4CA1071F87C6EC4C7BB3E7324
-70A2A64A4F7CE9F5CF95A6D87BC6C3B4EA37EEB6A03BC832D958DCA227380934538221F1B089
-7477C893D7148DE91DF96AC3656E0969266B3BCA01C310785A8AED34412AEFD755CA9D2C100B
-141B6ACFA53A40576393C364EF62AAE88E9A95A5EB731AD2C79E7EA145D7A799DB87D4458471
-D24A6F9978859D6B97254110EFA503E685E3B617C0141CCC29B3264F8EDDD5B0AACEA009DE0D
-943B47455E5DB8732483B28C98655FD16B9EE33181687C9E896E5995FDCC3CC2D5D5CD9D2FCD
-84657C635A512188EA14137DEC72C955FA40742168B38652D8F41ADB9E40C2257B0762C432C4
-13A18DF9046C871DE6D0A6D072A50F609E779D4AC5FE253EDF5C271E35F68F3BF9B86DCC92BB
-B93600D21331F68984B31060CE1B719255E56D887661A02C133522393D71D2F3B75B950C759B
-75120DF27DB30AC889514C445D044E80CA8732BC55A3CC4A176B6B49882BA164F63A9ADD3E10
-E868DD17CEDD681CB43E2D525492A6AAD9EFB32A9B63A9452DF1CFBB9E28F30C2845A6E79FF2
-93F54C7A5FEDDC823AA5828C62210E9E3D629C84DF4514233A74F5C8F86EBAC2CE74C00B4433
-BF3FA9D45648967E06C8C2F1CB46A04C7C0F7876EB13B644B701212F0AB6FFDBEF2B6A075018
-AD8C95A208265E76DF250945DF43AF2D8023BEB65B7F95EB7F2592256A5B00ABD8CED8F05790
-912B10FAE4586003493D3AE7EA4F2B428C5A697273AF966D94242C59B8A055D672438AB4B9A1
-D4BF75C462DE90CC77988E1448B083EB9ED01CE1B0CEC9B021296A23A089A930C8D7AF61AC6E
-13C8D5FA38C2714A0011C16F46B7FD6BEB9D0C53F9F9434C0776000976BCEB5B1F2490B9CA7E
-67BDB0596F19B963772A4EC041B961D643CDF851A6227F7C010D747018D51C26ABDD5F762DE8
-38FB9913D776015FFE3231BD7587EB8FE09CC0EBC867665C41B0B54DC2E0B13EB54FB2C825CA
-AECE7F142FFE5328596254276B21CC58F210931C1741A60CDCADE3F8D59E438CA63FB27A0D3F
-27581876093B8B4D42505E2E5B6A106AEA7F5C8B7F1D55EDBE621F1102D864F5E01343E88BF8
-91E844EB99161D905709A5C21E1168E80224C40CC27A82B4671FBB6ABF8A9B119617268794E1
-BA2E1296381E331561198C2902C9C90074D1BBE9A6DF64D1A854BE8ACE75CEAC77D8C88A5A2A
-BA4C87A12063466A28C9AD20D41C29A17625B23E77D6FC8B879780F8369359AA523C922CA1C4
-6FA00E8A8A5CCFF0A674DD833CD562B8D4009DACC966FD13ADDCA02DB8552D3E8F56FDF3BCB8
-29A5A194BFCC5B88B426CB54797E07A41002E67663C357FFC22D7CCB464F69FB899EB92C4923
-037C365167DD817EFCEF17D4E2FDD023DB97291FF620A76AAF779EF6C995747E5C270B33957B
-E50FC81343E41A9DE0B2B0DBE0C101B767CF11CE11CD79B3404069683F7115ACF9C19859FDB8
-9AC8E280EDCE1CD0BB32AB28BC03478E8BD945A1578D5F957ADF1785F5616C67A3FE87443E13
-FA40178FB97D4E6BA9FED4909FEA54B421149A0634FF5AFC1ECB5A1A30CFABDBA5584748AAEA
-FA1009B86EBE86A76C0CE9A2853A8B74D47844A624560B20908E0C8B22AB8312929439DD32FC
-B30E15C4E160327181EB849858195BDFBC8CFE34B240E875A113E6AC2EF5DF8D584DBD8734CE
-BCB7B91EC7F394E6909B973D99DD3AF009774D94F6E09C74197DFEA9F2A33477DFBC4FE41410
-935C28E543D686C3660088B578618A6BC0398427FB5A1CF463108178160E1C56E72495DB542A
-7B7C8B1AC88996BF3A1FEAC0CF92F74CF883C87429DF939B8AA578EC815C229DB15A33120000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-cleartomark
-
-
-
-%%EndFont
-%%BeginFont: CMR6
-%!PS-AdobeFont-1.1: CMR6 1.0
-%%CreationDate: 1991 Aug 20 16:39:02
-
-% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
-
-11 dict begin
-/FontInfo 7 dict dup begin
-/version (1.0) readonly def
-/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
-/FullName (CMR6) readonly def
-/FamilyName (Computer Modern) readonly def
-/Weight (Medium) readonly def
-/ItalicAngle 0 def
-/isFixedPitch false def
-end readonly def
-/FontName /CMR6 def
-/PaintType 0 def
-/FontType 1 def
-/FontMatrix [0.001 0 0 0.001 0 0] readonly def
-/Encoding 256 array
-0 1 255 {1 index exch /.notdef put} for
-dup 161 /Gamma put
-dup 162 /Delta put
-dup 163 /Theta put
-dup 164 /Lambda put
-dup 165 /Xi put
-dup 166 /Pi put
-dup 167 /Sigma put
-dup 168 /Upsilon put
-dup 169 /Phi put
-dup 170 /Psi put
-dup 173 /Omega put
-dup 174 /ff put
-dup 175 /fi put
-dup 176 /fl put
-dup 177 /ffi put
-dup 178 /ffl put
-dup 179 /dotlessi put
-dup 180 /dotlessj put
-dup 181 /grave put
-dup 182 /acute put
-dup 183 /caron put
-dup 184 /breve put
-dup 185 /macron put
-dup 186 /ring put
-dup 187 /cedilla put
-dup 188 /germandbls put
-dup 189 /ae put
-dup 190 /oe put
-dup 191 /oslash put
-dup 192 /AE put
-dup 193 /OE put
-dup 194 /Oslash put
-dup 195 /suppress put
-dup 196 /dieresis put
-dup 0 /Gamma put
-dup 1 /Delta put
-dup 2 /Theta put
-dup 3 /Lambda put
-dup 4 /Xi put
-dup 5 /Pi put
-dup 6 /Sigma put
-dup 7 /Upsilon put
-dup 8 /Phi put
-dup 9 /Psi put
-dup 10 /Omega put
-dup 11 /ff put
-dup 12 /fi put
-dup 13 /fl put
-dup 14 /ffi put
-dup 15 /ffl put
-dup 16 /dotlessi put
-dup 17 /dotlessj put
-dup 18 /grave put
-dup 19 /acute put
-dup 20 /caron put
-dup 21 /breve put
-dup 22 /macron put
-dup 23 /ring put
-dup 24 /cedilla put
-dup 25 /germandbls put
-dup 26 /ae put
-dup 27 /oe put
-dup 28 /oslash put
-dup 29 /AE put
-dup 30 /OE put
-dup 31 /Oslash put
-dup 32 /suppress put
-dup 33 /exclam put
-dup 34 /quotedblright put
-dup 35 /numbersign put
-dup 36 /dollar put
-dup 37 /percent put
-dup 38 /ampersand put
-dup 39 /quoteright put
-dup 40 /parenleft put
-dup 41 /parenright put
-dup 42 /asterisk put
-dup 43 /plus put
-dup 44 /comma put
-dup 45 /hyphen put
-dup 46 /period put
-dup 47 /slash put
-dup 48 /zero put
-dup 49 /one put
-dup 50 /two put
-dup 51 /three put
-dup 52 /four put
-dup 53 /five put
-dup 54 /six put
-dup 55 /seven put
-dup 56 /eight put
-dup 57 /nine put
-dup 58 /colon put
-dup 59 /semicolon put
-dup 60 /exclamdown put
-dup 61 /equal put
-dup 62 /questiondown put
-dup 63 /question put
-dup 64 /at put
-dup 65 /A put
-dup 66 /B put
-dup 67 /C put
-dup 68 /D put
-dup 69 /E put
-dup 70 /F put
-dup 71 /G put
-dup 72 /H put
-dup 73 /I put
-dup 74 /J put
-dup 75 /K put
-dup 76 /L put
-dup 77 /M put
-dup 78 /N put
-dup 79 /O put
-dup 80 /P put
-dup 81 /Q put
-dup 82 /R put
-dup 83 /S put
-dup 84 /T put
-dup 85 /U put
-dup 86 /V put
-dup 87 /W put
-dup 88 /X put
-dup 89 /Y put
-dup 90 /Z put
-dup 91 /bracketleft put
-dup 92 /quotedblleft put
-dup 93 /bracketright put
-dup 94 /circumflex put
-dup 95 /dotaccent put
-dup 96 /quoteleft put
-dup 97 /a put
-dup 98 /b put
-dup 99 /c put
-dup 100 /d put
-dup 101 /e put
-dup 102 /f put
-dup 103 /g put
-dup 104 /h put
-dup 105 /i put
-dup 106 /j put
-dup 107 /k put
-dup 108 /l put
-dup 109 /m put
-dup 110 /n put
-dup 111 /o put
-dup 112 /p put
-dup 113 /q put
-dup 114 /r put
-dup 115 /s put
-dup 116 /t put
-dup 117 /u put
-dup 118 /v put
-dup 119 /w put
-dup 120 /x put
-dup 121 /y put
-dup 122 /z put
-dup 123 /endash put
-dup 124 /emdash put
-dup 125 /hungarumlaut put
-dup 126 /tilde put
-dup 127 /dieresis put
-dup 128 /suppress put
-dup 160 /space put
-readonly def
-/FontBBox{-20 -250 1193 750}readonly def
-/UniqueID 5000789 def
-currentdict end
-currentfile eexec
-
-9B9C1569015F2C1D2BF560F4C0D52257BACDD6500ABDA5ED9835F6A016CFC8F00B6C052ED76A
-87856B50F4D80DFAEB508C97F8281F3F88B17E4D3B90C0F65EC379791AACDC162A66CBBC5BE2
-F53AAD8DE72DD113B55A022FBFEE658CB95F5BB32BA0357B5E050FDDF264A07470BEF1C52119
-B6FBD5C77EBED964AC5A2BBEC9D8B3E48AE5BB003A63D545774B922B9D5FF6B0066ECE43645A
-131879B032137D6D823385FE55F3402D557FD3B4486858B2A4B5A0CC2E1BF4E2A4A0E748483C
-3BCF5DE47CC5260A3A967CAC70A7A35B88B54315191D0423B4065C7A432987938C6BEDAD3B72
-AD63C2918B6E5A2017457E0D4EBC204B031F3FC6C13D7DA7277A94BA018E9998B3DD888011A5
-D7C4204989F30F908B95533BDA845746B673AB71EA5765A0D14F4350707E47C8276305B28513
-CBE1BB0DBD269A53719BDA46E536685DDF78CA0146B6B93E760256B74D939D4E35B5E77238F0
-4C92298DFDD188FEEA30E053EEFBCBB52F2011772B3AAE39F5805597BBC1E8BB75A446CE0140
-30F4F2F0F49F9E962EE4A1024A746FA92A3628DB5270732B54E43FE5ECFA524F127E5FCC788E
-77E66098336AD67FE4CCCAF0253272D5DF79864BF4B734CB9A5859D557D8BC11B8E00221EBC1
-2E97DE4B1F466EAD83A4C894709363BCA9040410A52D592E34EE40CC7E5EFA920546B981AA65
-9513A24B1B85C221A1875B62D0B89E57A368321B8043A5B094E0379760A443D632892B14AD6D
-19DACC8C78093243AD67E6A308E56E6B68412EE690B10DAC6E17708754A00D51FC957B500EB8
-0175716EEF4B2CA1EF867614659BEE3F2B7319E97B6FDF1EFC847BF3CEE3156F72F21751DA8E
-5FB6898919E6799820D3DE0642D756E09D6FAE4FF08DD3DEDA3173BFF4BB11F79109C97DDC05
-897AF709EA199A90FCEE8CE4C7A3C15B18170C41C04DE2D3FBA8F34296A95B8E1E8DE3739B17
-273F8F2C85E914615C31FA0DDA3022E6C4280377C7FC05D7171BAF721953616AFF57C215B922
-C58A7EADCEFFF122216DF93857BB5CF4FDBC70A42B604EC02620E28530F2896C96EF1F89F205
-4BC547825E74D22A0E029F61A24352D29E63BC4D26CE4802E718056EEF9D2C9AD5D440844945
-17023CE9CEAF70CEB3AFA574C8BB8B778397072A13B9D13E4638A2C5A4EA6ECBCCED04122D1F
-FF56BCDED0B2CC5EA98882E0F801E665C38314E603DBB1A53BC7E2A9C858CF8BF7DE0B53CAEE
-93246DEFC55A08C6DA35C9D2D380E19AFE4B3EFDE32275CC1A98F6A6B7BFC2D87BC40DCAB876
-EACEF88528656E5B520CB780B5A7EFE48F3DAAA5C53465510320115D6B0A535D4FC517853728
-37238858AF9E07DE961647B8B768FDE8F86622AB37630BF0DD6114A6A465662E368FF60EF240
-512A6356FA1732E4F292E00A7210DCBDB3850DBA647150D3757EC8F79C7EAB1A4027E22DF1D7
-A2A07360ABF76753BC1FBFE5B9D5F6048E4848380B1CC5E7998A63A118AED949AF6FD702DEF5
-96AD934CB0D462E5F248DD691053E47EA698176EC6F7A18CA69DC6E3D646C3E4E89B30068E4E
-04E1399A36CFB42D51123C6DBC226708FEA1E8AE38C699A49F9289CC7F133347A301014AE436
-D7BBEF41200C2ABDCA5EB576887707A4AB7A62FB6484092BF2B944F00D0EEAFBDB36CB5F31EF
-36CA1A01DAD6C35FC63D175C461FDF98E57FC4A9835A4F3A26D9D789050A1DE31F5555D368BD
-CDB3E0D14491603AD6F603850161EB2B094458A17B7D5B268A1066E003BEB5FE3E2B5026F85A
-DF28582D4A390F30A46335B4920A97F98F96F87E46127501F9F2A5BFF9CF0DD0FF53413745BD
-685F573045D84CD1D435FDC9B6E37BB7F8A2B3521114DD5FA77CE938D44478D9CF3E4ACEEB48
-08797FF2998338C16326B1178B22A78D19ACABBE1F52F2399D1EB7B18F1699C978ED9E11ED60
-B71F26E3C411CC9DA0E0112ED45A9ECEC71EF6D9055B1ECD8BA74288E3AEE52A9BA8B4816E43
-06BB18089BBCAEF3AE42A57277FF1BC6CCB4D6F02C61474904FF1DF172B5FECBC641A95311A0
-47356B11EFE6EDC4415A16FFB86EA5ED193E5BFA78D42F5E2D3F39E630D934660351BD48FA5A
-62FACA95BC1B4ED7EFF27A796208FFCF13616CD9797A368C66BD9CD27698AF716CE1E3A7A9FF
-F1A50E7E46F07411E91FC4FE30D2338E4412E709226F29000A9DBDF55C16EA8F0C7717072201
-BDEEEF887415AA4283141CD661085E0FCF2D6880B03AE978E8D7B0F735D52A78DA3CEA4777C3
-CEAA4BFB62A4D15B0DB96BDE97A0CDD4DC0D5B1371760559CE42411C5C4F0428ECC4109E92B6
-117D63B4057125EF3FFE6B5B40639450300F0F98ACCC48B923015CEF850D943ABD5AD61D3216
-EDF2363E5D305D351DA0499827E2A6D32685CD8F03FFB30C90AA594F914E41EAFDF892012BEE
-0776E5809FFDECADC2AAD840446ACC20C36BBBB81509FD461EC2EC3D754FF89016987CB9D96D
-901E7F0311446531197A73F1B9F9DDA3F20C2AB162AA2AAE8329C83FF25613E55F2581ED496A
-B051ECC231BA8CBCE253B6797C4D7CC9EF2C0B9F298D06EF8A8BD5B86FC61D8A5C43F0D08C19
-1BAB04932AC6A8E17AAEDB9301F9A63FA4E240935A58517016C9C81AD821021CB10F81BBC2B9
-403AB735DA0AD1104B4BC8751C8683BF22F998BBF3F0BC319B93C54B1D1AA2D6598BCD6C47BE
-3911D705472A675164D7A84A301940159D5A812872A3610E01CFFC3E9A215837960733E651D8
-B976F5B634A638551E0BB6D33930A7ED7B33FB62F545E10CE512B15C80A256CE37F2427F5578
-086809028C46AB299F34F56D690068F7A9BCF6F5F576816099E6BDEDEA4A219F3420E3E187F2
-38227F67D50CD0B49A4943C71893CEE0C1D79E7C31E8139D0B78D1F84AAA50A643EB83FADBF3
-68FA219F422ED543B5D16386019D437B1B855C4949493BDD135AE276AAC181E60524CB92F86A
-CD9410B1406E115F32DEF16B22797803E8B7A8DEA28E356A2F4B3BF34C4EB44B6F5341D7657D
-64523DA0220B0A7AA4A852BB8BD235479902CF9972AE0D75D9D4BB032CF0A92EF920F01E102C
-2E625BE813930ACC811377B9830B9F64F1BB5E3B0BCA62746A93E58917C2854478F779CC30BA
-17C132B79A0EF5C6B89435C28E31514685FA3AE81BCA4A440D241C25EF91A486F510284F1B39
-144F17721958FE1FC40F63FAF05FF31A1B7FB06E211379FF045F4692204CB3F556217F888A26
-F3BE16496AB58C1FEBED7C055EAE7D799C2B21E1D2EEE0FB3A5F7B3D2037B06D25B7C78F90C0
-4320AFFA3E0563E6626BA15689C1D3F7384F265E207958D6929F46B82009B11919B2D76F7CB0
-5AF2325E6633CAF0599C152C94BE7F2AA1DE202F3F676FF46EB32D1D0730D2B67B69A2E3B40A
-C2AF72614525525FE42B33996E2036E3EB4C7EA04A42C675B044C71B612F4CC8B06A311416AC
-D9087AE8618F86E807CE2080EC38E62F3E099759A9AA0554C1863546970BC458B15B5F8E8E9F
-B2833D967230868848DFD1716D8E91D37CA307840EA487218E403E2F77007C9467E5418D6308
-906F16CD196A105CF2069F1C2480845A6C18809E07BBA1BEE955B27514A4C99FDAF19E92A70F
-29732FD46636D4C9EF5B151B14ACBFB424B90A281F3E0E8EC5E4613F1AD73BC89B4C56EF33A4
-61DBC803CA529104926FEAF0F1A133D54D9834C1F13CBD9D57A3F2A34B28F555443618ADC110
-4BEA2AD095E09E4C56FB0624451603DF33BDB7440DF13F34206B85174B4131F52B906310528A
-B74C7458BFCFCB5C613451BE6D139FF3B20ACC1D9F79F94E54110EBE128F227731EBF5A58350
-23BE82BC90679F38671D24E70B0AD79CAB6A25F25D2CE0908CA1E6A8A33BD3C7C8E3D4D805D3
-242F33EACDF18DD7D97DA7C6D1F8357B6D7FED22E045EF5EB6EA872CC67084FA0472BDBEFB19
-3C3D494FB26C2548D8C27A7AC9B7F5BFC07982AAAEFC7818323D0A4FABC0D748B8A9CF7EB882
-3D05CE71AC4C3A8CFE18111C865A1957A28ECC041F2D8D4EE37EA4CBA429443BC14AAAC9F50D
-DD460237163BC11C23E45ED061AA7906DC5D7FD3CE3771A7DB949FC818E6D4CB097F67C2C16E
-60AAD774A323DF923E8FFF13F71B3EC691A355B7457E7FBCCF89FC2A4CC444E691E4E37DFB66
-E70D3E75389EA4E869029228D6106D5D9EF41567BF9AA220D569CE6103E7DC737C7B0195D7C6
-20FBCCC0DE6476885F36CB3EF80591FAEC849F7C827E4F73B29F4300A28B0D5E63311390A324
-0CF31C838C26744A285B0AA596AAB377D2999B7941E99D233E879B678AF62997887D08F54246
-BF1F5430E984D6BAECF9059CFB08F261CD144C31A90A8B39BDB7626A0B716D2E7CD811B2B71C
-6F250366319B36EE2D27CC87278F31CA61E35496AD84F628B5E89176F40EB4E8DDD525C8829F
-7F8AA5F33A5B44D4B3BDB723A84DA15878BA19B01A52BBF6D8C77209BFA6B4E78A2AF9197BEC
-6F34FCD3757D2AB00DEE8785212E9FB7B72656593ED6CE707AEA70C51852B7878F9353E202AF
-2FEB35AD71A2FB898816D2EDCBE5C7BAC71F43885DEA4DE71C37C2B68EA7266ECEFEA1C31100
-7A819F822E8027F01C9F5EAECAB699CCDCEE52AD6D44D1E0251C63D914140C45851BC838F8F1
-955DCD0EA727BFA30A2B0B24FF993C971CEFAE7E4584FE64C3C04BF94AD831617C2803B54614
-32E6642450748DFC654A1131DFE92959BC04CCA28E46F2E89DC6DE210C00D107AFFC8910C975
-D7CF231923F30E086162AFD98128A334B4F9F5EFADE134F195CF0EE34972BE8D5C42596C0757
-EEAA0ED6EECBEBB4556C56F7E106A1B2AD69BB2362D6F042DB73AE4F906B50C70127F14B7807
-3B8ECE4827D40033E358BF1C078598A6DE6E381BFB13789C87E95B1CDBEE1532C9B890F9DBB9
-B9F5B4E88C2EB87045DA271A26F982541541EBB93245E5E79DFD6C6EDFB48AE1B1B4619D6B10
-2583788D363269A3DC2A265CE0E76796C99AFBD7846AD9ED23DFB8802109843C5F67E21541BC
-7367DBD8E2C76B7CEE44407D1CE964CEC96860B988E130F78A310734CBF9293BF0FB5A871187
-BA52C4B2129CA5ADAB1C8F61D833B9C231AF3B8003924772CAC46710A612F443C1DB6E0EB5F1
-7C333F46BEF7352FC5AD620DE199071951A1277A368D7ECC273E778AA7EA39FCB727C960A340
-28D2F080E3C7C2580DB81C57A4B581A41CB3460183F9A06DE0EEDF7ECF8D131839C9B36F416C
-096225FB88E1B04E37E002B50A101EB109FCC0C94701B8E69F97451DC6D3F6E027715541A70A
-A9A6DE4AC8A49122EA016FF43160ACF30A53482D9D0FBAE022B0F09663ABF6CE64F0E4AA3317
-0C909BE688967FE2B5A417FB11E4D9373BC89EEDB4EF0686AD6F0227F140C67A6915E6CBE4C6
-8EDEB48619BAC37FFC4F5C268EB2166A2A66E8F4E43B18AE31137F5E8E83771D677F98B9A975
-EE2B373C49C9B716AFBBE9DF2E43D5BFC622DDB1AA54D7C1D324C1067FC6C203E0244F975DD0
-337938DD4FFABFE7C8AE52576F6C97E68AAF0879FF75B512186174E7CE0E4EA6E778F2B756BC
-3EEF7A1936BD4C4E06E4AEC74EDD876EA3BA792F3CEDCB2906ED2314CF222840E69F90C33920
-C3222FFE70035C0E90C8F490829B85FBD818A6BE174F82BCFB12F260955F0F6B5DC20803376C
-5FF58F3374BE1EF8F1B5DAD5BB91A32369B83E5C0DAFECFEFF5ED82D1F31F593862C7411C0A3
-076872083AB327219AF9BA7C86197CE7EC3E36D80868681C4420920E8392B7C1C73F415CA94D
-EE900704F33A5028542919901A382EE01F4C7DEFBC9E4B2059F7385D3EED97425FB302ADD76F
-FC7271CF934E9DB0511C3A86D189D11AC71BCD1716E37CCEB279C8E775FE86670BD5F40C7FA4
-664A856139F567DE5992D1C0AAB36CE53D39D7FC57F7670E416BA5CFC361BA0E1ACC95594161
-DB7FC52E92F2B5E1CA6B7C98B5BE6F1687BE51870EE0D3356BA4EA2D8869E6A76FCD09CDF9E6
-9DF80A193CA4F361F3C11819F50919EC23CC01B0ABA9A5507DCA9B27520855CF6F0F6B27F3D3
-A0DE55336C0784D3E6A75D8F873E841CF6DEA7E297B5749257FB2C598847B90331EED218E23A
-8A151F973DF9777ACD8C5E4F2DF4C54E95E271653B63DBDC75B8A8BD4256B78A4C8B8B264AEB
-1F00B51713E7FBF6B8D1AA45CF94C6F89476BD28124B2ADB58D616DA5493F9694F856039B5CE
-801A2468FB29669805F731E48DF97CB5CEAFE1EB6D963DFD7391A3CC3B272151E20195E7E181
-99F361912253EB81DC9CC439AA20EE5065C553F15FD29D8AAEBF3B61FF8B721F6930742B17B7
-B6A04F622453A3557A58331972DBC436847B0465A5693C95F3C4C316C8EAE4BA464C7520FBC0
-6844B2BD46350EB29E380580E34D3BFD82A61BA3003AED19BF8B2C06AF1FF80DBE1B8198D920
-1F2F6A256C40E1FAD376BBACAA06919411E4C07B8D1027E2005B986745432B6C9CFFC13EB294
-DCB17B604FE12A9D233308D9A62926F3C8B0D0AF6CA2AC06FB17820D419D56F3E125FB50D679
-BF19EDE34B7870D33176036FEB4544B331DDBEC7482F8C0C5698890F83FA03295593BB254E0B
-2BE01825E5049A4381DA26AD13130C7231F67635CCF4E02909D8B5F1EF887E0DBBEEA26FD7F7
-836DF4CE154CA6C712195CCE21DD8094BEADB8964E12FD3F6E359FF822AE5D6B35E477263D38
-3C1A64171B731B22B3F2FC9B9DD1E7E3A5AEBE55A72FCD266E9DF0D7B29A7C89C5ECA9FEEFA1
-1ABDC05AACF9E561BFA1793A5DA13AC8A01F4F6E3921DCB95B8D1F317941F717DC9A94B9C5DA
-B50AD38041E1530FFB377CB391F3CF8D92D85421539180EC91857BB7F7AFE1FDFF71D39496D7
-8C88DD4C18842900D3D886B278D5024AEC128901394C23F6EED66F1E5AD8812AC16A32CDBC13
-593120F1DE685C66579B015B7F8116448E7BB5C96DB6F15675C9BAE44809912D678FBFC6602B
-DAB5B0CC87D76978BEC68B47F48D599423B57827715CE9C08F0136691AAB20197AE0F2211953
-454CA523002E0E47696DC4B74E02D0EAAF6275E4D07B99B8012BBE5DF158D77D38D8961DA320
-343343FE12408DDDDF6CB84CAC70E44D04C6E9ECFCE8F8D595BE1CCE9F5847F410E0DE250B75
-EC48A9EBDCFEA5F4E44697B3350C635086945BE544755823AFF7FCA4113E9DEF0F0F8BE4E58E
-1B2C2D3CDB2F38BF23B2ED94596790927F5C535117F2A522ED99700D9E0B2372FA1B2E4B09F2
-3240F51E35E269A824A8D0CCB286E6219A9BEB516FED7C247D7A6C6B12B9B2D42A5564C64257
-5F6A93406167C36ECE06CFB4F3371C80D3B7D459A5F8015771151A62E9DF015B7893E5D54241
-BE2CEEF584D6F5502B433A0654625C0EB71F2820D2C8B4F4A2362E603DBD104A821EF860C26F
-6068E63F6DCA31B85C0A0163F644049593E30E5C793D7304CCF9D749C35458F4748A3A5C89CC
-D6349FA0DEFF2D520070770C740E83BDE55C8B6D8DB5E853D1C9951C3A92DD0DF49E1F18EF91
-7169DF8EAFC35B7D8A9F9F546DE5E5BF27F2DD5B2B242D9DA9A37EDCB9B3F0D91726941CB2AA
-F60E51C36DCE821EDF1DC3B071EE86DB4B3E071E08328F6204FFE21A3DCE60D7FAB0BB8B15C3
-B7D2964F4F4DA6C07FD824D747FD70C4F9C88CC0B6B5A653921AB61302A78FB3D1CAFB10A956
-8CC195E4767C59A1280F735AF516E3DCF1BEBB500249DDC756851EB56EE5BAFA767C4B1418C7
-76F5EB5FE36C4CE286877BB0178C6FCBD91734A2B5C043355AED6C02777D4DCC1240862447EF
-B7A429B6A3CB5D44ACC72F05F1AD6295FEB6B013CD7BCEB970BECAEDA465E325832B0F3997D7
-1E524810EB7E7FC002FA0FBF9B3165792797D5E4C25EFBFFE0D22125A09C9804EA51F7C3023D
-753198FA2BE152F3C134B7FAC9DCD9159DA11AC352999E503AD93DD80E2DEA7781D95DFF18C3
-D77170B99E2604D60C57C906FF1777BDEC756866224EE13C40C0C8D3C4424CDF59DA148DDA9B
-117F961EB44AFDEA66B2FEA3B2D4B1A8E49C347AC693191BC67D7794530C7AC0461BA657151E
-C66295F2AE05CE660B09945480AE821EF1AD822FDF34C193184FFC54AE01DAC5599CCAB04039
-61CCBEC949B2E6FBB67E0100969C6B3C546825AE7A41BA05A525B47A774B955D3AD85863E954
-34C541D34FF9AA6ABF0B0431AC5CA1E103BBA6C672BE9EB6A7AB9E0FC4561453B6DA7723BE87
-DBDDE416318028B8F998C3D8F029196F997C62365E197F1B3D87A2FE08B2FE3EE64C71D8D584
-2E71990F4EEDDA0A8614361715B5DB83EA756D6337ED1708B3005537DAD6AC1DD1F3812D48D8
-F8C259FCBE6285F4002072B433977837371D6EFA8E3A07D126E1757B09638018884C0F45D977
-C9D44412CDA894F986226C62D40ECAECC05C141E522B1CFFDA2FA4038A062B5AEF8C511C7485
-56D8D73DB4A04C3B2581FBB2D6E3264CBD5DDC471E6AC973CEE1C72A5FF0B0689CA8A9422E55
-CFF66D1974DA6303DCF158A4D406781E61D5A8A7772A4372EBFA49DBCEA766FD3013D5474392
-DB3BD0DDF81BA91C1B15BDCEFEEBC9224AAD41621DCF01418D03586D7563DB3BF66FECAB2CB6
-8B98D37411D9AD9241A0664A160757FE5B4AAC793D8D89FB42A2F6CC02FFC69A70C8160860B2
-D1DC9263572F49A7B7EC780E03B8D7B084DA88410A62638034E3A149279B4E14F2B93A1CC5DA
-39EC1AC4F63B8DCA5DC50F27A21F1C59517A28396C719C0368458ECFFD0F76FF2D2E5E11B39D
-8430771DC3AFDECFCF4F6F2DC55030E9F584CC37435EEACFF0D208D246CCAB0A03A2CF47C4A2
-161B22B39E67B4F63CDBD166F85DEBD55C126DB2A431582C056BA62E6046A159BF636657F95F
-28A7C95CBD718BCFF1EB2673CCAA16958B40A8C853FD8753E9C8CDB1D1BD6A8D27F38F4A34D6
-70C7D0DB01A9E56C1B38BD831822CFF0C60D4CE0B7FA676E0D8B3425BB5428958F4C22DAA969
-FE2917211EB8A48FF3506AD193BDC48F62E41CBF01C53AE686987966A25CBC7C931433D7FCB9
-0204C93EAA7D58BD89F9DFBA683B0D9482A2792EE83CF72DF5596A065A26CD5919625BF86549
-EAAF1D0478E4E023459D747FA11C92BAE21399B435B4E4493BAA96C24E60AC2099DCE54CC78A
-720C5EF49019B52A6816CAC0E9425F52C4E78F2F5CF1B9517A8C069323881401DA0735D3C9E6
-87DF1696DCDFF01BA592A21C5B26E3B1866DA966B929253E3DAFF6F6245D1D4AF667B9EE5753
-D549E2FA1B634CA4AD4F607FB4DA16B99AF607452940777F817BAEBF8550215DE86DB2883C30
-05C721370E072CAEAAA60EC9471923872427A3BCFB47C045696498DACF29F383CB452899DF93
-FFEC5A95B9D647D9AD93D5C0C57128891EAB144916051D2625085D9182D0B3BD14ACBE369237
-67C89450CBBFEEC999ACF0BBD1997FC0028D0F35BA39F816A524446A551C45BF595A569A0279
-02B02C982269016839EECA6CDEC7E1AFE0A54500E35553390A5A83F71718729B76783CCAF83F
-68F1E640FA713A5423736E8AE37C43D19581AF2963A661D0683A06F4AA318D380427C2BE040C
-5FA3C08BF471C22BB980B38B3F36086E29BE00076055E61B28C85E2BBC843CCADEC358F70B8A
-1581927144752187117A0D803869907AA431BFF15ECC8F9DAC4ABB32D4143EFB16611C3B93E9
-B70FCC9EFE8D579F07AB5DE7CCAE230F95D94E1F6362391DB1028CA5792F499E4524B73B0C99
-36397F5539859AABDFA432D7B1A19639F3AB1730DE30EF1AAEE9C0D38A40720B7EC332BF158B
-0D04518E97A819E29EEB458524DE37D9FC8E46282E56E8E7DC7C97EDCEA00C1072CA101497A5
-B7695526452151EBB838C597EEEBF0C045F82C751B68B756352C745B211C81573941454C0108
-9EF856DE24C4A5C9378048EC7AA146D7C7BCF51B5EE0B60FBB1A18471629CDC5D6077EBCB190
-2D707D6C5FD755476677760ABB4D1ACBA85E7D4931A2AEE47FAF850B858D7BA8BD98693E067D
-A80D7600388FF2643DD0246A16279CEDE00A8ACD56AB3F80F46D9BA434462D71D3D5DEFB0E2D
-7DAFD54847CF47B5C54E27B34F19AE5D088EC8D8239E8B035230166CE9F75274BD54F4FE650F
-274CA53D51091510A022430A3904383850A7E2D102C6D29F88AE9DB0490D52519DB5E417E782
-DB711AE0A166DAC0EAD955ACF0738F59A6E7EA3C773BB38328C5361B8662ECF1C9B97BBDD050
-D910EADE82243D3E24887AC8E9CFA037D9EC4E158514EBEE11D9EB195ECE349454EDD04C233B
-5FBE7B5F0AC57A87142DB0379763A2AF5A1272C60DB7402E4BB7521F5C6602239381ED77DD2E
-8C0891AB913EA0B9B93A44F3A8247B493477C47C256BC600F775C3BB1A845FF1B23B2468D028
-21A5714211A1F2FE6C6F16011A21367F6C282374E8B41C4FA81C575A73141337FD2AD2B169FF
-F0DF66F4937E7200D7B9AB78BFF5976AF42658A44F68DBD3F45862C3D86C51F4F105473B578F
-3DFD5A970CBF671DE89798625850FD4919A4B86B40B7A034000319EEDAF25C0360A633FD4A67
-5A43F06405DA42931D0AADE616CE5BC0CCAFF8E99CE138CBAEC1CB455A52E2F1AF704EE218AB
-D2FF0558953910A726D4256CE6D328E25D3498F18352DA653AA617B1A5C05CCDA1EA3EB591A2
-DB9A1C5BEBAB22D60619D2E6B6B371956E99FDE940C3B9F2AEE0B5C4FA0C902CD93F9E095B0A
-8AE9A32F0360FA22D7BBE965B08DDB1A7A2130328F478D375CF3A734A99FFF582166310AF3B7
-9418093B979046D34CD50126608A5A843D6B708269B626BE70F0C04B046DBCCFD4537995608B
-7A8FB52F2A4CA6004F0789DF7EDC11F4BDE697169DEDAF3A029B755091B4124C83E35E818B8C
-957775BFEE9CBD4A8F7B98547F6DCD4697F5630295B779906608B0DD2CC875B1CF7F1E2E2F1D
-DB45C8936D44D1C27F9BD8BD8F34AE624AB5251DAA505DAEFD130560EF62CFD400DF4F040542
-463D72A00CA86AC84D18FA3F3C56401CFF3779C733C814A69A4DEC3E055BA0A5B25761FE9EC6
-9F17DAF9695A2D1602A41E2E912ACA588F87E99549D670464955B78FA8BBFB582D52C6A9AA72
-B075AA9F2EE7E66DED85340E60AA552389D41E38308577C164E66C03BDDFC4742AA25C5A9654
-44E779D5BAD9822AE25D08C05E9070AF64480C7BC4323B2DEF5BD78672DC58699F0B47ADC58E
-A397CD341EE81EB212CE34580EF980BA11C7BD014111C5F30F5E2AD8E3293BC199DC173A3DF0
-55E241A0B9AA021AD6273DBAE03A5E536E256F9336A752BF2B1352E7159C32504797007014FE
-BD4EC5EE9D8AF97AF7C0D87EC5FF79148F033A828675A173983EF4F018D9B8739AE9A3D2473D
-8CD8003E3E7E268DAD511402532C2BBF49E654FF4BF8D1119D9E3204B8A82120445E3AC554FA
-BD3B42684935C5CAE7E7FD391A38AFA7A2E3773731C517C23FFCA677DA93DEF41F52BBEFE748
-B2B8A35F6BE0BB16B35E7A6EC070A3BB3B3E996DC2E3D4B357626DBD3C44F4614C7725DA3877
-849C1B883E509663B07AE4E5C638072163F979F97DDA7183A8E755E1B1FD9D492555E2C1AB32
-EA67D6F481C9C9C4AB4A048E28FF3FDFD1C5A0890E7A45A4A3A110D9B7AC57EC6BE60ED6F7A2
-5C110ADADCE3401836FDD7A73A645AB0BF866A49F93E81FE06E01B8CF3FE2D415BA1BD2550E9
-53D8005B01DA3390ACCC4048697D0DC8814D35A7C894635DB482CA710CAB1BE9929A16663663
-68FC930973F1492F3E0FBFB23A47AD644C8B40292A7E4B9F41913CECFDE6431307426DB8CE20
-7C7F8E8E2E3252C3E4C414D0CE387C7294BB324A2C61F6412C14EBD4455FD13807639CB60BBA
-38DC1B55A3FB37F100399A9695AC8DEF1A870B753DE3B3BD76DC583C0A27324CA6CCAF7CB240
-5622FEB4A3E0712E4504D8B63063D15AD6D7AA0E91465EEE3AAA20B16BC5770EB10C1494F4EB
-B3AE2B5F6002AD9FB749C687DA4D43E45AC4549FEC2D18E0766EA47835992FB8C73B39FD38A8
-019D8BDAA8F08C510021065FA847091C4F82F73727ADF79F03941E37706460F517DFB7988E04
-92D7138AFC2B0EB7597481BCEB2CABE6D9D2D5D5F99DE33139853BC170642F1471331FA532B7
-6399B496B8633D7EE8B91E4E1377748FCF36165ECC639896F190FE5A4C5C58BC6655A46BC9E5
-26159E10B5C62816F217B43A77E875B8E00CB7C980DEDFBFFDC9530C716D5A06FC5D571E1B52
-01E6AA0D8C9F7E9FF20B969FE1900A5A8C52C329256A393D92B08E705C3684DE948973351AA6
-99738EA10C550B73A54DA15685000FAB0330CB6A508B23B0926405E383EF37D4C9F71A49E223
-3F46C61B718A516A1F3C77D28B1235881FF67B044FDEE408F03BD62B9EBC06965B378CF2DCDF
-34B3A8B15C783C2AE45AC758F55C0861F833F40D508E53A4E50E2A27F009DA0FBA790FF2DE70
-2EF2C43F606FF293F7FB7EEF9FC5125F2D8C9393074BAAB8B544BAF2D8FC48098F3E4E2F7B56
-D8B389ED066B29D9A561ED5B9362A54D48968DBA73EA1FECB371E34950EC27F8506051A8C7F9
-7A7BFF6E7A7D98AF3FC0F17B2FC4A77DB59DF49D2A71CA281CFE90EF0DF7821B791D4C83FAC2
-62CEBCB096EBCED235E157A332F6ED2744195AB30357DEA9F6C5B1B2BF7C8D43EC1ECF62D13E
-31CB68A1F8E9F5B3D4FF92458057AD1FF7583C2AC5C8466E691BCDE83F608D1B79EE6CAAC20D
-A4E4A643E4E2D60AE28535186CB6BC80784B521D2DC8101DD5908558BBEAEABFD021546D75EF
-19FA4BF86C85BBFD9ADFFB6E815A66D7A8ED467D44692D9DA81275D533592CB826656BBCF67F
-89FD3DA4A53DE256CBD06F7612848B2288E976A01352FFDD417505E5B97F19CA4E8FA3BED101
-64CD8AB3CB5305F3E0279B474594F2E1C915CB6B3E8E2A7358120BBAD2A3B2D2E199558CBFF7
-6BF4BEE99F3EECDA0BD7C0BC1D132142B7274878B2E653E33889926D5A6A25728964DAE20590
-F1E53C33941678EED838118842FE14F421254FDFB14A6CFBB04F02F582244EF0DEF80153AD9D
-21B2FAA95EAFF02137F478B8B3072BEE9539E8AA1050B29EBE4312636DA000E043888972ECE9
-95B6316883794727817AD23C253ED315E0199C0D4BBAFB88CB5FDBDB978526A3E2ACD9B14347
-46E856FEBF0F3827990BAA0E7BE9AF759930ED684D7096A14F9A615C322E923891787619F80F
-A08D0786DFCE88BC794F911E31B96D4B3519CE1A757FA47C2A3FD140CC16D338CDD237212BB4
-691D8FC2AADDB60A2C072AE0E7F9C00CA9BB7ABC46CC59878841E68783EB2FBF9DDAD5C6DFDD
-FA720CC52889C3FB350FF122A6C02F8EF8EAD218212BDB8E61B93E8736BF6B14828F47BBE949
-40708C5B40FF2DD68E2ED11065836C33252FF84AF4296A71E274099ABAB51C210B7AC3C3D533
-CB328246B0CBD32A6F69C334FBA87FC76B88579804D39F8029FA8EBBF0259B33F12E7C45DB13
-01FA1FAA91377DBF284BEFBFED3B9D5B7D70407D5997EBB50431A7280B6DEA61B9B090663B50
-5DA43B9DEAA25343F5D7167AE7F4810282BE0394E4B36C733AF5E31204C2A539AFEC75311018
-D4CDE08230CD713DCB0F7DFEEF153DED21DA3E46C1BAAE7C22285D383DC3BD1D96F786D1C444
-3424E310D7C3D46B7A6D84A43A8AEF51BB2EEA0614EAED6308E46C64B6006B75F21C805A363B
-5D5EEF527162FB86DD6F1BC51A28953B5E059201E02829C4984A6E0FEC9BE90E40826CEFD22F
-A7EB988033AE50C3132C3CBD4881D3E561D46DE04AFD4788A5EB64EB09E1C8282CF796A2305B
-66D8437F0555DA80E7BBB67C379ADEE0C345D9DFB4DF3D6DF75DD35F0412808AB4A77FC9B327
-8539F6D0C8031DB76287B7EF34C2A4574237FB35ED635E7EC4F220782D9934FA7C505F22013D
-5D20A363D66C3190A17A03F61B9E44C1411F68F6435D06871030D1BF3FA3C1C1C827B8DF8249
-B29411C39006F1E81A58667E0B9B59656E5F67FD0FE246791E9C3578AB2B9C9BE2EDE94B68E6
-5C86579F4A577A6C3A474DC9D24ED2F5E02FE5B83A105670BC4571508FEF80AFB27160B5A2A7
-6A6576712D542C93FF1FE503E5E129A85577452FBFE5E456F5D24A0B3CD3218EDDB5F73E883E
-53979E1429C4DF32772116BA7482906CE4ED29389FA68B3491845FFF6D333EDAF99D76904E45
-1762D3D8ACB70B923AE492C8CA5456106D874D1895E9272046D9422FF81EABA7A8F15F767849
-6C7B0D58F1DDB22597B137CF6E5BBF889B53502D97E4503EDD053C41EF961E0FAA8D0B2F0387
-84D47C790E40D8488B49D9451E8DE49592ECEDA64C6DD80916123C348BC81331C6CF3F29ACC9
-D07E2718E8A25D90139B5886EC6264853CEEE4B6DA9475D32679864027E4BB4342DF9AC0BD1C
-A08C8138A779944986791A62DF1C2903331918296FA754D96F53BF043D04EAEB7627F2863901
-10848CEE1CF2E2D04AD0B618F83D554856948F4FE7C8E5D13EFB6A0E5F7E3610A8258F00F808
-C1F4932C545E047BBB6377466D986607D17C7A262A6A7D87D931A491073C65C691FF8BDCDCFC
-859B8E029183B5CAF6228F6BF65CCB94471E5673F4E8AC4BF7F724C9069D7590A6E6F3C877DB
-42766FEC506068FE1A1B7ABA6EA04C20849BFBE18745621DE771B38BF907210241BF6E07DDE3
-2419EB87F987BA49A646A9EAEC8194D13F8EAE89EADC2010B7FEBBEE297F7A84EA329823349D
-0EA5A639514952679C7AC723244C61DD31D106779011DA488CAAADE06AF2E2E658BE0CF325AE
-79B2C917FA5C262022993ABB17198DAA595951E96C211B67897CFAC55E35427CA16FFF3C531C
-5F503B5DE720C0C688D95FE0EDA1D54D4E11FFDAD3A225A7C46448D1F2AF3BF9DC0C9D0BE10C
-EAB53820A8DFB0FA5BDCAFA31F6F0FA6A71D84CDD572B858976415A39FDA955429532C7EC088
-A61A55F042179993E25722CA2DB950DE988BE65949CBBF9FCC1B9D748E2190616F18D3957958
-E8334E37BCC847431E671026615623B1705E2E66E489217BF6A0CA7C1807F25C587D5257F843
-759023AD9824DB0997B0E77EC11A5935223AD1E1C7CB832149BADFBE4858E2E15D6EB37C401C
-152BC177CAA18A7F80BC246FA59AF489B5E5A8645491587701A387B381D71265B4D59817C028
-8446DFCBDE203D268F60DC5EE0702DDAF52775D16345B17559ABC10CACE2096C4C38C6ECD85D
-88303796FA774393D70BBD2C9F047EF17030FCE5DE99B86E63C767CB58CBE6126C3AE173A9DB
-C52CD1B5D65D70A70319D307D4A2B8411C53DCD6131AC3C0C005CE38CE96D9B15DBC21A12555
-B8BF25AE88A80A32273DA1CBF1F1BED926330195EC39E25B1879A743F77B1EEE61979DB1D38A
-080998F63A3FA9BCCC9D4ECA20AD7AF8685E778B6A69A456808BB1308F989435347D2FCC377E
-6C90465FD4BE1D2DDB7965D1B68A4ED00E08B88A8D2B4C581428E11D1E3C959B0574A98645B6
-7873460D07D34F3E047C65FAB7279EB24ECAD4A1FBB9157CF74228989A9DF1C1D2436075966C
-D10D84CF8044E5476E23858BADA387F9306A44BEA5E85E3C792378B98E38877F7E4E463CB42B
-2BA2B293980778BAE0FF70C6E60D7E1076238C0A74EBF7EB88B45B7640FB547272E4B03D166D
-5E61C34E2092B8D760D672E1D2444A37E58472523E4EB662E7FC7673ABD3407A9C52D895CA2E
-35268CEF7B7B34570F1751DE990AEA9FF081C263C8A4B125570F74590286E7D76BE286E870B4
-0B3765F3C78263B05BBF7E1F130497D4A7F7AB3BD3EA5B3A4AD82C16923D1A93258D68F74309
-88AEFEA49ECB5FD53BDC43AF89C2BF27D8C3F2E152C93AC60F1C7FAFCB9B7F799FF195D3E615
-9C95AF099D43D29E6BE83917231A31CF344427DBEBC8463AB4054D378B7CBC530D3FD253D3A6
-9BE363DB9B67B44224D56AEBBBAB0C66B2764984157F071397CC054BDB0522BDE61BC278F518
-40044B6F4407849AB68E5404AF50231643F923DF3CC5F416EF892241A9E5AE445BFFD891ADFE
-427469C8E7C6B550D47A9C77442F166D50E8220EB5D29D1A2A9EC810D770AE5C1A468EA57E92
-D97A5369CFE68E328985730D6141D23E8AAA0F759818BC2759E83E8FCA6D00314FF73BC11AA4
-32080686B3F228F7723031F50AB36993D4D984C348B35967856CBE4C62C150A5280955ABBA52
-04E0D84F44649D2ED9ADCDE1151B02713D58448BC541EC3E4FF151506B0C476B4B7D19AC936B
-8D7392ADEA338B41F91E64CBDED7D243EDAE76208415B9B959BB84CAED942BCFD4B92F4E5EB5
-8132AEFD3CFA0122FD475672BAB659B320230772F0413E2C568A13961452B675E147EC08864E
-2CAE1C9BED8E9A0A13CFFB1630D4EF9D5B05DC1DC310D963DA33D3E29FC9D0CE91DB257EB107
-BBB6E60FEFE505C8967B0034A96B5C9046DEF13AB3E93A48AA3CD6D35B3EEB7AA58580CDDD55
-3A8AFBF91B17745E364EED82D791F6A4428A4589029DB11B207FCF8052AC5B7EEEB56C1F7BEF
-716B9137D6355D2B6E0EE26B93F11D0F0F83515947F84DE7BBDC999F5423A53F3033307FE62C
-21C36E8E3AF912237DE0C9BCEF1679E78ADCDC42D8A5E3021033A0AE7AE0D2966D3FD30F8B56
-B26013CBA3FDF845E184A0E091DC17A5ED14C05CD787820F79BB674F220AF4A324A770746DEA
-CA11145995F651D8B8D6F56451F31CABF022E79321B42A35A65ADBFAF351EFC7A977C828F966
-D40CBB7B59089AE418BAC4929AA941DDD0D119D4317B9E19D6660E03966A7D573484F8437F9A
-F98D16995970D8F406EEE9AF27E729D2704EC871C6BD9C06541EBF11B98E8BF6700B8769FDE0
-ED0484B728A05D06D0E86241AB1A0C214C391133BBEFD6B15C89F34D77EA7700FA864046DB39
-49CEB6D13BF0047B6879A8EA774B9C8BC26C7630262225E91311C2C093E299B33AE78B5D5AED
-39E9C62EFA0E2957BEFC2AB67A3846AA5D27ACC99B6A6422D9C90D564E736771D216D0904536
-986573DD22CB3AC60346A4C8332472B2D8F7C0E6A7AA6003A6954B7F5C66D51E7767DA3AE06F
-8A3D161691F682715757471A392771872CAC03BE5010C7C71FC52AE83596234E01313F24E8A3
-8D232B9344E3AF185C005963C9E74FC76597D24902BDCB30A3634F46F248A6FF08F011571A79
-22AE598D6CFD5DC20095B2103F72D09D1918B765AD7BFBFD03A9C0F95681452007413C53C830
-C1DB02B66336DF334F3E4E2F7CCE7C8572A1BDEF36452FB3049508C029DA2EE1070097F2C727
-BFA40B31573E950A147E7F950080C8A9AF9456CA6DFECEAAF6FC1319870A3549A6AC00BA6F8A
-EE0151ACFADAE35ADB2C0257225AC4F59CD21273812CDCE31D6EF716B8B888AD40559BB01736
-3E8DDB9686A76EA618712C8721EFD4CBE1ABD2B77B1B1BF19FBE122A9D0F81CCC1731915918E
-F68488ED2F5F6538EC33CB143E4F9B4BA5A73D3A4B990AB1F1CC2BBA16A354BD4401403594F3
-52A76ED2B44B3DEB518D1463E6393B0417BA79038A5F780A5DC0AE275E1C97130B2DD51ECA4E
-26EC6FD1A35AFA5EE2A780F18E1648FE91F6C163BF6CDAE4E176DC8880F01B26A234C7C673E0
-83CE8763AB04A1651C5EC7950BAFB7FE02C7837DAB1B2FE0F9C2C9CEF2E698A249A08A9586E6
-C7C33FD71F43BD924D970ED67C7E0133816890D0F51E200DE5A0101B694D44A19825426FFFBA
-E22B3E28C5A33D6EA4A51396D69CA91479B2D23205B74D581EF27C8C56CAB2008D2BC95E7BB0
-80D7BF7559EC0F7FCB1C78543EBE0CFFA19C89AC3F1E7420863206639DD7DF622495AD004453
-ED2A83171A49BF88FE5D469C5F179E81E60AC6D8F851596814BF16B5642CAE0ACD1DC77B4880
-7F07B4FF9930E39581B96A9352C2666D66AE292480E677558E1B09DB49B892D8222AE45A9DB9
-2CD4BFE8C103BE198B992DEA07E5C92EEFDF508CF04D0CD49D13791F94FA13D168F589B4EF44
-3AD24D17BDDCDE98DE741580FEBA43C437871E87CE94D7DEB6B9E82FD578C9C9400B8C072257
-A1BB769CC935CB313CD847F2B262CD0F8DECD896248EC584F23D251E9C8E31DAA252BB54147D
-F0DC4FDE29B12EC68AA93F3459889970B8D1B5A0F1747A771C181FB303CC63844CBD80E4B521
-2681D3B43F5BC6F200803FFD34D56F7D2EECB7439945EB57853308D707CD127C95316F5C528F
-8E4E2A0CE2E85A303801F80E23B99F892926AB7EA3500E92DDF9DA90E4B3D0E0AAAF259FB424
-6CB8968C468FEBDA0B390231F39E989BD772DDF4B10D0C1671F2129F3A6071C8BC854DC78BE6
-1508E87A705E154A84B021F431BB7A464B02F32451BFAE15744EAAEADC79417AE6B468358955
-A993F80FE88D56DF31D2EF2D2C57A0E0651BDA19F77359ED59B6DFCD1BC0283FC7345D4040C4
-86FAE52F0B41E0DC1D8BC60D6D6D5B82D9E3D06D03B33AAD4ACBB84D54CED85B6C8BAD5EE4AF
-91442BC17C7EEB9BCC25318BB9D0EEAA203461DEE2262FAAE87A6DB296465025024BE97454D8
-CD01D476BE55081BB42FD509FE4F272510E5AB5B62C71D2DF7631D4157C05724EA13A135CBF0
-66DC44645D59F2319827C42128474BBB920093A11070408E987AAF15ECCCAF9ED7E050E2F1CB
-70E70E05CA2F76EC08BD3CA0FD94504862F08000B20930A58D3CC87C464313C0EF9958B6D31E
-3EE86879DC4922141F0342553BAECD158869C51DBF69F0EBAFABC1376FAF1C564B384A6283CE
-2A95DED0F011642AC66E8FB34DE6962E1963023F8AB90483B6548C574A8885F6C841400BB602
-C33E30D7BCE3B3B2A0362C2D32A25A8D7D73B5E84B7FE8EE556E20478816FE0177DF7C6CBBCB
-994AB8C6743CF7A89D672F402E174C6B7F3315D51870EA9118CA090F4CCC68A566A463D51928
-3F7DA5522058E4B0A92F8BADD87FC18F9AF4780F0AEB79BA798954FB57F880D9A7E950A7910E
-4E3F9DAA340EDE950B69879A0B711DF81112917E5A8B7E931E75D6C664E887F3F6C58F6979C4
-37D8DB70817852E4C483496D9823A799B19710B3B0CF3B1204FD81BD0A4EB1CBC4284312FDAB
-D2ABBD3D45CFC4A0B039B8BA975AFD201323FB949B5A3D690CFB24E003EFE01DF3727B194D82
-E0B95D874CEB4DE2D6F864E8BE17859A33E4F26DBF22C59FBBEF4CF411592141141032D86F9A
-651A1F76C33C55F66619DA688383143D856E5D44DD64364DFE9824A74B5CC1D2E29BDEC81E56
-13A1D336AD6ADF83519FD50531E3BBD935FA4E5619DB888AC69B77BB11638CD3C2144D01B4DB
-E7E4825C0B95CFEAAEC8B5B25A94222D4B3D265CC4E3034C9366856D7123E8FA1478EAD71794
-2B36B526A98D0E6360E312D5B7E5F215AAA996632F5CA0E10EA00AD34C1E097166475F34E078
-7532400ABE560A190C0DE8E7FA8165C742615F359E780AA017EC56ED82C6E4C018B73860A72A
-B49F1F137ABCF30F4F9354E193452A03B49F52FDAC58D3C7571342BC77D3C32B49340665E36F
-2F929C3892534FE6DA881179F8654826C26FF4EC63DAD1D497265D11DD5EA3AA4A88E0F168D4
-F701FB7BD854B2CF4DB842FF00F002C23F87F86A932018B2C71BBFE000F4FF456FC4DC0BE106
-EDBBE77DC6D3F144BFC52FCA00A6CFF17D954CC0D9DD1270BBB1BAD4882765882244216CA939
-E0DD9FFFADA36CC6BD40D2029EDAC77AE825F470704B30CA49FD4104B2EDB1C5DD1A585DA431
-A0A5A445497DC182C094F21B7C64CB67E4BF8A151204EBBA635DAFFD6EE9BBF0EFFC4C2F6E7A
-E800ABD4CF7B4BDE8F8B45CFC1501E1F7A7F8C942BF820BABCF4860922DAF7F0E7D690E2834D
-2E30AF7D09434B0111278452B88EF44D5EF8D91DE2A7AABD8246354CC7827B4D32A224508A4E
-D6783C2AB892B047A206CD99B23BAF3BD0C633CD0113F6079ED17E0E7D2DB9252FFF80066FC0
-3A2B6629E310C736F2AF128909BAE7A166CAE629494999F16B0D9040F15CD16561A3153AF664
-B13BB923B90CE0F09BBE93CEA4C621F713756F358276FC72EA3B19363C3C9EBAD5D599F4C5CE
-9F1A8D7C75605C0916ACA60F241F9A9FBF3D6665D75D6263BE9C9139BC61C31D4F7C2CE47E13
-7AFC13708294A9E27A86666A37F17386599ACE191C43E41D406D732701B99A00A79B99694D3C
-2D2B6939CB4054C00C0F2F09A849AE3E372115A5713CEB165C9BD209EE3DFEFA82B1BC1CAC03
-8FB692571E0D5D044DBFDBA07F7D199902CB79FC719A31EE2E4AD7F1D50FCAA7C3E84B55A080
-D9DA54E3EA5408DD50FADD87CFD923F699DA9EE4AEE91525E4347A0A26002B97F324EB716D4A
-12D0BD7324E2F881F61406E2947D65A112A32B5AE17F06806D174C08D306D68774129F340466
-53AFD9D38FC13069B3ECE8E68ABCC734AD4213801E1641B5B990615C7BAFADD52AFDEEE145C2
-2574847D0A21D37B61E01C70F488E91F85A07E633179951B5386FD12D581D76390EC478148B0
-D819BF2E090B2A40708257FA3A6E02CD3F37B1D785434A7BF98EDFE85A6DE1CCC52D55CEFA41
-F81BA9F5FEEC35967F2D82FA1B1904F9D50072A4A2EC937148FE1BA22086C11A162BEFF80ED0
-C80AF897A1D2BF1E84E4193EE94324C2B06F317E17DE934578E62ADC63E54774CF1C9D162B33
-6311E8AF5C5478FA1E40102EFD17ED7D83C8F55694E7CC531144156460D7FE4754766E5816CD
-016A70F38EC92A3A54899278985CB868A7E213B9191FFE6D1B368CD50C766F0A75EF883C49CA
-41B0476EC297A676C09617914DB7CEC4AFE8A0F4A956C39C73087064D991828C5D92CF811E6D
-1DCC9165AA29B54000F0698B9DE7744D081ECADC30F7E9630384FD7D8D44B348536777E0EEF1
-3B52EDB21778E4219D26548B673681718A2F8DD08677AE6AE9B0724F2B7199D128447CEECBA5
-B086BDAF71873B5EBCB73B91E107C69E2988F5CF17303677416B79DDAD040E07D354562474E3
-0E9E76D675A00AE0ED24002E4D55764BC91397ADDD1369C308176147C1DD0585ADF03D5FFCF2
-F41A279349E18BAF53C97BA1693E8C41AC05C17739DDC7F0029E1C9448C2497E90E2A298E284
-77B6AD2897E4BE7A4F9F26A64A2A068E0D3DEBE2075024ACEE796FA218EE80C54635592BA855
-0504A100940BDC6B93AB8F327AD57E0975532B4764F641C389E5AC73459EACC95B62AAE4198E
-0198EC9E7E0D3583E000BFF13B0A04B10EA2579E77DDE7D2641BB738A17B169A96EE96FD6BEE
-C960704D1B3619B11AD2EC17133A638809CC4B2A0CECE459518250FAD46E66FD2CF14681250D
-94C0BA32EF9D0354EC9C5E232993418BE49CCB016F56EC3088080D9B49021CE762C3D28DA82D
-C6062623DCE857C701A5E3CC2325B4F4F05AB5CB02395E487B56F961D894722747A00C6B0F29
-FF35DF2C983454AEB8E75C8A3F0747DB78315EF273BC330C27BE545CFBF5BB46E18DA81EAD24
-7CBAF56005088A9B2C7F31B870EC5CC8E6FB234764091E7F57F282CEBDA7306158FF93F49FF0
-9031EF7EE383931360D878397281D7547719DF3384D6BF59A794EDA7867D8AC4E100F629E12C
-CF17D6AED1C6EB469F34AC6DA65748C108E95860EE4930114561C41F8BED9433636FA317E13E
-64854AC20AD41EA757BE07701D25B307A6D31BEF19561874256998A20C9B9E64CF3094E56CE0
-724C6E371279FBD1A4F561B351DACD49AFE5FEF3A0F78597BB60501C14D03C4E8234075C0728
-36042491448815FD92DAF7BCCE928C15DDA6F9AF1DFAC3212C1520397B50B3F8041728D4453B
-557D35D0326C4B1B5A5D5597E3B570B499C7CB4A70D3ED07B25F555298A176892FB3BC0CB7F9
-938CC74D5DCC5BC3DBC069F72BDA550BEE55C2317E2714F6E638B23188E890A68FE50F4C42F8
-8A00FD372F5C1EDADC0D689830B03AA40700921E0028948632365E207CDC18EB85AD4DF5F208
-BB868B26625C69ECD06E3B37AFCEB9568D7D1EB7933C037E6E6EA8D96AC3F9B8959D4254DE25
-90B2D5A5A182C36BC8ECDA974F6B1D23D82024A965B2E24A756EF04802C9116F60152047D477
-3CC7E40D8DCDD32D2990C5493F372FB499B74649261C77062B73337AFC3FCF51192D5317D70D
-6BB0EDF7522120588E797E87366D753781D8ABF9C401C5B31B2E0327BE199894F8B86DD6E6D0
-5C43B5B6AA0311806AD8E78B1C1ED16D2CE04E1EFDF37443509430A71CED5A78DDC0AF38F745
-C70F23ACEEF34D042452B96C6FFCBCB00127CCE2AED99E5663D8C4F1E53DEBB34D82F354C5EB
-AFB6DD0D58F6773239F0B28DE378A3B348C22C70F56C482FD6B05C9FE406B67B40DC3555BF19
-026BEF4F9A8340EF317D7048F827D19D39862F0EE656EEF75E33093B020B729203D35168C198
-7B755CF81A2161FFE86AF5B476B8AB7FED510BBE5A15BC712B7D78364A8ECFC8D5C50FFA2047
-D3B94496A7C69B3F721C67DD8DDDD0C0883994DACDCBFEA6CCB352B029CEE13C2112B3C55C7F
-D1998D193008090EBE3A49CFB52723197390A1CF2EB419402A0D55B8071192C0D3EA740A6DA8
-DE7F1255BE8B78371731F78D0153819B9417662C43CEA058176C9F93B48E21129015AB2BE566
-E40E1FD01B56222C07E2DD9F8D18343BD778EC4D31035E73AF36635B3B3DFCC358A953651E34
-814A209802F6AE860011A48B94DF7FF2B69D91AA4312F60B3A095F512C1926B5D075B6F39594
-5BBEEB60AEC7DC50FF235AAF1CB0C45957EECAB42789EDABE8CF7AD14EF84613DFD43189A664
-72342E85260B0D53038035184379E11822B6F0FE10824630C2479F07447CBA11E49F0A27B4F0
-BFAB9261593FE6A8FC3B5B9835E60F3CB3282DDF0553D00A08031913BA73EA6BE91F5F89F80F
-5432A86C3C982EBF203AC33C9A8051EE677890EDB22FB9F6E30E65B73A159539FADCF57C53C5
-094C722F563756B69F0074BC8A3E53807888DC0EDE2BEACC9BC9D7F35BDA3F49A7C28461FF99
-8F466ABD69DC976CE81469E4055DA48CB1F23EE8B13B2EEC14E8C68EAE3F69CC937808B5BB84
-0847F92B275A90AC86F9101FA6B0DF3698D1664D7C79058486D258C271E867878B8E16AEC6FF
-F29F9F34F3AA1353694BE5AD123C088D20D662D2235E2D6E35443008216097F732882C7EC502
-7B37444373766C5D9F28062E4FAA5BC91B9B13CC54348C6781501DDACBD3A1D2850604F8E3FF
-DCCD639ECD9C94FEA58BF0E94B0914BB908CD3D9ED63B303CAC191AE0559292F6B3BCD2120F5
-50C07BF2F150996E272F81F07A4C92EC61B7E64A4B1044F5D02B558C224DCCF1565628383DBB
-E90EF8795CB27D8E860B5BC8ED8385044C4CEF02BFD42C61C21A2F3FDAC0DC52F7B891CF58A3
-8D6D1C9D7F16A9E74BA046D2A31A1E1164029F78B52E543DBAAFD1A962793F9580E2491778B1
-1C0D322888C0DE763CA2D8B450C1F581B45B0541375CAD8717977734014735502F6A3A08546F
-E4199C12DFCAD3230376EB850E699A185BBB984739102154B5AAEF6DEAE24866A161EC8736F4
-46C9FA9DAE18579178C80A5552CEC076B27F09E23ECA68150EE7A31C6C558E6316B0967BA4DA
-9A41B83AD557BD366272D5EE64852A4A472EA17B4847D81EE037AE14B97C8784E0F18E938FBF
-28C22F2861E2564CEACE610AF634D99A5135FF9257A42F544939BCEE157E19F61072B734F0F1
-C8199F92090B579E33852AB752E9A08C3DE0BA02E89BD63BA02ED24FB2826A6AE3E13373B38A
-3A9830CC7AB674FB1217F7E8C4B5D60461616AFC2E72226FA66AFA2E52382190EB366F9EE999
-E0C82DCC5A76E7EAC0F2DC751244717483D7218E69424E60B99AAAA4B438510D4071F9ADA47B
-DC74D54087858FD98BADE79B49A6538FA2811923AB5652BC7DD425AC6A36A39666F7E6501A2A
-90885F510C797F9E87A7B323ECCE8A6E6A950A32120F9B17273FB84807AAA12ACD4E74D512C8
-B0EF6C5C0CCB52CD52FB4F94129188688B5BC26BEB3F3230558AE7AE1A7285C99BA3020BD1E5
-E4E0EFCB6FFC688CE07FE45CAAAF9A4417B027AE633C80C9D481B451858CE03C2B5E362277D6
-02B02EE792D2BF7790A8D74C811E005B77F31B78488EC060CC78CEDA644F7DFBF69FF3F5A314
-C899212F67A8BD5FEDDF543489A7EDE6E1149E214D0E3A1BB1F809C823A569096DFC0EEC9DBB
-693CBD2F48F732ACCFCA8AF5C0C65A1D6C33F758E982B6F40653A287980FB6B2B59ED7576EE2
-ECABF8B73675B35E1047A9C9EF565D6297AC4B5C79981EB59FAA9902E49B03CC590E86D1F188
-A7FF89FFDECA92D1E153AD3CCC014167ACED079E68AF4153E19AED5EB8289AB5D9AF8DE7285C
-0FABC48AA30E225B184BAA45E077C587F342ED1C3445C9730FC73DC29C70B4D95FF7024120BF
-BE4D2DEFF83512745FB7787EBFD86C251CD384047EE755FA5EFA46DA5D90EC423C9CBE18FB1C
-FE0F19E6F17883267A2A4F0AE7B20CE0F487471E3E0A1DA05219E3CAC0A6D50ACDD33E92F020
-FE2B0FD693A624EB5475CBB448BF0A430D61830B2702C73988B32BBBE55C747FE951F18B65EE
-AE41D17070C0F64AA5F7761544ABDF3892CDEE824571D87871030A2C7E75248193BF1A933AB3
-C389727EDB22529C1FA19422259B699E344A5FC21DCA13AF2DF4882FE69BDC52B5C1497B6397
-4FB094B55745231B4655536FFDA23A85989700F89C4D8EFF5D8450CAB45342DA446325A97BDB
-F4F61B20F5C5D98973685256E8C7E3D6375B586F051FB4DB2F266FB81A1E58A1E57AE4DC2538
-2F4FDAAD9E90DFE1C177D662321C8A741F9DFF26A965DBE59C5697F3CFBDFF0B3BC6C6AD433C
-EADFDB48BDE662B56369280DE596A12BF14340B80DDD7E51F78ADAFDB607D7EDA8114CA43CBD
-329D6958CC6400FFB32982CE9757F2114888B7A1E7748848C7CEC66B47D84A4423E0DEF5A7B9
-41EE67B61B0D58B81AEE40E76A6BACB2BABCA8CB6EA41713E3E73B1EFE0E1F420A4D6394CFEC
-4D486B66D6AD7E008C32653EF7759E03A3102723E80BA273F14A2A184F6929F61BB9B799E6A8
-5BA5C308F1293E47E7FEA53D7B64E377A95AF58431377966F7AB43EE0770591917359F0056D6
-1220A973703644B6CE37E36AF597F26165D6FA044264636B819D2A4B303C5C027162EF450C67
-571731E14D2613C64521F3C3ECCF41C160816DF99F79BE24A6CB5F3ACA9F6E153BDF47994AA5
-91EF9B6716853014255F79B348102D559643E71A38BE93CCFAAD1C27A7C85F0EAFC650EB6A97
-E6D020C20842D760F2C8C1973E4689B29BBD2AEF53BC192AD97DCFD72B423427C61069A02533
-C75994F95DF2CBD5C36D534619BC5194AD45E029DA3AC44B6876823170AD02E80322530188DE
-D08F059200DE33B37079FFB9C4EB8916518699FFB95473FD794D81B04CDDEC62927753DA0937
-C4F4B040BA9713623174856243665936C076A6837B034B9377E7CB76906DDB6D5F863BB4518A
-EE11F9C8DFF3E01BA0031C5A93F6F9C4D600FB7AAFD2552E68B53AB3A80719EB3D3353512B27
-4B0C919CC45FDEB8818B0B0F3B68FED2F60842AEBE5E45C917A52A6124C52B8F2115763D2D59
-32C77E0587B65CB6EB11D906915F2169F442F8B3876FBB061AE215999A02826F50DA330EEEB4
-B8954F7B01D94696FD24A949707ECF7DD0743E264E3AB60075058AB7BD2DB0FA7D16366B76A7
-C6312F9881F8B78455B839C3583CBACC4FFB5C6F3A71411E94EFFB61CFA185A739DD89A97200
-6CBAB68E2CBBC1ACC9E615C478542FD0CBF9A0E67389E1FF48AADE08E527A5015336B7694F28
-DA4921066327F71E0F4872AF8E991271C5E943912368FFE9C8CE3C87C9ECD584B24057B40493
-2A40D5FF1A4F58304E2534B4FA3DD0965DD2ECDE9A1ED6B9DF3BA5029FD82DBDC45E00552A07
-8CC6B4C8BB8CA371FA347F9157506723A891CBC45709F2DAFA80546B87DA9BEF73C7D237CA28
-9EA2D7944109466AF7040D55A4C9BC7D9DD87BB9E24FEFED827B6B9A9C5CDA9719A86F112953
-7A878A141301019E7FF1364607AFA76FCDBE192C51406BFAAABC6945BDBFB791649A9B66C089
-C94A97DFAB07CB65EC96545D00909DE84369A45BFC5DEFD2CF31C4EF43103DE425BC32FF3D24
-FDC0B87FCC91E594A4B3105E7D35CA227C779159C38347A19CC792B45F7728596F1E7606E912
-23B7005F5CFD6D5085F2CA33C3F60A08381A05DD70F851122C7051CEEB0B0A347BCB3C267667
-DCD66312E4837327B5A10E8218987B308CE4E6CC2120CB0288079434AA9C7023EA1A2F83A238
-96541951712FAACCFDBF030AC4A11008E0DF0D4D31442F5AFA5CA4E7CF0567EA2CBCC1B5CBAD
-B3A96809C61222CAB6A5F40AB0B0365277604C78993E72DF89D24EBCAC37CCA52327A094512D
-FE805526BAB27CC28C54919B00CF80648D3128FDFEC4D75075397A2AE9110B540CECFFBB7D0E
-B7D16918B665A09BB237961EAA35BCB74766355AD322FACD53FD15D837EC22057994E36BC2DD
-B9A7FD782D1887B9E4FD8CE2905E134409408CC4D6F7F52ED71E418E9DA484DC742E93798285
-9FCEE98196F6C9A74D61ED49EEDC35775D759744AC9C7B0AE822B3ABCC2491C5FA08F2118472
-FAD31D279C1DB66420D6EF048332CBF8DFF8D23365A21D937CF2D6359ACFBCCD4395B91393E8
-F46A6EBE39FB518104FB9C9C06624ED523C52B3EE930C68E48C75AA7564E9EFE4FACD1823D0A
-456B9D05E4CE1676FCED6A06CABECE7B28808FE9C39170D7240A0DB0888A27B3AA0BAF65F5CF
-EA8033728969073F76EF6AF80E693B563064A0257F9B06E1B4FADD5531452B33C7718B6B0474
-39A8B4B6C23683254A70BFE71465E9D8BDF8BE4FBEB81C815FABE967704C6CA69DCE1253DD02
-6CB66773E30F0437C325BFDB3E492489D5F05EBC12A96B16F191F3777F42A884577809D9294F
-273E81DA495428075F82DADFF5C3D769103B4F2FFA0D4EFEEF35FDF060438B0DA75CB2783439
-D9E213420389F6B0D87BB24409E75FF522421E019B727CF7B458A3FC59FE06BEB34BABFEFAAF
-9E5EC723945F9921D55978EF0E80C12F6BF7A458BFBE57069B049A34F3988A276F3D9671791B
-19D36456EF3F8702F12AD65BEAB9A68702132B719466ADDC8A7717E458E83488873320704A8A
-6A49D148675060915F93C7FE7D94810D9AB4F0BFD103EAEB3F5C23D371FB9F45F12295699E04
-30C7AF1FB76B9EF2F567AF729A640CE18FF94BB33A356F8685486770FE6286FFE718C9CEEFD2
-8917759667816EF1AE5107B18E16702E4A74B0D2F3F970E6419B0EA3AEBDF3D6093C33AEA645
-13ABB2F30122FA7E54B6F1AB31D21411A318D553E937E38C601FEAD0FF4AB6DBE46564AF51B4
-C6062BACDE2E6F8A832C4A927B5DF081B378D65DAD239FADC2F6CB28DFDCFC1FF4AD9A33BF66
-4C8EFC5786AE4D377F63ADBE5348FEF187A22B594CB75D8997F75364FD5AD8000214069CB258
-64B1CB3FF5920C84C24E6110E83FCD84ACFBE721C4339D12DAB0E1BDD55200AD517922CE8ABC
-4CBB102BE5EB135FF9F300EDFC3E9973D39AD42F99B63130982F3D36807A86FC53F787126185
-87524EE15BA2F62C76C333CB9676CE92BFC7E316BB2D797A3242E7AF3F9CBB5A889E1783CFD3
-EE74DA113D732969F684192C2D3E6937C37E45DB3B145FD337F59E6F73488806C9512BCF09E7
-8CB7EFC1012B0CF45BF7D4170FA7D64F3326E2293B9C99353C489D935AA442523BD1E19461A6
-1BE62B7DEA55A3CA98A77994DF5BE543EEFA6C6C4F09DCA7E0909715ABC776EBFB2D11BF1BBD
-611D73DD7A482586141B34689B78614C0CA3CDD84419B0A8301C3F443E32C6A4FE48CA82E194
-10133FDDCC075F2CB35D4EE79CA0445A5F4039B64A193D7729C04855A6C1EA49E3433534BB1C
-A573565A7FF0C2667B84C0E341AEA64DB7F09E250FE056C94ACB17CC259038124C22E3C7B790
-07EBFBFD4ECAABEB16847B3ECC389AE626D8173F9DFAA8A30BB90C51B10013BE0494D4EE5E01
-503D175C71DF17FDA410E4147E057D264D80BA710470BBDEB2586BD0D067C50153D030C1F9B4
-49BC7B043FB72E6DA5F4706A6BFBD1ADB647935DB7CE7B32E94744E722F0C2F0E2793AB987D3
-AEB7E1447FE554A08586FFB95AD172FE17BCC391CF8A685E2DD7886DACE6BD000FD9EB6773BE
-D9CC8C6A1DF5AD71496AE894DE7AD2B9282F765E6B62B84CBBEF09FD50DFDDA4D1CD214B4183
-2867C1A9B8A6FBAAE2E195FDC6FAB4846D718E703F05C1C6CACBAB43B7A0C2A869C538955AEA
-B23EFCDDE65AE1F68FFD1F949B1255039CD7BD5108CA7D0DD46766E298847BD895DA8A8EB5C9
-1765D00CCAD4E995E823C1F6B8709E1C02788EB86A1DCC716F9F266E106F1DB40F58223E40EF
-D3BA9D5735BDB896A62C0532980E6C3A6403853D6BDE51085F3272115FC932ACA8AD50714F71
-E07001BDA30A4745AC6F660DF69C8E60AB716120548A51A294DC12F15B1823EC1A9F878219BB
-89B8D73A2F5BE60FA65821B9548953FE8FBAD23DC0F3DEB62DA78AEBFF91B9FE078CE6967F48
-224100C740AEC1B61AC9ACCB469DB757BA5A8622828F6DA0798D617580A4E82101D72C1690AA
-A54B16C9DE189619E483FB1F0CB5AA8EECB80EF87227A5D1D447C44FBB5E544EA6ABA872D742
-F8B6561EDFDC7FF321A9811D867583275FB7D64E7EF9B8635CA874E68A23C7246E00EBEDBE1A
-301E5CBBEC6BC9EF7F4835E882F8449EB54B3215F8E637A0AD4984FD94AA125263584A94FE64
-0EBC08FB9285D3A09C125F23165FC861C9FA28C3C317B27A5E91044C7E92A17DE34567D3F2B3
-39899EFC9F6481F1C102747326180C4E44F46F5F216EFA7BC42AAF603914DB49BE664E049F8C
-D8825743A695F97D537D407717A00FA7B428BFD9F1578E6358D5CE70664699A142AD54D61DB6
-8E38314642881E3574A3955EA75E24272F66233173C2837C3CEF43F8D9005E229B2998C41FD7
-DF8AD7B37F506749F566CA1B8B2C6641ADC478FD064A01A40C1F2AD167AD245DF7638FCCBCBA
-28F7C838D4092463798F1484995829DAD3E349AE4FC9E1CCF7AA124085E38870A6FCE3B934DA
-A1D76A089B7C76894231A9696D97F81F4281C13A7A2FCEB4CB01BC97A5465E5B8DA52EB81281
-AA6642771AA3A71DDE64F6822EEDA1C0049BED704666A0ACD2124DEA40AABF7F094E2E0B5D89
-92FDBF1E997A1284C7DFCA7D00CAFB0315D9CB1D08065B0A1F854439EC8B823DB4771F5967F5
-92C96B7F5D386ABDF7EBF17B80A404AA0A52B88393B9E4A786B2786AB21B5BC1D0BD1E64F79A
-302FDD33B553CEE84B1201A7A7916468456785CF4F8C55267D1E21440C517C87EBE61F21A203
-C8727FA15FED6D48A703C291980F3A789816168336C18787651420188064B75FFEE401A85773
-CDA4C4F957B2957F7B76EED3B2159F5A3EEAD725889C8265F58C43005EDD5284B418E04FCCC5
-A6C1329816E95F278FF6DA40C5D73568C5808721734508157CC8E23F2C25171F2E2EC0BC0D9C
-363255A4EC0EB4FDB11642A96D598D0E1A82BF2DBBB97EF226A8D00DB5097A89C10B3955BCFE
-6FDA37BA4FC9E9E9097C774146ED0CEB61A8018E464A809738A2D85D518936FDE5977943350D
-C642F458EECFC8B744830D04A2A65AA12FDE497DC181DC2564A7F137986C594D540129D4D176
-9D3E7E241EE8CAF4A901E7FF24F70855D020539125F45EC99CA447723D1DEE38C146D7625E3E
-81167A4D8054976FCA5A7396FC74DA7DB6CDC5EDFECCDD3B07117A96BBC883AB74A79E4EF382
-082B469BE06A44B16B35E551A92E4E714F2DDE08E109D93D1B333B487B7B323564922D9D8F18
-7F440BFEC0D8F2D70A1B0BE46AB21B5B8B43E0C5E9CD0093EF6BE6476F47F9E412799275D454
-DE125751DDDDE43E3A5A698682D0AE0C77DA60C64A3914A28312A5D3EB07700A8ECBA1D2D24A
-110F0E9A68FA203458A51BC1BFA6BBB1298362270979B1661FC20237BEEB75B5E1AAF8B66CFF
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-cleartomark
-
-
-
-%%EndFont
-%%BeginFont: CMSY6
-%!PS-AdobeFont-1.1: CMSY6 1.0
-%%CreationDate: 1991 Aug 15 07:21:34
-
-% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
-
-11 dict begin
-/FontInfo 7 dict dup begin
-/version (1.0) readonly def
-/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
-/FullName (CMSY6) readonly def
-/FamilyName (Computer Modern) readonly def
-/Weight (Medium) readonly def
-/ItalicAngle -14.035 def
-/isFixedPitch false def
-end readonly def
-/FontName /CMSY6 def
-/PaintType 0 def
-/FontType 1 def
-/FontMatrix [0.001 0 0 0.001 0 0] readonly def
-/Encoding 256 array
-0 1 255 {1 index exch /.notdef put} for
-dup 161 /minus put
-dup 162 /periodcentered put
-dup 163 /multiply put
-dup 164 /asteriskmath put
-dup 165 /divide put
-dup 166 /diamondmath put
-dup 167 /plusminus put
-dup 168 /minusplus put
-dup 169 /circleplus put
-dup 170 /circleminus put
-dup 173 /circlemultiply put
-dup 174 /circledivide put
-dup 175 /circledot put
-dup 176 /circlecopyrt put
-dup 177 /openbullet put
-dup 178 /bullet put
-dup 179 /equivasymptotic put
-dup 180 /equivalence put
-dup 181 /reflexsubset put
-dup 182 /reflexsuperset put
-dup 183 /lessequal put
-dup 184 /greaterequal put
-dup 185 /precedesequal put
-dup 186 /followsequal put
-dup 187 /similar put
-dup 188 /approxequal put
-dup 189 /propersubset put
-dup 190 /propersuperset put
-dup 191 /lessmuch put
-dup 192 /greatermuch put
-dup 193 /precedes put
-dup 194 /follows put
-dup 195 /arrowleft put
-dup 196 /spade put
-dup 0 /minus put
-dup 1 /periodcentered put
-dup 2 /multiply put
-dup 3 /asteriskmath put
-dup 4 /divide put
-dup 5 /diamondmath put
-dup 6 /plusminus put
-dup 7 /minusplus put
-dup 8 /circleplus put
-dup 9 /circleminus put
-dup 10 /circlemultiply put
-dup 11 /circledivide put
-dup 12 /circledot put
-dup 13 /circlecopyrt put
-dup 14 /openbullet put
-dup 15 /bullet put
-dup 16 /equivasymptotic put
-dup 17 /equivalence put
-dup 18 /reflexsubset put
-dup 19 /reflexsuperset put
-dup 20 /lessequal put
-dup 21 /greaterequal put
-dup 22 /precedesequal put
-dup 23 /followsequal put
-dup 24 /similar put
-dup 25 /approxequal put
-dup 26 /propersubset put
-dup 27 /propersuperset put
-dup 28 /lessmuch put
-dup 29 /greatermuch put
-dup 30 /precedes put
-dup 31 /follows put
-dup 32 /arrowleft put
-dup 33 /arrowright put
-dup 34 /arrowup put
-dup 35 /arrowdown put
-dup 36 /arrowboth put
-dup 37 /arrownortheast put
-dup 38 /arrowsoutheast put
-dup 39 /similarequal put
-dup 40 /arrowdblleft put
-dup 41 /arrowdblright put
-dup 42 /arrowdblup put
-dup 43 /arrowdbldown put
-dup 44 /arrowdblboth put
-dup 45 /arrownorthwest put
-dup 46 /arrowsouthwest put
-dup 47 /proportional put
-dup 48 /prime put
-dup 49 /infinity put
-dup 50 /element put
-dup 51 /owner put
-dup 52 /triangle put
-dup 53 /triangleinv put
-dup 54 /negationslash put
-dup 55 /mapsto put
-dup 56 /universal put
-dup 57 /existential put
-dup 58 /logicalnot put
-dup 59 /emptyset put
-dup 60 /Rfractur put
-dup 61 /Ifractur put
-dup 62 /latticetop put
-dup 63 /perpendicular put
-dup 64 /aleph put
-dup 65 /A put
-dup 66 /B put
-dup 67 /C put
-dup 68 /D put
-dup 69 /E put
-dup 70 /F put
-dup 71 /G put
-dup 72 /H put
-dup 73 /I put
-dup 74 /J put
-dup 75 /K put
-dup 76 /L put
-dup 77 /M put
-dup 78 /N put
-dup 79 /O put
-dup 80 /P put
-dup 81 /Q put
-dup 82 /R put
-dup 83 /S put
-dup 84 /T put
-dup 85 /U put
-dup 86 /V put
-dup 87 /W put
-dup 88 /X put
-dup 89 /Y put
-dup 90 /Z put
-dup 91 /union put
-dup 92 /intersection put
-dup 93 /unionmulti put
-dup 94 /logicaland put
-dup 95 /logicalor put
-dup 96 /turnstileleft put
-dup 97 /turnstileright put
-dup 98 /floorleft put
-dup 99 /floorright put
-dup 100 /ceilingleft put
-dup 101 /ceilingright put
-dup 102 /braceleft put
-dup 103 /braceright put
-dup 104 /angbracketleft put
-dup 105 /angbracketright put
-dup 106 /bar put
-dup 107 /bardbl put
-dup 108 /arrowbothv put
-dup 109 /arrowdblbothv put
-dup 110 /backslash put
-dup 111 /wreathproduct put
-dup 112 /radical put
-dup 113 /coproduct put
-dup 114 /nabla put
-dup 115 /integral put
-dup 116 /unionsq put
-dup 117 /intersectionsq put
-dup 118 /subsetsqequal put
-dup 119 /supersetsqequal put
-dup 120 /section put
-dup 121 /dagger put
-dup 122 /daggerdbl put
-dup 123 /paragraph put
-dup 124 /club put
-dup 125 /diamond put
-dup 126 /heart put
-dup 127 /spade put
-dup 128 /arrowleft put
-readonly def
-/FontBBox{-4 -948 1329 786}readonly def
-/UniqueID 5000816 def
-currentdict end
-currentfile eexec
-
-9B9C1569015F2C1D2BF560F4C0D52257BAC8CED9B09A275AB231194ECF82935205826F4E975D
-CECEC72B2CF3A18899CCDE1FD935D09D813B096CC6B83CDF4F23B9A60DB41F9976AC333263C9
-08DCEFCDBD4C8402ED00A36E7487634D089FD45AF4A38A56A4412C3B0BAFFAEB717BF0DE9FFB
-7A8460BF475A6718B0C73C571145D026957276530530A2FBEFC6C8F059084178F5AB59E11B6A
-18979F258B8C6ED3CCAFBC21ACA420C9C83EEA371ADC20E038B4D7B8AC303004B0AA205F0413
-514076407216032FDD22E6219DA8F16B28CA12524DEB7BCA073CC5EBA65C102A5E85FD48E6D0
-62CD4283EE570A7774597E5BF0E3400B6BE72DB0115F3CB12DB70CE083722870CDDFADEE715F
-10F1FCAF20E06F3C54AFE5CA238539BFE2B596116E83F5371FF18FA5003D8543226CFD4025F9
-940365B392A858D27F078D3ABCFFE4A154E78C7692D1A32BF935967C64F01B24788FF8325D61
-145E2D4A489FD986FB7738E6B254522C77CA2797A504A9CE4676A77EBACB026ECA94DDE5922C
-936F8E90C43E285195B0C3CC5D6B8224F29F32A2EC82919726A52432EC85A04CE68682489FD2
-72CE28D86A678DDED99ED1A787735BBB4CAB90BEC8F2E43190C47C79EC6B97BD2A468E71B891
-3DA951E7829BC02AF100405A37F8582C644D502803505F51D1419043A15A3F235C71604E1751
-5B8916079E689C4AEA5D88956B20C43B4DE1289DBFFF32AC8DB0405DE8088E2BF93F33710D57
-7DB7A55DD4668E14980A3A395A4E1773A227740B7E6A214D02F0AF34AD82095ABE5381A94EA9
-70F55CD4F686DFCBDAFD39CF51CB96AA0C22B1D3FC592AE4668D5726886217CADCBCD3AC0FF8
-9E3E9C7B128BA0B6149FAEC3B48D714BF6A69CCE27E99421F9A1972FA495B0E8166A49D26D6C
-70F16F2ED8B839F5ABA59EE54BF30038FAFE10DA1BC15386AD068D3ED622441B7F76CCC40C46
-0195EF4E10EF9B971659AE29500472501A8013E42B060CE3456CE2B00A9FBE46E32F670146BA
-142630419456A14465C33DF3E50EDCC80F538844DDA73F57577AE369775B53D5BADF04D2A255
-755C977EE4A8DC15C11AB40BAA0A2C125D5225F5805E99364157F7CB24272C01E6122C28436A
-949DB6D0643E014B0DD4FDACD970F028BFFFAF437FB1342E1980B2F1EE9DF68DAF459F11109A
-DA0AF9B7DB45E26E7315F9085B851C37ED328C83AF0AB7B4054F7142705F3FBFF71C851DF223
-C3231B0F49B387571C2D6482793654D0A974E5C565CEE8A1CF1EA15D2AB97A71EA90AAC8CF47
-EF1D38248F274C4D95831A42ED55F3B973C7F5ACDCA7EDC66099993B88A318D7A6F9F0B04454
-7DF7F91BF6473C60F6B3FBE47F52A540CD27A35FDD49E37DF21775451B51B4777F0B96931DBE
-05EB5F771DCF83B4B7CA56F4D823B71C7AECE9436641CB1E36113050923CD998D630A975E577
-C08F525A7EC4E5092F525CDBB418E481032DC9144725713BBE2AAA7017252072D9261B3027E3
-4B8A8EA3027D094513A0DF3245392CFBD2359EC069A8BEAFB6A928F3079B96F5FE8ACF2D6C74
-59DDF50A4C1936E5B3B7154A86F1F741C17646A0726466946BAEF40E71A7ED4B041B3A1A36BC
-F787BA8438F103DEC7BA98C9AD55987B0DA739DFBFCCF7C9E4BE37D265213E2A58A1EBACA3F4
-35B30B8BF16EC6156D94200F8F7831C39F96C6F516C1DAD2C8FAFFC3E648079FBBE66F4792C4
-BF9391962957D5DC55CFA517C33B530678BB83A814B0B75BCDE404E0B7482EE1891D08FC4E92
-07B2B5B45EE2653E97EC1D698C317517221283143699D5F84C795AC9C86583E996FB74B68EE4
-B241148949067FC6E7C47F2417B05E198A1C526956B7D4E45B9E65CBB2B735C89B7E331AB4AE
-ED4199C911A0C9822E4BF32A74C839E229D1C02003D6411A3D94BF2F337B675C678A058E2375
-1434894BC36463449E1A147449C457773F4F71CEE20DC792EA329CE06E4DD5B02C5D681E6FFC
-A56EEA53A82609798011C95B725831B5CD1E7B12805C9073CB4A539CDB82085AFFDE30D03E3A
-499431CF93FD61157028B1772C39877D7894B8B4C92F5585346060D24FBCEE6FFB3912FE43E5
-4810FB9E6D22A7CF3DCE1C1B8D11EFDFD75905F560A17D2763EB801C58FA36F486AD7A4EA06E
-7084CBFB9F7105B78DE7563AED097FD08F1D9C033FF2E77440034F9BDA34AC058643D2298535
-79BA289AED15C28FCDA7F17CE8A46563EAB701BC4D6E9416DC0F9F6430689212F7ECEF9BDF9B
-DC7AA6DFF997DA6689DD39D05B36792510583C7752E1BEA52FBDA27B5DDB900A1E16091B783D
-A9EE98C48CC11602CC199E0DDC8E661FD2D9A0A3766D68506F1CC4461CFD280D1967F0D308EA
-D4038DDC3E1A94AAEE33A9B4B02BB4A8BA6B513F2AF64208D7E2BC6D0822823413BED15C531F
-F6DCFFD9855038B311800675E149DD36A97366C609800282939EBBFE8C395FC05C11E0E2167B
-774D6344BDF1708930207FF2A30BBB6771F787E8AFF1CA2A922BA6776C2421FB917B3207E3FD
-82FA4929D67E73506A66D705B594CAA3EAF7C2DC5BFEC50277DB707F243D8D3E58960305ADFF
-175C6EC48882B9B03F8A2EAA01080BC4CE5D4D6221F8FC4F002B6395EB59AF2DD25A9E557F99
-5629C1EAD0029A60FA070B65DA55446D35B7950F800FC11CC9A2CE1E9AAE1531EA2EF4C29B7E
-8633C936C37767EE4D4024644E45F103F55076CB3DDDDF10F17824256F92330A98009BF62735
-0F7C75E40AD01E8664F7A583D18724C7EE86694A80A1B26337C282216B7C03FEA5C17AFED388
-E142B787C76BE081922B91AB0387C57E74F0451F99508E838BEBC413B042F1CFCAECBF6AF1D0
-E741570FF2BEB84FF62FE3C8B9647A9E627187B163FFDDBE970A4F10E2FA6F9F5AF317C27E04
-20523DF172BD4FDA464665B9EE4FF1C5D293F6D73DE5266CFA2A029290D002FFE60DD694F84E
-516902A18E87EDD8668813866C5028B93A5EC69A6B6A2A0D73B1816014F1C94E4EBEE5973365
-CBCB73F99B875A1D6F3C36C1EAC742393E4537D986FDD563ECEB2F8BCBF6D866E7A299303406
-52A43233023CE04DBE3C0F5DFFA440E5FC03C0AD020499CF1D8190F41D0A59E69C03437506AD
-9397753BC28F2BF46ECE62CFC72A836328184FC2E06A5E96D0C7D70FEFABB1AF04F7D0E208FD
-DA268C0A51B3F2700A2D90FA5D0033ADB7F6A74F14DE8D8F8D9B9F326BE53ACC5E2494D150C6
-3E1D8B0BDFD614D4763E24E989B64C9C3651FF7A94281FDE78450502059E287F63ABBC81A18F
-A63273AC96FB329F0E016049EF478A8E7D60A84732BE0FEB183EF52FC609142FB265C45CE863
-F37A056BE4224465E2AEFBD78799557BD73BDCA37F6E9B84120F72F076CEFF1023BE7C76A9F3
-67456665B736CEF66DAB8F8F147C3F95B7290F17C57BD98F4AD4CB4DB3088A6D2729DE734C71
-A58E2527B17B6E4CBF43C908EAC6DA9CFCDDFF0C0D3CC5609EB9D0E0A7CCCE29AAD4B1FD59FD
-89C2BB6F89E628A3FDC418968D883CB9E1C3EB6FCA345FF427A5E02E61BED3E7BE27D8DEB558
-33B0E4E7930CA3A9B79F17B9E0164DED9238B06B151E0DF430205B75F80F3DEE9216F0B93AA8
-0A88C20D01EFCA95B51F6FADE81B5A7FAC6A1C23D232FBA545309837646B7D9AE87041A9C9A4
-AB40F6E3109E0BFE5FFFD115576CC19052DDB58CA820B0C6E7ECB6C4C822EE45B57E44A14949
-4D72AE14582E64E508DD5B2154916A41C7B6ABFCB93FEA3624002DB4E8D222DE1D766BB65FD2
-416DF892536AF99D84DAB00CF82A4C8DA77E1B22B809B4AE351855BE3B3A1507AD597DA8DBA5
-2EEEF739236E63A2C9B0C5C8BAE39E5A6EE94DDD94FB92077CFF4289A9948685AB32B07D8DE9
-7D13110EDC842AFE2489119EB9C3D797218C184D445469AF0B92F965E2E01E52DC75DFB507B4
-FB2DAEBEC9607D8BA559FE5E27553C4518A348683FE4DC3B6CC3C52AAAE3B78FCA86736E3250
-CDE54194887E2C51846BE81B6F487B74F81590DB1426ABFF5AA9C89625F2927A657BFFCB3397
-A294A0F33A9BCB11E608DC791818FDFF40268559B3501C2B83666693A4779B2BF5372FC11587
-3F29C932DCC3385C06C92AE4FAB295EE05753D58C0B97634D25834CC2EB8936AF366EB32F3F9
-0DB5D36D5BCCE640149EE4652F920739105B8FFDE7D4B4CF4CBB58D5BB2EDD8CCE9ADA8EFF52
-1F63D50A0A0C5377A88129D0661611C6A0695073533D10DCC3F45A36499D1202E790F225F011
-8083725429098DF3B25C52BB562E814D799358980BBA585B87316F08B66E047C97491507B30E
-84E7565448C8BD73976C56D4507511C04915DB81F2FC10ED6FFFC9F32A372C077FCA1C74FF59
-ED97557C6D7CDED646D876C3EDA27FAC6A66CE22162DD555D15079884CD4F4D5105F31739312
-335792C12CFA1B735D3076BAC6B88DAAF16568D35548442D4408F6E439909974F30DF11FECF6
-E2F92CDFB1B1BF6375C14C12605CA4E7437DF887AF4AE98175C7F7E9EC65ED1ED8E1BDBC700A
-977A37599B4A931FC6D4D1BE586E6DEEF35C4480EF889DC09F69796A1A57509457D207FBD546
-B6462C3A5C86D1AADC685882DB1851032EB3C533279CA02687B57E49E19BD788D5C73723700B
-A187543C963CC23E0D9B44D61D46B4120D4D0FC31F506D8BFA54E3FDBE4B95AA3372904526A3
-AEB7D432CFD58311084D236180B828135D438E13A6D7EE558FC78BA652320C1792B653E30667
-E1FD6266FD234B6E36E91218C7BFD0271D024DF7556EF65776F6C61A2989D7A5B8161E2E87FE
-08386A60D2DA2C25256BFC36686370D7F1ECC3F51F54AED262DF73FCB10431ED7933D20A9DD3
-3A6E5244A1378F3BCBAB12E251276EE2466A5A26ACE704B4CC777ACF07E3E0387F52CD68B7F2
-29DE3BAC0FD00AF1A208B2ECDC3C43460A3CC2189DA01103B4A4D92BE094897AC82B19FC9BAE
-9BD0C4FBE360833C94C4AA0F3D26AA726653E9C44A59CEBE18E44061675CA673E03B104FEEDA
-09B657665520147BD472F374928EFF3610DDF57DD53BB76F627EDC5E8E4996629D9CEEE51B6D
-E438794E8EC39C6303C3F636E9145974946ACC9971831DEC1A4D4669B4A46E944617656DE100
-37AEF0033E806F6EA6DDCB79B1E0B198919786650AAD11B9E8C8BE43C290EACF4FF036C88369
-0A03AC0259E88D2BF103C9F20F9B898CD2917E75ACFDCA4E4FF718A6E4B697CFEF9A824F3D71
-0D49B03D57E7769AFFE6317CF9E203731100A91FB420D7CD7963DBF2DD1240067A1C30E51B65
-CE67E56144BF3D418FE37B9A03A2D8C6EE81F08233E4C16DFE7C2CA7A7E9051CA99D42B3A638
-58C65D08A91997DCCC17412118BBDBABC469D0E5135ECDB78219A42D0EEF9124944EACC30CFA
-242956F915489B57F1053B8CB75A0BF4ECD82F7E56E3E942D3811A0B6F9D3226D89392929FAC
-9A4BE59871BDC65BE5CDDE1C3E5E628436359C45F89A743FFD0E18DBD8BF27F5B5A5AFFFD60C
-0CFC8925BA22A1E80B2CBFD18BCCE3EA136E62100FFA34CEA526223DF252050C4F97DDC50194
-F070AF9131C7BD86B08E10BA972A1606E900B445364379AF1EE573E275089D65F2D995EC5E05
-52053473D361EA34FB85136C2B72F1A66A6325D40EE70451DDA2A3273DABDF5CBCAE1A12F74B
-A360C3C6B42F196C7E881690E17FA2DB75FAF2399826CD6A7DC2FBB7DA7BB2039C7BFAF1CAC3
-411EC997AD1C60357323CFF4CD71FFC50424D826BB1D683E1CB4EC96E931732D279A9206FB2A
-DBBF62ECF83DC40A30C69146F871AFF91BE60A7B7F2B8D798EC913ABC8773EC5A6C78BFB6202
-BB0C248474E10433202129D1882F58B9EEC0D7EAC67E70EFD58E1EA4E7CAAA43F479CC3E9CA4
-2A6D223B478AA65835FC3C9F98D3A402F8B2EDC0F7EFBBF51AB2D685A00758AFAE51B2F50132
-796CA5BA421F57AECDD348AF624FB3CFFDCC8C04C55F7CC9DE7D944584B9D2B475F5D4F8092B
-D0A0BA9F4C8B62043673483F22349B486D870C8A909F078E40AA177A696B0F8745E5896CA47B
-68D9DFB58B174A627C1E9ADEB310298E8316BA8461D7E037BB18B297762032F508417F181DDB
-C14A0CE19FF4FD1A92B623009B96C41F65A821FE3BA364973BE4CF4E7403E381FC18E4F62412
-42A62A842A4A378E75A9C9976AA02E41B296B812C00ECF13AE09D51B8DFE91906F349E09894D
-EA27A3BD9EAB18CC06BEE65233217AA830CCB8AE3004C3E4807E995AB1C6DF8AB02B4A1630AC
-75F7EE2774C8431958CDCD366CB1BF2A2B5DF33C36DFDA2DCA58B55D07E5A01046BE169B4B63
-0EFF450248F68E87598915F09F765CA426AB1EE0A1B7D9BBDBB6A3011DBABFC37AB2DCF1C02F
-95ABAA31B3D0582E15DFEBED3FF81A1CF2A10B43B9092BF6B52688D6025C80041CE6DB0FF6C8
-5F76C16BD08611090C8A7C77A383A6DD8475DEDEE73CEFD753F3B3E62D5601D5FA8D426CE39E
-C8493625E31A50433DE2FD555CCA3476FAC12EBE2A8A97496630AFCDD84CB40EAA2026A1AA6B
-EF4C015709D387DDDBE12FD61DF0C978F31CF77539BB1B2DC9DE2C1DFBF726E88D43CBFFE966
-650B1D1F2251310909AE586A99F99A20DC95E5D7DBAAB2A5A6D9013F8777601853529EAA361F
-F1B5AD1E38232318F7C119CC43AD877F191B3EF88F478BAFD79A36762F65D7625357B685837A
-C049785E5B6BFB5B0FD8B258E2C08B951131E0E9B88B6F0E392D47F0E2114360BCCA61BF8324
-5560C3AE1372AA5F266884EB5FB562E6CF5B6FA6EDDBDB4DC41056A97AED39317D88EC036701
-9822E1B6CA3E05C79EECDE9FB60DF4CFD23EE9B52722088E1DBAD078739A24796850ADED9861
-01B1241385C15D6D99633349FFFD7E3D10A36A7C02E95E6790B338946FC637C4D8C3B1F52201
-AB810A7777B038EA80C07B6680B6C3B2D3F9A1422F953153D40B7AEA0E6FBEB55D822192E4C7
-9157DCA0442232CAA2E91F891456A5EF9BFD19196456C9ECA06643BA0AA67FD6244DDE35482D
-13663B58743985D47435383161F53D6B5A2BA18D65E488C7327CC402BFFFCA0E6376B8373584
-42634C4E21605A66EE3EFE92955CDB1674587B1A61C67BE6AD98EDB9E051C6F8BFD992DFEF3B
-E772C7F53B4244EE43C421C29C288FB509C07FBEEE8BBB0E9A942543FD87E27DA72B87DF13A6
-947457EBC4E80BE9F831BAAFCE94E46BC48079079C585C857B94F1A9D42DA1133994C355EF6D
-784924808DF73344D3857C6CF17E8792CE32D9567887E40B72B6BC2DB0C009DEC82E7F6EB8AA
-F8D3EF89DBA6D60245A36F0F5DCF10C7337F87DC4E2129FFD6E5A41A8F4A55C089B72831ED70
-43C885A408DF2D125C34DCB5F8094A22E0B9DD12013B12D1CCF258A647E398099EF6846A9C06
-57C49F733A6C00A2828E429ACD13B32C809CF5FC053FB1A370CAB1F54CE04AEDD139B40F3320
-138420D92F821DF20E766939BA659A3447ED6DF505A480E7182B80FE1B44EE838D9669516364
-B313618781026977EAD1D9607BC082E6B72B6FDBFB8E5A921B6FBE6558D09FB1B463F16F744C
-1F8DBB9A99B8FA0575820D5B5C1D61CF57E31FFEC73A951D998A38F1808EC3C208DD908EBDFC
-D80A714FCA976613908C5B6EE5C730E18E0935F9490A30F636015D4DD6AB292304878BB12773
-23501644A0C7B6342468ED2D50B02A19F4BB1C552C1C43C94691B74E111641EB9CC01DAF46D0
-01E7D20B40440FC3A8EC70165E3F07D232E3F0B5A80FD713229135F0AE8972F5E67E131784C1
-491810D2334F28BBD09F8AD81ED26C073C51D068659B937BB62DFC554B6DCA8FBA70B0D32D20
-69A1911091BB586934323F6B744877F8C4AC18E02715D363E7C8FEE14B76A458F9E2320F5E65
-263EFEE1C791812332C298A0FF162FD8209C7B0B83A3086D7097F838B9E2347DB520C450986F
-03312E4389D6E9DB875252C0ED0D034F256620E206402C4B408532E3B2467AFFF2FDE4588838
-EC9F0B17489166C21A1E47EEFB34736B60E02711223CE4832D5BF73B18003667B18C70670FD9
-443CAD0E43C194A6D2F5CA0324B48FE48EC0CB8CAF2A34D2CCEE1D1C98F7B3727B9E80A00BDD
-A36C716100D4A6957D7FD91B37B35E13104907BA3F997D309AF545978FDC7B61A066FD579C76
-BFB54E875EACAF944A62AF9229CCACCE33EE17C5B4CF0008ED36CED940FA55DC2B2DF1DF644C
-DCD6B3AEA5BF366EC759ADECAE380A5D3725D55DD2DA4A344DEDCD9F32B693B4005214098917
-86D77616B52B975B37BEF021AAF99344F4B98F18BF902C3ED8FCD233535948602C9C679DC42E
-878EC8CB59B096FA761BF52DD706066F1CE4ADC203E1961FAD16644EE853B608809A1F9026F3
-29985291EC4E2BE7F7BC984C61BBD8A1F9C416B1064E48AA80F0AAF5C501A20190ACBF90371F
-74320E9CE15BE21B3C2C47F0D26242DF61448C14DF29C19A27967AF6EE42CCAD7DEEE8FA332E
-A69CD4F25E850DF787457461EB242976A691375BB5EA7CB3EE6A20BED536BCBF74482AAB2474
-A3085727CA8E528D1A7A861625F34EEB634F571E92D1E8A7FE936998AA73A66BCB0AB73978A3
-DAB8948FF1F90560F27EBEDCEAD7B00800006486150B1EBC62C34E9D0881F299E1730047D9D8
-2E923AD5A337FD62F6E38B4726AE982439D858C77B533DD2CA2B67FC7EF1405852855658B844
-543A765FDB1BDE863873203C8378F311E2AAE5B18EEFD89F4110FAF516EA0AE717A13C4F0FFD
-99800086B878533A51DE3A76FA2FB35894BF4248837D11401A0DD48A74737233872CFDCA656C
-C43C9467923ED7D74DAA59578091B2CEFFA84EA5CA2611084F85552A00A5FCB569F02973855A
-267DB4289D06E175021102967B18C7383A97E7511790D5646A6BDC6C96D0359B00F32C1332E7
-3843B4E29686C3CB4B0D3EA207A68B8E2EF9200AE3D5FFA86020967F33B5673D6CCAC8226CCC
-C848E818BFA401934590AFB9E56576053682024B6D8655372B0755D7941591A91FCB25B2AFA3
-6A60177C3D4A9ACC0CE028B34E2E7314CA3C4D4EA230188F466EA0C220BE286E3E3C7F099FF8
-FAA464360CD139E10CAE196300C47C483303F28BB55A6EB99BBEF2AB3761C34880A966771643
-862AFBC80E5B0918F57F5B87CEA8F8EFDF7115B10D09C11ADFF3D1571CB3D81BCCBB4B52D051
-E8A73B385AF3AD6C2C679518AF6028A9ADD52D8E51F67A703DF61D08D2B6CF2C9E1CB78AF8B9
-899F12B0523314F3B00362B6DDE5A5CB25B9D4BC1F39A88FCC1B9754246A1EE8EA0CC61DD9DF
-C7CE59A594787645FF54DCD26A730C9C10A46887408C5EBB2E5FEB7981C4FDE33EAA190C4C0A
-F15EAFDB68EBC3E292D09C5453C69A006012FD67DF7D1A81EF4D0C130FDDA7DC425A4EE13E11
-5F37B9752555A6ECF6C3CC8AF0A9BD273D3AD3C6CDD94B4CC291FB7846D09BA16E21E1855585
-87B2F3A612C462B5B0D22C658FD1A73671E3AC0FA9E84921A4A23D7CB05CEAD346C425D140E7
-DE3BAB101CD20864A59B9B30135ABFD06C9A02844E90C161A5F9A870762C76D6C8DF26BBB9E7
-2D74E0BC16A58272CB21393451DED00652335EDA5A56F711886CD240180F12777D4D123F2CA3
-138786E3912A4ED601E39784CE0521B4CAD149ECB73A2B26EFB9DEDF0A64E6DA6E9A2758BCA1
-D91F87F41CCFCD50AFB80868FE835F7E19D485D30AFA1E08F557D5350696E58C94E9620F0DA9
-3936206CA8E57B626CBE6F2657C078EF90EF40CAE517FFBD15694972440493DE48A5F8D5B37D
-471DD402851D2039FA0B303E514A93671D999655C4D284EE72CE0FAE9D6D844E2F5C9D1B2B38
-2D9391E949CE157F8B57677359D2EB6582D0EC0C81C09CA24983E83AB9A94357C71558ACAF52
-BD93AA243E9CAFAC44EDA43BA82E52BCC0F27403F80EBEC440089B4ACD42220E55BCF567B91B
-7C489E34323FD5702D530DA3482A664D6CCA2A524ABF42C5AA28E1CEF36C7932ABE07055DA14
-59268F0E12A88372DE890CA864C88C8966DD859CE8FD7062C6A7F57C3B593531E656D22FF37E
-9683747FDB86B48281CD0B02BA67AE89FB39EC2CC907CD46FC10C2787D25F08DC1AE62564CF0
-4406B7696B0D1CD186E7540FB0853330FEE50EDFB8DCCCF68F56DA99822CDF02210A2DF1717B
-75E70C4D4DFD586A2638872A6A8864F5F6A908D34C5358AF1D3A92A619E55BDEFBD9555229C8
-6B2D68FDB4C2A73694051ACD33F851763BF30DFB2495F1706A60C2415D8A61A396A8514F8526
-38F80157748F5032C4192FD514B8B348D5BC92ECDE2AD133DFD1E33CF10D2E816F00340A8785
-79A193DDEBEFE7CA8790C033E021204AB2EF500AB1E3A28C59878684E9D85479BDC0AF707DD0
-EB0DEFB0D2D4562F199208D8CE0247223621B79B9B8B43F33237BD36C0D29DFADF862EAE0B58
-06F34078A9F0585F4564CFCF412BFC3E378682E237C0B717F2F932B3926AFD509EDB7180F257
-A51760C5318870E506C502CE88175FCB42CEA1E7893BC6B7DDC56FC281F9CAFB88196B94AC83
-48940E3B5D2EC42F8034451D41F986A1A21C6D114B54575FE7F73ED5A6FDEA80198D47C7234B
-D73D8BAD359780CA6DF5ACBB45560359F6395626A24234FE9F2EE5AD2107A49ECEBD714EC013
-FEA05DADB4A2E007E11B2D6DDB91843D8C9B490B80555C82ABC182C42CC0C192E91E0A2D283D
-769B383DBE3FFD92C1FA198FE8528638A4B841FBD66D35A8E74C47891BFC2FCA1C88484A4E57
-F2EE27CFBD270E6534E4950D7F3C575C77048EE81FF91A0B88DF764BE110E0DA7710BFCC433E
-32C9627B78DA9B22CE054B01BC8A51E246C53A42E8D7E4CCFCD6F83726626B8163F49BAD91B6
-D3B34C0B1BB6FF3B4ED077CC7D3FDA6989030BAC4882298BCE8120F0A7C7583FBD6CB4E9DBD6
-74C170F4C2348590E05052D202CE01A94BDA02234622F40DD23234A2715CFFB5E93EB1521C70
-70C95DCACA76771A8123D59D3D67799074813F5A6F2495288F17788FD1987BE59805B8542833
-10D923DFC6FE4579CAFBFA83C5294342A85DA1D9F208C4DA75D414E6D09884DE3CFED1D7CA01
-BDB37193BFE82BA1F1482126BB992A7509836DF1A78D92267B29D4B593637DF8A59FFF35F051
-19F266A72E89974B5F7C361728A8138B56FBB253D280FFAF6B9AC389C22EA8640E26D9049CA1
-40807B737B3FB6D34ABD8AB4C8EF41AA01AA17DA83572A4CF9DFC0B07B6E50443CB8E8C34CCD
-FF94FC0B3DEE8183D405066B040C7EC26FE9D801BFAF416ACE99FA853BB947B2A74A50347729
-22C92521BC85216CCAAE54ACD4A4EED9D378EDD351DE7B1FA144DF43178D891344F0137E93D9
-91FE49F83BEDF0F14E7DC7661C3E9B3166386777D6BD42E991DD9E4B3A8B80D9E630B88610EE
-2EA6EAC599CD4B44E6F3E774A484159069125D3C2EF4FCFA9E27841500F2600E4C369911DCAE
-273D8A44250CA1F71D189CAB07EB7AE314A6EB991E40618D14F2AAA65ED76994FA5CF24BDBFE
-E652028EF2CFDFCB381C0F419945CEB792C8114BA5DC56701B2E180365622CF407B159E66BCA
-B0885C0EE81F98AEAADC6AE498A811DEEAD1F175AB67D229F7A36FCDDAB217DCB5AF80D0E7CB
-2AACA7637A2BA01AD13A02C377F80B721C4DB8C4B351F3E63B6487D8608CD94CE42252AA0161
-12B92B995B8B8B3E27BCAF0E228BEDEF6A414A80B3AFC8F13AABF679F968E124EA11DFC5B184
-B88847BDABB47FFEA1261723EB4AB08470C897EB0D0F7636A1C153DA7E3D9BB267810C6D6F73
-ACA7C1FA9179D269CF6BDCD407DCA40C0D48E1C933D776B60389C311F615E92888260157E0A2
-C69BE07B104FF0539B8EDB7BA5ABAC5B1BBE2116DA8F89F105F150EB25C0A68AF78C406EE3E7
-46ADC868545F68BFA72A774812CB2DCB39BC1B05EB63D14CBB4194EC041306F675DFD15AA40E
-72C15D8CF587CD25BBF695A43D52EA1B6FDAF7BB4083637F8621B495388FE6DABED5E1006F88
-6C15DBD5FE010613704DE33E5BB4C1A25D7739AC14B5A9D000EFA43C9BF15F284BF1FABB9127
-4E59D36DE639010C971D395200955A4FF92DA852C6316E40A2AFCA3BA516834807DA299E79BF
-594172522675DFCB09E9D184691A917E705679AED48731E6BD3E27C80B9577AE1255119DFEED
-ED481ECF519482F9527EBBFD536DEE543D232937197E9C36B5C28291ECC51F4772C04195B7C0
-F11709EFFB3FB56896C49BF96E9FAEAC5066E981A5266E086CDF4D1EFCCC74DAD0B04420681C
-FC03838FB3427BE6B668E57FEAEBCE681DD4E08CA563975A069A29C7E0EE222680CF2C705762
-01B767220967D7267AD804A935DC8983768B1264824E0901166BE037B760541F45471B5CC941
-556CD3246B531612B9F9045FB6A1AFD25BE9939C12C7759656AF9C6888D2D2116EAA28ADB749
-2E36B0CC0AE59579FEA191B7C89CCC3D58D68FE47D20BF0CC7EF3D559FDDB436A558F6DD5F35
-DF56F380377DFB2958D095132D781C74F46ABEAD24993BCA7939D45E6CF8D2B38E5B13326767
-D098D79E68B16C870B19ED89F4FE3BA122248660CCF1878EC5B872F15591B48F1F4B22180007
-F9C384C2396AD65334117D5EC05FE4F7F467A46B53580FDAA00ECA4EF7353211E44A6E3E345A
-28CF6D51FFF3F0E676DE1A6675E413E01B8411854130F67A29353646FA1E82CFB637527EECC9
-3D615B3019380C5D6BC1EBAB07E083E748C615A0357877F252200F5F0692FD3E0E8DB5881D13
-27AE6197700B26D762AF86A6A508FC63A3AA2444681E7CBD1C947D71B53132F4D95C493E624D
-6A9AC7B71D6C1D05D51568E2BF9D9516CB347F7EBCB95C4E57422C7FE8CCDD83A3CDDE4E4F56
-F9D3FA40809C1D398D25B02CE3A74B8CF22CCD3E59DDA0FA656F7A15E9FFD31280135877BF31
-41A802AA07B007D5FC3BE0A375DB42F6BE3FB0B44491C246D15BBC2578EB9419B60BBB630909
-B791267BAA2AA5DB531436D7AB271887DFA508F1C8A66DB8641BE8366D123E32455F8358F7D8
-4597F3C74AC243D7D4502AF82FA4BB866722182802D2948FEE0C25E76DE72C7B1EDF14BAEED0
-EE8DAEA5C1E0F81423D2FD6DFD66996D2FBCEB11F9A849449D2D01CCA66F3171E29A9E7F7FA4
-05848649936E18F64634D0EB0D8CF80C18558108F5FCB8502A5EA19591CFEE1436744D6A4392
-FB19CDF90FE741BC3811ABD538C901D215C0C2519DB41E02F2728447D53A30C38EF603EA0DB3
-CEA51FD498D9FDFFDA5B598F513F59735D2AE46C727AFDF0BA7BD5D967CCF2C48D69E11CEBEE
-FC1E7EB513ED593275F1021D05398D3D746E94F9E402A5707DEA14C33C91BA0CF9CF83B873EE
-E654BA6FBB3A0C0BB692F191CF4B6BBC03E31B3B7E11593D650A6102A57DDD0AA168DED25F86
-4EF8FC1A3877B23A930AD105E3607957C6E9FBEF93CCAFB373BE3EE0636F0CB4D5712E726109
-0E786CF4EE23B5C87CC0045E8E5C28687E6E3424DEB0EAF844FD2BD0630B83CFC47D0B17455A
-E67909390708184D4FEDC389617D2131B44F66BAC4E8F5F51CF01139D8D5B95AAD41A90E44B1
-2C99EF0941C9A045964DCEA1CB48F204A263166DB83E94F3460F4215595FCFA32303AD5464DA
-FA05E8A76AA4D4EEA890040E8B123889A91A1B59D3B8C17AC501542D1608418620AE973BEA16
-D78C924E8F5CAA40CD430D56261348CC8A8C3DB005401B8968CB2D363F1EFF18E535C2C6A7F9
-8A762348087817832017A5E78AFC366D305B0BAD93236C6F479A3308598A89FE471B2A16FCA8
-BC27CCA40E1C7371BE15F86FD18F2ED5EF1B284762EC42B79CD3E10A381A24E0556A4980C205
-A41AE5131EC2057F6582A7F55357062FCA89BBEF95CD7A28F624929CAA9AC79DEEE605E7DC8B
-9F8FAC570AAB6E3727389809EE23698D75E18E7432A9341631259FD3061ABE485B9582154064
-0A455419C18F0E37E48D8F87694E8447564A88D2A1975510B35289AACFF49462B4F2DFA291CA
-AA819DB9E2850AAC40479C7CFCD3F51147349C5158E03670320A6908EC986DD297B30BD97B12
-6E0854738E15B0DFC67A75E6AD209F1051389FD5A44F5E9DA45564C39975FB3C6DA82DF8D1CA
-753CFFD9B849CAC01719E6205A81D8EE45EBFCFE17DEDD50C9DB9D8B9D5420FE60FA0D7E61D3
-9A8217CBE47BC3C4467124D485F5EB354761EF1663A780E9EA190C9E1939B68157BD3B0157D1
-6416BC22B8734C587DA7ADBF303089FDEE1BA13E1969B9AADCCD265AB3C5052800728044C3DC
-508D26446299E9139A2733E12FDB94809B1F744B09ED031ED015902787DDA85296DBBD7546C6
-802E5258F1B25B589C02DB4A839265D0518B91F1B31C8815C35E233D9915FBAC32D4E17FB575
-3037061D06503A9BB7E8EBD49598865576AD3471E00816D7D9F1CC0C1E4D789FB461A59019B2
-358655F4A38C71B051FC5496EAEDA5800D0E942E069F8A0EEFE56B5EC23AC7BA0FBC1C20700E
-D3A296A497287E412ECBF545AE24445C0DB10A15DDAE75961010B2CBA6EDCEF313924B9536DA
-AC0B534B9AC4DD0BFD71157614B79D9A20AB32226C266BDDDD1BB9E4BB4A84214B42A9EF5D72
-86C53CE20D570A8C2AB73A605697E9DB8C5F22FFBEF6FC8E1C1CEA54FA54FEFE9F0E4D7ACAB0
-8770A8361B37C610F35A21417200F414871EE3050E162A6958D8DF0926F52F9A3A5895C886AD
-0E422F26C10ACFE456BF7D131D23692C38210032077F29EA31279CD1B84F8A01CFC215F3272F
-9B0A95FB186334F46B783803715BD8560F284E9650C43A76D2C5AAE4AC422407C3C6895E3DD4
-4CEE8E690477B3A924839ECC751113DA57FB9D9657C6901D62458523D8DE77EE9723430B92C6
-1FAA635AB5AA7A07D37835A289ABD9003391DF004C2F5E557FDE7967CA9FDD4207666C039591
-E42991F5119C331585F8A09D792F22499F214CDACE78225E0CD3063E0715ECD6FCC08F4974E1
-B64E72C6C14B4777B927F30EABBC1931954FAC89AFEA552E61A2BCB05BF39D3C45FFD2B6712D
-14B6DD1D9D0E55A9906EB20859F5C5F665C6AAA4CBD6753D972A0292F0E73990C4BB38AA087C
-2D8639444FD403F6F21684ECC87EC622433CFD00A95F1F8536A6137854E8B0866CAAC3F4A1D6
-249FCEA6ECCF3F985F190C5E6480A2DF4CC88A2391318199C50C1501343EB3ECDA08833721BB
-DA5FEE143D59E4F252AA9209678FBA9A2EF4FDD88D60E6D854D01CFE73DFEE31CE6277307572
-05AAD6082182766183C851F16F67AC2DC495789C6D2072101C285370437D10D002B0AA82C117
-2F876066072FC2530C976B69A18F3340C2D4ADA76A2A6E726ADB4EC1FCEB0A86DDAA00E878BA
-EDD1DB17C13ECFD3EAAE8F3FBF9E2B52345F0CDC60F7E4A0DD38394A6225A2CF1C96ACBDE37E
-21E9285136D39D587FBB6A6EDEE011C0C667DBDB8AFABABA6E6FE5D819703093EE177F50ED3F
-15D47862AFEBC145557650CA6960A77157A87A8F097927EE7694240AD78AB093AE72852A0452
-C1E00A0EB75557F8E831BA3B8082904BF04013D9BA1578C1ED4B60831BF7439B250B5DCA27CA
-C38D178AF55B2E49526CF4DA79DEB91299D477EEFEB4714815DC65979F8615FB22C3016ABABF
-068F75ED6248A0B6F5F66AB1FBB5CDCC66ED2AD1812BB297516540A06C7FDBBFB750CA1588F5
-35E02846752D4CF6D49BE39390978C4D40B3D047B0A0CD8050A454466C2839C9243E35AA533D
-2C6F1F0DCA98CE5D8256F73AFD914DBEDD8F498259A758F92ED0ECB66D89E3A70DE3FAFF248E
-B8B69C8BC6D8E55F0000ABDD2BF1E3DF610D0AFD5B7ACCD5DFBDFD15C23D384A1EB06CB6DFD2
-77BC0683619BEDCF30C1184C8BD4EB85F4178E9568ED130A48D3490B6DD005D5FFCCA2E78419
-C4371BAD9ACB888DDEE9FBA5AE1FF78C3C242939EEF2ABEB0BC1A05795C92BC71494FB6FE52C
-0D0A543FFD028CC5BF1BC8F7C6AC735794A85F718A153AE94916A2F351B4F0D94971081A72E7
-1C2ACE1BC39C2F0516793DFF23BB28288FAF6C8B7F513F98828353C242E6E6692EB8EE40CEB6
-7EFAC724E9C42343E69F88F1516DB9FFB00E8848DE608B6BA395088314AD29A380CE6A99803A
-659575C8DDF800115B440C4090B2389AC4FF25DB8A09A8E4A85F7888CF881CF5AD753D110F01
-D53A2FFBA14B6AC8CC658F460A3EEE63DAD11988FEC5BC068C97DB6B31328B0858CC91036BA5
-4A7B1AFCD6C5C08C8A9AC0A0A7D7A9CA47DBE39CCFF2835901494C58F00A8418B6993BDEB95F
-F5471E45DD4A7A74C5CDE2D321328A2EC5084750594967FA6E215A142EC5119FC6D157DB5482
-504E49D7EE9F23004354D181312EA25E5E9AC417DF4C6B32EC854222581C9364EFC054E6FA1C
-965D632433256BFD10ADB3A2233932BAB77E0BAB645B769FF1F0B8C9D3066E058F2BFB283068
-F4E5FDBACE6838D1DCA0E054EEEFC10F141E3C0DA08F07F5D1C48A92232B675102C0F79979AD
-B3F5DA791EC549628C14D3E58C40C273D0B6A4E1C7FFFA5E13CE66B7E9902D0343AAE5F0E50D
-FFD383344B653FCB4C88B2FA228F00FF037F18855B9A69217B4F60FA80D623F36AD3CFF668F3
-075520EA8FAAF5728A0BD8E29F6154C59BDA685F8183419AB1BC7F3824B901C3BDC79E6BEDC5
-C18D3EE9B94B716A7C5622FF24EA26BC23CB6F64CCDAD9B152258DF372E1A1A4469FFFAF2E7D
-2DF544A6799229C90143CFEFD9CF032E6FBD47C9454FC9BEE3209F87DBB1A52E09D905203EA8
-EC63075172A9593888111260CB61F11BEA52ED4BB2739BEBDCC4D6401B8749132925F9B8F5A2
-D291D1D6C4D735015A0DA050EE2AAE3038F7D85A890E9DE0487B5AFBD91EC22CE3BB1F5A7740
-6EEFAEE524CCB776BA90F2B1934380FA10CE565A63EE449001FB2A6641DB9A50FC46240B0A22
-36BDC7C1019290D967B871E35DA9EE791C791A7217994B3A403A1DCAF5BF8717A73585719BD8
-6486ECF0A39FF60199DD4C904468AF19DC26B95B134150B5F1807E05933D54E316B433DFEC50
-BD206B57BFE173219484178332E6EF618DA247261AC40F7F66551BFAE05D7D6127F550BF2183
-B73804B711FAB1383EDE3C463EA0C0CFD45FA674737BC06A266DF6569CDD0CED4F4686CA608B
-F503D3DC7BB78DC5DEC75A4CA8B260F7D0E20A725E1B8AE664FE131A4EB52B8B25B7658ED00B
-D9F3543413CEF6ADFA4513B1C317EAAC1861AE3D3573E18B48CC0581329E5698CD66C362E6E5
-7B4AD21413E15483A6081B4FB41009A9B6E668E2D3E9FEB62FDAB26D59FD660ACD7A1FF31FC7
-61AC6615880A076041BEABD2E1CBAEFDCE7B2659451CA2ECECE2507F52CD1CD866508F89E33F
-CE2D8DDC60B5FAA67B741ABD563FC47B416C7C9E921DF099A62AE03BFEE0931408992EECC454
-B768D552423A7ABDABFDD9828E451B062723F6E7B47A5BC51F58082E7E742C6350C290E98541
-34F5A116EDDE076FCF8A8F5C31B05833AC047B4ED07E52378BA9EC1BB2BEAB971868D794BBEE
-FFB1E793729D364BA1A07377D062E9F3965E4E88B6C3404F7BCDEE04D76F75071671C2D01745
-E52A5D354406A241CCFD75ADAE4BAB41463258629BC099AABF382A784A9AB5FB8D12DC7DCCA3
-1B432279B53F29CC72E1ADF288F49C0747653AD0D6014C59D02F2F525691ED7FB008D3393D71
-AE32546EFAD70AB793175F3DCDA2FCF38B4DBE8D19B5BBF83C5637A82845C512BE5174391686
-5643A25E68FBCC41D263A2FAE21AF027A359D747046D3AAF2C9D4F4FE3F27CCC767D80DE6CA0
-8ED6DAAF2D030EACCFA8BFFED9F6223E992A2BAE3EC096C2026678EA1F6EE604B5D317727F6D
-FBD358910F82B3070B751E80A7BE80A0D619D589E7123E19725A225CC09DF9CF16789C435FCD
-071CDC596C79A437F620EAA068F7D4CC5B125C8A30F21F569004251269839EABC217B9CFEACC
-A805A52B58D254EED7D4C6C72F54ABCD6405BF1F2DD9194C899CC43DD6302C62F6A157048666
-1D8F58D024C7344B0FF63B38225B236CA82272A5A4AD4B454079A5D2465883BF6BEEE55BE89E
-A95E7FA1712ED7D2DA6748EC2529C307783CB0857876D9EEBE95FE5A518D37C01849C906EFDD
-A83168A91AB5566899CF5C89D1BD72616959B07C902EFDB69681ED9BB39A94C19C3B8E1F1111
-4E5BEC7F77E1CACCAEDBCE3042BF218F1CD54848948BBFBFEB203D11DA207F0D348DABAB63B2
-4CE6897B23CBF906A6D04819ADFA44E0704E60E5D6CB014FD6BEA5DBF58817BB1DA194539D26
-860D021399AD8224D3CBEB34B3B82D371C758FD87D5B34E6C6C8BC9288D57960E95740767E35
-5FCB1293317925D8A692FB0946B766CB57E732DB4443C835D259E1C94D7D6A2CBD17BB3CA042
-DA6D6DA42C6582C5F78A5FD225A9AC2DBEE5D569F55576C347C81C91CD3A82B84D65CA2A7894
-324BC7B0B2A85829A3832DC2A662771D1167D6D53B2670888E5F794F53759B20A61C5678670A
-6E5C4363AFA5D38C824A83CA371C501B1F7302C0B87177AAAF559672DE70F7FEBDFDF577AC30
-5470646CE026A27FAEAEA88CF55A8A4DC872E75F872F2F16A465591CC1415BD46AEC545282E4
-CFD66FE372BDDA314EEB37738D43C1B97DE289673F443DA68D6C60CE8D5C027C4012C914779E
-3729F59A34C47219456459D635EAA63A56C2CA55C080C716E8544A3F254D63BE387DD6026404
-97F0D327BD7CA294F9D176161B5B0BD023781D7679380E002DC2E19D728210ECE5C7AB22B9A1
-C1140261AA1962939A6E6A345A53494F369692A9512A45022AE5475D754A48F359726A12B97A
-5165AD6A58437CCC3106899CB330776ED999DB92BA5144BCA9C444F7F6F5D362DA9F6CEEB1D5
-FAF4675040425E06B7C7DDD8B40D5D692CF14F1A5A4FB8F6DC3861BDF66CFD5F05BB12BBB8F1
-FD15E12D52E672D1C3C8CB514A75BFEFE6DC9AA762595FEEDB6CF5AA725A957E0897A1277563
-B2BDAB4B5A815B76C2CFEBD92AE853145CE0A58819C508858280C19EF5FBD7DAAA97EAFF96F0
-6A446055EEDFFEAEC91F7F838A3E88161F89590ACBFDE2656F5E953C2FB4B5D50966A62BEE95
-8455EC9105AC85C83F7BB77BF3BDB70BDFF29E8217466C54FE16FC4430BAEE46C420D82E8943
-EF299B1E09C2D443585E83A6121B8B537DD67722C56CAF855C4D5838C83CA87B2ED99C837B40
-AF624C91B0046C1FBE5BF3B692BE8030C7C9FFB8B425A7CCDB858EE89C43A7C48CADF31DD31B
-335D09A895DF509AEF5C30D91505DE675499572E6FC8FB27302CB4373DB374A5CEBBDA65A08F
-A2AC1D4F06BD23008F3F12F589A5BF58AF6C8C3B8766A501A82A329B98CC9778CED03294C24D
-3997ECADED04F2A97E2EFF2DE1727D016D27AFDFA2EB5332322C373C1A8348CF0D33AA781473
-566B895B0D10CC90CD3FF96F43C130D0922022DA6B1EB53AE6F9E4517990E3CDDB861AC49C8F
-44F543BE249179CFF10BA3476B6E3354E5AD2E333618E917DAF0FF094E85583998D5CF7206D5
-E3BFB039391A352AE42BBBDC0467D657B53D2C9AC681D77F1AFB311B566823B024BBC7CF146E
-05786685F6A7927713920BF81265EAF138CAC0D5AB781329B20D3C7EB73E34EFCEC43A94917A
-715EB30F6F29DF56706A478A02C333B4B943FE1EC1572B6B42600B216D8C4B9C9D46BA874DCD
-DED4F9DF2E6E58454FA85CDA1233A28E04D277F6F918B80DCBE27BEBF4E6605D7C8F3D8608A6
-BEB452200A52FEE8821E0618D1487770080EEDAD336D76655720EB61D40E8F5E15955661DD1A
-E4BB5F427ACE98B0F983B4B48F0BEA0E487AC197D028AE6C7562EC906133106D61AC3CFFDDAA
-0F25AEFAB9FF8E837A72AF67E6F176B451EFBE4F34902CAA5B969468C96BB535EAF87CFBB97A
-C988C59581985138A19B6A51CAEFDDAA6B7160D240EAAF9D4FBB54B65231F50F9C5BA0D476AE
-5B0C76637B4B880E5446BFE6BDE7B56B8E737E1FD641F279944CD7F3295A266D1DFAEC9A5885
-513B4B8732B8E67BB09ED3D0AA03D16630386FC67ACF35C882D578C0BEB9177FA46A39C3F97C
-54BE0AECED39AC138B968B5F01F0E01379610777DB04DAAA7AFB55023BE0CC8D01ADF14426AB
-57732B265B9DE5112525CD5CB6FE9A959DFFA2E7DA469B09603A327387C49E7F0867933B72F3
-A197B34C05F30089ABCC0BB715A21D9355ADCF436FE350C00CD8D4BC59E0E18B3492A979AFAD
-95B8A761E41999374D26A1C09C93021E695EBA9067471AA584A845074400CE4A3F04BEEF46B6
-AFF66EA761C45247675C4DC1D79F78B611A75FCCBA2AB6D1DA500CF96A013F4C1F28161E7964
-155411B0BC0A77C207D93B3D782413EB42CBFA714DE7A72C727B9AEC77CE5F0A58B62E239EBD
-811BCEA607FFDFB93A020416CFC34493B824CA9E074F7546CF1279AC40D0D71E982085214125
-CDB98A8C2386F7A92351A77CD0D0B75BD682A0C3065D273543FA230278CEF57BF1DB2808303B
-EE7AE7F604AC2DE80C2F568CB1F1900E4C37F125306329144A6A22CE746C8CD71100B0768F2D
-69239D75D2F0660A6423C4A19961E5F01E9C09959B4810B8F8488F60BB5FB2F566D74C63B631
-7DC0A97BF4C91DEC9643EEF4D02094863BF02CD55CEBC700A85B245D847153A4A779690D41A6
-09FA3C4CF28657FE5B8CE037FA07C9ABDC8918DB25FF354065C3A16D7192AD10D427327437B8
-93DEB9612A3032486442689F3FC9FDB5678C91A845D3504BC1E1B8ABA72CDB4E4E79C60DAD1F
-AC49894A75FB34387C9A951FA989A1693CA4B2636006188F8710E744C838B8D70DA7A492BCD9
-EE0D7A1E9DF1C0D6295FA32041E39BE478920C21143D9CFD04C60F8818796925A4D377FD9689
-DD2D1224AF05B124CCA30C7C221707F237BCB98165C58887E6B82BC7A939C0E246005B284662
-CA1C87176AC8C4703A5A7AFDF82294D6E6CD91006C487CCF462A2875E4E6126D055D0A88AD28
-04F59A58EA578BCC6B61F5BA75FEBE2F8C76B5273A3C030B0AEB6DDFA6A01E3275ECEA0CD835
-09FBCCA1D037A340BC88BBFF847E5088A213F492D971B2D8047C14F3C5931505D2E911E755E3
-974543CAD391D5BBB318C8230245BEB4FB59836326583CAB5433B2169575FE4F45463381790D
-CD51590F47F39D79BB39975CE037583193E37A6806335489FA0D2E47153A2F74BE9E86B0948A
-E0F507DC59727389DB6157A925CFB76304B0B59ED7F8202A24B412DEAA496F9309E41BAAA625
-A18F16B8AB629B898E36E85204E22874D86E73A787A715E6C565170B08ED7FC4861102754636
-CC3936F63BFA4929D1C5092C1CF8A3E2087B8D9541CCF334457353236378D9D427F114B53663
-54A15D54AFCF6A1D7CDEE04F36B68C6825569ADB3D0788E55CF5F2BF4DE021BAAB47576FF620
-ED9AE72A412DB43663461D9DE39523B26C93B3A35193D4EBA8E4833542A11A8CB896F97758FC
-9119E13259988174ED701C2B28D5F92F36882F920EC1E54318F019CF5520EA2EF70225C1DB74
-87DC8A0E94C4A7F24707F0E23292A5AB9ABE542D96E9595E19E0429B80D631074FD8E2D7BD7C
-8CC9D46923AEB9970DDCEA25A40F7F7A13C1BAAF57168C79C0AE9F05EE46D571D82F87913923
-E022E0FEBBC90BD9A3E744DB38131E542D0BF7188490F13099B62132920602C5D711D7081225
-E701DBD08EBE2071D0E324E70E35BB5725201BEB4D2FABDFD0F5FECCEFF23E435A7D4829982D
-99A88124748927447A9BAE9825B6655F5277D9F111CB7192F454C8B44782B8D8186EDCA1BF79
-B5AA1E7977A851653698F0425B5CD7E856566CD5800C2341D34896F3244FE11F635C9AA4106F
-2EA3068DA169F75D66938E71B0FCED3C1A54791CD1FE2C2EC414C833851CED804FAF35B2A200
-75C7A0859BA7F026FB0385D1EA8FD2278E3FE65CB153878B81FE64D259A4E2B13EED8C86FED6
-C0CF193D4A8C7100DFD88DAB1B302C1FEBD6872D0CF00BE1834383ACAD76B4219BA658C6D489
-C278C7F6EA7BB7CFAD2F4F5EE7FC7318720DF8705D25A807A5FC57B0BC36AF6D74829739BC9A
-748F6CD46143B4CE8586ED3EE3DB2D37FB7661796F5DA96CCCE37E4FD676F1BBB29C88DE9428
-61647C7B36AC54E5D8E42A8A8AE041F6CA3B8757E177A37B3ED4893CA670DF149F0B135C184D
-6DA0E3F93F67C0F1675B7CB9F9EBD861094BC4DA3FA0374B6CD25948F5AE6852D399C587B2B3
-A9963663B3B3F2B2D14B3B27FD4CB9E97D81EAC355791CDC45C06CA7C52394AE54282C7EEDAC
-9F8C3D429F9715FB65B926B5CD88015ECE358A9F3B9B31B2C18805A2F46126A0DD98E03A7FF4
-724322A8A2241A44BF9E93F31172101834282FF3CED110BDBF102993EAD3CD0C71F138BC7999
-DBE8AA3028096AF00FB84809F5BC40CA0C2CE67E061CF4E638DA27C71960ED73B8A5B0ED3D2D
-0F9F397A08D269FFDEBAF83B83FF97CBCA40267726C484E04CF7ADE072F574BC61A4AE9CA114
-3F884B8247AD84ADCD11163DEC94B58972450F4868CEE42262C05762E0458E534E22D7B5BB8B
-159AB97AA6690E6B1A1B5640F9C58F9CC28B91C3C52890A4026C2C76BAC8B7EDA68BFB544C06
-3DF31995A796CFF3628D925E0592F1E72EF6627B6F1B3B9865BA2EB3284DF3693C7DC0A2AF3B
-DDDA2A4C1EE50A7B0CCDE017998ABD9C8276AEE714F7458A7C857EE042102DDC5C49F29A6710
-40DC62DB365431E23F7E1BC4F6FD507B66E4423DB05819538D6917684DEA13E78930DD1C9DB0
-88D3BAAA6345BE54FEDA693D0BB7389B3A9AFB3B8AFECEFCE0970BA3D92722BDCEEE00264FA0
-559D936302ACA513794ABEE571F6665F397AAF1FBD723646426694F5345B1D0F48BEB93F223E
-D9C7DA6DFA75B8BCD7DD0AA2E87154E61DAD00F621D79E33DC2FD2FFDA81A5A4B90B85C0BD8A
-4C060A981C2CA7CED899F8C74D7B60FBF405FF4C389DBB528ECA3CA64F033FFD3E31B92854AA
-10BCC16CB865FC107546F7A694214C94E5A5B3B85A1D67258260398559E1B4054BD6F3FA4729
-875F077A755E4E02D71312906FE8624FA920625E5174A3FBAA5E96A61DF1CD9518519D843032
-DA44517855CB870F01EE360777037E49A3790F38EEC9E94C1EADBAF269278BA12A39E391D896
-0CCB05F63A1DD58978DA3A2CE42FA115CFC8D61B367395FE34B594F16B6D4D130079028C7C74
-42832BB9B66CD692738A3BE4BC85EBCD6EAFE50750545CA3B23808417C5697F93DD784E67313
-2FB4B1FF3118F6195194034FF68678BCAC43E7D2FF0A1468A3892B1D5A163DC3DDB65A556D12
-8A1B15C549B69FFAD02663FBE7DE7EAC20064675DCF10934DF83A26959AC0B04D9CFBCF6F261
-780FC94CF65998D75A809771C837C48A93404349EAD0313EEC1EE84B5C9B8037B7A7E996DB96
-BD66CE43922DD62BB77A7FEEEBFFE7670C57569166432932ED84044A6027C93068C225AD3DA3
-879E36A664399519ED173228BCA4BEFEFFF9BFD5159902E11E6B13643B521EAF920A9CAC2B56
-A6697B580610903778723510FE5CC92A6A3DF0D96A454DC06DF338ED075BCD234AD86D0631A6
-1AD8D8925F625C6DC9158F49FDB0F0517AD41440828F708C1DBE571506FABFCE22744C752624
-FF7993400278A8DF6E078F2273D2A0D3B9CF07841C1552E50C449B50940AAB3F5A1BAA1A0BC4
-EB701549E3FD01CD406144E3EFC8C20495C792CC798E9110CE6A439C097EE4E741DA165A6F20
-C21442F16F55723BF9C55A10919E4404987C444E30E0984B83BE5C053D4B1D422E367D92C4F3
-18FDB00386C392FDF99C1ED26FB6613B87728ED350A1020D5203E1D1A466B7721382792B84B2
-19D06B95138DA309EA262CA8E867A57401743FEC4429F28DE576959B383AAFB5DE2B23B200FB
-5233665175F45DEFF2AF9B865C9CCC8A93592FC20A2E6E10A38E35892605640DF13D09BB8159
-4F7E11C65E2EED1D48934E3B0A94A79A149FCDC373F4506079480D7E82104E3868629A5DD392
-ADB96E26BFAFCD270C5DA150E98384E657984F6B370B74A6EEA6ACD79FD9306F6DB7BADFF03F
-0927545F17C664230E576FC3DF6A208F2A8C208383290CCC3C4DD0574D3AC5BDC4F1B0276DBF
-476349F41C46C671F1B82C2A9610E645994EE34097E09E451BC27390E1CB9CA422345441A4F4
-BCD078C16DE180CD6494523C5A581176A8408A755F847D06F6A4C48C2660D67C6B6B18E3DFB6
-E9BE677C53C994152B35CF7C3DCCE4963750A1595215EA8F722E0771A63C2BCFEC97E6F98130
-9CD50E60A81C54C86A24BCCA7EDBB00B19BE8C15E5CE13FCE95DFB5EBF4EC4DCEA1A68B454F9
-098B9A75CC40942F3EAB27BC312D59E6A7CFE103D52B8C50EE3E9C4F2AB230ADD24DA18C7E26
-A1F6CDE55C57FBD187207395EB22F5B233B447E08F68224E79E101FB811082BEFC5D9D5B2E8F
-13BFDABDB6CB530A109ED613EDF59859ABDE44857A09679E7591353F60AE6D9C53D8DDF43A5B
-E6F2361E6E27BC4A955C9C1510F6D29B7D51666E532A3246DEC71EDAE02E5F309F60DCAB3E7A
-3D2CD95A95169DC277B7C585C5D5D9B9CCCBEC5434F9B067292FC24E3EA6239AC38315364721
-61921834A111C3343DC4520D9484EEE08A55EA4BD9A12A929EBD9810433A120B55E43B041240
-2437CC64524307FF4E606BB26DD078409791A5DD682B9E2559DC573F62B81B785F6355528CCC
-3514FA8A3E9EDF85C8BC481CDEEFB542923FCC4F47C03CED7C440FB8189A9E751E23E8D095BB
-00106CE0C38E2D612AE2B8A92379C78D59CC63ED37F43EC95DC0BE32ACD8CE66C27393C6AEC3
-22349B76D0349C6B55E74C67743AF556762C910A68A04234FB04BFB4BEEEEA798FB8F55C41E3
-E31338CF4A42802CD6ED0107A67FD0581C46C9C7A468A77502E300FD635969EA54B92E677C6B
-91CB4B4A64844FCCEC3736EA5BCB6D83EC7CA804B7F4A7EED268652220F549FB75CEF0AFA34F
-10E275E1C73007B42E7B5E8289142008872F8926C43D6427A16BC2C1B66B6D802F0108B1B9DA
-CBB4387B76170250E86CAE8980DBF96F6E774C8527EE3D04FC331430E18EEB6B865DF98332D6
-BB14C7303CBB3872A8942681DA74AFFB8AB6B377340AFA1E3388C9632D3380384BDF0E37570A
-45130B14CD9343C43B57254D2EA03A7410FADEA0E8B1A2A13F656ED449EC3259A84C51E5F701
-8B08C3819B0C6DDF5CF1073FBF4F0A6E3B1D018F1F4A443FB03A2914BBD7FBA81F91AB449186
-D73A82570C74374B74E7DA1FAE6657ED2F8D7582F3D15C3C9C953C316381CFEE50FEC00A4F55
-60ABB4E100C6ADC3E37BE6654C7037DA755DB1472CB0F8B6D347C04CD674CB53BDB1713710B1
-DF523B74DF0B426FC48F96891467E195FB094614DFF0CE43312EA13079D43542996CC5C1A131
-252FDC2FA550B69251FEF4E24578995EB8952A971A16284BC96E25E26367FF2AF5236DE22B36
-A0EFB6881ADD045408009CDAC60AD78B1E1051506AC99854865D5D8510DBFE71F9C591176467
-C403F965CA9B08E22E9D1EDE893E2C34B507B8A4376A6F3E6F5E757EB4496D173444C60D1AFC
-F98FE8EEAED785C09E31112EA27AA763AAFE140CFAF439D2FD0F3B86387865A9C4AA87880AAA
-84FF0D7BB7FCB3193933725FB8D987AF0E319A471825083B707BEAF840A891B480F815EA22DA
-E691377EF00B33CFD6EF571E5304F6DD6707AE9B8848E209813E0482F4DC9C47B96C2EAB67B3
-3BC04C04133DB261DCF40DFFB09B2A9490E33E19F981220DE96CC68F6DC4F7C0035D1C126EF7
-EFECFAE676C9F806AA8A6CD2A6BDAC7616AA5614F8896C4BDC15DA6235354B98E064E48E8A76
-921384A5BE910090BC90D7AF6FCC281D22A9D439855DD1422759C7BED1141D4BCB48B3FC5342
-4269CF9362038AC9A04758F9FF94E76D96018861F14AC920D080CD0AE97DE73F6A9145C7DD35
-3B1AD58F591A3BE8BF38EEEA8705A38311DCA34E9B7D63A3B15573F4208F2E43E871452FE806
-2C2665E3E25BF0CFF6D0939642E428E5D08BAE9E285230AAC838C56A2D447C3BF63C1D7330F9
-E1498D10DAD3DF58EDD4ACCB49E988C2F6F0202DACA3B8036EB8F61F94E36E40ED92BAECA8DD
-79E99E6382BD0178886FB83830D626F97DD14D33F2BFF91E50D7C17634E2F543A15B5966C465
-6D0AE3075004DCB4A66BE2B51E5356F925D8D5B62B5AEB4549E1264EF0CEEAF8C97221AF2EA7
-7BA6A3D3744D8A88819B80A8E59DAF8482433E1FB5B6F6237D757C044FB8E12CFB5C64139728
-9D3798874BE006294E31D32903DFA1F66F6CEF645FB73C4B0B14A73534CB9C607D8F048AD48A
-D4277A978E1160BEEEA75580DBB7791BA8AA165AE3BAF2B224954B8011E55C1FE6363FEBD606
-B538A3375021CBA4065F824F7F99CE1B551F1CEE4E5C9328C340AAEF2FFBF6F377AAC5EB9419
-F450F9A201F887C967969413B6EE3DE4E2335EC521D3ABD8E649BB08D521B11C16C247EBB046
-1548C1262BD0861A62CF99E1F7FDD2AB71233811BEAA49CC1F2B6AE58F56F05D284034BCA61F
-22ACFDC9CF7A03ED685A9B3AB8A160B08217A5B4DDF98C953C29CC82E9152562BA62242618A4
-FF7F8739C0CED2213BF14CFEEFB0E07A915854C045BF48B3AAA7685B18417FD1F4FF45AF18CE
-1C217F2DF4379FC8C810CCF104A38DB09EB8A6930BE3BEE910B3080089A0FB17E95144F422B9
-62CD95E00EF418DF04D75C47E24DE299A6B8A6EB89EDB0A2B554FD8CA33B04A9C7E33178C5DB
-0E23FB15123228F193010426A239AFA02F18A3BB58A63F66D545D12BEFCC740C6878153255E5
-ECE992777CF4A612A091A9C87D86D4F3FEF60B84A309D63F2F2E918C59E542F72518932B1C43
-06117411C3EA5AE5DCA73B0C9B4A01B3629960D4AD3D70BFAA9AE84E30CEED14AD32C8EAAC0E
-B40B6499E28AB960F6363F49B0B27DDDA32ED97A33517C06F4A3CA5B7BA3BB538814D48C6F27
-C49623D5DC771F3277A6AABF83FFEF730F0942309D15FD971C4D1F47D477066D77434B85227C
-6499B40BE44473EBAA3EE101CD76F3CE7ADF0A32AF501D7514A8ABA0AC2A6E352FF90AE947BC
-961C2D1E01752A7F34DA5545DDD27F1734120A68FA77FD278017E5274972E0021E36DE869361
-36A7AAD834BBD5B8F844CA5A9A63DAEC3C79E2F5BC7A0D3958BB337F61F613DDD591B06B37A1
-1100EC29C1F525D198A107252248F27DE3C86E98DD92B7D74E18A621713B40E12C9433C29442
-3F2ED4D9B94FF6B05D5C10D157F0009326D975D2298379986D8109969C8C3A096E9E8F0D8C47
-EB19920D4CF9CC7C4CB80D5C9DD202856438A91DFFEAD2422D0FB09599BC99E8DB8FE6D6EAB5
-B9D777C32401203A48E7760EE537C84F9887FD7688D0CC469C0C91CDD02C1B9B55A1D10747AD
-6B9FBF0BF3A01763906BBF16B047701BAAB00E11296541AFD1C7BF643DC015477688FFE9EFEA
-F140281883A47BE37C213BB560E95D7EFEB6ADE69AF336708096C5F022540CAAFE7AADE44F40
-7969A4E4913A41E0E5D00937F07F1796CD2DDD2D5C1771AE65BD23115B635CCA24F0324D2051
-B864C10B23A9109104E79A6B3F8EBA781FE514ACAA45050355791DCA56901B2AF1E2D2F08373
-4765B4A7C1599A1CE2D64A7496A5CDA0B9A4EF903C512E21B65E57F7FD31670DB7DE11F671C7
-03E152AD7A73E635DAE6E0FA6045B36990055C9066C8DD999DD8CE4DFD059826E6C202CCF8CE
-1D2C0E2260441BA9F9EA847F34C0D5B8D15B662F6601EB5FA5A5004F5021CE4AC285BBD74AB8
-90E2C08CBD405DAC4DF27565DFB4877F6B04BCEAB38CBF8169BAD84983FDD08231935FE86D75
-60928016FB7FC8595E24CF57C80C6926FE17DE0E42FFBBC84153C49011E07C0A3A28A20FF216
-27C04190C455A6D8ACF89B031D4D93C57FA8115FB1295A6BAD8CB03AFF8E25B04CA7A1C1E59B
-6FECFD35A78F00570A5A8B87551D9674A14553ED82D502AB8B61A0C3DA780E5B93973D685BA3
-CF2C52C5A927A281EE5BB30287F2C47A5F2D243FF41254AFA16C9771AD70CA1BC151470249EF
-1EC96667BC0D457794177281351C1E11F857A75F55EAF79BD18C296470AB3266220742A57F3F
-0235C268F84552156691A5FCA685740D1C81D26870FFBABF0D6CBBCB53DD148364A68191A3AC
-A5A07905D8B68B33CF4DDACBBD1748E094615419FA0B3E03234F6253DCBFB55E6DFAA02C54C9
-4F9444A2E27AA3AE0A0418B02D36BBBA9218380EC6AE37AEF57FE2F74246A301005BFE238EF0
-04914D457F286337C0D6B2B67D91387FAA4816A9900878DFE1D806B3B3D068A9BDEC10410D24
-6C687DA5F0F0188F48A038752920DBF96ECBC2695E1AA22683230B599F9C69BCEEE52B501EE6
-E40C81062D0DA506213C8B67B8D16B8E70DF4881DB58FFCACD173F341FAE5441401A9957E959
-2A5600A3B060D34F4462EC500A18D479B5E4A1C5D367E7E14295CF8C00F4D7930DA68B32670B
-A19932387AA1C151009B75659A2BB8CF86AD1BD60E7E210F846A18DF580D292593C506596E89
-8DD343705915389AC9C63ED33A76334396F37159903CF38EF8E7179DD64A3E5F04E769865483
-4DE6248737C33AC280DC82CEC0EA75D632F8C7CDED29B60AEBD51DDE2EA51E3003A8944FCFD4
-2BB0C5DAB49FF183F27E6E850A3AE818A56E3D6E3B8E59622D0A7A93BDDD3A91855B466265FD
-75F7AD0E1E4C8DAD51674D504E34F150FCBB532101D012E7E9342903BF522DB9CD96178A5DA6
-84A7B0CEBC0297E1659AFB54675F3A739ED4628D87EED9F323EB508B6B161B8DA5E9AB8AF38E
-66C8F394D101752C1CCEBD0924BE0EA21AE6F3AE40FD78383DABD973DF1624B2A44505C2624B
-812A476DBCC712F36E87019645F06FE6E6864EFE50F22F4FFA3A8944BFC5A1D8BD3C58D160D9
-3F1BDD3FC02858F62A5C38C7F1CF543EFED0B65D64EE041B33CBAF8C62A9EE917EAEAC56420A
-FA5716003527B050729BCD4E96766D2D1C37742EF967ADE9FF3533A9BAF92CDD1930E8CE3E7C
-7362286A279F104FF796898C199FC264099E598E9E271D9682674134C27C0E854BE6567A9064
-0961977414C2D8F2111C5C690003B9309808FA2C8076263C0BBB44E528D4FD0B3830658A4CD7
-7FD55F80F51F97BABE33CF845C992F674FDFB9132B7FDB72BB782438784F826CE5D2C8EC99F0
-E71169AC57840BC87D33A164F0246B80FC62E5F81EE132E2C8E0CC44B57C1A9D90A3AB7C4797
-575B871C72E5792A6CC486BC75876A670BA675BCE2FFCF00D9DAD7040C1D1D3EA2109AE474E8
-4664B6ABD6FBA45D3AF6F5E1EFAC688369E4F6F0A29BABA0BA5964A02E036A820070C21D8111
-34B2528EB31BD6CCA1D507F23A6A8B2944E86820683EC6BED9D783DBA6F7FE3B96E947F8D500
-92E4D02C78D33CD5E2CBF39D9460B779889AD1CCD23344ECD89A40D26E027A797D58215F13A0
-5A2F884C6198749B9A8ED85C8070606EE09CA143EBCE206554FBC2C97139B3E08F302761C96D
-13B012F7D11F49A6F567F318C657408209C7C756CB899E2CE24175048DEEDCD6F547ABEE87E9
-80CB04CAD870345526E3E98734C523F2C11CDA86D334FBF256CEE166478FFE08B284739B1F83
-853C60C860BF8507C3D57D22FB67BDDC502B41137E366E9744BF23921D8B142956ECB9E8E1D0
-5A225EF57187065B5D5E5DC87783C120AE01041D27636ECA28406DE2BBC37D1C4EF4939D69AB
-09CF2064262A479D68DDE9EE6EE319493B81F29DBA8BD04A8D2F9E8162E101B22147AE49A122
-E926F13BC2517917E92CEDD953324E4B6B315174F8432295837080BAE3C5FF4E2523BA4EB941
-2F5E81B959B36DF64CBE117872A822E619BD601A438D787F29C007449AAC181A0DBA8ADE3108
-6EFA516573470F09CA4250FD74A12E5AFB2CB6ED26EE8F79E3467E602C2CF625912550F60F34
-210C47E4E8A91650D68E38AB15C6AB7BA0A751C2059617BED4916C7AC528FC7B000B3C8A2AA6
-DF0F3B1F50F19F225900CD4C1835670409D18707A71BE92C3FFDD6C76839413063F00D3E1140
-51802B180254786DEE98D97987EA7B747AC5AA18A9412C21BAFDF5241ABD632E06E54C2A2CE1
-713C628F22E3F901E0111DE4932D9F7B9794AEA516E4B2425E411AF7AF447F74A1ADC66060A0
-8C6EBC13D6350B92A503B287F372D1DFE516318EB1E10E53210C834F28B638A72DD4522B5BCF
-5138CE0A2310A5E2232AB5EE30B6C4781C65B987DCE4A05B6D4E4BD0D1AEAB431487C673FA2C
-95628B91EAE7434B8CD946F4B960154DC50F22E73CFF8775D252F40065213896F2139E7EB485
-9B3898007B77045D4C9216FE95257F9AED69695AD92E559261EE7BA6EA2107631240071FC538
-A749C450DDF7DAAC3C4BBDA96D8EF76A16FD219895FB6A097BEDD0520DE1249AABDC16AC19C4
-06C598AC5FD03435DF4D5CA0D937096E4F82027AF67A7225B919DE9609E41D27A827AF156E63
-F4AA5EA3AA82D8EAD99B567D2CB5CC4FEBA00A4518B5A045CA5A8F5F67DD7BDDCE90616943AD
-0FF8B01573E3A9B4D09BD3114BDDC4D45B755D076C7D8176573142092C0839A01217BF83E30B
-DB4EA635586C36810E1B7B473DAF1289CB1F842B78A69E9F9AA2E7DB1C1267FF3191CCC1FF86
-FEC94CBB2342B97273E8C70DE6DEC339BE73A347BC2DEB43ED44DDFE8B722001071AE544BAAA
-6C917806843984984FA68D1AC774DCC4B0CFDCD9FC9498FED9CF0CC605DA253806866651A93F
-217ED2BD05E0D08942F04A47E1A646B26C2FA847E32AFCC88D037C1C20C9AC29C167C10C2F5A
-673A7155A80618A49C6DEB8E2B6F018A4A7C4FB554C06D1BA2108A45FA839B03FBC8F0C8F1F2
-427106F72516465D15DF6E6C2C22F65B2FFF0C32C84E23B083E3489CBD7223D70E837C6C87EC
-00244A09B2C38BD981C92FD5A4193F7B655004C5A51582AE2491E419922BC82A9D70674087FA
-8948DF8420E192780073FDFE6FE220E699855828A93F0CBCBA7897802B2B4429F6716F9DE0DE
-1404001F15028EC9367B2ABB4BC1BD7D8524F719F00A9C8EECD7742050926762A767BD9FF873
-735DB93F391244CB7E10119A1A024FA076648CF3E70C9AC8EF1A81E6CB992D6891ECA008F995
-856B4CC47F25D4E1D39C626CDC14134082790FA445E3ECEF485180171E573EF8C5BFAF2DAE13
-05275BBE070C3271C8291ABCFDF5EEFB19D91D1EE8CBF84930ADC8C809B7C7E3C7DD14715425
-947405CC936865A7D234F7C8FC2B9B5512818E9636622E05BEF243CBE48E9888B3471CEA221E
-E6F386F48045D8432B170442A0704AAF6C12DB808CF27BAA98D31784A79C13AE3C46EC27798B
-7722E852C151A1FC1D66A1FED137B6E09DAC0E8BEE800498D9FA20E8CC10DD5A0035EFFEAC3D
-2476A5BD5D4CB8008AB6FF586DB951CBDC319ED4B1390FC46DE8986E5BC18BCF36F900BCA97E
-4D503C6B481B4C3EC751D4F9F32F04ED794990D5AC273E238F3BCDAF71322C0E0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-cleartomark
-
-
-
-%%EndFont
-%%BeginFont: CMR8
-%!PS-AdobeFont-1.1: CMR8 1.0
-%%CreationDate: 1991 Aug 20 16:39:40
-
-% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
-
-11 dict begin
-/FontInfo 7 dict dup begin
-/version (1.0) readonly def
-/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
-/FullName (CMR8) readonly def
-/FamilyName (Computer Modern) readonly def
-/Weight (Medium) readonly def
-/ItalicAngle 0 def
-/isFixedPitch false def
-end readonly def
-/FontName /CMR8 def
-/PaintType 0 def
-/FontType 1 def
-/FontMatrix [0.001 0 0 0.001 0 0] readonly def
-/Encoding 256 array
-0 1 255 {1 index exch /.notdef put} for
-dup 161 /Gamma put
-dup 162 /Delta put
-dup 163 /Theta put
-dup 164 /Lambda put
-dup 165 /Xi put
-dup 166 /Pi put
-dup 167 /Sigma put
-dup 168 /Upsilon put
-dup 169 /Phi put
-dup 170 /Psi put
-dup 173 /Omega put
-dup 174 /ff put
-dup 175 /fi put
-dup 176 /fl put
-dup 177 /ffi put
-dup 178 /ffl put
-dup 179 /dotlessi put
-dup 180 /dotlessj put
-dup 181 /grave put
-dup 182 /acute put
-dup 183 /caron put
-dup 184 /breve put
-dup 185 /macron put
-dup 186 /ring put
-dup 187 /cedilla put
-dup 188 /germandbls put
-dup 189 /ae put
-dup 190 /oe put
-dup 191 /oslash put
-dup 192 /AE put
-dup 193 /OE put
-dup 194 /Oslash put
-dup 195 /suppress put
-dup 196 /dieresis put
-dup 0 /Gamma put
-dup 1 /Delta put
-dup 2 /Theta put
-dup 3 /Lambda put
-dup 4 /Xi put
-dup 5 /Pi put
-dup 6 /Sigma put
-dup 7 /Upsilon put
-dup 8 /Phi put
-dup 9 /Psi put
-dup 10 /Omega put
-dup 11 /ff put
-dup 12 /fi put
-dup 13 /fl put
-dup 14 /ffi put
-dup 15 /ffl put
-dup 16 /dotlessi put
-dup 17 /dotlessj put
-dup 18 /grave put
-dup 19 /acute put
-dup 20 /caron put
-dup 21 /breve put
-dup 22 /macron put
-dup 23 /ring put
-dup 24 /cedilla put
-dup 25 /germandbls put
-dup 26 /ae put
-dup 27 /oe put
-dup 28 /oslash put
-dup 29 /AE put
-dup 30 /OE put
-dup 31 /Oslash put
-dup 32 /suppress put
-dup 33 /exclam put
-dup 34 /quotedblright put
-dup 35 /numbersign put
-dup 36 /dollar put
-dup 37 /percent put
-dup 38 /ampersand put
-dup 39 /quoteright put
-dup 40 /parenleft put
-dup 41 /parenright put
-dup 42 /asterisk put
-dup 43 /plus put
-dup 44 /comma put
-dup 45 /hyphen put
-dup 46 /period put
-dup 47 /slash put
-dup 48 /zero put
-dup 49 /one put
-dup 50 /two put
-dup 51 /three put
-dup 52 /four put
-dup 53 /five put
-dup 54 /six put
-dup 55 /seven put
-dup 56 /eight put
-dup 57 /nine put
-dup 58 /colon put
-dup 59 /semicolon put
-dup 60 /exclamdown put
-dup 61 /equal put
-dup 62 /questiondown put
-dup 63 /question put
-dup 64 /at put
-dup 65 /A put
-dup 66 /B put
-dup 67 /C put
-dup 68 /D put
-dup 69 /E put
-dup 70 /F put
-dup 71 /G put
-dup 72 /H put
-dup 73 /I put
-dup 74 /J put
-dup 75 /K put
-dup 76 /L put
-dup 77 /M put
-dup 78 /N put
-dup 79 /O put
-dup 80 /P put
-dup 81 /Q put
-dup 82 /R put
-dup 83 /S put
-dup 84 /T put
-dup 85 /U put
-dup 86 /V put
-dup 87 /W put
-dup 88 /X put
-dup 89 /Y put
-dup 90 /Z put
-dup 91 /bracketleft put
-dup 92 /quotedblleft put
-dup 93 /bracketright put
-dup 94 /circumflex put
-dup 95 /dotaccent put
-dup 96 /quoteleft put
-dup 97 /a put
-dup 98 /b put
-dup 99 /c put
-dup 100 /d put
-dup 101 /e put
-dup 102 /f put
-dup 103 /g put
-dup 104 /h put
-dup 105 /i put
-dup 106 /j put
-dup 107 /k put
-dup 108 /l put
-dup 109 /m put
-dup 110 /n put
-dup 111 /o put
-dup 112 /p put
-dup 113 /q put
-dup 114 /r put
-dup 115 /s put
-dup 116 /t put
-dup 117 /u put
-dup 118 /v put
-dup 119 /w put
-dup 120 /x put
-dup 121 /y put
-dup 122 /z put
-dup 123 /endash put
-dup 124 /emdash put
-dup 125 /hungarumlaut put
-dup 126 /tilde put
-dup 127 /dieresis put
-dup 128 /suppress put
-dup 160 /space put
-readonly def
-/FontBBox{-36 -250 1070 750}readonly def
-/UniqueID 5000791 def
-currentdict end
-currentfile eexec
-
-9B9C1569015F2C1D2BF560F4C0D52257BACDD6500ABDA5ED9835F6A016CFC8F00B6C052ED76A
-87856B50F4D80DFAEB508C97F8281F3F88B17E4D3B90C0F65EC379791AACDC162A66CBBC5BE2
-F53AAD8DE72DD113B55A022FBFEE658CB95F5BB32BA0357B5E050FDDF264A07470BEF1C52119
-B6FBD5C77EBED964AC5A2BBEC9D8B3E48AE5BB003A63D545774B922B9D5FF6B0066ECE43645A
-131879B032137D6D823385FE55F3402D557FD3B4486858B2A4B5A0CC2E1BF4E2A4A0E748483C
-3BCF5DE47CC5260A3A967CAC70A7A35B88B54315191D0423B4065C7A432987938C6BEDAD3B72
-AD63C2918B6E5A2017457E0D4EBC204B094541F345EC367AE85CA9BD24568A01D3B9F8095F74
-20E6C423C414B3DCCE6DA48DD1C89A56D078E0D0E2F262A13640A06D17E44EE3866C3471FB58
-FEDF5A3B77294517651C16BDD7267D39A54E7171752DBDE63AC19BB4B3021CE95EB5FE67390B
-09AE4D9ED4D704A67443F55DCE17ACD996C1F5E023C9E5A18CBEECC3097F23763ACB86CDD7CD
-13381AE74E48495EC7FA520539D87F8A8DCB3C826275469B6800876A457E7D1E5BE867C7B1CC
-AD69742A8C9B0AD943482BF2A4AD0AED40BAEB69A0233BAD36B4CA2D2DA7322956C70375D152
-653500B2F22D2AB6990CADDE2DA14B4917F7515E64BC3D96BF775258FC7DAE4E42A4C9B6DA8E
-DDEC4A800C8AADC8D75E48CAE52137E05C03677F5D6A82FA46D9F2FC7F56D62E5C605A1B7898
-B8D1401C2CAC1A0122A2C8A7AAE09607F2C5F29293A09B9959399283BE89051452898238B777
-DB9830FF4318A298B221C4A820C700CA964FC99E6B1D9EB0BAFC39BE9AA9DFFA2FA326B2A466
-1E67261BB3DBC85844E76F9313AB533EE16BCB478166DE5502B68E6EC1EA6110F53206CF1CFE
-A419571B0243244B32EA38B828A01E62B03BCF476A32C2341CECB8A0FA2A3F3F339315D5F3AA
-7704408F6617D02462A058589B4FEAC5BB976C6B3AE317372686B1BCA18DE30D2D182553ACE2
-5C61478314E954B6395B87FE5D785DEC41CFE9156C730B8B2EE923E921B2A9D8EA7F0DAD2483
-077F3208C38A5DF9DFD5477FFB993B6219A8B2DD4249FB87D613121F07A75FB990864A3E3000
-66B22D0D4CD2908B2DA4B491736C976494D87C69A88A3FED1D179325FB28C5D3A89E6022A98F
-20DE1C232DE3DA0FCC7FA7904994B8C9DA982F68E3A75BD92860859D8052AF677F8CF937420B
-4BAAC07CCF01C5480725C696FDE80052C1AFBC0DB44FB7C5A10E442EE06333551FB0FB9F0BB4
-433C23EA481FE64772B91987D1B17F817469FCBEEF3E36D8082CB1198073E3D0FBD172B78930
-2B1D506081C1D8A373890B13C84FB09B4577E6F89874938CF4BF506AB5C8810E8763839EA8D5
-FA2CC899F0AD4BDE6429232C8966ED67F4AD57D8DC145459E3F598633A0626483517E6925045
-5C8E83C9ABF54794C3A02550F7E9FCED32F4463361B3D8B7C1F28D2028121C685925BD3A649F
-8B2039E43EA3B9DDA56FF4256EBA0B054BEE9CBE33E5CEDF2973A9B573B8780572EA8CA3C0B5
-B39A03D4728B5EDEEF715AC38ADBA0C0400DD4EDBD169C0DECF46CB16405A028EC9B29E94D4B
-C52D0E0E60563BD088D49AD7A56C7422F2DC4906BBC9DE186E0EF59251AA0DCE7BB5678CA2CC
-E2B2DD27F7E603237996FA4DC777CA7CA6678BC856EE5DC66E1CF82F3D97E233936539A08765
-02D05F8B6F0CF666A918FCF1B52DBC189847FAA5889637B1441173908472CDDC584EC335AE26
-0088E11FE413DA79C4208DE2336D0B562CD68CFF8B7CB75B035DA226D82FE5DBB9F03502042D
-F0B606AAAF1601A19B1CF53930FA8344EFB7CC7D35E8A04898877AEB59F58FF42086AE7E3A80
-D7599FF1F15647AD44C8FBA536A3DA566768A9ACD5A663A523ADEB21B0576E92D6F1685A2DE6
-FBC2D3E0200B28AF178F336731983FD6C4525C054557BBF0EFDA0E688B8B15A8C8E764420F2A
-B1FB6661109AC32672DD8797B2FE7638C03B94545B362674563795616140F19A271D3C9DF7F6
-93ACCA1A35A689396B13D41B4A5C65FA6223C8F09205B59BA0D50A73A2C65782697501507083
-73E50B8421ADBF6595ECC1D1C0C8E0B600161E000D1AD558708545FB47F8B60B7518C323A129
-2017BA2E0C14EA031CD83682D14AF8FE19A44FB02ACCD189B18B172427D145DF07C4FEA8899B
-85655BB8634CA535628FA1C6C1F35303DDC85611655BA2888FE4653CB23D502EC2B7D76EC9F6
-ADFCFE6A07B82B3C2F7D3020C84A56D73D6635113F64948A916EF0FF2309960BB0F6BCE4D9CA
-21443C1E9819D4E2DB41A107C571C3B186DF7034B0F6FA5E91AE45FD11C293BB64ED1C8DC441
-F0B6BD2A91D3140B9B4246434F34E8FE8F848A4C9E51C6CCCB93BC45EF723BBB3F091B6A040F
-55D9C38C63AA59601EAF905C89D28C9A621673C567DFEAE4634844C91693636AD4285BCA0B60
-CEF05ACCE4B5A20BACBE48FE420042758F26CE9497DF990370138EA5C389CED42A581CF04A1C
-D71B847A1A48DEF222A7D4EA7DD130188D42A886FA41822E489A25D86C99BCFED702198C849E
-C0A5BD6D40B1BD2EA749472F510601954828DAE77E29EF2D7A421674B6FE348ABF7D12FF4807
-6BAB30C8C631CA03BD98AE86434ABF760F39E0D75CC774139690F208F5A85A19C4D32672B067
-D580AFE555ADA638C055597F05AB0DB34255FE89EED026AD1643B0609AE2F8C8D284C2605C9A
-45DADFAD6324F09C8B02D41F4D391611360236BA3D0E6AD474DCB842060A1D389A00F62A98D5
-2EA6DBEF5DE250F3CE7EF7147728EAEC412B68547DF060FAB31F28C5F6F324C4910637D67CC3
-8AF7AD414E1396B129EBC8411CE8C7062FC178392418384A5DCAC09038FCEC9E32E30EA2EB69
-BE7ACA538C6240C672940BA4862BA68C9DDFDF907BBB8FD4DF794F37A53B605F42EDA33CAD9F
-A6C5B9B7E7B69D5CD4163C4C056B7665543425D4E6F5655287E2AFB442C40CEC29BFF0343018
-538276B7403576E78E9B8257D5BE31E384F2C72C21333BDB5D929A7D546610F4D1B928D96185
-1077F9E6AAA58B08A5A9B7A0EEFBB162CD62D2BB7936CB3D1C5895CAA28A5500DA21B8B89327
-B6269D88CA99661200E9CC867F848E4456EAEE99B4C315C88AABAEC708B16C6963E93E5041EF
-454F9B3AA7A7AFCD2D6B9E0420256BED6D20D3647B5E1FB8F43A7267991743D11E3DFE0462B5
-576226F9E5D1F028A263D9D6F90CCA3481352DE314FE39408175488D9A3DBEFFD0D91420ADA2
-CA4E9ED48B41E4ED72AE33602D140A5DB8E8EF7BF934694EB8FB3A6366278D053FC4DAF68981
-3478A6D43216DBDD1949EF64442776708BF32C80FEB2FFB6D893691A078B96BB928426310003
-3BB6790CA645FC4DCC73E6F2AD5FF4619AD53A5FA26F4DBCE2E7216F29C2527EF3BBA558ED25
-B0180F87E1E4C389E6CA8865825D12464F570CCA0B0AE11B125B3510E1D9C008AE04CE7A20A6
-36370CB2ADA145AA532CC3195CE9C690B002D70DFA3D52EF9FE7393A94EF97CFB3BCCED87FDA
-1C252718DDEF7CCD33FF26BE035C2EAA9DA8FFD48B8A1760FD9557255A82C87F2FB091462634
-44E9C7D4DDC561C949F10005D81357BE0A84C8D7052830BE77E2ADFCA059562622109B048F4D
-3A3EF26126ECEFDA884AD461329F44EFCE4D5CF0759FEBF50C826A634463A6ACBCD0BA8D267F
-E9EEE70F46F0BAF4CF89FB68D86D9CF61D164F07ED96E83DEC0965DF06AC119822E1CC0E0DB3
-22EA29883476A5F37D380FEF4ACBA611CA1698E1961BC34EFEE983F10D004943DD2CBBF7F6EE
-55BD5176CC2095575FBE9E1B81FAC37B168B93DA3C738357122B4DFC6E48B1A512D05BF5823D
-EAD06301A4C0FDBB33D4AC9B7BBBB2069D6308E8F4ECA675FDDB721830E63469C35FE0CC3D18
-B9E047AE83CF481620577D971E07B093B982B57241F82B12ABE4C4B5D63ECA37C8E19267CDAC
-819E305EE0A5A14A5BC2809CFF174066C331E293DA2A65CC88FC7C5CBB7ACAB36C2AC6D99F0A
-3B46A37A8321040D46E1DB19216F519D4C86CEDF455350C43360BDFE7362093947C3806134CF
-3B62978DE363C0678A8FC2D88F91599F43726CA8F7E96664EFD377EBE1A69A30E5749FD86ED2
-671BCD29D704FAB75A2431CBB0C650AA568F358327176775356F9248B0382248F8A6EAAA533F
-EEA32136C0E0B3A60898D331FD3CCE1E9E31811633D71A0A59CDBAF8A18E385F2ECAFEE87C64
-59C5EBF10929B1F06AA6A6DA13DE4C1832B388BF0EC4463A6036182EEC8E82AB89AFAF29B92C
-B20CB4888AF2E69646440DBFFD2738403986AB9669B22E9EC3790E293A8DD51B6E07E75CC7BA
-300FD2448002E168A031A8FDB831D768FAE6F02B1FB25229300787D1DA9FC6A4063B37181FA3
-18C94D1EB0CDA058238898353C12CB4C986C6637E5E6196D0FBEB533C62394384C8C63F75666
-87DFF6A9D86B4F39A7B113534D88ADBD831DEB17BF3D53D86FB4F75D7CCBBBF2D9B7EA192B39
-7E7FF6D99D491DAA967A200AEF392EB5D78285E10C06D3D45E77B1C99A96CDD0D118CDF88DD2
-18AECB19F37ADD787EEA063D0A4E508B8CF9D5F1F41D189AEF64B6933966865B2C1B598E9130
-D479E064A85B1187BDA04845DB6520ADB5C729A0F41BB4790547066CD7FC7496662CBF94A676
-5A2F402AD44BA5CA08006862A23958503C4343FEE1578488974B9C88575005124C1927B1F228
-CAB60979EFC6A7B67C97FCF9D6A3569A58315D8DC5B1B211E441AD48896629BD82443B6F7203
-522292B4E05823CCC881ACDD0904826C3B5B197F7226FC153023C45139C6B79A688015394586
-F6DBAB40A011E3FB751A677AF0F3B0AA5143E33B7C1349B53435550A2ADD2EA4FBCE0BB0D136
-A7B9F31796B285985DA642085FF687410846F8864C319DF9AA1A8CDA45BA19DB48818B8D7D31
-BE15A3C7DB12971E1F989B508EFF04284AC3DF07EF5C9AD6DD8E71ED0EA21E06C2F8A3195C94
-E55CCE84A96A1D95746EA4DE3EA06E9901AAE28E289FCD992AEBA64992F94E084800CBB7668D
-D13342AE31836B7AEF1929D9301590A1699286B211BE4FEF0AD9025410617D0177B225481BE4
-757EBF860C9DFDF830844A7E60880C354EED9B18E3617D69E36F65FFFD07C4B88AD3FBF59895
-435174F6614EF44ECB3C1491340EFD29A3C0B73E65D58FB030AC046A0AEC3F76E57C7DC3E81A
-5DEB683C6D5960199608D42626A9A961EEF94961BB416A34321DAB1D7C934903053AEE2A22B1
-5758FA42D50D773DD65524CCBD896711EC275F5949BB8C7CAC20CFAE3EDE952167E235834AF4
-9AAD10629B4BFF06115FC91D6826FCCBA64FA7E767446D856701317C43C0520BD470AD63E135
-B05BB61A41352C0456C677D5954B4445582617BC14458924EFB580B84BCFC1072F3A967B0F75
-CD738EB760DD679F127B0E29B2BC0295C2410C3B0FE374BC162D12415DC446CDDF6B41D157DE
-FB3D890DADFBCE3AD2E85657F372977B392DD794019AA2F60BD405D8357A430F5F240F4A0B53
-0704D96942C3BEA32F673BE3DC03BA68EC3A23A6917C98376D318B1B584611520838E6DF1169
-1160F17AA6F6FCD770B24022A2126902F1B59F81D43F77EB906DB9F946CE8D0D796015389E23
-D08DBE766F6F3990D563CCDD1100362452AD57D2EE8A79F1013FE184F86D29B75FCFCD020936
-4D35395378CF3369937633177F70B418F1F5665E492D9FA3854D8DFE6CE949258B30FD709421
-44C6F71BA4EF0966A059D05EAA32CDCC48EDE88F690734BEE59C2464D7D5E77BEA22CF9A015F
-DA90E1BB0753236AD5262A8C647FF2398B05C929FBA120EBE631990F323B8970C3772DC1A514
-32E2F898F46C145C817D4CFD0DC50F167F40ACD1A0CA601CC8FC78FA6C92F745A2A83CF9B946
-95D42957CE061659A3966A9054B67A016D3534A5676C0FFABC5EAF819A03A10C84890BA479D2
-791ED120782FFC2CDCC9DBB65064F39C75711744AD58D9E653F8F650EF3FA079CB4A7AA51B14
-750E53E6751E0DE4F81A2746D01BB8F6D28D17626907E26CE5ED22EBD396C782882C1CF842CB
-C12561083BFE65888C85F913307CFD0BE1159C4B018716B2D5734B3FD9F93B1B742A95B7DA6C
-7E565DD6C338B5B934E36E5F2C75930DA96A2AA72CD8A17A12AE2FD589357FC9D237EDBEBD4A
-72FAFAEE8E0E81D1FEE25CEA488E1F953EF574A0D0A6F7B6BA600760D12E75392FE4D0744DA7
-D20C3337B691612C61C28FF9A8F163B7C119DA771705C268EAC8DBD5EBBD6011F868062C45D0
-A717DADEE306A05400F70DFC17FDEB9758FFD837B305736ADE36665DC4867F1AADE11D7F1270
-C9E3F0771DC5D1D353113966598DB9B0FE3BF774F35856DAD07FABD7FB435E56F9303551773F
-CAAFCB14728000376B51AE5A3BD7894D4E5BBAFB85FE2E310C89FDFCBE53B2A3871AF103E64B
-C3E4034CFF11873CF2A024C35A674F56BB9A940A2FCA4D9182C911D8A70E4C89D9C1F26FE550
-343301A59FF812F848D5EB5F368D21AFA35AF0248C770AF8FEC0D35B73574C09115246BB8BD9
-2E2626B1B5D26FA1F279D34A74A8BF225961BA088DAC3E0D1F6CDCE0A31CBC3A71698A0AC2BB
-E5FEBF615CF83F45A6F489B8C8D2180489B1517208388EE4247AC1200AD7BE2CBD217A2009DE
-01BF8475736154BAA12DD2A13BB79CF23F7F0685920E86045D332212CD6BC0DE24831BCB869C
-C957F05A9B34E8E30D69B2F1D8DC1F6F715D9A72704900704CA0ACA8A2F1E69E5D341D37548D
-7F9781EDF1F835E6D74EC6BD4F51622EE742F7AEB1E038B0ECD8D044E9286CEF378B426F2D75
-83B59B727209734133BD2966B2DFAE60E2FA75ABEEF01526AED91E9B13BACEEDF746DD1DFB23
-F9BC9E8F142E4A64CF17718492F4820F9F7A6D8D6DD4B70D2FB9FEC0B4004D2EEF8687D87747
-9C6C76C6593EC9C6C3F13CDC1C9945D957F1AD96970099A8CDE95FE2AE87703EDE55FD3D1916
-AE3CCBA58B28CE8A17E636F26DF4A9F2D97BD968737976636F96629518F48E902E3346F2CE91
-066B135F00F823D1897EFBD62280E88CD7D94CB39D7BE7549D9F12ABAF4E05D40EBA5467248E
-D9F5B7F24787B3E4584F21F84D20B1B061693E1CCB47D27C2B9CB910B441812A0FD068D8CEB1
-3B74D7648D34DF01EF3A159EDE49D522456579F668D121A0A1FA4F905C7823517DB38B709430
-147B31ADC4082BFC9BB3E8B8E255B353A7568134AD4D2B7558FB6805B4A690AD84106ABC9186
-97F0A29FF273650E25151749F1D2DAB75485877D6A2431574FDB0DA2BBCD9AEAF2B53C9C44D0
-C61FD984149E9738DF8DAE42007777B6B11E61FF7FB9ED8425778F652CF780A76721336BFC84
-294F73402CDF3415564DD1A625F7E07768824B0A61C4468098A09EC610E49F3F386D1A1085FB
-740EC34E05775F4E15C04EC78DC7BA69AAA9B9623FBF9D2AD2C762301F2B9A38A01CEDDFD6EA
-DFF04905EE06FD862D4BB76C84A874372D1B8E05F2673E2A2E6D2E340DA24595BAAF998D8B8E
-28B056F8AAE93AC3DD30B2D4A55F9F29DA5A8E380A5BB827CEC521064935BD22DA2DC60C60A5
-5A29D624349660B8A79DFFAD28B795670F246B9E79A46E26F6C4AD633388E9F9B74E15389931
-D23DD27F8899A854D1A208C816EA1378EEF8D01571FD8D84271A70F857CACBD22519F41E180F
-49EF4FBD9DC7AC7AFA85D3A2BAC7E5E4A6077D97C20DBD9950EDC9373159279B63F33C3EC041
-434818912F952C872467FE30396DAD1CCB294E114823EE85003E19E7E50D6E4614F94732E794
-7269A596164E75BAF289DF397FD82925B8618DBFD4B50F35E49DD3AAE8FBACC3F841290D94AA
-1CE3BCFA22AA2F5A9C11FAC6F1C6D006E6E4E1DADCB73A0935737F8097F6289EFF0E71E40775
-7518D0BDE2A3226A0290B1A1C270467FDF89F75DE64635E527718F69D35B5A3D17EC6845FAB4
-0F3074720A79EC69C9EB160B07F12082D129732C794316057A215FED554DA3A6B5DE264E3DB6
-C65D576F1E796CF4C9D1778DF1A19BF1F931970CD4215494CE07D004908CC921186630A59238
-42BA7780E958D2086D4357F039F227BCE9DF0D5B3216487611088BF0E4D1968D8122A230331E
-EF553094E109BB8BD94FBEA7AE05970EB8329E6502C0003DE6D3CC8007CA714C0CB5ED0802CD
-D92B2152C7A1E112C219EF1BD832FF6CB055444A1267CF4CF5342B350F1B9CEB81C8B36028DD
-0FFB85A62B308D61E5E814778D2795A78E17AEF05C7DEC1B285246130EA22A5C66501E135F9F
-9C75E4AFE4DA46F25F503EF477360D98AF220740CEFF7BA8CFB269E952431C38F12788F3603A
-158966DC2A702B6BB81553FA29410D2CFF0D7B2D04965922097B9A29F02DA5845E957785EAE2
-782A960C234E7DD814238DCC2BBEDCA1C361180E7DD186382CAD3C37F2E98A0EE597EEAB1A92
-6432CD5349AE271A44BBB1A4DBBB8FA487CF5205B52E97D35B18C3674F63FD37F2C8D9391DAB
-65C2F5632F329C5E2EBB8F330A173809C3403D4D9D9F46B0B1A35C11F22914BBAF54C63B1922
-FC916CA1547AD3CF9D14C8C4D3D0C491444D9BD4CC261B4476C9AA053C5F292563107F9A1B9E
-684094D23EE400AA8DBBAE8B4A55458A7B64CE424B65406B5BCFBEF87CFC63017B2B24EB9DB5
-7CB7C5F34CA5C015BD7CE1233293A636E9C3780A1841EABCAF0E95B27DD58A361F285E808459
-9B3E9D45A6EC744978D3522741F405DC393C97290237F2A5C57E6654BACE9DEB1E188B630A1C
-3D6EE53C78432E868F65DFB8A98D7E70B5F47372A5C710044C6DC7A97544394E0C97022B0428
-B461A252B4092261E149675720FA054AB1381DBD90DD5D66DF4CBEB601446ACDACB7EF081C0D
-2BCB78EAFF5F4AA0E35503E6700E900D60E937B69E6A5D5387EF7FE0F9EFA8BF391DF347E266
-EC9AA0B53290D4379F7FD8E85B70D75D503092A1A1F1CF9AC27F373651D377EAD105582341D9
-959FB8217422A0DC51D625B52FC86DCC71DBD87DBC9FA19A28ADAEA69EB8DADA4C634BDC836D
-554A2DB4055B87CAAF748A86E46F4002E568DFE9EFEF5E458CEB3EEE38DB8CA8C2504C24298C
-AE7828E111FA88DA2EB0C0B89F4C33E19DE07FF6BE6B1DC3BB3E741B7962CB6D972CE5AC22D7
-385716AA3512FBC54562DAF8AD1A31DF0F27D69239BE9E69E631EEB385F58FE6415A79942F2E
-21DE6C234DB196B0CFCF402399F8FD23C927937AF9C04D212AFF84EF854C83F770DF33DD38CA
-6CC333899E53591651D529313A21E6160C514594949EB2EED1464CE8F94072E0A30DA6FF947A
-EAEC653D108B1562412C5628FC0343BD4C20660942A52C318C05FED691ECAA7A357E744BD74D
-1D63A624B78001B8FCCF25EF1D7D66AC77DF64FC897E18A2E09709BACDD2B9F7D28207A0DAAC
-E692F7DCFFB736F2B221EF0615C229FD7E2BFEB15292DB89A4AD75646C18208ADA06B8D1D460
-5E27134F7B8B3F4213AD94F2B859B76E27F956A2DE59134758E2311FBF1263E06557D8E60132
-72A83DFCFF3418AF0CC11EAE26A930A817D8F4A070E33E2113D628E1B60F0D3745899CAEB3DA
-47849DA317FD6E3E41B6FF083122A078BD7CD9B5E734DA2C5324F40E591DB3E6C7A4080D5867
-CB165B1F297165DD0D7E8298FD0DC49A875E5BA808864F77085D8BFF2FC7B0EB24FF76924883
-31DD00A9D559D335404F959FBF965EC7465D8D88DBBA774AB614F34C5539793AD285127BF561
-4B62B5A267FB2F0B19B2C321513819F4143CF8209DC46D7FA41A82A0FE630699C8BC8EC002D1
-9BCF20D7947C4160BA2477FDA33DCB3B7845110AE81FBC5D9E9E41680417379ED8B00E745963
-0B41FE131B6ED4CC23BEA4844D36B8D9C69812CDB5DF872374B74DB41BA549A35014FCAE7D84
-BD213099EBC8ACE36DCA75FB4528683BECC6C76884F7E9D1B2733404F716A190AFC5918C17E3
-9FF94A2E20786B0BA6295349D32F897D37B8BE8969C9E6082C7E4E7F0D362DFBB4CDEC1313D7
-A0EF7AFD8E5C836149D05CD9EE36D6B55E7B1AA0B4A5B717468C55945171C6EE16DB2690BC4E
-AAE05AD3A83306BD13E83E96CC47032D0B26904C2B0F1D8B46622A1AD31FB6A98E7699AB02D9
-B550AFC66BB053A4F187595E4EDB59204B9BAF60BB36302A062399AF3D3908A2BE4417DD61A9
-1F7EB5B46016826C56D89CE458F489429D18066EE241A5E90177BFC9B3FADC31EA7530FBDBB4
-736B2F50F1293FD0DD6A16FE2CE8502ED811F6A79A29A0018983372687691540272B9000CF3A
-F955A616E28846EAD7A881C48421FBA911FDC00988812C2A464E9CF4B39DB276F6E27674D386
-A2E1FD9F84310C65208B05997EC8FBD5ACCF3180D4825B966665511E3C3C4960DCCC0007087D
-D34D12E1E2A17F1FA728CDD86604D1B54CEC9C1456A367B38A860C737841FB8B9398B45BF92A
-ACE74CCE996454E25F60636E8963377B6225B75874754E98BC275A4123CD6A7FECF22ED29577
-716B77C5D02005BF12BA9F4FEBA9B9496CD8EC6BCF37D7EB97488E6D2527106A39678AFAB701
-2B1D09289CE5A8635615354043DF07FA012AB15AE99395E5B257EB0AF5DA3FCE40D6025B5B69
-7FBF8608A2AB29485B4B51DEBCB31E7AEE991A2AD7B3A6AA5C021B7771AF882D42E0635DAC51
-296E7396656BEC79C3BBF2F6E4A9A42B6D39DBF547583D2DD071AE814242CFA11EC3FE049D16
-737E95C7029E3F0DB1B55C43FBA4974BA2C1E6AEDBCD3FD27181B64B70719D45229907BBE323
-74BD80390D43AA4754D1773B6FB37986D4681311052A830F5FC5F2714E5B37E166D4CE6FECCC
-65E808786F1F21009F1F644B5A6C83CFC1E07FC5E354281D4D5E8AAF0614314F32F253077FCE
-348A27A9E77084653E214C67F2E84CAF157A86953CB2D3F779C98D00CEBF27518102FDB063FF
-80781BE5C9E04A7EAD723CFEA7EB295C6B19C5CF60B0A7F35F5D0C45659FBF21340178E8C8EF
-19BAF6A832154A6A75EBD98548525F60645FA5F8F055BABC2B763BA69BF8A25B7A89815607EC
-86E8D2AFAF7E9DEF9C6AD0D6228472B51D8A1F79BCD143697D87281F935A95AC2CC68264618D
-0BF20486C763532937C725A18718CBA641239B0CAC338957CF485BA46680E1D955F4CC300BE9
-96BE9DF331E7765198B703DD0FD3E0AB3098E0AE259F5F4B771D71BA9040623EB910AB23DB1B
-E4CF53640775EAF2D2EFCAD7D40096D9E61E0AAF082275329C91BDFC15D13509B7353B4BBA05
-BE9D6DBB6AD813030499EA6FFCEFB65F751EAEB57261E1CFE1596781A77794078272CB4B4503
-FB7EC38E1FF1C482593EACB002393984F97FDED6A6502642F2A2BA5A651C1156B2FA102AE383
-CC75131BBFC8630BF69117DDEBCEB09718F6707A4648C27E8F296C5854DED70211D676A496FF
-8B42D90A198DAB42111C546EEFDF8A5E80A958A82A55D00F281799DB1A4AFB4E056DCC67B04C
-3119B135D7D8509D24BC3F750188940602791C7B80F059724779B1D8F4E633B7AD36323BC7AA
-EC8AEDB88D8436B394B70421237B0668EBF19E1A3BEE48C68B6DA8D8AF2CACFC37F7260B9631
-BFD9C4D8B04204E318655F00ECE95185BA2A70F41A0B9FE02405F69B911EC26AAEEFBA08336A
-CB748C4F2EF98C89BC5A659EEBDFDE0F533EAC1102559040DF07DBBEE4CA7204C5C6CC0C4A2F
-A4A0459F2D56B02DB6BF4720666E63302326FA5821C4B383E374048FBB9BEFD871411FA4EF55
-4C66B1D84FE45856B6A3CA29F0807B2DF2577BB74E34691C3BFBAB37CD02171E7415D99B99A6
-7EC855E171BC3FD37C0ADF21E95CC8EC5CC2FFD02397D813CBE09DC77CA2517926B07C6272E7
-B13AA52E51D9716406002FCD87D8B1394CC395FDDCDF516F2217D9245070A0A66FE3FB3627E5
-79E69BAE9BE1C4F05B5C429FFFDC57528B367680D1326AFE7E2AC90664D8B00AB46015B889B7
-AFE99663516268E8CCC7D73CD6CAAD3048C26E17B2CFF9D2A90BD621E725307E5F806809DC30
-BFA0512A67E5FF9BA9D89919F42FAE667C4942FF499FD963AA31F6EC6BD13DD1A5CFCBBAD4D9
-0503407A8A783199C814C5D64D494987EBFECFA6D872771B17CA4B86620AB7C6F7F3140907C7
-751F658AC88EEFE6A11DF86DC2763179F43F0E70C9A2A0AF34D410B93726E0D8D0AA51634529
-30A655B18F0762AAE63B6118FB8F15402CB2508B690680C28D717C9ED2F1574C1EE736B6EA4C
-F3AC3DAEE3800C9DC2E964239050EA9F798C4C221E6669545789EFDCCA1C99681E5F46E080ED
-1FF29EBFDADC2A2E052016B05403ED338E30BD2E75D1D0EC8E69A8C5662BF5BF606608CFD000
-9EBA439299A1484F4A1B83CE56F771ACCD25F6CDD53F3AD0B10347A5D2807461EA476F99BE42
-E7D8ADB4187ACF81D37FD12C5043E89B434D608110B93F62B0DA92CA15C69DE0B2AF8134AB03
-4E0E2212056C6CD9F4A8BBE2CCE2739652F2B10B0E3C6823025D34024EE14275767943F2535B
-DAFD2A48354A3F857B47BE5C261784C9A28CE6D2C8CCDB11F5A1C645D414A66D42C04E44A362
-D5348C965AD2821A367B0B4409960AD6AD95675E901047326CCEFD6D5A6ACE1EC38431D5BFBD
-DA51DEC970E5C9D62D13BDC1DB2D21193023F62820925FAC9D18ED3B70D31D1E6DE644B82589
-F566A5E0BD4BBFE8D4458C093BAEF91341B0E5EAC448E117651D951D936E0AAC5C95DF94E2DE
-23DA401F456B8B87B06728EEF20B1C1241518B453E1E7A98E0D9DF0D80C13574DF581F1BAD7B
-DDB6FD8EEBFDD0745D9FDBB7635D52FB980FDEC50730E26E386FD3856D5F523314204C59DE51
-78A825925B380AE0CC15BE8F7075D91D3B25BF932966818EB96C8BA078B77046011366A664EC
-94F0608C70B7BF5E5064A463941ED70596872027BB9B52D30708571B8B7FEAB86BB78BFB9086
-03590951E7A2184F4CA4ECE102B0DB283EFAC4E895F50A3ABC74F709AAF6094E4DDCBE25035F
-E642D6541F19692A6DE2647226E19D02F03A070ACCEC0283F2B948831978C199B473A989C8FB
-56E90CD76DE2D620B0E079B8BF7DB49DB5259F3EB8657D44DCCC4413AE9D8586529297810B4D
-6AB61770B6A4794953700D3AAB30D5632103B32C8561216EA7870B7082AC23459AE125BAF8C2
-722250A616AB1DECDA20933CC23EDE0CBBF63939532654D72A03EEF34423A929A05B81FE5216
-A12CA1D72ACE7F607F1731B9D77E0AB6C939EB454D3B08035F3195E1C30DE06A113D7F59467B
-AD2AF5EAB05E7C18D72C230009ABD98945D14B98A01027C2A77042FE44B194763744AF4C0EF4
-0E56ACBCF0EBE17B89049235F883195F23FD2C99E9B73910EC6D74DAB321F73CB9A0D30E8222
-CE4E55CAC430919845B8E5372DE9C78C40DFCDC2672961664FDD7E7ED1C3431FA2CAE688C839
-979088A63096AD8FD317F72A3227D980217FCC7A5766D24CB5674431CA72603EF8974ED7712D
-72B27D8D9A09A2376A3C4D8FCEEA7DA8530640A2800B59F0940692D6CCE311131E2FAFB22CC1
-D78BC29A248E6810354BDA8FE3C5E33950E28C745C931A18631FC2F0BBA2D5346EFCEE2031D2
-69126639975EC36876B725A8DF106C67D8A1E01CC199D73D87A5D561DE5316420CF79832D9D1
-BE05460F40E4471FE9763B6AFBD4B279B3C5907D1D3E5352DE935E83F9730104026157E6E7C5
-79F579B122B34E9C0196DAA6D8069EBDB5417DDEA69A78F67A29D874286F476BF3538B5090E0
-0B2B478685D854880221550E5DC8362030C2E4D9296D27E437237E77C8479FCED1DDA6B6BEFC
-FF3DCEAB0FF6C3CA5182060530E7BC753CD6C253F4B0D0600637D83CF924F46B89102C528D7F
-E2932F29DE0FB540DA93F7942E570CC68D47395D13B05305FDA1947C4DFC51CA1C85D46ACEF5
-1551A2BF333F323122662404841955B4536E307F6617B5AED7C6D2F0485B6FE732A3E1C8BACD
-1A461C8CC6464A2792080D3465983FE4CB906142A4F27081009A0A05CF7C920F9B658B42F6D4
-B4D598833EADEFF61832DA30F823E66DBAF5C587B5BEE8072AF2A6C2078EA4C61AFFFD11F60E
-37FE2EDC0C82E25AAE539DE1D2C7AED6B1F646E55174E0AA5C0869ECA95E8143E6AE9CEB2F54
-5E4BB021C67DADC3F069BD74E3BEFCF9E922EAD3B95BDC4890A035CB7B0A3EC2A6F1EE4749A7
-357D42AF30E426453864A9358D57D1B627D9684186E8021001DBCF1D05912A3C9F91779A701F
-D1F25461062579B8F1DA485A977269B43E85A0F884CCB5B0313866061733F8B5C7CB2750223F
-36697D98FB9156F4CA6DC3C77500F5E2BD24EE972397DC3B44C7EFD7D4CE7E2A26A104B41045
-938190C624A8CE124532EF0D8C30D983CAFFBD5E0A3F2785D70AB7A3E2877D674B0E42193CB2
-B7927714CC27E9B41C051BF095BB16CDFD27A1E1FBBE1EC55ECBC27F8B6547A990DCA70C3434
-D90EE4EEE407C1EC4A70E5F338FB33AECDB331C819A44C73E1F2E32D985973E57B200D251509
-465CF9345695B1FF889E8D4524386D9E4BE7865B20204A851C3F452635FE96A5EE175D9405BB
-C5B890769BBBFEC3AAB9F676C39ABD15B165195EED75F1A1BC41E208AA717559F64C40B91C42
-460A7D2883510F1FFE7536369C4C8570CDC09086CD94B95648A7E368620517DE82E79821695C
-9E68FFBF23D2DEBEE114011DF080D8C2D071F1AA93D8FF3BC575D64DEF447C7C4715DD7A7C92
-8E38C162285E0C1B9C14F8DDB5B179A3D80C50085E9BB5CF1AC9F9FEB4CA63DE673838A49BB2
-67D19E974FD2CB6268D0D00347BC58E604BBAB1D3FB2D5038CA82018171F0BC45A8C76F59BE6
-3E04FBFCBB239107B8D847B7673E0368F5F930A5491F4467A24FFDD75531B00873D18BE85710
-8B6B92C46EBB11C3058CFA7F81928C9C689DA541444B2FDEF7D09591924314A48867C5F95314
-CFCF0BEC2CFF236FA4AD8EE80F6C5CC9D11B02ADCEA4D93FE45F08BAAD35DD1EDE63644E8630
-4948114B58AB5B1CB3A9AA1A971D74D827149161CC451C595087D45E0012AA3B84568D564C1E
-C251A566CA3B1CE29A79F1563C99253DC2010A8D93066E550EED86058A3080DEED8695D35271
-4DFF2F49C1EE50B70E2D0361CBC7B87108014E4E24949D8F7B0A8E6D34C7C7968EDA07B075A5
-E06DCB1195D69F3707289DE014E51A8C59294D8197C71C08977CFC0D3918AFD70D497C4034C2
-9B282FC91A86257F57F9FBED017AEF98B7503EFD40DCD0AF6EE0F9F6F9669CBDB76D6B348AA3
-30D794D69FF3E5B17C64D848725A496DA4AFCF47632D2B2A7E113B3534C1F8874DAD5F9652E3
-5E1B55C8F18013BD481641455A3BFF6107561E602B41927E2C9140E9CE1AA888E35B53A2783E
-159787A0655A82F1DE9252ECFDAE0ACC6C733FE048A4CAD4A3CD0047CD663B73C5BDCF680309
-DD661A87E9B06C982E444E2403418496F62C8E953372B384E229BADB583B97010B9E38EF32AB
-19CD5F370D9F10ED445A6DDBD0896D009CB8B41E9740E272FCD916C9F7968A22F1C331968468
-A1DE0E023B485C6E43D2A5B6FCAA776E8EDAAFCBA4B2D03D832D4CEC193C6951BF7B7CE9AA93
-DEAD21D0708842426F9CF10431B3A1661C2A5E01ACC8A4D35FD65C9B9D6BEF4F9021C0F61900
-992DD7EC016B9EE153C2D55AB4AA455536CC7503A9086CFAA6AF30AD23AB03C7131CFC8D2F7F
-DF9F549B9B63BBD8424936E98E2756FB2A8019852F99F542DAB78B09F271BB8309B45732D3AA
-98B061D0AC4266B20DCECB8DAA74F90F9D5145E64C8DF6B5D349301AB6611788CE1381A27546
-79F4931E1E57DE6E723AF26553C8F19362DA48B5D6BE3AF2FB5B3D8445421FAB4227D20FF107
-BC0C215F2F9BC72F121896BE0E9ADAE120A8AE0E15C58FF83438A85DF93FD234DA94EADDE7EF
-B05E40201057F9241A60D6FB6B4DE9A49135EB7B11FFD0E357DCF9271A101F32859270630FB2
-2386A91345FC510D17E3B1EFDFAEC388A6FB39814B90DCDEE9530E06864A3C2C791A2718FA6A
-5900A4520BDE40B44C0DD8E62BAF7150918A4D885BF2F7FF570D6CF9D50F813095A6E0A76B1D
-4BB6649C0F7B73FA28B13EF83442B812A5D809AF1F501321E9615D222BB8BE89505DE76D6528
-03335F8A466D27EE52D0090206B54244C020E4E1D104E5CA339A4D7F38CD112FA42B35394F03
-62397A77059E7521B3A8D5CDBD50BB3726839CA17D95125EA9EFD1AEC18B7CF38CDA15C28165
-E4D10FD8D997B16B57A9D68111C25B38F149704E15D285B6B3EFDF9A1375334C13D63BF50AB3
-6993D7451CBFDE7655BCA3364E6C2E054F19F4FA23CCB64F22F5EA69562F03E77CA265B5FB94
-C7A8BBA6A414826A05EDC6C5588BC3080661B6C2B041FE0F0CF45C66931AF69DFA5462CA0F01
-413B8D4CEE60A06FAB03EAF16A8DE1EDCC8BFC0A9D5ACD749300B466268F7FE8BACDBEA08D74
-98D371852A6CB40964F2C678098DA227DFD914EF2F4587EC2199F3A441FF96916A03DB3DB362
-CA99BD6A8B2FE189A444E8991BCFB9AA7BF3CD0ED8F55FF0427EAD09FBF6C9B9864B6847D166
-675085C09295FAAC2361D7270E4D7D202DDF89506774265C744C3C74CFDEC75F23851F1D9961
-9DF897F2E1416D88BB73C4DE4D7AF4609A33A24032FE7BDD64716804F5FB7F5E6D943D67A72C
-2D3380CD11DAFDDCCC52D1577E3F42801A1C2BC852989A985F31F123428D62E4E422077E67B2
-9720C7BEAA0CD2D6A1B3296C0011CA1DA764DB0DD6038DCAB552C4764DF3349C35BCAC50A5C3
-ED58F5EBF81215127929A168E287843FB068542F4A8E0DA821922E90176F3F200FFCA7BFCD2B
-CFFD1AC611D3021F2BE43D4200CC0080395F4CFE815161B1EFE4BC86E012018B1AACA3EC178E
-41B1324BEEF95F01AE5C7FCBEDD799A0F0673D7988F3098CCEFDD4131B95FE216F37621967CE
-EED39CFCD0E8ACFF3ED052E10B22EE910304B2C60CA8F1BC2347B886C15D222713EDB2CD0774
-8B49AA2A75640AA731A00E9ACF344EBF88EE94FA673A3E26ACBCD30A0F7FE98878DBE4ACD585
-16484276D8880EA5335884609E19F843910444D43BB77C64D4F2FDA0C42E71CBDA70D0C49086
-A5487E31A9B8EC5399F6D4087EBCBBFE631D6D57583055BDB9A90506A9D7D5F9E298B6989CF1
-51227C46A2E0C36AB1E2C91B5130102640B4A41C32D33E7B4D8381525F23CFB27AB236BDFE37
-DD2199E0CC951CC6F0BF9CCA0F42AB62242F78A89D9D303837E49D713C28FCE2618573A33C27
-2E5E77B8CA385950A24A3ACD3B178730DBD9615868ADAA5838C42A7DB016B81E092827060CF2
-3F330FDEDEC7C0AF990F6A3283C294BF824D14CB4FF79A31895A5F35A1B81F90C4465D142FC1
-BE27598E49F1CE9185A1F46793FCA94A7D7760C99C31E8D33D640CAB15AEBFDD9A5BD8A520B5
-D61CF9F893A0E5EA62D0A459F977559614600E65FCF5E06E9FC93304AF3E4F41D8021E625AED
-86C7762894AB6B065F9897624B044C60F194D68C01C7B958D493DDAB54019D0E52AEDF25CAEF
-50F2863406B9B04FED44611DA947FFEA887860182454B4C229316656A7C4639C475CC7126435
-19CEE03691104AF60A6F5530029585BF62255E7F19B52A5035BF6E70A02E3996F22CBA14177F
-E5AC5FE4B6C31B8B3EB1A3A5DF45CD6B300B028FBB2F9F7F790B42B37AE8F168CD4302F88D16
-FA16ACE7522A50DB92B20A8DE84038BF447E336D00C1BAAF2EBB4348E2ADAA12093241EEFDE1
-D7BD8FD415D8449E104CEBD9D9933CD1440C897509632C2787D06D279E08D81D924E3A47337E
-B29207BA9EBAF4AF52E137F33F2F40ACD7F932AEB8BEC328BD9FC7E20EE4C499522B933F8907
-0724BC735179910E0702472BA32F165710AF72F34FE861822409D7A6E1C2E7A040D08CE4D790
-E129072EAD4CDBD5C20C1E4D789FF5BB74AED8C721B5A0798D3F879FBC74E73922D5C031F295
-98E9C26A67A3261F20C7B5633AB76DC3CDE4535A78D31B615906B326B364DDEFCA3D1279C980
-DF1B607327B08EEA8172B879F7B974512A158E9BE93C925EE882D130A30FA3BED0184037C0CA
-BEE96554D1F4627F5BA4F9740BFAD28252473B2EE7B1CDB35CEFB2A7167159AFBDD8F5C1E110
-67112AE4FA45BFCA60FB2138FBD530FF80EC723A7B91EC6399D0A09C010159E0F990BCD6F0E8
-A87CA402859E9161BEB1174067EE22264222442524CEB6ADD108CD380773FE7CF53210F311CA
-C05873A86C384C297D9D7D0D0B9263623B1C96A0D67ED33914BAAD1D5F496B2E335F92C824DD
-EBACD11BFC8767F8DDF330E82DF1EF978A85E22E6F4F69D7C01B87B2DB876C31406D8ACCEEB5
-3355ED66693E71C98B7CA74E8FC4A3D4ED5D7A5A7A0B3A48DA7CAFE8491DCB1DA2B28DECC4F1
-DF70CB55FDAA7BE6BAF4C9F656865E84744F437310FF68785C86BC076D325FDAF143E29D71E0
-1B1A217CC1EF03BCF42557B293B044AABC8E526E7369EEE789C32559744FE85F3D1DD0B6F243
-61C7E98E46866F32AA576C174642DC4BC733CACA810F6BD48F43BB74C5F39B34881B3872DB8B
-2BC3E56A51CEFC665D79E810CD26C7DAFB8AA333C68951DB0AAED85B6428478CE4AFFDBFB73E
-4A534E177DC5B3386DEBD2D1779DDD361978F8A66B0C2415DA5D7110C9F84212D203B62185A5
-2934D48F9167D79EC5383C53164A4E304A9B3FE6E678C632A04E63246C212A5AF2D954E32C09
-8AF13AE7B6E24B1C4F1A6E75853522C53BF7B635D1E4722AEA1C8DA241BE48314B9C1C2651F1
-936AABEBC986482B997A672F9EBD08F1363562EC76412C17A16356EC2F8B9B9AFBC42ED04AD1
-E64A9284D84932E557D8F45C9851373A130467B8F184419BA1CD59D0321B78022736E66B6DAC
-E7F2DB0464B1608505161C8043F9DC14B8B0222D65BD266FE8548024E49C6DB0495FC7B3C53D
-A30B45CE76FF49567761F2A2BFBAE2476533A1D06E84A52A72F4D53B8A5A399B8A0F4E7094D5
-8DED07D9E3AF6FC7306F311C1B4CCB576874655EDDAA259E8B52B4B16A6D142C8A2719EC7284
-DD98624D745BC2BC6155178B97BF9097169239A61A2329E8A0EC84E5F87035F550B28CB69D02
-468B6F2E8F32CE0EB7939275BAEA470A92586473BC2584E8F70F1F42D9CB72B3F4A303A81FC5
-36C70D3280C356832ED101437ACA7FBA218C1E93A63D54C16644E75BE0F59FC03636A8400039
-492831766AED50AEDC061E2100D044229534C6FF8B9D8EE685547D511B8D6D3CB743AE02FF35
-F3D095A4F25CE66ADA8B2F574E2CDE7DDE3CD44883C7E2932AAF31B9FE3B06F331DFA2CD3A73
-958ADFD29EFE32721123DE67B0793056EE05461CEAAB6678FF32DE6785691D0ACECC89EB20B8
-25E5EFA4237EE8570010A4B66D36746E6FB9A12F9D1CEAE577ED14FF9ADDBDD4380B38377489
-650F024964C68940793E2AB479716264C96A28A75983579197EB04301A47E2670B2EDED6E6AA
-DB2E3EBDBD5193580C3006EB00BDBA268D695B8CD084900093C66F65D6B4A2E5E3CB9EBA935E
-EF0A1C90E72843BE124B24AE09C4FB2F67A3B021B6DB3997F5C2C58AC1CD24265814F25FADAF
-56D4019B63EA0867F2A4A746FCA84D008F0FEF1ABD87F662DA534EF59DB37FDC715148D52C2E
-4F3D110B8A0C097D6929022F9B4F41411C38EAF4A11DC7B1FF4C5FDEB1DA1DCB2518F464EB16
-514C93E1B4BCDAB6A0227A1E2446643FA6181BDA489B726B1C4BA4E1F721CFCF0413F4487DE5
-61035C52714824FB8A991AD2BAB5FEC9E342E1BB99602C891BCD505304FDA9A8038FFBB0280F
-05E7056DA44E3F726B36DF147DC9F408A5AF2F61040B1D843FE1DA8773AFBBD1399B0EF9C97C
-E83353E62D78120B176EC5CF579E4A4F6DEB0DC296C18BE1CFB25F5F7F2F7BB8A7F3365A0FC8
-D2D4503FFE62288EA33A5D2B6AB492D587451842B6A83BC4EDAB37D90846F111DA5DEE94932F
-66D3C6F0C7D515940FE6ED5EF717836F55BE04EE14CCC33E601D5F37ADD307512FF800BC2CD7
-2601E97070FB6312700E06908518C8A3BAFDDC2DD4A45F9E57342FE3B5198B860AD1243B1AC4
-E00527C8A74397AEB1930FD08D3B0A34612E747F12736C26764163C2F2AD10157DE6A8B0EEFA
-BC40752FD2F767FC337B4ACB8015C43CF76608B18C79C2E4433DF8CD1B2CE53897F970A615F8
-3ECAE566DEA55B366FCA70F3C63CC3AE60FE8807212B0BE4C7418EBFB583A03783C24CD0D299
-43E6B02A4569E99E77818A4BCECCB066C5C66DD1AB8121EE933E1FC8042A49ECE66FF1286D77
-1D27443D566397349F7DE05CEE81EC4E3E1D803E2C7520153B541E0A1F341F8C0A1EA6E6DF0F
-F20133EFC4AD274C26C4B41451F78B01FD488453FBD0C4C8C3F11545777A3B9E59560EDDF512
-2C376F3F3CB18DC114DDA9633E99F943DF6B5F2C8E077ADDBE9F4ECEC96A5564F2C0B496B10E
-3A0C7F1FC197F0248DBF05BFBE63BAE600BD9B82FBBA9A07EF929408DE62EEEAD7032437C635
-F7216FDAE5754FAEDDBA1A8F8D9A67FE3065840A7E2D088797FB306E6C6995DCF6901E6F4F1D
-925858CE94556488DFC28F2F704D891ECBBE0D83078757AA2ACC890DC6B1C7D59BB21755A3EB
-AF07BC61F1869C947B5D0640A1662E4FC84D4EBBA1E185FD74740462D1086DC731CFFA325D20
-04AAD7BBBD764A157469F11DDA236622F2D14A57C5B0927169173FB347E7DCEE4403DD9B0D93
-FF3E82761CF241911FD2783B25DBA8F54CE7B6CBB16E12749A043F9EA634F0741B35B37A6B09
-93A763E796336FB33D35318EA517A7E9DE777E0A7015460D68ED2A2C3996D5D374FDBDB88309
-FEC4F66FB2EB2D90CADC7435F3757B10AF1B1215F25F093FB26955EA22103BCC82FC7696F112
-29B6BEB2BA21D9D415F746A6B18A42D9CC3316388220230F31354953F80786311D2940B42C43
-3DEC8F126D91F4F2302F1240D55F132CDFAB90DBDFEBA8E49FBEE19DB7EB90B81D6511E64045
-955899B94D6EBE62F7DC06A47C531E52D396DC967B9C39E20D262ED3DC817258202B60F9F33E
-EA3A37783C74B088107A9B263AC63232D4899C4E5718622B40FAD50EEC80E10ACBFD8D007D2C
-41481D91ADB1DE1D17040160BF26BFC090270D539A040B9A8D3EA4A96341721DC1B86F592CCD
-3BC1C42672852FF6A9CB6C14CC49B244B5DAE3AF384BD79DF85E8DA20139964D231183262CB5
-CFAE5FBE67866AD6F44C185064222337027440979593B346933644B9B9D20922F1BBD7DA15CA
-0910790ABEED836845E47C49D2F0F8128908EA81772FA823B830E725B015CC07AD1FDF14B7EA
-FB78B7EA9B0BFAAEE914837539AF5209C3DDD962215212E055F97BFB1AEA72A91804E87D3EDF
-94DD1A6757F1BF7123B99392274A8DFA077CC8932538F776971C8D2E098B07ED29366D94F88C
-C5B96226CEDC18DDCF91D6FD40E7EBE662045ABF928D30327AAAF4323DB228A88D983A1AE202
-ED41A4DAF1EEB2129C629446D0C0269D00B2B68CFCA3650DD0F184513EF5589D6BD041727072
-FCAD10F72647F9558D901B80D3159CC6380E43FE53BFF227A0310D1D571F24675579383C5994
-FFDE1781E76224197AED6F0D1EEB09AB34C240F3DA63513F7D20E6CC46C61040943A05D9E9C6
-3D6FBECF348144246E70F46CA71CB7D050CE254A51D0B2EFFC25BDD6617DF05FAE99B49A593D
-B19EF6110DBF01579B7AAB8F918CD8B3AFC6A52CDC1F3896257AF1EDDEEAB6E2F082901BDAA1
-17D07BF5134F6454FE4D70CD6A14F055435F6C07DAE9DCD2301B9AB692BB2C55226F10024883
-15E067FB59BF0862C27AC2CC601B54CBE0E4E36CFACE2BB339D340AAF16E71256299A648ACCE
-D93CDE98E3904872F121664B5B5DB6A353707749BEBB7ACF97A9D3775543CC34EE3B933676B4
-DCCF500B76E674B18DD7DDC11BFFF012DFFE99A9EB564350A0AEA2F90E765E9ED3E6A2D1FEEE
-D4A0ED77BA3A18808272AF43376D47B738C62B9E59D65008B00AFD4A8012D97C4E322A0BF20F
-A8647588BE734708031B3FF0673F8DF0C7873ABBC98E6C3EA75253F66CF2B2131D50280AF892
-40B5BA40A2EF0BDCDF2355B62FFB1C3026A243528F35B111C80F7C7768D471B600B4714E2336
-E76ED013C6A2C1D6ACCB359AD2F9BEE8533DD1570289BCE268D529B85DEB189185DDF400CDD0
-384CB10D14DE57BED2F1A504753FA9342412590A3AA059EA9384913831B1CE239F689F5E3778
-179194E993D7F233F2DB5FA8E416BB861A2D0773C764F9D30B1AD046266B4BC1ECBD7ABD7450
-35C1A20798967099C2D318E2D3395391E5F9D68B9A3E3975183397EB6BD76A8E313CC208CEB9
-D759A0592473B515F533238EA720DA3DC56EA09293B20F0C68A7C820783F97CBB7CF064472DC
-3F1B518307A54066C851DC3DD57EA8F327ADC8D52A6C2CB848E60DE094F66270D52BA72A989D
-898B46631C08300867B8517F2297302374076288D529D37ECF0A66DCD39C5432D9B8C200E92E
-0A0DF4E9A75A49D7485EBBF5E2416C051402F6B11E916B02BFBCBFC19C5670C5273953C8CC5D
-22364F4EFCF141736FAF79F0E2AA008EA13B52378879B91AD928247175AB5C5532D99CB887D4
-97A9397B578516BF7C087F710A4ECDF0522EA79643CAA25A77C42AD6C6FCAE5759E5875652E4
-D2A5AD9C205B814C688930865C50C73AAD74F3C1F03845D34CE3B63561A16728DA5147F1FC0A
-A8D03BF94FFECD0556C1940A3752E44FF07E04D266D317466A801F50D7BF6CEF16FA9C027DC4
-8E75EF0CD7B3B168F6B5589D5D50C5CE5C2F0916A8B2011A5EB6108806697508D0E202985F0B
-CE2CD90AE1BE517ED29A181B907DF261F87ABE40604315DC0AEF297D8F0E2847D2859F96D074
-2087864E890CF3C72440EF481F630CC7D82CC4BEA21E8D312EA256EF406420DE7FF76FF3C14E
-706CECD2970DAECAA3B96B3C0DDAAB8FE72F99DC0D8809F7AB67BD3AB51EC174567E3B7C8679
-7E6F7C46E22ABFE919A5F70279082BE6609B919DBDC50EB4BF01C1DA386B0821FAD2369EDD1C
-F5CDEB2063F9A0DC033EE15312BC11FA73ABF2CBC15AC252EF8AF0E4CCB3887F9744691EAE27
-01B71B6D179D642930DDE788A412E9B93AEFC2A1A3D08489927DCBC8EB014C07774EC976B97A
-BB64859168358FEA0E66E763025866DAC885C992EDA6F53FD455CEC494B7D76CD1EEB555EA07
-71A8607382EA5E934904B64B10258D60470A28478EC03988FC221DC323BCC12247CCE2130129
-5C1878C60B523ABBEB9506B0B88B07C77A457140D813A04CB9E0928A309657465244F338BD17
-A15CC0960C755C8EB5DF34F0A4E3A1954FDA82040FF5C484E90DA96A618C3659DE9A4320EA63
-24E33CF0C76CD5D85A8A51DBD55B016EDB1AC824A814E3E158B1699A98BD927D58834011D09E
-423F04523402781F00972C0E399B2F24676E765A812CC6193D65D353C649EC25B92AC1D87484
-436F67744A44533296A819362D5FEC42726F242A1C9D898E26949F1B0AD072EB2DBD186DC326
-A091830F9311C68AA077914269E26815AF0A65AE44D38146C8DB96401EE0275AB3CA1AB578D8
-9FFC880D51E3E2604ABE30E0D536040A81F535607C627D7555C3C5761DA9E4B55106FBDEF4A1
-9E2AF9B67E04100D0BA2933B1938251B81B95A16D70A98CE18F2FF9F8075F63261337C79D209
-F85799D051C64C3979C06ECB3C0EFF5C4B88A8C41DE0E533CCB38C700D563108BF35216A132E
-E6B68F14A6F97EC4CF915BE5F9CEDBAB9519BC61D419251A8D6490EA7566C5A2D32B59D5F2BF
-3BA354321F38E1F42A7B204AA49B4312AC01C24E668A5DEC5062FA165B08AF954403868B3072
-4DA0FA8D96E9E460676469255781EF3DCA04D973F89A224DAA2033683358C90FED5447D95F84
-A2D83357F612AAE7CE77B75E049DEF29C58804030AC9DE20FD767097D69DFA818AAEE276B5FF
-16F83D143C4560BAC8FB7EC6744D126A9F6ED0F6AF01F776AA27E9215AA76BD6A656FC11B3B9
-B325D0476944B4F0319B3B4E4AA9DE1CD69300ED2E711360468602DE6D70858F7264DE7C1E8C
-7792ABF5CA3D77242D36819079A429EB1A6F52EED138A3D0581F74948964B639CF4732EF1963
-781A27AA4D0E41F9721E760E1D3DD339CFD01C23799846911C47273277D39375D205413F13B9
-12548B747E0BC01350EF6BC24E708AAF802DB81DE0F11F5A517DA13A6521401B7EE95E13411E
-0250159427D9AF4CED49143FE214FD4C8E60A0827A5402076A70E28BDA9702A4B50A195A5F12
-4ADB1DE49B4FF904B1115F9385C0FE5FE698AF968D87DE72CC5A042CE0D77CFFF66DE6492CAE
-B1EB0956B98975AF882D19AD7BC35477008FF1FB2F16E12B00BA622EA9D36003D003B06E2303
-516E37EA86FD6CB6DF3C0C4C88C972667C8EADBB55AF7B1BB5B63616730A34EB3531D754DC6C
-761D48C91604EA8BD809F81614AAEA4FBD71EA9275F9BD0CB3247038185621429DF649EE6307
-77F35957C6013330333B22390886470EFB75DD46D095232AD56D1AB73B8B285BF0DFCD9F3759
-60717C99DCA4E85063EE10804363AD83E034F7096844DD939C8F5AB2B15D8C54E43F0C781919
-6EC3EAC50BDCF240730C15E72D20D53D54CD94834504746B7B15614D0E49FA6188685A059372
-FFFA9690D6A2585A3C78980ABCF00B7BAE97223A9788C1237DE074ED3376B3E864473AA6C9A8
-2C4DECF82832CD4DC440F5F4FC71E71DB79CF6F44A5F5BFB754CA78AC361C60469429C3F2908
-E36A20AF935E5B28122027F772687670973BAE148EFD299C362128FFBB4F20EAEB3A4016C257
-FE852D86AD98FDCE4452364FDBECA4EFDE760C1C9C1478EC68F3577CBDD741576686229D1E55
-44A4E7BA2DBC1326F298941ED25BB0BA240C6E500826F48D57F6261ACBBC3D748DA02641F1CA
-59D2AB55F3D4864CF4DDD114A7E76E0CB8206D978B8483B6E47FD0560466F003B5FF85F561F3
-01026E0DDD816EEA60B26679192A2C1B4CC3612E0546492750768F838D3764B08201B8807806
-D6511B295BCB852D888756AF2306C75EF4EAB2BC2B35C93F88E7EA1CD24B6E1C2E409A9187D0
-51C8D98DAC097869AB13ABB7A3EE8A2A883C347F1D4C297C96F193D50327D230205BB1C997F9
-AD9C11707B19B49F289B4778A411F22399E401B437394E771E2B94D73BA650D9A1FB7F7733A8
-D4E2A10FA2F63863A6692CC612C564F42D8B1DE59146E42CEC5EADD19EF97601D3079827DEB5
-44D0C743E56529CED1AF3871F854E5E1DA4FAB5D61A1810808947A3B578A277FC4C85F446AAF
-6DFDA767F2B4289F97AED8FA1B03400CC79F6BF98944AF81949C70863684F22CDC2CE61EB07E
-4D79CFAC9D088B43FBC0D833F98B03376AAC61674DDEE39A7498F2669D91C26DE5B38EB7BD3B
-108DF6A55C0B0B8B73987934657C1FE75AF99247C770D2299A2AC2F56463D0FBEC4586610932
-3E41F482826CFD0AC3C09725E88CF849A77FCD001179F02794911B9A3EC4EC492243E5DA11AE
-D2DE28AADBF3714C017A533BF9B6E80ADE6921D418FA46CA09E66291F4C50A7A60454E02B211
-6ED7A53D0FB7C1A7C8FEE48133CC47EC8C9F17C8D5D96DB0CCB07994F0269CCF5E7BA08DE086
-63004D2E79DE05DBCB67E4C6408E07AC72DBF35ED62DACF37746111188D1573034D227AAFFDE
-9812EC5F32D4B4F1123B93BE2DBBE2C5D92BC24A0AB1047861A2B5C678BA8ACB01C919028196
-FA2A9DCF7499582573875ED638744AEFEFC7906847105D94735577BB8B20C0FA7287674E32FA
-9471480A21A2DCD6E7686286C30DBA19CB0EDC8D4441097E237C17414A9FCA533E239D8A6DB2
-06877E243C149E5A956C803878F529F2AF8FFDEF2796B6F0487BDD3FECF22B8725B9DA6A9FBB
-7FFDBFE0770A3E4C4375BCDD8EACE581491954700DECD136525AF7D01EF3776A382E23F510B2
-3F2FB4478818811F8485A1EF49350FDF10429D9267D1071361FA53176C9962BC9FC49488289F
-93780D38F5980F66328B1A066F723FCB8989A33758CE7A090AF733CAFD292E9068FD39727F1A
-563EA7EEF3E29005C3B54C7BFAAD968867D186A5BDEBED92C2E1DE7F4D614C7E7BA10D38CAC6
-05FF800FE1F28909B0736B4762707E8DB1E75F1805D41AA9FEEA89F52DFA3E14AFC44E8CC753
-8D88CEC3E74FD4DAE424F8D8E5A8742011B39F21F118D550C0B6CEA21E88A605FFA17B8E4462
-5D6B2E4C1E4D2B90DA3D4BE844459D3D3BF0272B81D301BB8595033267BAB4902C3D90FAFFCE
-743764B1E0D2A4BC64457F1669C6BF72483299C1A6A976E3C5C975D55ACE50FAD15A27693B1A
-E3713A028EBD33187E9A35B30665A32F8E23B37561D4856683DD2D99FD941524600C5A4075C3
-65DEB013DF4561E845877AB16027DDAEC4E543CCCB8127709E0BD78A0D06786DD5AA5E8435AA
-ABEADB6BA8D93BFC6422DA4D6EB8D6688FBFB85C383FB80325EB7E3C8027E0E3140E9A1EA001
-A70958F0D2AAE8D948FEE71B47D8BB5367116180C97DB48F0B6D72ED7951C5D2AD85CEBAEE7D
-F3A67FE10260FF901CB36249501343484A8E99378287D3017E828CA6A0A7ED377FF3D4863D01
-4DD0B81E4BCA76DCB20B243B1884049C54CD3D811F6EE4242BF75CF8E6375769715A21F6F018
-E225C5B8DEE383ED36479656B0700A4CD42910116ABD387DE79C542B01B2E966DE3926CAB980
-47949B5857604B61CF2C35E02DD94859C6CD3D84B647811ABB297C591F9657A468F7C01C55E7
-A6848BCDAD3EA11137959C93BAAFDF79569DCFA3D0554D34A9DD91FDD27C8FE17510070E2A7C
-5606CFAF9593B91E347696C937F9ABE7AD3813BDFCADA0844A7875080BCEF9CE30CB5318E26B
-37590414941C1C5861440F31F82CFE52DFB5C2E5F70EC8C6B7D3D28F246C6A9A3A3D2A743E37
-2809979B43F66A8263036ABF58A3D77B955B8651A490A83498838596D0D9726374763BE85D6C
-28E3A1AABE8692B02C9CDD93BD1A7E70FBDB0C48B0AE9B74849BEB4353FCABDA757FC03E313F
-89F23083EC01AA2A16F96DD05DC16910F2E58E9EF9054E8952C9474BBB916B894AFC91EBF6AC
-5503D3D026AAFD92760B7A8D4FD160D5511ADD7705925B32C194FC86AAABE152EFCDD059811D
-66CC3FE59B7AF178FD12F9A9F3B004994C41221AD2876265865D87242EF1EDE53B8313197A7C
-D59D239A293E0E6A2C3E4DDEEAAE685CB34125D4030398E67F1DE99EAF60F70F72CB18F14BD8
-A456BB8845D95E74D9AF94F011CCA1A66F6934FF5AF19ACD5F0DE1AF753598B5B73EDE6BCA2F
-B7DAA470F6751CACA86789A80BF146D4E55C4FE07E56E673C28457AD4CB41FE9667BE478DCC4
-0259CCDE65A93771B8060DCBFFE988F5FA689BCD1D818BDD9E975CAC4922ACB2FCCA08144087
-0255C4A1F5E1A600538518EFACFF09945B1140AE65CD51BCABEE1DD9F02D104B196035301454
-D88D455BC4906FB4B630323903B7B340A921BE0DD40B5EEE67E2DB724B5DB94F634B51DA8070
-A425B229A8F38FAA27EDD0BCDA4171C4453BDDD75C80C59D38D000A0E1A938B9FC9A411B1057
-123BD593BC9E94B9FB0A675E1E9FA30786D0B2FBA7503D9291433A81EEF85A544C20D6BBB411
-44F67BF7BBACC1629EBD882BF1FB1F371CBD3FC5FB0ABFE3C0DED921C46A782F983A6807B94C
-7641CD03F748343E6A9F8D4165A41A4ABE7356C68173A2AD8F01F1DCBDB25DC4499F25E6AC0E
-D8C060412E3CE35F4D6EB63EE2A5F37D6E59045D3D31C5CA76D7A556F0F7D169BD84008355FA
-C3A7290C969243C9EB19F3023A329A60B3F8419179DADEA384E5829EC35FDD0A6E287C7B47E6
-8EBBB7B2F212B40961FC5D8160BC566C0712D56C27BB3AB011E63D5C8F086F30BF1F38F1E6A1
-716FA1A57ED4CE22E37B783E22ED9BA9DA10268A88DCBE8CE83FC905173E143D621FD7C8AD65
-5026EF658C0981897D26B6149F510E940309153D0310DB254977C6358BEBDD52B876852FFA16
-D638FD9CAADA1B2B3CFCE4668CA1D6A9B7E17768FAFE86D52301D94B317C4C5C0FF293FFA437
-160FEBF933635A3DEE76813C122F461C0C086CF7B6D22218B46E20DC9168B9EDBCBCC54C9DB8
-D81382544C3B2DDAA4BC8F8301D4342F87C9129EE76E3B19D19198B358649A64CD97564379C9
-2DBD9F3BF2F83461A73606DD35C438B2BDFC9232AB362D65F7F200957E129294B6FBCFA77369
-DEE6BC2295C5772E26C80D6629401067E2ABD847ADF75EEDF2F23ED260D1625BA262DA871EC6
-04ED4B09E693227AEEF56FE1AD183FDC01E7128BF27CE53058724441CB1508D1B38314573B84
-0D36CF45E4C79CF9BC0155D8FDFE9798F5B214F393C4311236D83E7C334EB8BE5DCD338837F0
-06B8E58433613CA6C865457DE5981A48CCDEB8AE8D1F779B6CF0EFC318CCE90DA3A55574F753
-568DB629B5DB2E442B85BEAAC8F7ABBBDA5B7D057EBB1638F91E9C8F3CD4C31A78E8A658254D
-EC9B30B1B3BED3FCB4677258362448DDEB51FDE97C5362FB908B5A2A43E971E95974D389DAC3
-626D7403C6EEC87F9975FF34C520506B910242EEF3858AF95D2033A7654504EE312811FB6B84
-4CD32104DEF6B0C0269C4922AE89A4BCFE1F76537C7D255FB88AB1E8DD0D19B9FEEFABA58414
-77F4C178A737590DE343FF2F3F6F78903226079EE74C653E72ABFA7F4E2319292F20BC0CC03F
-A65457311644EB7FD01BC607A8864C86F67F1FF11FA0881A6F50253BA814FF8641F8E7C449CD
-5A20C2A66492767D2A5B9C455B4410508804DAFC454A04E24621BEB24C9098CA508543E1E202
-C0F5E9681F25D3D95122DD4CAFDBF1708C13EB989F753E38C46B4605C31420F33F80B7190D49
-2EA65215F9FCDA923D90DBB2881A1F7B85C7A7ABE8D436BA7C7A95743B0A05DDB02796D44AC3
-D6CBA2183C29A102EE0FF23A81B724ACA0B6A02DA31AA07AFFCBAA66268973FD7E2830F5B8BB
-9DFF12BEAD9EE1E7EF77E966F4101A544FA9B2603E019C26F40C471F7E0197B221E18E875475
-807E6EE6309081BB9256F7221B6B6ED384AD840FB33D7098AC6469485EEA5A6D788E4BC07E3B
-24A9619C47351DA760B0959C8015896D58E0E07EEA3EC1A2376F2EDCC34BA3314E9DDA39DAAC
-810E2DA2547290D79F6B483A82CF46EE3B41DE189E9EF46070B2870CD1E8CB3C7019B2C93EDE
-69B37A5939F7BE8F22CBD56D532860316F4F607EBCB9A3931AC02E7A44C275160E8B574375D9
-4CD346CC79A1CE4158DB0E90A727226934B963154FF4F5550509F4B956DCFAD74F0D4F16229C
-2FD62BAA41828C437AE3A00F5E1E6420426B7054C1436E8007FC86326C1036606B539F166024
-4AC4E2172E281D279532336BCACD45E9CED79853B3A52E0A202CC077105D42D821DA2FBE4141
-99A04CD35F6B09E154412C7BCFE4AF6E04754362A220D88076F7E1C6F8BF87C9288AC4DBE820
-7C28FC12C84BF01B1C49B65D58308C16C3231C78AAC46C77B2CEB07128ECB792381A79FF1B32
-53A18504666C7C368C1349643EBDF2282D945D6CD5DDEB57F78AE5A3A15E6093E6FE20C1EC4A
-EEE034868C37F78100B114E03BC443079F09C860C00BD93C471FDC04848E18E5AA02002FA8B7
-7DA968A4E781895C1A6EFDDA1247D1AC39E4E157375B67FDD03578412BC61F297BC1B363C287
-F1A17D193BDED6D525BD24C0C25D34B45B7F72BEB48F9C93A24E774F2533C331B86D09D315F4
-0B7214CB635CB4DDF9B15DA4124F7BB80351889A033846BAD31860D51BDC98181536B197E7C3
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-cleartomark
-
-
-
-%%EndFont
-%%BeginFont: CMSY8
-%!PS-AdobeFont-1.1: CMSY8 1.0
-%%CreationDate: 1991 Aug 15 07:22:10
-
-% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
-
-11 dict begin
-/FontInfo 7 dict dup begin
-/version (1.0) readonly def
-/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
-/FullName (CMSY8) readonly def
-/FamilyName (Computer Modern) readonly def
-/Weight (Medium) readonly def
-/ItalicAngle -14.035 def
-/isFixedPitch false def
-end readonly def
-/FontName /CMSY8 def
-/PaintType 0 def
-/FontType 1 def
-/FontMatrix [0.001 0 0 0.001 0 0] readonly def
-/Encoding 256 array
-0 1 255 {1 index exch /.notdef put} for
-dup 161 /minus put
-dup 162 /periodcentered put
-dup 163 /multiply put
-dup 164 /asteriskmath put
-dup 165 /divide put
-dup 166 /diamondmath put
-dup 167 /plusminus put
-dup 168 /minusplus put
-dup 169 /circleplus put
-dup 170 /circleminus put
-dup 173 /circlemultiply put
-dup 174 /circledivide put
-dup 175 /circledot put
-dup 176 /circlecopyrt put
-dup 177 /openbullet put
-dup 178 /bullet put
-dup 179 /equivasymptotic put
-dup 180 /equivalence put
-dup 181 /reflexsubset put
-dup 182 /reflexsuperset put
-dup 183 /lessequal put
-dup 184 /greaterequal put
-dup 185 /precedesequal put
-dup 186 /followsequal put
-dup 187 /similar put
-dup 188 /approxequal put
-dup 189 /propersubset put
-dup 190 /propersuperset put
-dup 191 /lessmuch put
-dup 192 /greatermuch put
-dup 193 /precedes put
-dup 194 /follows put
-dup 195 /arrowleft put
-dup 196 /spade put
-dup 0 /minus put
-dup 1 /periodcentered put
-dup 2 /multiply put
-dup 3 /asteriskmath put
-dup 4 /divide put
-dup 5 /diamondmath put
-dup 6 /plusminus put
-dup 7 /minusplus put
-dup 8 /circleplus put
-dup 9 /circleminus put
-dup 10 /circlemultiply put
-dup 11 /circledivide put
-dup 12 /circledot put
-dup 13 /circlecopyrt put
-dup 14 /openbullet put
-dup 15 /bullet put
-dup 16 /equivasymptotic put
-dup 17 /equivalence put
-dup 18 /reflexsubset put
-dup 19 /reflexsuperset put
-dup 20 /lessequal put
-dup 21 /greaterequal put
-dup 22 /precedesequal put
-dup 23 /followsequal put
-dup 24 /similar put
-dup 25 /approxequal put
-dup 26 /propersubset put
-dup 27 /propersuperset put
-dup 28 /lessmuch put
-dup 29 /greatermuch put
-dup 30 /precedes put
-dup 31 /follows put
-dup 32 /arrowleft put
-dup 33 /arrowright put
-dup 34 /arrowup put
-dup 35 /arrowdown put
-dup 36 /arrowboth put
-dup 37 /arrownortheast put
-dup 38 /arrowsoutheast put
-dup 39 /similarequal put
-dup 40 /arrowdblleft put
-dup 41 /arrowdblright put
-dup 42 /arrowdblup put
-dup 43 /arrowdbldown put
-dup 44 /arrowdblboth put
-dup 45 /arrownorthwest put
-dup 46 /arrowsouthwest put
-dup 47 /proportional put
-dup 48 /prime put
-dup 49 /infinity put
-dup 50 /element put
-dup 51 /owner put
-dup 52 /triangle put
-dup 53 /triangleinv put
-dup 54 /negationslash put
-dup 55 /mapsto put
-dup 56 /universal put
-dup 57 /existential put
-dup 58 /logicalnot put
-dup 59 /emptyset put
-dup 60 /Rfractur put
-dup 61 /Ifractur put
-dup 62 /latticetop put
-dup 63 /perpendicular put
-dup 64 /aleph put
-dup 65 /A put
-dup 66 /B put
-dup 67 /C put
-dup 68 /D put
-dup 69 /E put
-dup 70 /F put
-dup 71 /G put
-dup 72 /H put
-dup 73 /I put
-dup 74 /J put
-dup 75 /K put
-dup 76 /L put
-dup 77 /M put
-dup 78 /N put
-dup 79 /O put
-dup 80 /P put
-dup 81 /Q put
-dup 82 /R put
-dup 83 /S put
-dup 84 /T put
-dup 85 /U put
-dup 86 /V put
-dup 87 /W put
-dup 88 /X put
-dup 89 /Y put
-dup 90 /Z put
-dup 91 /union put
-dup 92 /intersection put
-dup 93 /unionmulti put
-dup 94 /logicaland put
-dup 95 /logicalor put
-dup 96 /turnstileleft put
-dup 97 /turnstileright put
-dup 98 /floorleft put
-dup 99 /floorright put
-dup 100 /ceilingleft put
-dup 101 /ceilingright put
-dup 102 /braceleft put
-dup 103 /braceright put
-dup 104 /angbracketleft put
-dup 105 /angbracketright put
-dup 106 /bar put
-dup 107 /bardbl put
-dup 108 /arrowbothv put
-dup 109 /arrowdblbothv put
-dup 110 /backslash put
-dup 111 /wreathproduct put
-dup 112 /radical put
-dup 113 /coproduct put
-dup 114 /nabla put
-dup 115 /integral put
-dup 116 /unionsq put
-dup 117 /intersectionsq put
-dup 118 /subsetsqequal put
-dup 119 /supersetsqequal put
-dup 120 /section put
-dup 121 /dagger put
-dup 122 /daggerdbl put
-dup 123 /paragraph put
-dup 124 /club put
-dup 125 /diamond put
-dup 126 /heart put
-dup 127 /spade put
-dup 128 /arrowleft put
-readonly def
-/FontBBox{-30 -955 1185 779}readonly def
-/UniqueID 5000818 def
-currentdict end
-currentfile eexec
-
-9B9C1569015F2C1D2BF560F4C0D52257BAC8CED9B09A275AB231194ECF82935205826F4E975D
-CECEC72B2CF3A18899CCDE1FD935D09D813B096CC6B83CDF4F23B9A60DB41F9976AC333263C9
-08DCEFCDBD4C8402ED00A36E7487634D089FD45AF4A38A56A4412C3B0BAFFAEB717BF0DE9FFB
-7A8460BF475A6718B0C73C571145D026957276530530A2FBEFC6C8F059084178F5AB59E11B66
-566CA5BA42B1911A5D7F1BF343015EECE988B7A93BCE0C7AA61344D48AED9C92C8698D4B7C99
-51C87D103F2414B39E1437F9D2E50C4EE5F218F2E6716926A79EA978F13B1F855345191DD7D3
-1D8F82C2E3343C7A5894D95BDC492C28226834EFCB5C12FEA36AC5CC430E0AA604961E34888A
-DF6C1F3954CBC2498E225D953CF5685852162346F4745A2A7087D5D7AD486DE16D2CA8E15CEE
-26E012671BA3BDC7D95CC8C98BB774F508625E968AEE27FF7D1A06E63BCFB5AA4876C3F8F13B
-30CCCCEE73C3CAF4E70D98E6ED2F422DBB4950BF789680E064150995941A9F4DD68A57594984
-7A7D012BB910BF03A42555D1BFD222614F67768FD3F07D99A85D8E700E13CA99D0D283F56CF6
-173CFBA5261209EFD6F5158AD8E4C089A5D5D57F527AD5E8766F2C6F1726B4FE94A5A08E2193
-162F626A61005225458A25BCD2F6CF2C23EAC904FE806ABEED47B5797C3F85E92A50B78EDE68
-77A20B259E3A0BC70D30C7B6DA98D54A811F20D1BE1756121553E2E953F57B3A61984D28FB37
-FC275B0C9C0A081AFC60B60954AB43C4289E7A24B52A49E61F2E14F974A155885C80E3DE0D03
-0F60A6F4FA81B130A9C6F0CBBF379E52D1A8B4582FA6DFF68876BD2BF8CB5B9A029C17B0AE95
-C8163B75701588319B8F65044BF416A68DF46E72130366B7A903123E4B14D1D570C6728B8D1D
-4BDD6E0B59914FAC1E75DD5EBF0B84C5613442C80106300585501F247EC10E4D4D6D76307656
-0D8611FA998C293688A78606200DD036FA03DC20374EE821AF8C244945399193B8D461084A67
-25645E6BF1519AB690FF2047FC32301995812C4486B4812B34FB7E51BB807C77A0F3BED05EA4
-909F6531B8767C97E930B953703B1AB7FC944B20B0AD93B3D0E5EE27C7F55D64FCE361AF7BF6
-F196244B2A842A15A3A76FD12926469E1979FFF2E90CB5C49A74637A7E76241CE36A679626C8
-7141DC2B73567490EB87D9D88697915BB515931A1AC0FA3ECC1831BF72F59C14826D3B032846
-03786A667C754AFC8293D06D2B03EB236687607842062D483B555E3C666A994BB427659972ED
-FB51E7DB17B898FF9E44C648FC2723FC14D5F9DB9473501F8EFEC50AE9C6E323804355ACDDD6
-8F06DFB9C8ADB5D0361B28A0F0ABE450811F296A139DA21394DCB7A2EF39666242B2FF46239B
-C2BD4E117C4209D1B71EBA9F194253AACD7DAB126A23D2E63B4CDDA42DB0A48ACC9CF1D3858F
-D3EF70A9CF518268AD91AF2AEA4C5EE5D46879A476FB4EEE6A7658649CF91310A728E3243399
-C0E7EF756603755BF9360614EF05D44BFC50BF841E2EC12608F495B2509FE3252013DFFE7091
-68F38FF5A92C0849FD83F40AABE0F8DD7557F1A26D93BF9EDDE4CE48AF630D52E863A6A5E35A
-0C4206A6BD29D33F92AF37F4E9D8FF6FC38A3A44894800E7088EEDCF02E1F6923847F241397D
-78F0383815EBA58B4BEE1AE08FD1A7B9C19397A8D878BE95DC79AA2A2D7DA32FB47E332C5958
-2616F813E18E3AA12C3BDC1C1CAF746FF3FB249507DEBA615B9BD3A822D9F995D07B6A4E9A30
-B74C142C3621AFFF4738D2183F91E798E14ABC13E79AF89F2484A0F92A4B69E1DC35D6B0DA09
-494DE0D950CB8B23C74246B2185F85A38ABF429AB5F808256642EEEDE0ACDD159945B09FD93D
-6436B00A4C6DC31A990B97EDE6F89F508A52F618FD9D6BB7622D6413B093A1E5677FA1302A96
-F3592C89DA03885D5DE2D7ADD51706A204E8189497DF0D5DF0640A140699EFAAFB9DE8962596
-D00F4A75E2D3CEC8F84E4AC4E852E0EFA300D02394ABE2A57D2F965AD9CB1B07270BE4E833E5
-5BA50F9733386785A0065E5DE094E3BD6DC1B0A5543D3C4E39975EE00194194CAEF6D9B8883C
-39D9D126BE0C0CC5C68153EDD6F1F7A268739945C96BDACED9207A99E69C5BBF60558C0BEB6B
-9535313F9083D22C8657C0AE80131070FCE9C475941AFF0A3896FECBB3CB9C70A1A4DFC7D95E
-DDAF694E507327DA8584E6BA6ACC0DB4D0D949616F244042112B6F1AC25D470931DB53603C03
-E3FF90D0FC49EFD41B1BA74B5C4490B5C7866FA8CC4B243A6E16FDBA8BB7FF7254D1116A67C4
-97BA7DAB91A3F8BAA52BC4D9ACEA57EEDAA9CFD1C60E2D5BCD83467FD2B7D8E6B4B81BD896EA
-5227C2EAD781B3C25F8D69F414F1EB1BDA73FA14FC0C909ECA25792B78003D3E2AC50DE70EE2
-4FFDD7111F5B1AA6411A5B98E6E63BBADE9049CECA7ABE47F63C112D2AE89EA87774EC32B1CD
-2A0425504ECF9CBA62497E1CE55E6E32F592948198AA7200880C84517BD898BC77D3C7F459D3
-13AB5882FAE7A945AC46E5743201076F3601DA9011A180436E36BF3CAB9B0B28822C8EC9EDBF
-E2316419B1698077607F7A60E3A2CB102CDFC00EB5635AB3EA9C069C03E7081ED2E8CC32DDBD
-C80A9DE5248D3221D0611BE5836BD02FACF0263B297A79E467718AB7E1657556AD428D0263E9
-C597474071F4445F5634FD2BE799AAD511719FE169D0830D47191C01F1519EA5CA0722AC3550
-7385CA3972C83778821F195DF4376542F2F5B5C3AFC0177A47DAE02E5C582C7E49BABA24AC3D
-DA4A9B2FE6393CD77747BAC66AE9571781DE8A4B862380872A097A7B68B3EA30BFB54A55E8C7
-608F707F97A871DD639A000AC29DA673F65FCA6AF06B7D27F6D31F709EC7791524D5B38FBF73
-4E99AF406B98D7C5242746E09CD288683540861BA2E9A0A67CD2B6499ADB4F745490010B4A69
-00546D22D9974D5599F809DB9FF8043FFC4745C9C212BFB55365B4B6E270A0B3B80E19917E68
-12CA6AC3F7C31D137DA4F48678E3C02C5AA5182509ECA7D2845B039BF315499DC33320E2604A
-1D889FE7F7DD47164A7F8133E4E6E59EE7F5A27CB1B3B378A579D31238B787436A4168DEE580
-8E10500871FD13FF47653A2B20ABAF14501B3F50792079E7E636251AAFFC5EC9D85EA533FC37
-DFCE216C7E1DF73097BDF32FBDF22C83E4F62F4794052C1CF7B8CB28F113CD776972E7755BB5
-9A2E3E65C16C5FC58541CDBD480ADCEA324EAE854868ACD35AB5960260BC0D2806E7319768FD
-82F15F95FFBB21463DA70D86F6A60D581D2911056BF0CAA67AC1E9494E7488168831D316A6A7
-A6A13CB00C633F6342CCC2BD024FAFD0108EBD9D4BD724AC6C2F8A02CD1FFE2D8C7ACF9B0C8F
-2491D1CCB9A9082524BC6FFDAF45A13E30D4C27172DD89715D9BC384BF9E9825EB4DE86A7020
-7A063EBA2D74F89D6E9534B3A3A48A18006F673929300E357EE555BD46A894ECC220F408BC4A
-602D064B829FB4D500F07D0F470321C037B20693AD36DA3577466043C331A8554D99F0A4D1D5
-D594F33D5FECBAC2C5254B5D3F9F7166D51EE54EB2889D5FDDC1E502C385F8271ACA8D3725E6
-3564EDBB9A024E476C12CA7EB357D8E3A90825D93D393A1728B199A49444EE4335EA1A49D81A
-CB500EB23015AD31FB582ED50EAA8D22F11F82C2BF9ECB09245AC95E20960C5F5450D1D31F99
-ADF233C710F7FB6B69D34F773B1D3F09BAC7A76EBF9F272D992AA2F3BF31E3B79A67B15B3FAE
-2803E80529E80F6E5951CE8BFF1C7341B9BE0C3CA698931ACF65BE17F3B838FA6503F2D023B5
-20374ABB9AF57E1B0E4DF0FD45834D3CA52EFF6E6694EC13FE174E3340F232257592CCFB4A10
-B2AAE8922E1188CE51F91758455ADB101DDDC0323D4EAB915E7157D17FB878D1DE78E170D925
-2EFABBC555E17CDA4042719A79A3B0BB686D685C4FF238520D4252B7C45DA9E8861DD1045663
-0BA57CBB0789DF66D28F95C74E56DCE67DB92B50BB3FD8ED194D98CDA99534D91EAF70213D0F
-85FFEEF5F31F397BFC978631E22ABE1823A6C9A9DB2E236DF70F925E2CCA09D8F9195411EFC8
-9A0DF5D7FA9D54E7032F45595DAC03A9F8133F5E0DF9A14F131704B70A23A322621D1DFB1F4E
-E7B4E6C2DB61D578BC643FBF6BE526A6FD7E2B13E031AE28B3ADD02905152575ACBEA7897FC5
-F2A0322B04D61EA0D3AE06280EF1524750CD6520C2AC1B3E95AC1994476465B1266D76C687DC
-A1A964964EB9B7BF73C11BF662C1ABD2FD8D46396A350C929CEF16D1527B7064545CA46AF3D9
-52B57806610C78E9A913C49BAB0BAA6A55CCCA1F5A1B49CFB4433CDD9D53DA50A10963C03E3F
-1FE85FA9E4CBA745495C14EA573B7C1D43F06FBBA9DD68A6394494A26DF614AF170D784A2525
-4170F81FBBC63E1D649BDC3FEC774D539E1EBA4B0F151812F972B5C5AFEBED014BD6336AC8E3
-B996630CD7593B7E175A292FFC8B8F1BDDB20A46E145A92DE00F4CB2EBB2783C8A81B85FD86D
-D3E712C8369EBAE8300DCF3AB99749A7A38D0E19684372EC86BB6E24689E26889400F56EE2B2
-272D921E7F4D7503E1D467358A4FEDCB04F1791E033E51BFC71353B59F1262C995B199A1ABAA
-82523864B113FBA4837E795E0054408A7508DF67EF26372E48C8B97DA284C9CA1E9C86088D06
-8382EEE68A8B5148378EEB4DE078757B43F8FAAB9FED1166F26A280446D0DEF6799043DAD5CE
-1BC78AB88393BAAC9ABF41986E1E135E43B3ECCB9209D14D672F3AA69D095C0D71E9DFB0157C
-D7C6C384B9F697D601B2DF7240FED6697B9D091382626F33B3350F5F1BA163082F07E29F3D12
-BE4429AB90DECEDBA1401E74192406D77EB4B83B00E8F23A2747621562DFCE76800449F04002
-222839D02B9D58FBEEBE216987B6A22E83963E4011432934B5FA46402E506DA43DB2FF98B857
-20254BBBDCCDA351777FEF83CD71FEA6BA9E578D7FA3088A3C76B07113CD007C521BE5AEB573
-7BB127795ED8F62C7D487E2866558A87A0964D1566CBED024AF1133F44B055BB64F801E5F2B0
-978C22C03802F7E5052EFFA90450D4B6DECB944FDCFD6779AE081EB977BD46942B2A9E441421
-B9C6C742E2EFFC369D6CC6035B089A280DA0B92AD6D1D0D15DA48673D464A8CB6F88E633D7E3
-C726CB591B12A9CDBF095DDAEAE3D8BE5804BDCE4507FAD2F6DB65F438ED69FCA1B496E869CF
-FC4DCD474C522E3D8DCB39CCA6554F1E685B2AE34926EFF6D497D2FF958A78CBD351288F2CDD
-D9A4DAF32C82A950E3E1F582C62B497814EEA3B4A68C58D2F7469AF81FBAF4849316878E3DAC
-A19B3C5A3FA05E8E001827DC47FF993EBB6B2C6E03A5A936B316B96705771749210AD629D7C8
-42587C8AF2B7C8B170A85788E4AF426761002E9147E81A0D1F380D60FF3534597BCDDA3E3337
-47A07D6FD44D0992CF23978FF89FC8737D4B865C43288B41F3C69950A53F567D2FCF1609E283
-76DE1DF173BDB21A20A2B8B827C24DF520AA5F9780E3A80DC614931A87E8DDF7164ECA85E696
-F2E91885C6D47A23248F7BE5952866364F2233266EE6D5D45ACC9EBAE93A094CC2D5D329D75C
-526E84D67D070E7FB7630D5DBC6A69F686300DAAD7527D8F4C7C42CED676A628346549A3730A
-37CCE926F022BFE59E33730D2061C268C62EE64F61AE0C19558D995EB3F48FE3BB28A6A7B82C
-9099EC415D016EEDD099A08D413DF0247154A6B12FE87045BAA3782A18205D0B99A20F803B45
-5A1B8D2E99D07A5F54A770C8F2023AB4354416B2CC3E272212DA4A48D788C9B84A79B00521F6
-67316BB9EE5286F58F95CA48BA0FE5A337B833592FB3C289F62AF33F8FFE607A7A9E84428590
-4E6A7B87106530D615C1F6287CF67DCA2122B3669BE7F40213E0522FD146674285B7585065E4
-65140D2F967BE3346536C149E338575E743A532427C4B8B806CD8A32B9131E2695AD605C7383
-90521B5304351759002D6FC029C81398BD1D0EC8BD162FD9AF7DCA704A68269B48A658CA2551
-6DB543B0F7B138FFA992855D4FC3FC55AA45599404379F2289AF03933D4CF20ED7E16692E23E
-966181C384EAF9162526AD7A2EF78B9297DAD86C7E054CA5311C37F0B4719D278A23C2B76E36
-296F6D1037FB7C86B4379761C8A554829FC14F051BAEB654B367B0A4BB8D7A3A9B39F31A0A4C
-F170A2160A5EDB3B8E402D7843A64BAE0D7E4837491A361314503CF59ED7DD34DFAA0FB7829F
-328A6669DAEFD493A7BC3CEFD479A308B8313B6BC7B092D48921E0A1AB22BA767999FD8F05EC
-D1F7B0917B2F8FB49E2181C269FA8F420793C97C6B569C5F67DED3B64973173FBD1CB20F00D2
-9FE82E0395F9B45F374E78FC574E0C955994AF5013E49CD0E6F1DF8ECE8730D312E483642D72
-E5320A58FFC362D9599C9DC8BC3D9409D6364D93E5BEA94BFE88098E44C427EA5532999E3EB4
-B40974F778BF4BD4695197733304945351EBDE387906DAD08E38C5A201BFAF1A84261AE477E4
-A6A2D2927DC1067D1D7F90556B43CC9865A1054A90CFEA43DF47164877550CE70D3401FD5B7A
-CCDA1DA72C0D14056D5E4058CAB469B0F64258DFEB16FD795833EF83F1D472A9EC10911185D7
-689F8382DBEDDE4C589C53087E88E42F736683B6072ECEC4BADF1A4469A2349E443596EFCAE2
-FFAFF8D2766D412E7063E66C1B1DA608D8D34BB432C10A929A0E9F79F3AEAF506904CB55B061
-DDD26DB1043BB35596138E8F8D93A7781D72BCDB3B6424C0AADD6182EAE23D738EF46E8F64D0
-C95BE4C15A2E8FCDBA2746D5587D1C351876C567C18225CC157454AFAD386552BB17DFD45B48
-ECE74BA8F83AFAD29EF74AEE37DB8D912009732FCF07FE85168E515DB6AE65808E4E077EF6EF
-278920D63B6553D18254A485981DBBDE35A7B1ADCE54412B9124A86DE8403C836BF8B906BC87
-9C6600AE8529335FB97ABA4D9F6D037292E8546876194219112CF29604B7EEA26FA4F4D425B8
-4B4F02C6D9F71B74C0AEE12CCD65D95410073019F55DFEC6FEBE974A06101E3A158E532B3CF4
-44BE9FF0F88F536CDCC4B8292D5502150FC168BD1E9A810C6D1AEF6DAE6710353BD51E3768F4
-E5FDA7E94359F240B3ED446631C5123EACF57ED0A3CDDBA2EF840284D226543D00BADBFBF5A1
-9A301FDB9C304BBA2FD388D519B4E5628D4A4CCA6B761869B99A3DD14B8DB0A34A71ED531A31
-BEF15F08DA020D2228853A0ECB75C0C358BEDC5A7035ADEFBFDD2DAEC1353678105057DD55B8
-F656C1EE01F8A35830C04FEA0CBF16C609A7C9CB4917CB78033F6857B95310D74A63766179FB
-C9AC247E52745F0A69A3B897751F4026D4D271FB071094C35619160D09B77BF1BA9060181D93
-ACC12FB8BAB71A28F8AAF678E3CF9D3E5ADE7F967B8675546799B11842B6F6C269280946204E
-F91E951D2CADBCA64B7593724E6E8F03CBBE6C17084363C5C102505DB1AD41DDCD04BC9C2D75
-C91B33A4806BA85E4D91B0A604BF529D64B709499399B69A4C477082EA684E040A9113619774
-C0CBCE89EB13124DFE7179D20909D18C263B611CA17E8B33E17BD4E66D938721A600EE36DA78
-B00C5EDEB7F3410CFEA7CDACF544D3CD4ED9565C8033E1E51D879E4A35C1EBD7AA9DE54EEE32
-714D17712B8217B8E45A6C3DC4891D5B2DF3CD3F39456CE9145713D83353F0D8060EB9C4044B
-99AB988DE9C778D943581E296CDD0534FF5B128157F4CCC8DC7516EE4C60B9F4077F0C64F540
-A2B708816B007A5B0E1630D4D8521BD20C07B344EB1AF54AED67C4CE5689E46E4D823F256E63
-7992B4D777D9BFF2A84CFAFACB4C0B737B852F53FEB78E9F633069841E5334F55D92E32EF65F
-D1BF8163DCC71C1B34C835CF8E5E86A9121FE9DD840CF83E3FAF4FEB34031467C0D9D56A90F5
-6D68D4EB47DF03FC92D98D497BBAE9090BB48BBBAC835E5D11394A960958F377544FD325A31A
-DD62EAB8C62741685785321F41C30BE89895BF98D079A906E0A1223BB22071987B74CC8495BA
-DFB6252542363AA115BB7002C2FDC6BD2C94CC651AC8B82A2731D628F91528E2729D0A404091
-B2F9059137F577A1E4753813BDE2A3D1A6E6F71BE0C5D5287B1EF26A9CA6F805653BA7806C2B
-508C2ED8233347E718A3F7CE8F75B181AFF8C4E2DFB7092D46C7D287EB9AA7B77DB11E8A00FE
-786799D9ECA658E41D86066919FC092906E1F08D7A6D3988A483B81FF928EC0E67192545ABB0
-F83C2E915C982A5AEF723F6E3DC13BCCC6DC36D22D1029D872B5772DC232605C1EDEA79EC0D1
-81B7C5A9F07DBA46B91D426AFFA597C500D3E1776347B4B74216A1DB5D9761C6211FE00BF6CF
-0E907C727728CD61A7C494EDF8D675A224F4A35FD5571C1CF24C130BE2A71973CD7FEBE3AA5D
-B74B09946CC9EC436644863EDDEB88F7125E66E82822AA76A8EBD2375CA95EC1BF197A05966D
-6D14C90EB346F454AD0CECB1A81E30397868CF3397C197CE4272983157EC9CBA19433D7D8037
-275CA7681270F589CC94B6549F991BE3CCE62E6D92DFA7413A3EE7C064D1B98970DDD6FA0AA2
-5784B004E39F86425A088899DFCFBC1654D9BD1E096074E48DF9D3633D06F0E7826F65B9BEBA
-B3D9CD614B487EB18D8283BB62CD78D6864F6EF6A7183677C15461DD160D7B80BF591AB41B17
-9B697C5FABFF32108084CC5B775CFF73307FBBDBD0CE1882CE720466FA1B31FE6633BC6384B0
-C9F1352D0103BFE76D26171CB07EAEAD65335A0162CF0A1117607EAC59607497337620128FA1
-3660DEB22F72763E8DB17D5B06EB2E91D1FC56193826AE3C43D09DE2486CD957630635D5D6ED
-0E2DB99BC6BED4A31B09C59EC549E10DE783654C5B185A31AD2AF5ECB449CADE7E35E53966CD
-DD84F485720DC93D0F825AB2F5E71E0EFFB0734F26EAB3FFB07F2EE55BC50731A248295740E7
-4F7B9787D88082F15B58EEC492C64FC0138AE020FF448A918D08D7D4780D49631E0144C33698
-164D73510B7AD2A4B58FCB7617943E3468EB10C88DD9E544A3D01938F9B97D342671869A985D
-789EEA94FB5A76EDD5B8AFB1262248CC3A2D1F6B53F3FE7FF97392E7CCBE99A2A9CACB7AED56
-37F61057FE6B0C2E573C09467DDF06A03C1C313D0E6DF7D0B98AF845DD5B4F2FA2FEA3ECC9A1
-BD7D49E84F63C913FC3832381895F8F9517521D18B220F3A8E303DA842736D9604C17CE7D935
-F5ECBC3460AD37D9A83C5FA812F2A815C3F875E209FA655F5277D01642279FE809E29E73674F
-7C67A0526A09524D9AC61B3DA9E7CFAFE10F754E2B949FF69F5C82CE817B24F79455D1C1106C
-6F5DAA8640998EAAD2AAD0D6A187DFC24DC61CB49DFED8169921AB11FEBAC87457B5F957482C
-CD79DCE7A215FC413E2E7E0BA25D531A9980919F0FB0A6EAC3EA7A6581E633CB2A2DBD9C3862
-6DB97C94626C79BCE7083F7D041B24877CF53CDD3D2569B9A6FC6A82630F8CCFD7FB03668E9C
-7CEC2995D76A39450876E84F5C9A27452AC2FCAD677DF0F8E5E2412E7B41B2932D99CCC482B9
-826B16CD4CBED83CEF79B6DA562877AF0CEB4E7A130FA53D8464438A41EF80E81C2445DECA1F
-C517715D6AC71305D99DDEDC5D601B324F8AB1D053726F8601FD0AB0380AE0FBE5DECA11D170
-7DA47D4DF1B15FA2E70D220E41D11E5D4107CB08F1C1C5A53B1DA0D7AB81910CBEF5E3F2107A
-2288BFB86409B367AB37D84255B12BAE222BFC4518E838178926B32BA1B59E8C773CAAEA9C5D
-4113AFE9F46B7E3FB068C14E7C74B00738AA8F3CDC64AEB539C743243780AAE97200EACC8950
-0A6E1E414047AC3590AF318E9018B0A150D1FBD17109C8D454FF7DEDBFA9939A708C70AE3F6B
-3049FA07CB5A7748D9EC4151D39A495871175E8ADA874A999D46008C5F96C3EDFCF6AD1B38A5
-0E27090FEE82C1728980E0C8C0AB58B0722EB226318A30B6CB4203E61EEE54B1409B758CB05F
-A6E0ACB0240392512D9C397E8AAC6EB74945CB0463CAABFE5D07C69F29996A8DCDB5668C1A5E
-1EA55C6DDC3B7B19605BB2E10FE6E1FA553C617C0C4E0721CDF1C818FE809D3655E3209229D5
-E62E939950EDE691CC50E6DE4796C22644F6E7AD0150A12B5A8992E403DDE0E1D5B681D742AD
-6DFCC3D0E4B5996698DF593DC1F9DD131E694D05921D9D8E27C43DF7E86EDC1205DCA2A2251E
-4902E1F9C94E680E1E955001699E375B26E1BF9DB2E74CE5FF2F61D154C6D5988D7EEAD5841E
-7CD5389760B9C1591C5F0E12C86319B5A19E66BC63EA3363EE0E9E0CC1BA6E3DB1750C6B32BD
-84278A2D9960AE13B75BE750928E8CFD28AA012FC3B02A2949E2D99EC7956E86FEDE302360A6
-34E417EEE92746937CADA4BBAF39273F08AEDD063678882EDBA6AF97F3A2E0ED909505D76819
-FB8B79687B60BAE843942BAAB1E257E6DFE048BE16B17A7F23258B4A657F085970FC81837834
-36107DEFEB83751C82E86F5245A21743A1D4B535B66785410C54BF4134AD1830054456A65313
-88BC0B2920F64C11F69E398B8C1024849A1BC1A47013944F9C1C7E176905BBF7CBB2E0FC2FBD
-A025CA4921B89E0F22996C37981C43806D0A4543264C301C3F2EAA935EC509189ED58E8FE039
-D605C3A8A299026ACE420E185A2C92C8C7EB6BEF4224078EDB090A4932F559F7B364FA0FB07D
-587A4C9804067BADF70EE718A755F74146BA64E13CE3F1002304163B91ADDCDE0E3411ED45AD
-8C85A244114F5C6578163B26CDA0D752F88BE5E79613F893E645250ADD7A6622A7D6AAA87E21
-92FBE906F4E0231D955E3C607B7B1F1AE65F97ED0455E29F56F12FAA16918B480CC047F1790D
-7E7C0A04101628DE2CDD5EBDB85BB04F61E6169328555F4B4E26FC29FA10F25B089B600E6358
-01DC4CE6EA86044EE5B8C7B15676483CEE2BF7D896AD035AAEB0DC3B9A871E2A9A0E2C4EBD20
-3880D40F9F5B2ECAB16294B432E6A64FCAE9E3EC9C6C5CF65D8AE8FF7C310221F7114D09415F
-B4B69C1D269D6D17C0EB9DE73598D2DB82B9415BDFBD3A17E0252E989BB91C78321EA8FB87EF
-075AA94B3E990993862DAF7FBE073B7D2DC36C2EB23B59FDCCC18A4CEFB4159B70F6687A946E
-5BE9B5E726EE036385BC93AC41ECE0328EECF154CE96B56EE5913A265C378680B94B4A1D43CF
-47717E9117589D14704102C42AADBD2E05E4D2278744648FE93C197529F9547DDF4069B4CF4A
-B57C3A116658FF0391EF71ED3E067A3914BE2677973E96BB21057798732221EFE36D151C5946
-E5B9C2FE6A5188E53FD9AE52E7045DCC0351C8CCCE6615EE938BEBE2C211F57A5E56A781B538
-B124B7D760E66A039F99A0A1BCCAD03003C53321013ED9C75EF221E579660787945642B251FC
-9A9924CBF3812C3AB213BD5049275BCC02A9DEA546C5AB5C0128C09A91186788BF98EB6B404A
-83ED58D220E4DC7F19EDF7CDD955C26F33A2FEA8ACD4556F741BDDA4F747B1C3155A1E9F17C2
-19083D87ADD59EA6ED1EF4744E0ABD805F97F2F67C5CEC364C2B6DC98D20535795ED5D8A40E0
-7F0CB7BBA105BCB030E3D244C24C460647CA2FFA069A130E4985292265F208BA62006515D5A2
-478F41956F6AAA61631CAB5DFFB1EA88F4F921BF30B31A96804FDE43B6159072D64102CEA317
-810E35E10C209CD23482B93D14C044CF6CEDC4371D0F92926D4F4FEA11428D4A598A6258A773
-3BDC1F39046A34FCF61691748587B65495D49572334E4FE7D8DB7214903C87F7C5A74250703A
-2CB917B53F9F2A39A161FAD80049AF7D0FC179C9D686427C4D102A79A5B56035588C2E621E9F
-F325997194011757C99BB10A504C83C68438D3F4A9D5C6D21BF22D23A718F75FE848CC8735B8
-C6133D03FE1BE69F91B8644EC94C3DEB9610E95C3C20AA2CEF534AC20A4AFA4D4A1F4FEAD6D5
-608C44D64D63A96A3EA466460E1D61613D5AC99A3643674F5CCB859A20CF68F9677C98C5B6BE
-FC10B23A43716F05492635315E907C4DA68297453E16CA81F9BABD02FB86F0418DBEC99CF4F1
-4D59D563398254279A0B0D3730F388640310DD228F0B0C9A93F7D8F2B2F5706F7C20E76E4DDE
-13D6DBE7BA6D347E5783C69CED0350D3E6F6503FB93F3C87AA36849625A6001536454A5667F5
-851F63A2C6C352E87C8C1B43B25EBC0EED72BAC7281D58BCEE92053C4E76163417744E955D48
-076800B829CB6025FFFC6DFAA987240C9EA28C7D115B1B6C034327E226FFEF597C38D463ED5C
-481B10E08C320ABF221EFF3C52FD2908BB18B383FCB356E62D9F78ECA4587A0E080024B19971
-144445CCB1FAAEB35624C972A72830C86D9CFFFD2FB5C0214F441371CB7B34ABDD52AB1B436D
-D9E4BBDC50239C456C20008F78DACE2E2AA0FE2FB5B77D1E768C6633D19DDC9E77D82937C3E6
-6976F74982A9B09F11F8264C9BAC1D72BFF5BBB3BBE1C22C65DD3FE6680E0B77154F2751D7A6
-292B4DAE8211EA08072E238C46156D4E669308B2D2C5E08129F5FAFDAAE599C58EA397CD3037
-305F52B51F03B266A7ECD8AF86E6ADC9E98614A43D870B0CAF609AB4D4D51B0BE2BD36AED2B8
-CBB2733A875BA672DD1790A93CBD1DD892A1DF843F1BF22049E0E96D6EEDCB2272B06579C0B9
-744DA658862E3C2898DF75AB97DAAAA0FD235F4F8CB7D6E82E4A38D023342453A6FFA709F66E
-E784C3CC00AB79FEE9A98175299176AA6550F320A298894E946C3C6BA4DB7E43C78707A8E36E
-C0CBF8C78D133F4288D6E4583287C2ED7353AE694530078844A8A6A1165188536E69279A41D5
-D35AD2BCF2D8BD6AB4F5936F705E8B0A665EA760F6CBA476B5581B1B0FD5EC6A3ECEED59C862
-D2F790A61A2E02AC1BB48BC0CB247BE24E9D9FE88ACC0AA270B2164939E304A056C877795D8E
-C80205C109799A5BD15F7485830281E2697DFA7DB1CEAF0443F685ABADC42BB8F29D5C1D6249
-01776B5F6055D1C6497C079EAB97BCBDE040C37C7250ED61E97BCC971AB71CB87A79A998AD78
-5C5762F8171A28D85A1D0D3C32E2B0508FEA24B45710A14008395140124F6E87EF73E4C4FAD8
-B52EEF0E66F62215E0E541D699034D46B1752CCFEB047ECEBDC625C32E8E3AF5D71543516017
-65C46B24CDD42E661D2095C3D15F1C3BD3E1FBFA4751EEC1BEF46EEE1EA8F925347869C53213
-D51568403A8A90F7BE1A360C485BD66D30315275F0801EA18CC762064F7D5F44CB4AD70685BE
-AE02B18B2188BE315B87B8AF9ACC3285355C3D05BB3C392E1008E58068E34EC7D4E8972BD5F8
-090A2A506E82C650895FF5AE34E5483112623A00CF6BBFA4510243B8B7E2C6AF9AFD098D02E2
-EB1D07FC2FD3E37154457B96FB4550B2D7CFB06A4B607DFE469D73BD9537C73337B5D9DCFAB3
-33987208094CEBC299F0BDCB70C4BCE2A1CAF04402F1748A63A733BC5DB817CA185C35654021
-35F247705AF28A13F8DA47CFFFB68DEF345C95202B0933FC691F439DF83001CBA8878A7EB2A8
-9AF544AD9A7068BF48EFB1D201BA2C191F650E84E8909456EDBE3864CAA7F55079E928692E01
-EE90600D0E9E207A52500FBA0EBDDA50A6773A9B441FB81BCA7215D0276EF8D70C00769EC709
-C214EFE1ED92148349CD8DD8A0CF45B46A9959AEFB2C42AAB1D0DDD68880B107BF56F2CDF87A
-891DDC95AEF319B0A9C66DBB5F07B08D9E783E19546CC9C8FD963637D85FAB62285295CDAB0B
-DBDF348CC819090C32AA094B3B0FC4C27A096320D373258EA68A2A508C12AF1D81368528F554
-303AD7D7E06ABF9B8925D57D694494220D6CCBD8ED04B4C57FDA1DAF70C066A16930FA34159E
-296633DAEA111D246256B3D1EFEE6C00A7BBFA808731D0E54666FF407ECE2074F2A93850500F
-F3A5F0771DE7C30E7B94CD39AFD6CA1D777298E2472D0DCDCDE5BB5DE16AC1F0A03364DA00F1
-24F836E79DA442749A0C6EF703617494F9424112D9766760E83BE50D1E6BFC14F26BDE4C83E4
-BD304E99DC3EA4EE8B04A654808621DD9038AF320339FF09314BE0B1395E9BD9B0A0CB9A8B39
-965EA7DE8AFE1815B53402DF5385C674A5B6104D4358D52B57D95CE13092CAD1D02A4C89F4A3
-B934861264CD55910FA13BE467FFCD0EE03A6D5D35FB4DEE5A4C57740C7F2EA1C7849B17FE54
-C5782571DDEB8534EDE38DE11A278D50DEB22581E24B93D08F4A8C7F3C548943B2601B5D073C
-7BF8AAFD94D8859A38E1B5E659CCA1EEE4EF023FD05B94A2774C13B2CBF461265C9B7E8DEF13
-AA61BE7F29DE723BF1D3B05E4077787F8BBC48B610FFCD20DACCABC1F79676E74F7FA8A741AE
-A7F0AB8C341893C8A5489CF725562B14E44C5B6A526839D4B3772D49CE406FA34A776DFFD096
-41FB81AA9C6E4959F16B404E3C9BF5193623C4805889B453088E67967ED2C82200381F3794D0
-F041E922BF79B9D5ABFC0EDDF32152D312F739DB794010791779C69280A9B0FA675990BCDA5A
-84FC2E9A5D10C15354354DD84304313491CAA58CD1F2169496F90F14AB18894640D21118D9E2
-FC2E8BC621B196E83871ED12B9702127C7BB5943748A6D8507FF124FCFAD6C40DC1ACBA3E1E8
-8BDD3D880F0E6239985F351A5CB862A2DA92C93BB6E7C727A6AE4E4F84AEB0E07B86C057BBBB
-BBC08263B822B7125CCCFE0B304AE38E2E3901A6ABF94DF34BA3BCB94E2DD5B52048F79FEA9A
-B6D3FB01FEBF28845EC0507BA993E55BFB8FABC33E33348FCEE38DDC5A6CB69201A60FC89EB2
-0D5F0EB6884E21A64F3F70229943FEA58AD614042AD7B528DFB4CE1E5598D02363EE56D7AD9B
-C4F72062E65224E7DAC416F78A0E5702E0A4E2E007221A09D03DCB7F42A2204DA6F32F1ABF63
-53E64CC59372ABA05709EE9C2E95F10BE48401741B141CF71A954F05ADACCABE5E608773E9E1
-A81D04D2B967E468321FEFD24ABA05E63A844B3025213CA6BE493701A486AA84E284CF29B8DA
-76F99C6E35AA40328C9F7DA6A7C2B7E4E100DFF57F996C06EA233918354F8DDFBF517357AE06
-57DEC94EC1075086C7C4F1DF9E5656BA07F65E8A90659D4B4C4ED13AE75E66EEC8282ADEA7AB
-E0A064C4CAC8E4B7A2FB5D99640ABBF650171B2243B22B53B0BCB67DAAD4EB7F2470FBC29134
-53BBFFA76DBD0FBB4C24E62E2C476E0989DEB4C0BC5B57098616770B2B47E4EF1746266E964C
-87F6101EBE9AF0D5EEA9695E29B3CF3ABE98EAC41AB4084E6B442C774F4A1E15C3DCC1402195
-B8662893106AD7ABBE5143C9F3A4E7C7B2F36660B91102030F74826DBC04BAEF6152272FD2B5
-493C01A9E3893381FF183006B9964E9BF4715837C2F9551DA9BA46D3B82948FF4A48E5C7D3A3
-F9680F4362A739E6704BD991F0EE095186F4067753B7981AD6F275D3633CB443514DF9B8DCFC
-E18F2A3AB3D9F60F01155C06CF5CB859646D3F310CB7D2C932EABE87423F2DAE43306230A497
-8644AD3F3A59C322D012E6BB6F0975625CE5149C42E7131C3FE843F6751914D13B1A63A0815B
-6F1453E18A7C4C0638A2261584F7C60869A9346F832A9C031545B6C5685224144D9B32146FE6
-A8E5F36557C028081D2F7638854D45F9D50159DA707DDE276F1A93AE5DDE443542ED50A93454
-BCA1B6D675780CA5A8E9319E0D16A8A77A208E84990CA84BE3D912603790D40C66919B4F3C6A
-E827E22B7668C2EDB597A3DAB0CD2367B108D1CCAA39B9DFDDBF4B5E695E8ED22BAF75F1B385
-628807679BBD1054A7B9907A75F1E7FABBBA92B62D8B070B362CF38EBD4374FBED4B616A390C
-91BC00956DD97D97FC14E367E1636FE8DD8D20700BDF000B31D9D090B549EEE154EE63D07FF0
-37F3112EB79330B0C72DFBDBF17B436CEDEC023B0C773481FAD03E12F1490912013E243CD020
-AAB1451244D98AF01DDD7CE3B650E42E5BBBA3D93A5E4FED991DD6659422E8A2DD55600500E2
-ACC4244B3BA003A67D52EFE3B58C95673E70C2B14EBE1C828294298ECE960A2B9FDDC99F8437
-D1EA61C83A9A9CACC8CA84BCD48A52C72166CD725404AA6A67584D430C1428B445B80E7B7D68
-F540ACA94FDEE666CD1336BE06FC42F0826F0FF5823D5328A253C57F6D4548344E5C265E9972
-3108F471EC0B9321E06BC8884F786D4A56BC0ADF54513C8F7A084FDBA70A707FD937F4C1F4DC
-77F53ABBA64608A74F146126EF9C46C88AD22ADE2B5A65CC770DC588DBC10372D65C15132C8C
-BD27D38F8C654550A95F957C8C6D107FEA656D9F5252AA8E257A2430864B19FB3F0D78AAC47E
-664869F037F8896C566A3A253E7C1FC07EB1A85B692815D153571A972B4AAA046B1F13A6769A
-3B862B667832F6938A2FFB72ECA03C34B607923125389187F6795C775CB754D0DAE5E6C79FB3
-161AA7D71894F462EF8F0750F9875144B429BD53C6FE1E029A69745ED66929C1D6E55D4240F5
-87EECDAB8F6F2163B224F5420205144C6D52A252AC9518F4CED732C784EE5BF3326CFEDDABA9
-4F25C721B380386B6C2A3E7E07C1D81142C4D928C24E2C48F4F617B3FCEBC579F49D89691DB2
-8346621946AA0503238E4C6E80214EB6998FE2FABAF5897EE040838FDE840352E6761514929E
-85DE7EF9484CA75D101B6D2EAC0ED7DC20E896A0979AE3A5B5D83846ACBD88D8256CA987D147
-F747E3A1CEB860C8486578FB361AB53C6D78F7D7670B18F4FE388A44F3CE5BB6F5BF751BC7AD
-7775D21BAFF538FFFB036A841008667DFDA5B4DA834C1192591B7D479DE5BA0FB95627C0AE7A
-7E46DD8B3BE8C7B952098DD8BE9C2F9C77ED077C7EF973E0210C68613BD95B8BF0B1FAD95C46
-7CD606FE9906D2604531DC6148CA75E23B9970FFA17064BF536DFD7539B0D2E8C4BE97609DA7
-ED6B78A8D0A1F96650344043B06546D4699EACE1F7BD6071724FCAF6E095F5EBD9F14DC86436
-D89BE68586D53CCE108E7628EE4C1453E1F62C2455F02F998BA9E2DB8887E3D5611037406AC0
-28CD5FBD94AD7DE9F2F514311040FEA41CACE74915BA6F19FC33C0FF6116BAF3B2DA9CC682E0
-3F9A85128FD578218D43FF7A153DF4325EE24281BDA2A49A0A8B9352D92A7594CC4614C30E09
-1B23011006DAFD0120847511C4C6C2748F2915AD400452733948E84E7B7F91A82D4D02CFF68C
-CAF9B2B9ADF2132AC50D44B5EE3E5E21F7AAE0C2595AC9AAFC0780BDFD6BFB5316541020ACFE
-67B04D11F3016139F5E2453E9EE0CCBFA049C2168CCAC2E0D81B546FB593DEE069482E397AD6
-704D6DF8681C7436D86CDA2718E9ACA0249DB0F81528C0D20F46B3BF74653A9BBCBADF94BAB5
-829D789011505FABF205B871689B9521925BF24D80718AF437114A50D7EBCD2FB1FF37E0A163
-F4873DFB24143DB275A1E9CDFF76142E6F94EF13B8B6CD6F40717D6460435894012C24E98DF5
-0C906A5B2AACD26EBA7B440119F18C2D938C32C36B82EFDA90CBBA99196AA5AE862E0C1F4985
-E120B80B9CDFCD5EFBDD5C80A7F3FF4700FBAE9DEF98847E6A56CF4B609AEA604EFF27C8A633
-4C437868F9ACF73202420E87C92EB76A10E79B0ECBB2A3CA1768801872AF0443AA930B00442F
-F801B05B39CBABB684AE80B72A8515C88EB2B46F39614E5C998687479B057AE7F686912A8845
-C358DB36F8859F819FC36147E1D277DE0BCF3CBD13097B7D9BFBC1EBB590D93CABDD12811DB5
-DD5E70DEEE50819E871FA62B97F37B09BF828FD02352EC2DC627B4C6F911BBFBBB4D7F1D9EC0
-9D36B0606DAD9DCBC0CAA84872668C90728CF434560D767FFC3D1BD570652B80089FAFEBC335
-8D1AEEE0C4F459A5875D92FC676EE3D9AFCBD9E319069467C89BDB5EB058BD85CAB28B0F6315
-A99B4421951AD063C27F85ADA776DC0AA6311EF1EEB51411A6FC35D57FB8EBB19A1EAC5348E8
-AE4D0D2DC9957B15D4C5C2D8FA0D59C72AD0A03BC17E19B556AC783276A003B51CC055E1A430
-348983D8A2AB7473A1C7BB64578BB9306305C92776B833B6F6F36356338753CFCBB761AD4C25
-34F6066B6006FBC1AF36C65C71AECD2BFFAF5A43EA462EECC9FE1A0DE404B643836055B81987
-7367B821751A2A8BE09D1F0E8D8C07FD7A9282B555A9F31CBD01B5DF89AD8ED817758578D4EF
-17594E1B08F2648680398B8CE0D71B97300BE89E060D2B6F60B6826773866ACA2A542885E893
-CDB35E55EB49E0D38A78E2936EDF0A0F227DD4358B2B9370A0B990AD007244BC924D14BFADDC
-454680D28C6B85BDB3717E772A853A7B98D40AEA92631B8CBD9A68753586C69D6E20681E795F
-2F2FDD84E3FB8E1C44F1AC77C850F5689B40866363ECA7062F7C6DC0580A5F201E5E14F670E7
-155CA663C3D39C893E3412A020638EEF5D54217666FF360E37773BD707FD4993F88D842B6537
-9ECC10661343D3B12CDF717C9D13E0EDD64AC4B31B1E764A2BCEBFBE143A1797CEE91D2E50F6
-45511C4F585E00C3A51194815C6F2056A012934FA250E94AE9057AE4DF502BDFBECE97D8420B
-CDFB9ABA6E02EC9CA027156A87BCF95724757ABC2CF3AA2EDBE544AFB8356C0B0C88032D02BD
-69C1C49FF31F14D40744714EB378ABE6D7251453A15287831AE4069CFBD8893C6044D7AD5E5F
-8C2FF860DF53EF5373A1D65008E84A3972651564325DD556EB5C2C4C6D3DF6CC881996D25F91
-1769A63B82E913015C64C468B4813601DBDE27D693984AC53E84EB62DF3B12121AF3654DEE6F
-D29229F61A987F2F5511DDB8F963588EF79A1EA33CFAC9F2E687C14083A1897F5F587BC5EA51
-CA7E584EF7CC30007D0A9E07BBEC716A9FD51FCB85F2DB432B6232DEB1A3D8268CE153C98C33
-8E8A7234E101D5A92708E0C399370E8FCF656599239015C7F2CA935A533C5411D34AA555A09C
-E662B210A65C2FC7DFE9B65B564B99560879AF882D6DD96C9B3174DD0144E7C1DEA0789A6989
-D1CF1A03B855947F77E18E9612D1A5C2674DD39E9E637B623520807DD96D8DAE5B5CF2F68890
-A0B2C81202357719428BF854FFAECC7E52D4175C3D489B34FB0D4C58E371207F9C7ED7CDBC55
-44FE687C46BFE974D5F6CAD0D5437B4D599F31761D5434DAA1226FF9ABE51E0FBE6FCC0EC143
-390349C058E34B429E4133CC930CAE0C9FE114C3ABDE3F1BCCE09A9927AA487B4A9B43F5A5F2
-BA9D948DF9B0F79EA9A114AEC7D0305FC805996B85DFF17810A0ADFCA6E6DFC31D200A038106
-BDEC7E3CD6515E553A48E1954F3D1C7198F1A33095420A45BF580B9D626D135919471D5322D0
-B70E3E7997B8C6AAED4C504E9D00AFA9D2356180F605F03E6931738A1DBFBFE82EC63193A94F
-ACDDC0AA090EF016FB5A0420DB26037EEE279125D09F19EC92A28A6CEF6F1901B8F86DF979E4
-8CE7F7F0E16A985F47B396520277708139960E87DB0D9A3E829D1D4799B5D7185C4FC5028DFA
-BC4F08960E3C76BA09F23AA5D0BA9D6C8C55449AB57895CDEBB6029CB15CD842D30CB0B997DC
-CACC1C3513EA25D29E0C7C6BE745382C450D3E2442F97F7790756ECC60C56E55BBB622004254
-8DD7605B8F2BE654859BDA2E270CA9CB3AA37CB4A1F071C84F75656DF83367B9B2B6FCC43864
-275A34E66D44B05ED3009989CD7368B7599FD86D9C5CFD661F0F7233B595551590177C11AC61
-F17824E6DFE8516925F229B9178822818968507E65FA3CF7E446E285C4336B7007C5AB399275
-6644199D5E8CB724B876791FA9F1C33716E9B47F5B849B4A80FDA40FBCA80E8EE57CCA4A22CA
-C478A78A1E5712382F93BA09A2A9CDCEC60081589F41BCF2727804EB5AC6F540C3E41FC4BE27
-4F90FE4275A323594F275F93527CD29425C319F98750BE42CA276B31E3C81843CFE850A11357
-E9BDC646C17209CC5AA8DF558839BDE80E1E0EE7183A05CE800F780BA596176E00CFD95BF8F5
-3F2479C0CB7F727E232BCAE8DE54956C43208C26611548B542C1A1AE9AC45F79CFC682804F87
-32D0D1774D6F5FBF110508BB2231AC33885FC386AA4B180DBD051C17F097417F6EBDAE790090
-B314C80F163CBFF42B546188ED29C54CAC565D271AE6AC73D102DBB7A8DBA02B36476A4BDD0A
-819F77B3A7BE17E2576102F08D337B2A2E8899848561F04ED57B5777DA86CF3A341B68903D55
-7B38C25A5B619499D7AD4CF4C0A59B402D624AB87321F5A1180DB6321E4F18EBB11AB2D79107
-AADC1D17E0B29938AFD818C8D9726F7E086029A06D80F7F4436BC4B0FEE65F7CD5DC68B56D0D
-DCCCD57EB33086A7775B52B83ABCE3BB6C50B61053CFBA0C70E3ACEBC72B4DACB1093C9454C8
-42AF791C097C5CCE5056D4131487C48C821F685CBF17CE9C95BD30BF195768334D454554889D
-A05B693A5363B996BE62776FE7B30384728E1EAAAA489004A3C3C55BB9A8EDABA708B66049FC
-B181BD1B79B86273AC11EEC22E3793BAB6253DCD3312E286669FC90CE3C69F3E70FABA8D1E7D
-92371CA34424F20B2638DF486DB4C220F6AAAA97D60A4AB8F06E9ED04F80559FDC27D861AD6D
-A2D5BE433D1E0BD9333820AA23EDCE9E36F26DE9ED876342275CF85B9FBC34793502E668D104
-075682B8629B4C3CDCA8B0AE3C2A7BB114AFBC1E96787CD8937D35A6BF5B366674AC2FB151AF
-96C55F7A301ED8E1A681091AF8F0219EEE29E8637B55B90962F121800EAB46946DAC303B44BC
-29AF2FA4774E9B930CF6BD42C5E06EA9823F697E28F52C77F8989693A4C22D4C7F0A05E67DF6
-1BD9F3FBA554FE8FA350CA2F22C96C67F35731E07EB1681331D0865BDD165C535C251F8A17B1
-DE9C72314B1125A97B2B1119C8983E8D72F60E6D11A3961148805BC150CB433A1ECBD18B6213
-9551EF1B1D1DF95BFF5964A64B66479491E000FB294F7230FC8E11D2B382D1E11B04761EE84A
-0E3E0B00000470A6C101D8B8F4E107B7A4C2FC3AFB4EE51D175250CB6C1383A100558FE4B656
-D3A017967BE2905E1ACC871BEB04725D174B93F88A8EBB629AB52200E61648AFCBF0E2BB46D6
-3793B79CE0CB627448708CA138811D8C5990FB01D427E9614ADDAA6C0DB329DF406CAA8614CE
-4EA19642D2192BD523B0F33862218821839759F846B8D934EEA9D7010730379FDA70C87300E4
-709ED00B97251C37D3A10C121F4EA5A08B5720EF8A0E6DA35FB853FFB967C07BBB4EE1F1CC02
-09DFAD1840867FFF032AF91CEE26AFF4CB77E68D18C623A5086D6A922AFD2516B8A6B96A799A
-31AD682DC85BB9AD35481DA23F503F16A53E1CB34DD397CBB97A72243C2878C4BFC51C693348
-2229D49BF016434D50B68778EB76FD89EA014273716C7654A56B2BE26EAD04A0B95397664F68
-BC312A178ABFF2D60251C50BBE1BC8FF9B8B10C84073D050ADD84E72B103FC5DBDFEDD13AA73
-610E2378CF2C0239B8703897C31FAF350AF5A17DE8324E6A1DC0346FB0DD8BC735F128E75414
-01D4DE5B5E3902DFD3CA54FB4125AEB31F91DB9CA8237F7D88AE2F5AC6903561DA92E8EAE076
-284F907EFE6AF78599CD0B48361E068053D5ED94A63D095E083F18617BB263907589276D25A9
-7104518550CC3E452FCF973C61AE45CA4F44C3CB388E7450F396D0B8A62ED3F56A0905152547
-A4EBE8F464624BC1DB345697DA9057F83A776E6BC8AE2689C8AB8FC5759406C2A9AE34D19FC0
-5BC17FAEB82076DB72EC93A5F08505F01CCD7D4DCBACEB89946EC9A5447FA893D3A7F57C5E21
-D6CC93C3003F15104FD11AF5E10897B0F22DF722CF313691EF85BD5CE0D305100D5E100925F6
-7D486D8777EF15787913176B8830F03513ECE9416E6C0B15C2173D7DDC20D2026754A3485306
-DC3B243B7779B7E3CE753144A130D848741BDB5754DDB9F0571F16364782B8B463F77B6C1C1A
-7080C90F9B43649FAECD00EAA8D0D1D271CAA53585D60E363DC329F855582CBDF1CE5352C28C
-2FBB41C18F7DF0F3805658455BE45027F42C06F8D9C801D6D4A0E72C7553336D31587D20F94A
-A15B15691F7BEAFD2F6E844163BBEDAA80106E507E68C6767934B352AE63D38DB95647D2F44F
-60946CDB265AEDD05866004C2D0AA198D98D2E8F455EA057F26CB8F3F6337EF79597C5333A28
-67185F17E4F97C6165D92CBDD5FE946FC4CEBD5D6C4ED5975B7EA97012CCC6043B3363C522FB
-BA7EB7A189138A8647666037120B913CBAB9A405974B031446F08CAE6B9B7350DA3840F2B4DD
-4ED56944ED909C761AB342F16EDAF50D2E9064CF2C51817275A6EB08F64EC9A3BD38CA24784E
-0F5A4D6A14BD41BBB55D97125BD4DF34C4628494BE557F488C84B849AA3D224C19F522BE3815
-1A9F9318FDA6CFBABBAE1C8338A6C486164F8EF1BD2FAAAE01A7A58A05DBE945ED13ACB556C7
-1EFBC946CA8CE6EBA90EA429C2476C4C01167ED98DCA0076F7451971EDAD96D612C73B199688
-0EB2A60B6F8F5B7B4ED40CDE14A634208A20E889487F2ED74B7458581A54404E078E47105125
-F30300225183510B73E82322DDE19BF3FD8FD7AD8CC94811337E9BB099AE40A6CCBC81A68F51
-472F51554A445B6A99B1A4C28E920C4F9E68CCCB1F7FC6E9DFF89E611E0858CA5BA11B053978
-96ABAB2002CBA13CDBA5E2F55BD63F2EDA4286C1BA8C0AB5DAD9022208E5A70234ED8EF0E1D8
-52890548F089F588324E56F16FDB67D35F929FDAF560CC1582A8B371C8419DAE6E7C852F3BEE
-62B62E98B268E71C684F7572F11F0061A7BBAC2031AFC2243512A09C121B6323EE6B218647DD
-BD73BEB1CF3E58309C57A1BFCFCF969E4973463A164EB592E03F87BE50A7DDBE88D87157D1F2
-00DA92E0658BA92FD97BF9024D3CF634B4280B0C13E42232980BE78973C8156FB5C941A5A116
-0F0517283CD24A20E1E5CB68BFE936E3DB92B2E27BA18444F84233BB6C8F1088A102631A38E0
-98B7C7239311AAEC87EBF15C32E2B70BC9452AF087703D373353EE81BF3AC3AF94F17854A852
-502C07A5F8E4642EC5A15E2F73BD70F50EFC145198D755EAF2B1D38AE56C2D920EC8583E6BC9
-CD30C2E98638B9A77ADE825AE575F59570A7B1292384C80D589084B7198556C25B71B969F895
-46F4257ABD1E3E81BF1B5DB808CF3E8FCFE8DD59ED2FB0EA59DF88533BFCB0B74110948A9055
-FD58CC5CF757F578C48CEC499CF42C1054CC27618CA636DCAD9D07BF297018484A1D5342B0C2
-39F7C67D24188718713E6B7FFD34416753B66F48387F902912F7A291C86C8C621988219789B2
-47DBB66E5BBE15B9F2883D7F927976EDD9A8044FD025E64C7627DA9158C4F8D630C54A6C83EA
-9F64FC0BB94292428FACA748AADFBE5D45DD254A481F791E84E6C9BAC78A8D09D7FFD1268868
-1215928D0538245B743FF62DB32CC07F41C68CEAF06399A9800A8525333DE29FD4F8BD38B75B
-426842E77E6B192643EDBAD0ABC38DA56FCFF529BCE4515B5681980B18D9988CB09A9E373A6B
-1EAA1ACA1C1A1F778FDA7137DDE500C257F4C866CBF149E6FD276EC262766EF7F769DD922243
-55306193006096A96B112FE5E713948BF9DD461F4315DE8EBB4158E29289A371815E8B0778E9
-31C2B666DE328EF6072AD8222BBA38C609FA709B30CDA37C4ED546275DEEB9EE0EF08C721BA6
-0927E413BDA92B29C05F360B4BFEE15725E1F933E1A5FC82B4AC88FE256F236A4AB8950BFDA4
-4334FC4F6449031D800845B2C0691AFFF8939144DC09C12610B54BCA4775E5459C2759A89EA8
-F71CF869224449B69E6DEC871E39E98635AED90ADA7102ECD351DEA4718D852DA0C544652B86
-F132BD5C9FCF216917C4AC561F25C35EBCFB4F834C690916D45824281A46FDA3E982F343DB61
-9E760E64A70FF4FCD4E0E834DA799404139E1126F514D4ECEF828E8FCDC67E8CA1F0EEA427E6
-D19F8316416399898DB531CB2345155BEEE6F0039A58FD75872C738F71BADD99B57D77C603A5
-1B6E2DCECE8C886E36995D8A190ED111B68BC70B1AC4FF97CB3901B86542DC8A39F9FFB49BF3
-AFA8647B2107D4EDA3AC3EFA56E9CEECDE18023BD66E24521778540AE588A6F53671ED68CF8B
-20B50F9BA0048A2BBCC2B720AC5BF5CB9C86AB599FCB876965CC489C64A30657CCD138E2AFD3
-49662B6D18B8D0A84A0FF6039619CAD4513785840A647FFEB4EC5B05A0931FDE46A3AE39E25C
-41363BB19F635CF1A18DD7C52953B1F2BB686F93BBB3068D0E1CB3E49D4832067D0C99075039
-41E7C27D3242B27ECA29229DBB555AE2C85B81972DC2DD0C2355C031EC92CDB03786EB693783
-08722A6EBE4D9EAC74FC4C4B3DBD6F5DA21ABC08E82BD3D46F9EF2E8E11442736D19A044CC96
-3803BC8AF41DBD44D113BFF945F1A72AC6B9D6CEFAF21BFA49B1B817CA185C242783805F0960
-BA5F8E6CC56D9F26B6240705C1E83E78542BECFBA34FB5BA0BB0A851AD88F2AD0A88F1CAAC63
-B18E5338FB6EB1F4596031149EEAF066ABA0FA2F7A08E1D8EA0488D460B05936A234D52CB84C
-778F4A9D1B7D4B519D69697760C371F7EDA46A3CE391EF95ECF7BCFB64F1629FA13DAA7ACDEE
-3F0D00C87EF732645A374603FB559EF0CEED5E9A83BDB1899469384727C301B5BC10CA59CD60
-7E30A7B856D4E8D12EC0A7813839640FA9956E25EB657D830EF70580260579048EA425D22C6F
-85D787C85DAF8773008CA943E98B18A5CE6408E0A28FCA308A3281D9FBBB70BB44C13A4B2A93
-C6EB83CE5E8A33DA1B4183A7B95965CCCDFD6E7DD8C0BC32864D7BAFA91CA254D829B57E59F2
-D5E5711BE8357D9B7E8AB702763DECF43742FF913FCF98E7584D22CE9FD8909A4F1958BE5367
-CD78F4F42B09D2AC99E67E27491EF05B1934A8DF9781D686C4AEB3FA76B57578ED6B06AFE55D
-710915D5DBEFFEB23DAADCCB7683F190A93C1139EA8D24C9F8A9B51512387F8E9705EA9BB286
-1FCD3D47C184BC25EF19C8DE5ECA8934532DD7C64595BBD4ADE2FF914F28EDE275BC0B34955A
-F6C97E9A408D79959DCA636456D45C636B1ADAC1AC3874ACCB7E7C9A0A1234F7765D98A07BEF
-C5F1E2FC1D7CB90681D8306ECC3B31A26DC655D02CDE4810EF6A289C01F2E4F922E0A3EB607E
-53D7231EF387D1E9B370BEA637ED19E0432AA132C9B115CDECA6B5592D5648551FFF020BE84E
-1273955021B4025DBC602CBE279D59EFDBCB937F5FD449D8B637AB40E4859AA9EBA889F1C6A4
-43A594F238E8E663ADEC1E1E701D37BDD01D80ECE8A6F2EF59CBF3BBF94C9B8A61B9739A9278
-AE2940AE38A7B300FE58659906803CD67B3531A54D52DC36728CCE61DB8CBC3EB9317D6BC37F
-94B0E7232AD8F15BDDCA08C15803064BEAF7852AC4AE167F60DE97FCE39C40B68108BB14475E
-B184350A5415F7123CA34DB5B525FDB273C695356483C7E7DA0119DCD25C64936BDC95467680
-564B9C98F2D7A5672DD0E241374366ED9BCE31404A27DAF14F0BCE9E356AA7BEF5166A194312
-9644C12ED4EDA4014F5BD8E0464F26738DC0F66A7366B968BB69EE911EB6AFDA00C4521AE550
-9D7CDACD470868765966472C778615CB7F2FD4E03B334FBEF0A3C9CB4F1722E38FEDFC41611B
-E423ED01E37BD1A68CA57311C87B001508687BC6FA25FB91F093259B6A9D09DCC0047C6522F3
-1DBD42B79840FBC1D07B75F32DEB0B7803ED5894A9D3E4F463B702A8910FA1DDEBFF9D9E0E2A
-BBC89BA097D7A2FA0B99A5F3A265B3B3E5A30BACB931D912EEAFB64D67C858621AE9E94DF7B2
-7B7E886810E39461FA9E9E6178C02CE134E8E135885B381B381FA9D8E684A4DBCE48EA5970D4
-A96A74AA05DE6FF8BA30CA4C94826EE9D46DC0736B336499BC8409D1E5148C3AD185446E1D79
-AE18EA4D2F41062F759B507DEBEA62FD6A85941AC69CC008C4123FB27D819363C336582FC944
-1AC4060773AFA370C33656A86A7284D484B41148BC6B039C5F0E09E496968714E5E53AEEE4FD
-C15CC8E2F853A60C9FB2839A007FDE21493874CB1947E7A2FF425AF3BE23FF3C3F14E6ADFE5A
-09C5D405C7F3C9522930499C281D5DE05A8847338A2A09B07310C73C0CF7B1D47B676269A398
-555725A99C5368F01956CD2B121C555B2F90EE2A060A700A8A8BA1E6114481E44259E8C91886
-526894BF4A76E74F96D4EBF0832BAA40FEBAB9D8777D97C3CFEE843FC608776745954547F68E
-6F748B7D37A6E2A8970E24BFF4D53272F6B45300B35150FF8B223B1727A32AD18632E5C51142
-01E019CEEF0C397D91CBBC4314224ACCA8DCF02BD0A37C61A0D6EAA4D594771F46EEA452F059
-5599DE49E7DE1127D8088F0F1277A5990311E37459C6EA56A836A61964A2526B4902D438338B
-03166ED5A254589C3AF94D13FA0A7C741D4B2C6E121B7E772013A129B7971538C8FA1D755AE4
-70A21E8417EAAFBBA9BF92A09C43618F989BADAA04DC8EC279CE38D557776618C36C5A17A4D2
-D53D8A2FFB37197DBB1BF84F30404DB77D021C222CF1CACD6B48D35BC8CDA2D0DA99778F5EAC
-CD2F0F961E19DBCF50D923F939F95B437D6CF9532EEA2C4FED87FB58AF7282A11B2823B4983A
-8EA2ACFFF8FFEC16C922E8A8982DFF5A027DAA9E08255DE3DEB1C579C4F9174634EED315BE22
-7D0BF437812FF93E225991F533E9161D1001DD1F3AE31CD02F297D20B0A90466699E9659E8E7
-37051C305675E2F0A84A06EE2A81378A51632271A27254053E1680AD432990EAD055BAD97DCF
-59A9739A78EDEFC151C5FE1FCE8CDC9663D31DC378B6A445D85E1765238D185ADB5CAC0EADC5
-499C120AE08A628FBAFD5E39C4F20B9FD2439C0A8FF7B14C3E6CA0AB1D380ACBAC5712063DF8
-29D0633C42C042A2FCFE971B4EB4780F6473ADA604397FDEFCE615304EA31F969A30A28DD768
-3DD8B9BB8590A562EF64065A9C6A88E528EA6FBB2B4D22F16295BF3F8AC0BBFA21D3CE9FEDAA
-8BE788E56825207A0D5C5D385557F9EA9D2BD78B3F70B822F338690DBAAA06B822BBB4883892
-DEA8AC0A15D9EC1DC3880D07E44C777BB3C383D8AD466AAC8601B6CF5137EEE7A68B19BD1FA3
-CA527DA1C83B975118B8BB12F1BFDD37EB4235775B0FABECFA677988CA670A75955E56395469
-FF2308CD2C61578C57732E29214643A9AD155343A36B1D8203A94B2837B141C4FDDEA6501A04
-48292AB6145D20E712DCCFC98AC99B9DA205804598E6F3F6B3259352B1E5CA9DA09403DAD3E3
-19DF352718E02609D6EFE2F7546B556A6054A93D7BBF4967361413C610DFEF54FC553E46C1D1
-10C00EF4F66C8DA601083697D19CF3060A5572BDDEF708B3D2C38BE0F224B5AF05F234498810
-2D47BE7A2446232DFA32D7281C6E94D0D8CFA4824A5E51E758446D69A7BA3AE9EC068529E8BB
-7CDB44EDF86D2DEF84BF9E1B79E080CEE156D17CDCD6A96EED43CBDCEA3787182DFC2A31E9FF
-0C5219B25BBB5CB7A8FB9FBC507C537D0A69A5D1A9D82417A52812A2FF6DA8E63C30A5BC748A
-19CB04172D9F1B577B81EF8C973527D8D239508643C6F39BA406B5F11D7FE6C6ED366E91698D
-80DCA0C7850A313DF91EA9CCEDFD86310E371EC44935BDD9C06A3818880606BA66981386B0B1
-09640CF0247EF8138C28A61DB9C6008653C6DAEC7D9A25037EA6B1E9E987E58DC1AAB19E3A76
-B21A35E653578AB830995DF32E5D8ECACA14102E47AC2C47A3AF86CFF8C5B057F64D015C0D0C
-728A055A2C791355F064CD1C304C32CA86F82A83081A2807E076DC8407088AD35AE083D4C329
-918CD0298D1BF86098B58AE4FD0440E56F360AF452B671C0875E3867323F421927AA05C7A8C7
-A17DBF50BB8D47B4A099AAA136FC61990650E5E3E9D37AF224E98BD133BCC3347E4E85582602
-E460A590050038E4FF2A2F8F1FBEDA5FEDF2D15B9DADA67D88CFB3F09E1E0B68A239DB6013E1
-91181122DAE2B175592D61545A2FAC0B130116718AF463C894639E85024DD0B4A8808B0C25F7
-10E437203F7BFBBFABCCB557327699BB811F5085701EA2C19C0360B312D49DA83E455D3A44AA
-42C3E30D099C756F7A55ABFFFFBA78ACCC98F0794CE235FFDEB19B69EE66E4CBA1F1A803AF4C
-79CCAF80336DB931F1D8B46273E87370970AEFFA3E2D4E1539DCA3593D6A2698477B18169A32
-D1FA53DF598A9986B91BA38DAB5DC202ED9D39B46ACC1C810F631E5DA1609527903E000FFAF3
-0F2C741FE39003E5436F3BE032E3F7FDAF6571F8E2FACA2FF441DD6AD44C007CA6A09384288D
-D1DE9A1D568A92E0AC4C0AADA19C2B22F5F0472DEB827513453169E26728DCDB4B3D9925A915
-6357B48DDEBA7990E8ED0A51FE6BB33BC8E4724B2E68E1B19E4964FB828B0639BF8508EEA5F7
-130719EB1035D62DD06BB05A341D8CAB69F9019ED410BEB9C07A0A2122017FE145BF358D1ABA
-DFCFFAC09AF140921C1808A3E7CFFC44E258B1B6875B16AA9591DEE3FC9451055A38BB2587F7
-F84CC9C278A30540692AD42BEE08CE7A8BB11F89BB200FEFF57A8D10FE39B6FD01BD3C63603D
-739E33AC7F3FC9A921570B86742BF033D627AA5CD1102AC4A9C1F65A786366086940B0411BA6
-594E2359BB232BA0F853B64ECC030C0BEB2EBB49FF4C7F8EE8E5D7AA0147B8A5926ACB3DFFBA
-1CDC85AE69CB51F10AC286DBBFBCB837E74564C45CFD4BDFBAA163B9DEE5265024F1A20504C0
-420D934B4E5DDC42BE87FC8D1A3502A046A51B2A7D4DEA75F69FB5EF451F56F2865490D8D0C9
-028D0F9A268A2606AB2B55FC6113C2951F05E6A3CBD160B3B823727B5D5E6A04B7E4155992CB
-951D36A0F67E3FD567F8B01695F9746036EB59A32EE7DD5D862F98FA3CD9C3C030056DF01604
-7D25D7F42B2383EF29F176C71BDAEDC458B6C8A89B60035C44CB56BEECE9413074DA1C8F5FB3
-785DDDB9B9AEAF0786326F44A95ECED93785FDB3AE78807F6191C55D0A62B4D02F45DECD9D68
-A44241DB47C25D3A68E79EB36668882E0F19EFA3FFA1487ACAD689BD0167E417C0138AF983F0
-951E8113EBDDB3F5234A12D155470514660F0AC601BA23D92A1DB5FC2B3956BBF07D92563BE8
-6E03C9B10CC8149EE4588F30BE99A2007F2A755E56A1CFF200B1A29BFAEF136994FFA8CF3A47
-C8F19AA03AC9696D642D5A1D2A79813E517EFAB086232A019AE8F3AE5A56C71DFE8637996BEB
-E083F82A74448F44EABF84E38BCBB38FDB6BC4CF5D06A470CFAB8636F9F4CA6E5CEBEF0072D0
-39A0E96FEB5EA0F2B3BE59AAD7EB1DDDBC2A82B3AC17379AA309AAB4EB1878E4C24291774824
-5EBC4BA543EF0E45470149CC8508095D6B73ECEB2A36F7B77DAE43DC6C41E4305F99D3F503CE
-F8DBE43849B1C91A74403D2EF7DF9DF2ABBBEF13E68BF8DEF9076B1F981971ED050A71B70354
-B324A873069B2088E5C7DDC39DE69B90468E9079F743336B5049374DAF5295B25463078F8243
-8684794A4329A3B6C2063E5D01BDE8B55DB87C56DA2AFAC294022A33B1DBD5A8C2B3A156E832
-15CFCD2435979A3811B6303A9BF7E6DB2BB953727A26A45B447953D5458B7DF96BF8750B4C57
-DDA90DE912F0FB04CE3041656547E0E46FB91B4637D9CEF96138D423F3EF2E4CD1D5F3DBDB53
-32E4B7585052DA67A8FDE90CA8739B645AA29A4223204CA8E14F1CB360C91DEBD1715B8F9592
-45B37F96D1E79894AE8C4E2ACA48E9BD4421816183E62EE5571A18A6F65DE68C4AB9FE341185
-999FCCB18297EDF55C6EA99C050C67A5E24AFD8459F6CC50036AFAD04DFD75CE9D2587EBD661
-FC9A665072B0D693A0FA25EC01D905A64EE18A43332F87271CE349F9FBFF58F785EC63966B86
-EB1FCB749B637A3EF75882ECBC03E96809D57D35CD846973282AE0D354F6C5DDFC8CF92BA25E
-DFD75593CF693309A1AE38B793842A7DDD0F60DEFFD7707DC2B0951C8733E2AE6D964D5BBE1C
-FF5C390097F05098DEB1BC717DCBE40836518CEFB0E6995C1B6846B19ED1E9888EC6D904F260
-744B18794B64BED634D6596B3D383DB12C8EDA8272BE4B836893B5A6DF5A502CD8D710FE5622
-D77D369D329E327BEFF59984C0362802A950B24B8C47D77B7574157D8D957F3DA86982C3F78B
-3F5E3A2A935B87F9274FA52E460B7E9AD8AD7267F4C670C0DCD915F63939FEFFDAB8C9F068F6
-D53CE40F0F8F0208C1B7E609E4B8BB7CA161F0D40FDCCB8B7D73A1CC663E6BE4C15CB3CE7409
-71266E7F8D355E49B73D34BD500CE9FA93EDD5C71C9483C1102FFBD97DB6F5BAD077406D4869
-7F16F9DD522FBD237B8C9F77DC95A5173E4AEC67013E96760214E3F069F83673B6C9F64039FE
-5BDDC965B000B9F33B9B2610E51F82A0BE36D0D09DDDBCD4EE2509C0014EDBCB49E7F9AB87DA
-7906259D88F02FF6D7AFA264F1D736646566D9C890598068E7C725A26F9399ACABF8DB7BAAF2
-FC973225A0A8E31458AEDEEB49809C0D165E2E92B6D1931746E77D56C3135C3EAF2FCE7B6A70
-2A451916946116FCEFEEC611A164E70700331E8EA71E16F51A0519220A791D945728C5E1A228
-ED254A33106BCBE285AA78EF33FFBFC25B21BAC78D758C0479B769F2DD27010FD5ECFBCAF6BA
-953E7E4AE24B652D4D2E45A91AC753FF7446936C0C5AEEF1934447FB6B3BC6755D2A3AABAE2B
-41554C9B539A6756AF2DCF6737934D63F5442F37624E7736A6CD405C3C34C3680CB50D8FC80F
-5102F2FF74770ED2D34ABE11B659AA4054E7F6F52989324FC1BB330DF7E40B957CFFA3BD7465
-755CAC876CEEB324743B1B7C71BA25C57D21E1D9E87377166D3BA77C46758D6D925B4B00F898
-1C75AD7AEE6B4E20336CF779B6CE436D633809C469923F100D550EF12796A14662BBCA7498E1
-333E2E2AF88C37326B10D699E2CB8AC02DC9A001E2E41A4E7268E06BFC13025F88620EFA27A4
-95A4310C77550B978F1B6113A3C923B16892B42CE4DB2000163F94C64939DCA5711E0B1C46A4
-997EF5BEB061BA76D13A620146ED21637317E3467338415AE1BA3CAFF3D4CEA0DA75621B315E
-E13CF84B5C458957E6905B9F50710C6E2CE772FAA2D29A87E8468FC8519CBF2E7F3B89C14715
-25FE9B5B21A4D7B458B79DA6D40D683B0FD02BB1E5EB8E9B78A88BAA4F420EE0BF6473C35CCE
-DECBABE217771071B138F1419CFD1076537AC6BB1D52F44DB2B01DC9EA6FE01B2402DBD54E2A
-69E631A9E5A8E9D39FF9F9AE9DFEB38A1075AD69AFEAA4421AD036C2ED7E6783FBC7FBD8D10F
-0B1E19BEB4AB360E2052023ADC3463D67B938CA80BEB7A38519BAC591B6D101FFE78EA6A397C
-1E37D2809C0F9BF2FB5FDC9651CC6D1EA0CF473CE12183C0375532E84EEEA60D940C0354DF4B
-513E7272DB356AA408D36EB5AE41BA7229E55CF7A3E8586275DC49386D9DCDA5005A35C15497
-AA66FF8B76DDE02F1C7819B25B1EAF8CEFD9E2DF881647D12547434A9CDC10DD5F57E2F7735B
-F351079F44C58E9B6A6E50A7AF0A09F441B6F1D6A3C7DE172BEA39225E18A30EB9D283D2A775
-BF15CEEF5E32F96E57C9CC6B325A770000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-0000000000000000000000000000000000000000000000000000000000000000
-cleartomark
-
-
-
-%%EndFont
-TeXDict begin 39139632 55387786 1000 8000 8000 (bind-wmm-2.dvi)
-@start /Fa 140[344 344 1[443 443 443 2[393 2[443 4[393
-11[738 2[492 443 2[541 1[591 738 492 2[295 2[541 541
-1[591 1[541 8[443 3[443 1[443 443 2[221 1[221 2[295 295
-40[{TeXBase1Encoding ReEncodeFont}28 885.568 /Times-Italic
-rf /Fb 244[604 11[{}1 885.568 /CMMI8 rf /Fc 71[295 33[443
-28[443 1[639 443 443 246 344 295 1[443 443 443 689 246
-443 1[246 443 443 295 393 443 393 443 393 3[295 1[295
-541 5[541 492 591 1[492 639 639 787 541 1[344 295 1[639
-492 541 1[591 1[639 6[246 2[443 443 443 1[443 443 443
-443 1[221 1[221 27[295 16[{TeXBase1Encoding ReEncodeFont}53
-885.568 /Times-Roman rf /Fd 140[474 541 3[677 7[405 541
-1[541 16[879 82[{TeXBase1Encoding ReEncodeFont}7 1217.66
-/Times-Bold rf /Fe 192[796 63[{}1 996.264 /MSBM10 rf
-/Ff 195[640 60[{}1 664.176 /CMMI6 rf /Fg 139[458 4[589
-655 982 327 2[327 3[541 17[655 4[921 78[{}9 996.264 /CMBX9
-rf /Fh 139[332 4[498 554 830 277 2[277 3[442 17[554 4[719
-78[{TeXBase1Encoding ReEncodeFont}9 996.264 /Times-Bold
-rf /Fi 149[284 2[512 512 5[626 4[683 57[1024 33[{}6 996.264
-/CMSY9 rf /Fj 194[796 2[284 16[398 398 39[640{}5 996.264
-/CMR9 rf /Fk 134[442 3[498 277 388 388 1[498 498 498
-719 277 2[277 1[498 277 442 498 442 498 98[{
-TeXBase1Encoding ReEncodeFont}17 996.264 /Times-Italic
-rf /Fl 134[497 472 2[549 340 418 431 1[523 523 575 836
-5[470 1[470 523 2[523 7[760 760 3[732 9[537 74[{}18 996.264
-/CMTI9 rf /Fm 145[529 4[244 4[529 1[529 98[{}4 996.264
-/CMSS9 rf /Fn 139[523 1[523 5[523 6[523 1[523 99[{}5
-996.264 /CMTT9 rf /Fo 157[439 8[594 846 27[796 284 284
-46[656 11[{}7 996.264 /CMMI9 rf /Fp 197[234 14[621 43[{}2
-664.176 /CMR6 rf /Fq 104[996 28[442 498 498 719 498 498
-277 388 332 498 498 498 498 775 277 498 277 277 498 498
-332 442 498 442 498 442 3[332 1[332 3[940 1[719 609 554
-665 2[719 719 886 2[388 332 719 1[554 2[665 1[719 6[277
-6[498 498 498 2[249 332 249 2[332 332 332 36[554 2[{
-TeXBase1Encoding ReEncodeFont}54 996.264 /Times-Roman
-rf /Fr 207[243 44[424 3[{}2 664.176 /CMSY6 rf /Fs 212[732
-43[{}1 885.568 /CMR8 rf /Ft 252[470 3[{}1 885.568 /CMSY8
-rf /Fu 136[879 1[609 338 474 405 1[609 609 609 947 338
-609 1[338 3[541 609 541 1[541 6[744 6[677 2[677 1[879
-7[677 70[{TeXBase1Encoding ReEncodeFont}21 1217.66 /Times-Roman
-rf /Fv 134[996 1[1439 1[1108 664 775 885 1[1108 996 1108
-1[554 2[554 1[996 664 885 1108 885 1108 996 13[1108 37[664
-45[{TeXBase1Encoding ReEncodeFont}20 1992.53 /Times-Bold
-rf end
-%%EndProlog
-%%BeginSetup
-%%Feature: *Resolution 8000dpi
-TeXDict begin
-%%PaperSize: A4
- end
-%%EndSetup
-%%Page: 1 1
-TeXDict begin 1 0 bop 0 TeXcolorgray Black 0 TeXcolorgray
-0 TeXcolorgray 0 TeXcolorgray 8949 2148 a Fv(Specifying)497
-b(r)-36 b(eal-w)-20 b(orld)497 b(binding)g(structur)-36
-b(es)10991 4558 y Fu(Susmit)305 b(Sarkar)17857 4118 y
-Ft(\003)19601 4558 y Fu(Peter)f(Se)-30 b(well)25691 4118
-y Ft(\003)27435 4558 y Fu(Francesco)302 b(Zappa)i(Nardelli)40012
-4118 y Fs(+)14933 5910 y Fr(\003)15413 6333 y Fq(Uni)-25
-b(v)-15 b(ersity)249 b(of)g(Cambridge)27576 5910 y Fp(+)28253
-6333 y Fq(INRIA)f(Rocquencourt)p Black 0 TeXcolorgray
-1 TeXcolorgray 0 TeXcolorgray 1 TeXcolorgray 0 TeXcolorgray
-0 TeXcolorgray 0 TeXcolorgray 1 TeXcolorgray 0 TeXcolorgray
-0 TeXcolorgray 0 TeXcolorgray 0 TeXcolorgray 0 TeXcolorgray
-0 TeXcolorgray 0 TeXcolorgray 0.25 TeXcolorgray 0 TeXcolorgray
-0.5 TeXcolorgray 0 TeXcolorgray -672 10757 a(Representing)271
-b(binding)g(structures)f(is)g(a)h(central)g(challenge)h(in)e(mech-)
--2000 11864 y(anizing)366 b(the)g(theory)g(of)f(programming)h
-(languages.)g(T)-80 b(w)-10 b(o)366 b(f)-10 b(acets)365
-b(of)h(the)-2000 12971 y(problem)408 b(should)f(be)g(distinguished:)g
-(on)h(the)f(one)h(hand,)g(one)f(needs)h(a)-2000 14078
-y(speci\002cation)322 b(language)h(for)e(e)-15 b(xpress)321
-b(the)g(comple)-15 b(x)323 b(binding)f(structures)-2000
-15185 y(found)250 b(in)f(realistic)h(programming)g(languages,)h(while)f
-(on)g(the)g(other)f(hand)-2000 16292 y(one)409 b(needs)h(an)f
-(implementation)h(technique)h(to)e(represent)g(such)g(struc-)-2000
-17399 y(tures)327 b(in)g(proof)f(assistants.)g(Most)g(research)i(ef)-25
-b(fort)326 b(to)h(date)h(has)f(focused)-2000 18506 y(on)272
-b(the)h(second)g(issue,)e(limiting)i(to)f(the)g(case)h(in)f(which)h(a)g
-(single)f(v)-25 b(ariable)-2000 19613 y(binds)276 b(in)f(a)i(single)e
-(subterm.)h(Notable)h(e)-15 b(xceptions)277 b(are)f(FreshOCaml)f([3])
--2000 20720 y(and)236 b(C)p Fo(\013)q Fq(ml)f([1],)h(b)-20
-b(ut)236 b(neither)g(pro)-15 b(vides)236 b(a)f(complete)i(solution)f
-(\(e)-15 b(xpressi)-25 b(v)-15 b(e)-2000 21827 y(enough)355
-b(to)g(deal)g(with)g(the)g(v)-25 b(ariety)355 b(of)f(binding)h(found)g
-(in)f(the)h(wild)g(b)-20 b(ut)-2000 22934 y(also)249
-b(simple)g(and)h(intuiti)-25 b(v)-15 b(e\).)-672 24040
-y(W)-80 b(e)450 b(ar)-18 b(gue)449 b(that)g(lightweight)g(annotations)g
-(of)f(productions)h(in)g(the)-2000 25147 y(syntax)358
-b(grammar)h(are)g(suf)-25 b(\002cient)358 b(to)h(speci\002y)f(comple)
--15 b(x)360 b(binding)f(struc-)-2000 26254 y(tures,)366
-b(including)i(multiple)f Fn(let)523 b(rec)p Fq(s,)366
-b(OCaml')-55 b(s)367 b(or)-20 b(-patterns,)366 b(depen-)-2000
-27361 y(dent)282 b(records,)g(and)g(e)-25 b(v)-15 b(en)283
-b(Join)e(Calculus)h(patterns.)f(The)h(\002rst)f(annotation)-2000
-28468 y Fm(bind)171 b Fl(mse)247 b Fm(in)171 b Fl(ntr)472
-b Fq(lets)366 b(one)g(specify)f(that)h(the)g(v)-25 b(ariables)366
-b(de\002ned)g(in)g(the)-2000 29575 y(set)346 b Fk(mse)h
-Fq(bind)g(in)f(nonterminals)h(of)f(the)h(production.)g(Here)g
-Fk(mse)g Fq(ranges)-2000 30682 y(o)-15 b(v)g(er)289 b(sets)g(of)g(meta)
--20 b(v)-25 b(ariables,)290 b(which)g(can)g(be)f(speci\002ed)h(e)-15
-b(xplicitly)290 b(or)f(by)-2000 31789 y(auxiliary)312
-b(functions,)f(as)g(in)g(the)h(Ocaml)g(or)-20 b(-pattern)311
-b(language)i(fragment)-2000 32896 y(belo)-25 b(w)-65
-b(.)-1385 34370 y Fl(exp)1170 b Fj(::=)1108 b Fl(x)1791
-35495 y Fi(j)1647 b Fj(\()285 b Fl(exp)348 b Fo(;)455
-b Fl(exp)8485 35072 y Fr(0)9068 35495 y Fj(\))1791 36619
-y Fi(j)1647 b Fh(let)285 b Fl(p)-52 b(at)662 b Fj(=)569
-b Fl(exp)348 b Fh(in)284 b Fl(exp)12875 36196 y Fr(0)15387
-36619 y Fm(bind)h Fo(b)p Fj(\()p Fl(p)-52 b(at)93 b Fj(\))284
-b Fm(in)h Fl(exp)23035 36196 y Fr(0)-1385 37726 y Fl(p)-52
-b(at)1301 b Fj(::=)1791 38850 y Fi(j)1647 b Fj(\()285
-b Fl(p)-52 b(at)378 b Fi(jj)285 b Fl(p)-52 b(at)8398
-38427 y Fr(0)8980 38850 y Fj(\))6009 b Fo(b)284 b Fj(=)h
-Fo(b)p Fj(\()p Fl(p)-52 b(at)93 b Fj(\))228 b Fi([)f
-Fo(b)p Fj(\()p Fl(p)-52 b(at)23256 38427 y Fr(0)23553
-38850 y Fj(\))1791 39975 y Fi(j)1647 b Fj(\()285 b Fl(p)-52
-b(at)378 b Fo(;)455 b Fl(p)-52 b(at)8284 39551 y Fr(0)8866
-39975 y Fj(\))6123 b Fo(b)284 b Fj(=)h Fo(b)p Fj(\()p
-Fl(p)-52 b(at)93 b Fj(\))228 b Fi([)f Fo(b)p Fj(\()p
-Fl(p)-52 b(at)23256 39551 y Fr(0)23553 39975 y Fj(\))1791
-41082 y Fi(j)1647 b Fh(Some)285 b Fl(x)8584 b Fo(b)284
-b Fj(=)h Fl(x)1791 42189 y Fi(j)1647 b Fh(None)9452 b
-Fo(b)284 b Fj(=)h Fi(fg)-2000 43886 y Fq(This)422 b(speci\002es)i(that)
-f(the)h(link)-10 b(ed)424 b(occurences)g(of)f Fl(x)544
-b Fq(belo)-25 b(w)424 b(must)f(v)-25 b(ary)-2000 44993
-y(together:)-170 46730 y Fg(let)171 b Fj(\(\()g Fg(None)p
-Fo(;)342 b Fg(Some)8565 46730 y
- tx@Dict begin tx@NodeDict begin {3.87498 0.0 5.35074 2.67537 1.93748
-} false /N@x1 16 {InitRnode } NewNode end end
- 8565 46730 a 1 0 0 TeXcolorrgb
-0 TeXcolorgray 1 0 0 TeXcolorrgb Fl(x)p 0 TeXcolorgray
-120 w Fj(\))171 b Fi(jj)g Fj(\()p Fg(Some)13802 46730
-y
- tx@Dict begin tx@NodeDict begin {3.87498 0.0 5.35074 2.67537 1.93748
-} false /N@x2 16 {InitRnode } NewNode end end
- 13802 46730 a 1 0 0 TeXcolorrgb 0 TeXcolorgray 1 0 0
-TeXcolorrgb Fl(x)p 0 TeXcolorgray 120 w Fo(;)342 b Fg(None)p
-Fj(\)\))456 b(=)p 0 0 1 TeXcolorrgb 0 TeXcolorgray 0 0 1
-TeXcolorrgb 455 w Fl(x)p 0 TeXcolorgray 291 w Fg(in)22145
-46730 y
- tx@Dict begin tx@NodeDict begin {3.87498 0.0 5.35074 2.67537 1.93748
-} false /N@x3 16 {InitRnode } NewNode end end
- 22145 46730 a 1 0 0 TeXcolorrgb 0 TeXcolorgray
-1 0 0 TeXcolorrgb Fl(x)p 0 TeXcolorgray 1 0 0 TeXcolorrgb
-0 TeXcolorgray 1 0 0 TeXcolorrgb 22737 46730 a
- tx@Dict begin gsave STV newpath 0.8 SLW 1 0 0 setrgbcolor /ArrowA
-{ moveto } def /ArrowB { } def /NCLW CLW def tx@NodeDict begin 0.0
-0.0 neg 3.0 3.0 0 0 /N@x1 /N@x2 InitNC { tx@Dict begin /Lineto /lineto
-load def false pop end /AngleA -90. def /AngleB -90. def /ArmA 3.0
-def /ArmB 3.0 def /ArmTypeA 0 def /ArmTypeB 0 def /AngleB -90. def
-NCBar } if end gsave 0.8 SLW 1 0 0 setrgbcolor 0 setlinecap stroke
- grestore grestore end
- 22737
-46730 a 22737 46730 a
- tx@Dict begin gsave STV newpath 0.8 SLW 1 0 0 setrgbcolor /ArrowA
-{ moveto } def /ArrowB { } def /NCLW CLW def tx@NodeDict begin 0.0
-0.0 neg 3.0 3.0 0 0 /N@x2 /N@x3 InitNC { tx@Dict begin /Lineto /lineto
-load def false pop end /AngleA -90. def /AngleB -90. def /ArmA 3.0
-def /ArmB 3.0 def /ArmTypeA 0 def /ArmTypeB 0 def /AngleB -90. def
-NCBar } if end gsave 0.8 SLW 1 0 0 setrgbcolor 0 setlinecap stroke
- grestore grestore end
- 22737 46730 a 0 TeXcolorgray -672
-48467 a Fq(W)-80 b(e)257 b(will)f(e)-15 b(xplain)257
-b(the)f(process)g(of)f(gi)-25 b(ving)257 b(meaning)g(to)f(binding)g
-(spec-)-2000 49574 y(i\002cations)279 b(by)f(considering)h(a)g
-(fragment)g(of)f(a)g(grammar)h(for)f(System)h(F)23639
-49685 y Ff(<)p Fp(:)-2000 50681 y Fq(for)446 b(conte)-15
-b(xts)448 b(and)f(judgments.)g(\(In)f(this)h(particular)g(de)-25
-b(v)-15 b(elopment)448 b(we)-2000 51788 y(choose)382
-b(to)f(ha)-20 b(v)-15 b(e)383 b(the)f(v)-25 b(ariables)382
-b(in)f(the)h Fj(\000)f Fq(of)g(a)h(judgment)g(bind)g(in)g(the)-2000
-52895 y(judgment)250 b(body)-65 b(.\))-1336 54571 y Fj(\000)249
-b(::=)-340 55678 y Fi(j)499 b Fe(?)6164 b Fq(Tdom)284
-b Fj(=)h Fi(fg)7403 56785 y Fq(tdom)g Fj(=)g Fi(fg)-340
-57892 y(j)499 b Fj(\000)p Fo(;)342 b Fl(X)157 b Fo(<)p
-Fj(:)q Fl(T)2964 b Fq(Tdom)284 b Fj(=)h Fq(Tdom)p Fj(\(\000\))228
-b Fi([)f Fl(X)7403 58999 y Fq(tdom)285 b Fj(=)g Fi(fg)7403
-60106 y Fm(bind)f Fq(Tdom)q Fj(\(\000\))g Fm(in)h Fl(T)-340
-61213 y Fi(j)499 b Fj(\000)p Fo(;)342 b Fl(x)120 b Fj(:)p
-Fl(T)4086 b Fq(Tdom)284 b Fj(=)h Fq(Tdom)p Fj(\(\000\))7403
-62319 y Fq(tdom)g Fj(=)g Fq(tdom)p Fj(\(\000\))228 b
-Fi([)f Fl(x)7403 63426 y Fm(bind)284 b Fq(Tdom)q Fj(\(\000\))g
-Fm(in)h Fl(T)-1336 65163 y(Judgments)331 b Fj(::=)1330
-b Fo(:)171 b(:)g(:)-340 66270 y Fi(j)499 b Fj(\000)455
-b Fi(`)g Fl(t)264 b Fj(:)171 b Fl(T)2993 b Fm(bind)284
-b Fq(Tdom)q Fj(\(\000\))g Fm(in)h Fl(t)7403 67377 y Fm(bind)f
-Fq(Tdom)q Fj(\(\000\))g Fm(in)h Fl(T)7403 68484 y Fm(bind)f
-Fq(tdom)q Fj(\(\000\))g Fm(in)h Fl(t)-2000 71288 y Fq(The)241
-b(k)-10 b(e)-15 b(y)242 b(b)-20 b(uilding)242 b(block)g(for)f(our)g
-(de\002nitions)h(is)f(that)g(of)g(sets)g(of)g(v)-25 b(ariable)-2000
-72395 y(occurences)197 b(that)g(are)f(meant)h(to)f(alpha-v)-25
-b(ary)197 b(together)-55 b(.)197 b(Thus)e(for)h(e)-15
-b(xample,)p 0 TeXcolorgray 0 TeXcolorgray 27224 10757
-a(the)327 b(link)-10 b(ed)328 b(occurences)h(of)e(concrete)h(v)-25
-b(ariables)328 b(must)f(v)-25 b(ary)328 b(together)g(in)27224
-11864 y(the)249 b(follo)-25 b(wing)250 b(abstract)f(syntax)g(term)h
-(for)e(a)i(conte)-15 b(xt.)32494 13528 y Fj(\000)284
-b(=)34499 13528 y
- tx@Dict begin tx@NodeDict begin {6.14998 0.0 8.28621 4.1431 3.07498
-} false /N@X1 16 {InitRnode } NewNode end end
- 34499 13528 a 1 0 0 TeXcolorrgb 0 TeXcolorgray
-1 0 0 TeXcolorrgb Fl(X)p 0 TeXcolorgray 157 w Fo(<)p
-Fj(:)q Fl(T)-78 b(op)61 b Fo(;)38884 13528 y
- tx@Dict begin tx@NodeDict begin {6.14998 0.0 8.61678 4.30838 3.07498
-} false /N@Y1 16 {InitRnode } NewNode end end
- 38884 13528
-a 0 0 1 TeXcolorrgb 0 TeXcolorgray 0 0 1 TeXcolorrgb
-Fl(Y)p 0 TeXcolorgray 194 w Fo(<)p Fj(:)40919 13528 y
- tx@Dict begin tx@NodeDict begin {6.14998 0.0 8.28621 4.1431 3.07498
-} false /N@X2 16 {InitRnode } NewNode end end
-
-40919 13528 a 1 0 0 TeXcolorrgb 0 TeXcolorgray 1 0 0
-TeXcolorrgb Fl(X)p 0 TeXcolorgray 157 w Fi(!)42860 13528
-y
- tx@Dict begin tx@NodeDict begin {6.14998 0.0 8.28621 4.1431 3.07498
-} false /N@X3 16 {InitRnode } NewNode end end
- 42860 13528 a 1 0 0 TeXcolorrgb 0 TeXcolorgray 1 0 0
-TeXcolorrgb Fl(X)p 0 TeXcolorgray 157 w Fo(;)342 b Fl(x)120
-b Fj(:)45280 13528 y
- tx@Dict begin tx@NodeDict begin {6.14998 0.0 8.3451 4.17255 3.07498
-} false /N@X4 16 {InitRnode } NewNode end end
- 45280 13528 a 1 0 0 TeXcolorrgb
-0 TeXcolorgray 1 0 0 TeXcolorrgb Fo(X)p 0 TeXcolorgray
-78 w(;)341 b Fl(y)89 b Fj(:)47699 13528 y
- tx@Dict begin tx@NodeDict begin {6.14998 0.0 7.42424 3.71211 3.07498
-} false /N@Y2 16 {InitRnode } NewNode end end
- 47699 13528
-a 0 0 1 TeXcolorrgb 0 TeXcolorgray 0 0 1 TeXcolorrgb
-Fo(Y)p 0 TeXcolorgray 1 0 0 TeXcolorrgb 0 TeXcolorgray
-1 0 0 TeXcolorrgb 48521 13528 a
- tx@Dict begin gsave STV newpath 0.8 SLW 1 0 0 setrgbcolor /ArrowA
-{ moveto } def /ArrowB { } def /NCLW CLW def tx@NodeDict begin 0.0
-0.0 neg 3.0 3.0 0 0 /N@X1 /N@X2 InitNC { tx@Dict begin /Lineto /lineto
-load def false pop end /AngleA 90. def /AngleB 90. def /ArmA 3.0 def
-/ArmB 3.0 def /ArmTypeA 0 def /ArmTypeB 0 def /AngleB 90. def NCBar
- } if end gsave 0.8 SLW 1 0 0 setrgbcolor 5.0 3.0 0.0 0.0 0 0 add
-DashLine grestore grestore end
- 48521 13528 a 48521 13528
-a
- tx@Dict begin gsave STV newpath 0.8 SLW 1 0 0 setrgbcolor /ArrowA
-{ moveto } def /ArrowB { } def /NCLW CLW def tx@NodeDict begin 0.0
-0.0 neg 3.0 3.0 0 0 /N@X2 /N@X3 InitNC { tx@Dict begin /Lineto /lineto
-load def false pop end /AngleA 90. def /AngleB 90. def /ArmA 3.0 def
-/ArmB 3.0 def /ArmTypeA 0 def /ArmTypeB 0 def /AngleB 90. def NCBar
- } if end gsave 0.8 SLW 1 0 0 setrgbcolor 5.0 3.0 0.0 0.0 0 0 add
-DashLine grestore grestore end
- 48521 13528 a 48521 13528 a
- tx@Dict begin gsave STV newpath 0.8 SLW 1 0 0 setrgbcolor /ArrowA
-{ moveto } def /ArrowB { } def /NCLW CLW def tx@NodeDict begin 0.0
-0.0 neg 3.0 3.0 0 0 /N@X3 /N@X4 InitNC { tx@Dict begin /Lineto /lineto
-load def false pop end /AngleA 90. def /AngleB 90. def /ArmA 3.0 def
-/ArmB 3.0 def /ArmTypeA 0 def /ArmTypeB 0 def /AngleB 90. def NCBar
- } if end gsave 0.8 SLW 1 0 0 setrgbcolor 5.0 3.0 0.0 0.0 0 0 add
-DashLine grestore grestore end
- 48521 13528 a 0 TeXcolorgray
-0 0 1 TeXcolorrgb 0 TeXcolorgray 0 0 1 TeXcolorrgb 48521
-13528 a
- tx@Dict begin gsave STV newpath 0.8 SLW 0 0 1 setrgbcolor /ArrowA
-{ moveto } def /ArrowB { } def /NCLW CLW def tx@NodeDict begin 0.0
-0.0 neg 3.0 3.0 0 0 /N@Y1 /N@Y2 InitNC { tx@Dict begin /Lineto /lineto
-load def false pop end /AngleA -90. def /AngleB -90. def /ArmA 3.0
-def /ArmB 3.0 def /ArmTypeA 0 def /ArmTypeB 0 def /AngleB -90. def
-NCBar } if end gsave 0.8 SLW 0 0 1 setrgbcolor 5.0 3.0 0.0 0.0 0
-0 add DashLine grestore grestore end
- 48521 13528 a 0 TeXcolorgray 27224 15191 a Fq(W)-80
-b(e)247 b(call)f(these)g(sets)g(the)g Fk(open)h(binding)f(sets)p
-Fq(,)g(and)g(we)h(can)f(gi)-25 b(v)-15 b(e)247 b(a)f(compo-)27224
-16298 y(sitional)274 b(calculation)j(of)d(them)i(o)-15
-b(v)g(er)275 b(the)g(structure)g(of)g(the)g(term)g(by)g(look-)27224
-17405 y(ing)347 b(at)g(the)g(binding)g(speci\002cations.)g(The)-15
-b(y)347 b(are)h(not)e(al)-10 b(w)g(ays)348 b(allo)-25
-b(wed)348 b(to)27224 18512 y(actually)285 b(v)-25 b(ary)-65
-b(,)284 b(such)g(as)g(in)f(the)h(e)-15 b(xample)286 b(conte)-15
-b(xt)284 b(abo)-15 b(v)g(e.)285 b(The)f(v)-25 b(ariables)27224
-19619 y(bound)377 b(in)f(the)h(conte)-15 b(xt)378 b(can)f(only)g(be)g
-(alpha-v)-25 b(aried)378 b(when)f(placed)h(in)e(a)27224
-20726 y(judgment,)250 b(such)f(as)g(in)31714 22390 y
- tx@Dict begin tx@NodeDict begin {6.14998 0.0 8.28621 4.1431 3.07498
-} false /N@X1 16 {InitRnode } NewNode end end
-
-31714 22390 a 1 0 0 TeXcolorrgb 0 TeXcolorgray 1 0 0
-TeXcolorrgb Fl(X)p 0 TeXcolorgray 157 w Fo(<)p Fj(:)q
-Fl(T)-78 b(op)62 b Fo(;)36100 22390 y
- tx@Dict begin tx@NodeDict begin {6.14998 0.0 8.61678 4.30838 3.07498
-} false /N@Y1 16 {InitRnode } NewNode end end
- 36100 22390 a 0 0 1
-TeXcolorrgb 0 TeXcolorgray 0 0 1 TeXcolorrgb Fl(Y)p 0
-TeXcolorgray 194 w Fo(<)p Fj(:)38134 22390 y
- tx@Dict begin tx@NodeDict begin {6.14998 0.0 8.28621 4.1431 3.07498
-} false /N@X2 16 {InitRnode } NewNode end end
- 38134 22390
-a 1 0 0 TeXcolorrgb 0 TeXcolorgray 1 0 0 TeXcolorrgb
-Fl(X)p 0 TeXcolorgray 158 w Fi(!)40076 22390 y
- tx@Dict begin tx@NodeDict begin {6.14998 0.0 8.28621 4.1431 3.07498
-} false /N@X3 16 {InitRnode } NewNode end end
- 40076
-22390 a 1 0 0 TeXcolorrgb 0 TeXcolorgray 1 0 0 TeXcolorrgb
-Fl(X)p 0 TeXcolorgray 157 w Fo(;)41619 22390 y
- tx@Dict begin tx@NodeDict begin {3.87498 0.0 5.35074 2.67537 1.93748
-} false /N@x1 16 {InitRnode } NewNode end end
- 41619
-22390 a 0.5 TeXcolorgray 0 TeXcolorgray 0.5 TeXcolorgray
-Fl(x)p 0 TeXcolorgray 120 w Fj(:)42495 22390 y
- tx@Dict begin tx@NodeDict begin {6.14998 0.0 8.3451 4.17255 3.07498
-} false /N@X4 16 {InitRnode } NewNode end end
- 42495
-22390 a 1 0 0 TeXcolorrgb 0 TeXcolorgray 1 0 0 TeXcolorrgb
-Fo(X)p 0 TeXcolorgray 78 w(;)342 b Fl(y)89 b Fj(:)44915
-22390 y
- tx@Dict begin tx@NodeDict begin {6.14998 0.0 7.42424 3.71211 3.07498
-} false /N@Y2 16 {InitRnode } NewNode end end
- 44915 22390 a 0 0 1 TeXcolorrgb 0 TeXcolorgray
-0 0 1 TeXcolorrgb Fo(Y)p 0 TeXcolorgray 512 w Fi(`)46931
-22390 y
- tx@Dict begin tx@NodeDict begin {3.87498 0.0 5.35074 2.67537 1.93748
-} false /N@x2 16 {InitRnode } NewNode end end
- 46931 22390 a 0.5 TeXcolorgray 0 TeXcolorgray
-0.5 TeXcolorgray Fl(x)p 0 TeXcolorgray 405 w Fj(:)48377
-22390 y
- tx@Dict begin tx@NodeDict begin {6.14998 0.0 8.3451 4.17255 3.07498
-} false /N@X5 16 {InitRnode } NewNode end end
- 48377 22390 a 1 0 0 TeXcolorrgb 0 TeXcolorgray
-1 0 0 TeXcolorrgb Fo(X)p 0 TeXcolorgray 1 0 0 TeXcolorrgb
-0 TeXcolorgray 1 0 0 TeXcolorrgb 49300 22390 a
- tx@Dict begin gsave STV newpath 0.8 SLW 1 0 0 setrgbcolor /ArrowA
-{ moveto } def /ArrowB { } def /NCLW CLW def tx@NodeDict begin 0.0
-0.0 neg 3.0 3.0 0 0 /N@X1 /N@X2 InitNC { tx@Dict begin /Lineto /lineto
-load def false pop end /AngleA 90. def /AngleB 90. def /ArmA 3.0 def
-/ArmB 3.0 def /ArmTypeA 0 def /ArmTypeB 0 def /AngleB 90. def NCBar
- } if end gsave 0.8 SLW 1 0 0 setrgbcolor 0 setlinecap stroke grestore
- grestore end
- 49300
-22390 a 49300 22390 a
- tx@Dict begin gsave STV newpath 0.8 SLW 1 0 0 setrgbcolor /ArrowA
-{ moveto } def /ArrowB { } def /NCLW CLW def tx@NodeDict begin 0.0
-0.0 neg 3.0 3.0 0 0 /N@X2 /N@X3 InitNC { tx@Dict begin /Lineto /lineto
-load def false pop end /AngleA 90. def /AngleB 90. def /ArmA 3.0 def
-/ArmB 3.0 def /ArmTypeA 0 def /ArmTypeB 0 def /AngleB 90. def NCBar
- } if end gsave 0.8 SLW 1 0 0 setrgbcolor 0 setlinecap stroke grestore
- grestore end
- 49300 22390 a 49300 22390 a
- tx@Dict begin gsave STV newpath 0.8 SLW 1 0 0 setrgbcolor /ArrowA
-{ moveto } def /ArrowB { } def /NCLW CLW def tx@NodeDict begin 0.0
-0.0 neg 3.0 3.0 0 0 /N@X3 /N@X4 InitNC { tx@Dict begin /Lineto /lineto
-load def false pop end /AngleA 90. def /AngleB 90. def /ArmA 3.0 def
-/ArmB 3.0 def /ArmTypeA 0 def /ArmTypeB 0 def /AngleB 90. def NCBar
- } if end gsave 0.8 SLW 1 0 0 setrgbcolor 0 setlinecap stroke grestore
- grestore end
- 49300
-22390 a 49300 22390 a
- tx@Dict begin gsave STV newpath 0.8 SLW 1 0 0 setrgbcolor /ArrowA
-{ moveto } def /ArrowB { } def /NCLW CLW def tx@NodeDict begin 0.0
-0.0 neg 3.0 3.0 0 0 /N@X4 /N@X5 InitNC { tx@Dict begin /Lineto /lineto
-load def false pop end /AngleA 90. def /AngleB 90. def /ArmA 3.0 def
-/ArmB 3.0 def /ArmTypeA 0 def /ArmTypeB 0 def /AngleB 90. def NCBar
- } if end gsave 0.8 SLW 1 0 0 setrgbcolor 0 setlinecap stroke grestore
- grestore end
- 49300 22390 a 0 TeXcolorgray 0 0 1
-TeXcolorrgb 0 TeXcolorgray 0 0 1 TeXcolorrgb 49300 22390
-a
- tx@Dict begin gsave STV newpath 0.8 SLW 0 0 1 setrgbcolor /ArrowA
-{ moveto } def /ArrowB { } def /NCLW CLW def tx@NodeDict begin 0.0
-0.0 neg 3.0 3.0 0 0 /N@Y1 /N@Y2 InitNC { tx@Dict begin /Lineto /lineto
-load def false pop end /AngleA -90. def /AngleB -90. def /ArmA 3.0
-def /ArmB 3.0 def /ArmTypeA 0 def /ArmTypeB 0 def /AngleB -90. def
-NCBar } if end gsave 0.8 SLW 0 0 1 setrgbcolor 0 setlinecap stroke
- grestore grestore end
- 49300 22390 a 0 TeXcolorgray 0.5 TeXcolorgray 0 TeXcolorgray
-0.5 TeXcolorgray 49300 22390 a
- tx@Dict begin gsave STV newpath 0.8 SLW 0.5 setgray /ArrowA { moveto
-} def /ArrowB { } def /NCLW CLW def tx@NodeDict begin 0.0 0.0 neg 3.0
-3.0 0 0 /N@x1 /N@x2 InitNC { tx@Dict begin /Lineto /lineto load def
-false pop end /AngleA 90. def /AngleB 90. def /ArmA 7.5 def /ArmB 7.5
-def /ArmTypeA 0 def /ArmTypeB 0 def /AngleB 90. def NCBar } if end
-gsave 0.8 SLW 0.5 setgray 0 setlinecap stroke grestore grestore
-end
- 49300 22390 a 0 TeXcolorgray
-27224 24053 a Fq(This)215 b(phenomenon)i(typically)g(occurs)e(when)i
-(some)f(of)f(the)h(binding)g(struc-)27224 25160 y(ture)307
-b(is)g(visible,)g(b)-20 b(ut)308 b(more)f(binding)h(will)g(occur)g(in)f
-(higher)h(le)-25 b(v)-15 b(els)307 b(of)g(the)27224 26267
-y(abstract)249 b(syntax)g(tree.)28552 27374 y(When)358
-b(we)f(can)h(syntactically)g(decide)g(at)f(a)g Fm(bind)g
-Fq(declaration)h(point)27224 28481 y(that)195 b(none)g(of)f(the)h(v)-25
-b(ariables)195 b(will)g(be)g(pick)-10 b(ed)196 b(out)f(at)f(a)h
-(binding)g(site,)g(we)g(can)27224 29588 y(seal)338 b(the)g(sets)g(of)g
-(v)-25 b(ariable)339 b(occurence,)h(which)f(we)f(call)h
-Fk(closed)f(binding)27224 30695 y(set)p Fq(,)249 b(such)g(as)g(in)g
-(the)g(judgment)h(form.)e(The)i(notion)f(of)g(alpha-equi)-25
-b(v)g(alence)27224 31802 y(is)324 b(then)i(the)f(equi)-25
-b(v)g(alence)327 b(class)d(of)h(all)g(terms)g(which)h(dif)-25
-b(fer)324 b(only)h(in)g(the)27224 32909 y(identity)252
-b(of)f(the)i(concrete)g(v)-25 b(ariable)252 b(within)g(the)g(closed)g
-(binding)h(sets.)e(W)-80 b(e)27224 34016 y(gi)-25 b(v)-15
-b(e)274 b(a)g(de\002nition)h(of)e(capture-a)-20 b(v)g(oiding)276
-b(substitution,)d(and)h(pro)-15 b(v)g(e)275 b(that)f(it)27224
-35123 y(respects)251 b(this)g(e)-15 b(xpanded)253 b(notion)f(of)f
-(alpha-equi)-25 b(v)g(alence.)254 b(Correctness)d(of)27224
-36230 y(our)305 b(de\002nitions)f(cannot)i(be)f(de\002ned)h(in)f
-(general,)h(as)e(there)i(is)e(no)h(widely)27224 37337
-y(accepted)281 b(other)e(notion)h(of)f(non-single-v)-25
-b(ariable)280 b(binding,)g(b)-20 b(ut)279 b(it)g(can)h(be)27224
-38444 y(sho)-25 b(wn)275 b(in)f(speci\002c)i(cases)e(\227)h(we)g(do)g
-(so)g(for)f(the)h(important)g(special)g(case)27224 39551
-y(of)249 b(the)g(untyped)h(lambda)g(calculus.)28552 40658
-y(The)357 b(Ott)g(tool)g([2])g(implements)g(this)f(language)j(of)d
-(binding)i(speci\002-)27224 41765 y(cations)296 b(by)g(translating)f
-(them)i(into)e(what)i(we)f(call)g(a)g Fk(fully)g(concr)-37
-b(ete)297 b Fq(rep-)27224 42872 y(resentation)344 b(within)h(v)-25
-b(arious)344 b(proof)g(assistants.)f(In)g(this)h(representation,)27224
-43979 y(abstract)395 b(syntax)g(terms)f(may)i(contain)f(concrete)h(v)
--25 b(ariables,)396 b(and)f(terms)27224 45086 y(are)408
-b(not)g(considered)g(up)g(to)g(alpha-equi)-25 b(v)g(alence.)410
-b(The)e(\(Ott-generated\))27224 46193 y(substitution)375
-b(functions)h(assume)h(that)f(the)g(substituted)g(term)h(is)e(al)-10
-b(w)g(ays)27224 47300 y(closed.)372 b(A)g(wide)g(v)-25
-b(ariety)372 b(of)g(programming)g(language)h(theory)-65
-b(,)372 b(includ-)27224 48407 y(ing)257 b(a)h(formalization)h(of)e
-(OCaml)h(without)g(objects)g(or)f(modules,)h(can)g(still)27224
-49514 y(be)385 b(done)h(in)f(this)f(restricted)h(setting,)g(since)h
-(typical)f(languages)h(do)g(not)27224 50620 y(perform)419
-b(reductions)g(under)h(binders.)f(W)-80 b(e)421 b(pro)-15
-b(v)g(e)419 b(that)h(under)f(suitable)27224 51727 y(sanity)320
-b(conditions,)g(the)h(operational)g(de\002nition)g(of)f(substitution)f
-(within)27224 52834 y(Ott)332 b(coincides)h(with)g(the)g(de\002nition)f
-(of)h(substitution)e(respecting)i(alpha-)27224 53941
-y(equi)-25 b(v)g(alence)379 b(abo)-15 b(v)g(e.)378 b(The)f(proof)g(is)f
-(on)i(paper)-40 b(,)377 b(and)h(partly)f(formalized)27224
-55048 y(within)249 b(Isabelle.)28552 56155 y(This)220
-b(fully)h(concrete)i(representation,)e(is,)g(ho)-25 b(we)g(v)-15
-b(er)-40 b(,)223 b(clearly)e(not)h(suf)-25 b(\002-)27224
-57262 y(cient)272 b(for)f(some)g(e)-15 b(xamples.)273
-b(An)e(important)h(problem)g(for)f(future)h(w)-10 b(ork)271
-b(is)27224 58369 y(to)254 b(de)-25 b(v)-15 b(elop)256
-b(good)f(proof)f(assistant)g(implementations)h(that)g(respect)g(alpha)
-27224 59476 y(equi)-25 b(v)g(alence)260 b(in)e(general,)h(not)f(just)g
-(for)g(closed)g(substitutions.)f(One)i(could)27224 60583
-y(en)-40 b(visage)363 b(doing)h(this)e(indirectly)-65
-b(,)363 b(by)g(translating)g(into)g(languages)g(with)27224
-61690 y(single)376 b(binders,)h(or)g(directly)-65 b(,)377
-b(e.g.)h(perhaps)f(by)g(generalising)g(a)g(locally)27224
-62797 y(nameless)249 b(or)g(nominal)h(representation.)27224
-64093 y Fd(Refer)-22 b(ences)p 0 TeXcolorgray 27224 65532
-a Fc([1])p 0 TeXcolorgray 387 w(Franc)-344 b(\270)49
-b(ois)280 b(Pottier)-49 b(.)504 b(An)279 b(o)-13 b(v)g(ervie)-22
-b(w)279 b(of)f(C)p Fb(\013)p Fc(ml.)504 b(In)279 b Fa(A)-27
-b(CM)278 b(W)-81 b(orkshop)279 b(on)f(ML,)28644 66528
-y(ENTCS)221 b(148\(2\))p Fc(,)f(pages)i(27\22652,)f(March)h(2006.)p
-0 TeXcolorgray 27224 67967 a([2])p 0 TeXcolorgray 387
-w(Peter)248 b(Se)-22 b(well,)249 b(Francesco)g(Zappa)e(Nardelli,)i
-(Scott)g(Owens,)g(Gilles)g(Peskine,)28644 68964 y(Thomas)300
-b(Ridge,)f(Susmit)i(Sarkar)-35 b(,)300 b(and)f(Rok)g(Strni)45092
-68960 y(\020)45067 68964 y(sa.)571 b(Ott:)300 b(Ef)-22
-b(fecti)g(v)-13 b(e)300 b(tool)28644 69960 y(support)221
-b(for)g(the)g(w)-9 b(orking)221 b(semanticist.)322 b(In)221
-b Fa(ICFP)p Fc(,)g(2007.)p 0 TeXcolorgray 27224 71399
-a([3])p 0 TeXcolorgray 387 w(M.)187 b(R.)f(Shinwell,)i(A.)f(M.)g
-(Pitts,)h(and)e(M.)h(J.)g(Gabbay)-58 b(.)244 b(FreshML:)187
-b(Programming)28644 72395 y(with)222 b(binders)f(made)h(simple.)320
-b(In)221 b Fa(Pr)-40 b(oc.)222 b(ICFP)p Fc(,)g(2003.)p
-0 TeXcolorgray 0 TeXcolorgray eop end
-%%Trailer
-
-userdict /end-hook known{end-hook}if
-%%EOF
Binary file Literature/quotient2.pdf has changed
Binary file Literature/sewell.pdf has changed
--- a/Pearl-jv/Paper.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1819 +0,0 @@
-(*<*)
-theory Paper
-imports "../Nominal/Nominal2_Base"
- "../Nominal/Atoms"
- "../Nominal/Nominal2_Abs"
- "~~/src/HOL/Library/LaTeXsugar"
-begin
-
-abbreviation
- UNIV_atom ("\<allatoms>")
-where
- "UNIV_atom \<equiv> UNIV::atom set"
-
-notation (latex output)
- sort_of ("sort _" [1000] 100) and
- Abs_perm ("_") and
- Rep_perm ("_") and
- swap ("'(_ _')" [1000, 1000] 1000) and
- fresh ("_ # _" [51, 51] 50) and
- fresh_star ("_ #\<^sup>* _" [51, 51] 50) and
- Cons ("_::_" [78,77] 73) and
- supp ("supp _" [78] 73) and
- uminus ("-_" [78] 73) and
- atom ("|_|") and
- If ("if _ then _ else _" 10) and
- Rep_name ("\<lfloor>_\<rfloor>") and
- Abs_name ("\<lceil>_\<rceil>") and
- Rep_var ("\<lfloor>_\<rfloor>") and
- Abs_var ("\<lceil>_\<rceil>") and
- sort_of_ty ("sort'_ty _")
-
-(* BH: uncomment if you really prefer the dot notation
-syntax (latex output)
- "_Collect" :: "pttrn => bool => 'a set" ("(1{_ . _})")
-*)
-
-(* sort is used in Lists for sorting *)
-hide_const sort
-
-abbreviation
- "sort \<equiv> sort_of"
-
-lemma infinite_collect:
- assumes "\<forall>x \<in> S. P x" "infinite S"
- shows "infinite {x \<in> S. P x}"
-using assms
-apply(subgoal_tac "infinite {x. x \<in> S}")
-apply(simp only: Inf_many_def[symmetric])
-apply(erule INFM_mono)
-apply(auto)
-done
-
-
-(*>*)
-
-section {* Introduction *}
-
-text {*
- Nominal Isabelle provides a proving infratructure for convenient reasoning
- about syntax involving binders, such as lambda terms or type schemes in Mini-ML:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "\<lambda>x. t \<forall>{x\<^isub>1,\<dots>, x\<^isub>n}. \<tau>"}
- \end{isabelle}
-
- \noindent
- At its core Nominal Isabelle is based on the nominal logic work by
- Pitts at al \cite{GabbayPitts02,Pitts03}, whose most basic notion is
- a sort-respecting permutation operation defined over a countably
- infinite collection of sorted atoms.
-
-
-
- The aim of this paper is to
- describe how we adapted this work so that it can be implemented in a
- theorem prover based on Higher-Order Logic (HOL). For this we
- present the definition we made in the implementation and also review
- many proofs. There are a two main design choices to be made. One is
- how to represent sorted atoms. We opt here for a single unified atom
- type to represent atoms of different sorts. The other is how to
- present sort-respecting permutations. For them we use the standard
- technique of HOL-formalisations of introducing an appropriate
- subtype of functions from atoms to atoms.
-
- The nominal logic work has been the starting point for a number of proving
- infrastructures, most notable by Norrish \cite{norrish04} in HOL4, by
- Aydemir et al \cite{AydemirBohannonWeirich07} in Coq and the work by Urban
- and Berghofer in Isabelle/HOL \cite{Urban08}. Its key attraction is a very
- general notion, called \emph{support}, for the `set of free variables, or
- atoms, of an object' that applies not just to lambda terms and type schemes,
- but also to sets, products, lists, booleans and even functions. The notion of support
- is derived from the permutation operation defined over the
- hierarchy of types. This
- permutation operation, written @{text "_ \<bullet> _"}, has proved to be much more
- convenient for reasoning about syntax, in comparison to, say, arbitrary
- renaming substitutions of atoms. One reason is that permutations are
- bijective renamings of atoms and thus they can be easily `undone'---namely
- by applying the inverse permutation. A corresponding inverse substitution
- might not always exist, since renaming substitutions are in general only injective.
- Another reason is that permutations preserve many constructions when reasoning about syntax.
- For example, suppose a typing context @{text "\<Gamma>"} of the form
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "x\<^isub>1:\<tau>\<^isub>1, \<dots>, x\<^isub>n:\<tau>\<^isub>n"}
- \end{isabelle}
-
- \noindent
- is said to be \emph{valid} provided none of its variables, or atoms, @{text "x\<^isub>i"}
- occur twice. Then validity of typing contexts is preserved under
- permutations in the sense that if @{text \<Gamma>} is valid then so is \mbox{@{text "\<pi> \<bullet> \<Gamma>"}} for
- all permutations @{text "\<pi>"}. Again, this is \emph{not} the case for arbitrary
- renaming substitutions, as they might identify some of the @{text "x\<^isub>i"} in @{text \<Gamma>}.
-
- Permutations also behave uniformly with respect to HOL's logic connectives.
- Applying a permutation to a formula gives, for example
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}lcl}
- @{term "\<pi> \<bullet> (A \<and> B)"} & if and only if & @{text "(\<pi> \<bullet> A) \<and> (\<pi> \<bullet> B)"}\\
- @{term "\<pi> \<bullet> (A \<longrightarrow> B)"} & if and only if & @{text "(\<pi> \<bullet> A) \<longrightarrow> (\<pi> \<bullet> B)"}\\
- \end{tabular}
- \end{isabelle}
-
- \noindent
- This uniform behaviour can also be extended to quantifiers and functions.
- Because of these good properties of permutations, we are able to automate
- reasoning to do with \emph{equivariance}. By equivariance we mean the property
- that every permutation leaves a function unchanged, that is @{term "\<pi> \<bullet> f = f"}
- for all @{text "\<pi>"}. This will often simplify arguments involving support
- of functions, since if they are equivariant then they have empty support---or
- `no free atoms'.
-
- There are a number of subtle differences between the nominal logic work by
- Pitts and the formalisation we will present in this paper. One difference
- is that our
- formalisation is compatible with HOL, in the sense that we only extend
- HOL by some definitions, withouth the introduction of any new axioms.
- The reason why the original nominal logic work is
- incompatible with HOL has to do with the way how the finite support property
- is enforced: FM-set theory is defined in \cite{Pitts01b} so that every set
- in the FM-set-universe has finite support. In nominal logic \cite{Pitts03},
- the axioms (E3) and (E4) imply that every function symbol and proposition
- has finite support. However, there are notions in HOL that do \emph{not}
- have finite support (we will give some examples). In our formalisation, we
- will avoid the incompatibility of the original nominal logic work by not a
- priory restricting our discourse to only finitely supported entities, rather
- we will explicitly assume this property whenever it is needed in proofs. One
- consequence is that we state our basic definitions not in terms of nominal
- sets (as done for example in \cite{Pitts06}), but in terms of the weaker
- notion of permutation types---essentially sets equipped with a ``sensible''
- notion of permutation operation.
-
-
-
- In the nominal logic woworkrk, the `new quantifier' plays a prominent role.
- $\new$
-
-
- Obstacles for Coq; no type-classes, difficulties with quotient types,
- need for classical reasoning
-
- Two binders
-
- A preliminary version
-*}
-
-section {* Sorted Atoms and Sort-Respecting Permutations *}
-
-text {*
- The two most basic notions in the nominal logic work are a countably
- infinite collection of sorted atoms and sort-respecting permutations
- of atoms. The atoms are used for representing variable names that
- might be bound or free. Multiple sorts are necessary for being able
- to represent different kinds of variables. For example, in the
- language Mini-ML there are bound term variables in lambda
- abstractions and bound type variables in type schemes. In order to
- be able to separate them, each kind of variables needs to be
- represented by a different sort of atoms.
-
-
- The existing nominal logic work usually leaves implicit the sorting
- information for atoms and leaves out a description of how sorts are
- represented. In our formalisation, we therefore have to make a
- design decision about how to implement sorted atoms and
- sort-respecting permutations. One possibility, which we described in
- \cite{Urban08}, is to have separate types for different sorts of
- atoms. However, we found that this does not blend well with
- type-classes in Isabelle/HOL (see Section~\ref{related} about
- related work). Therefore we use here a single unified atom type to
- represent atoms of different sorts. A basic requirement is that
- there must be a countably infinite number of atoms of each sort.
- This can be implemented in Isabelle/HOL as the datatype
-
-*}
-
- datatype atom\<iota> = Atom\<iota> string nat
-
-text {*
- \noindent
- whereby the string argument specifies the sort of the
- atom.\footnote{A similar design choice was made by Gunter et al
- \cite{GunterOsbornPopescu09} for their variables.} The use of type
- \emph{string} for sorts is merely for convenience; any countably
- infinite type would work as well. In what follows we shall write
- @{term "UNIV::atom set"} for the set of all atoms. We also have two
- auxiliary functions for atoms, namely @{text sort} and @{const
- nat_of} which are defined as
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
- @{thm (lhs) sort_of.simps[no_vars]} & @{text "\<equiv>"} & @{thm (rhs) sort_of.simps[no_vars]}\\
- @{thm (lhs) nat_of.simps[no_vars]} & @{text "\<equiv>"} & @{thm (rhs) nat_of.simps[no_vars]}
- \end{tabular}\hfill\numbered{sortnatof}
- \end{isabelle}
-
- \noindent
- We clearly have for every finite set @{text S}
- of atoms and every sort @{text s} the property:
-
- \begin{proposition}\label{choosefresh}\mbox{}\\
- @{text "For a finite set of atoms S, there exists an atom a such that
- sort a = s and a \<notin> S"}.
- \end{proposition}
-
- \noindent
- This property will be used later whenever we have to chose a `fresh' atom.
-
- For implementing sort-respecting permutations, we use functions of type @{typ
- "atom => atom"} that are bijective; are the
- identity on all atoms, except a finite number of them; and map
- each atom to one of the same sort. These properties can be conveniently stated
- in Isabelle/HOL for a function @{text \<pi>} as follows:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{r@ {\hspace{4mm}}l}
- i) & @{term "bij \<pi>"}\\
- ii) & @{term "finite {a. \<pi> a \<noteq> a}"}\\
- iii) & @{term "\<forall>a. sort (\<pi> a) = sort a"}
- \end{tabular}\hfill\numbered{permtype}
- \end{isabelle}
-
- \noindent
- Like all HOL-based theorem provers, Isabelle/HOL allows us to
- introduce a new type @{typ perm} that includes just those functions
- satisfying all three properties. For example the identity function,
- written @{term id}, is included in @{typ perm}. Also function composition,
- written \mbox{@{text "_ \<circ> _"}}, and function inversion, given by Isabelle/HOL's
- inverse operator and written \mbox{@{text "inv _"}}, preserve the properties
- (\ref{permtype}.@{text "i"}-@{text "iii"}).
-
- However, a moment of thought is needed about how to construct non-trivial
- permutations. In the nominal logic work it turned out to be most convenient
- to work with swappings, written @{text "(a b)"}. In our setting the
- type of swappings must be
-
- @{text [display,indent=10] "(_ _) :: atom \<Rightarrow> atom \<Rightarrow> perm"}
-
- \noindent
- but since permutations are required to respect sorts, we must carefully
- consider what happens if a user states a swapping of atoms with different
- sorts. The following definition\footnote{To increase legibility, we omit
- here and in what follows the @{term Rep_perm} and @{term "Abs_perm"}
- wrappers that are needed in our implementation in Isabelle/HOL since we defined permutation
- not to be the full function space, but only those functions of type @{typ
- perm} satisfying properties @{text i}-@{text "iii"} in \eqref{permtype}.}
-
-
- @{text [display,indent=10] "(a b) \<equiv> \<lambda>c. if a = c then b else (if b = c then a else c)"}
-
- \noindent
- does not work in general, because @{text a} and @{text b} may have different
- sorts---in which case the function would violate property @{text iii} in \eqref{permtype}. We
- could make the definition of swappings partial by adding the precondition
- @{term "sort a = sort b"}, which would mean that in case @{text a} and
- @{text b} have different sorts, the value of @{text "(a b)"} is unspecified.
- However, this looked like a cumbersome solution, since sort-related side
- conditions would be required everywhere, even to unfold the definition. It
- turned out to be more convenient to actually allow the user to state
- `ill-sorted' swappings but limit their `damage' by defaulting to the
- identity permutation in the ill-sorted case:
-
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}rl}
- @{text "(a b) \<equiv>"} & @{text "if (sort a = sort b)"}\\
- & \hspace{3mm}@{text "then \<lambda>c. if a = c then b else (if b = c then a else c)"}\\
- & \hspace{3mm}@{text "else id"}
- \end{tabular}\hfill\numbered{swapdef}
- \end{isabelle}
-
- \noindent
- This function is bijective, the identity on all atoms except
- @{text a} and @{text b}, and sort respecting. Therefore it is
- a function in @{typ perm}.
-
- One advantage of using functions as a representation for
- permutations is that it is a unique representation. For example the swappings
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm swap_commute[no_vars]}\hspace{10mm}
- @{text "(a a) = id"}
- \end{tabular}\hfill\numbered{swapeqs}
- \end{isabelle}
-
- \noindent
- are \emph{equal} and can be used interchangeably. Another advantage of the function
- representation is that they form a (non-com\-mu\-ta\-tive) group provided we define
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}r@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{10mm}}r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
- @{thm (lhs) zero_perm_def[no_vars]} & @{text "\<equiv>"} & @{thm (rhs) zero_perm_def[no_vars]} &
- @{thm (lhs) plus_perm_def[where p="\<pi>\<^isub>1" and q="\<pi>\<^isub>2"]} & @{text "\<equiv>"} &
- @{thm (rhs) plus_perm_def[where p="\<pi>\<^isub>1" and q="\<pi>\<^isub>2"]}\\
- @{thm (lhs) uminus_perm_def[where p="\<pi>"]} & @{text "\<equiv>"} & @{thm (rhs) uminus_perm_def[where p="\<pi>"]} &
- @{thm (lhs) minus_perm_def[where ?p1.0="\<pi>\<^isub>1" and ?p2.0="\<pi>\<^isub>2"]} & @{text "\<equiv>"} &
- @{thm (rhs) minus_perm_def[where ?p1.0="\<pi>\<^isub>1" and ?p2.0="\<pi>\<^isub>2"]}
- \end{tabular}\hfill\numbered{groupprops}
- \end{isabelle}
-
- \noindent
- and verify the four simple properties
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- i)~~@{thm add_assoc[where a="\<pi>\<^isub>1" and b="\<pi>\<^isub>2" and c="\<pi>\<^isub>3"]}\smallskip\\
- ii)~~@{thm monoid_add_class.add_0_left[where a="\<pi>::perm"]} \hspace{9mm}
- iii)~~@{thm monoid_add_class.add_0_right[where a="\<pi>::perm"]} \hspace{9mm}
- iv)~~@{thm group_add_class.left_minus[where a="\<pi>::perm"]}
- \end{tabular}\hfill\numbered{grouplaws}
- \end{isabelle}
-
- \noindent
- The technical importance of this fact is that we can rely on
- Isabelle/HOL's existing simplification infrastructure for groups, which will
- come in handy when we have to do calculations with permutations.
- Note that Isabelle/HOL defies standard conventions of mathematical notation
- by using additive syntax even for non-commutative groups. Obviously,
- composition of permutations is not commutative in general; for example
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "(a b) + (b c) \<noteq> (b c) + (a b)"}\;.
- \end{isabelle}
-
- \noindent
- But since the point of this paper is to implement the
- nominal theory as smoothly as possible in Isabelle/HOL, we tolerate
- the non-standard notation in order to reuse the existing libraries.
-
- A \emph{permutation operation}, written infix as @{text "\<pi> \<bullet> x"},
- applies a permutation @{text "\<pi>"} to an object @{text "x"}. This
- operation has the type
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "_ \<bullet> _ :: perm \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
- \end{isabelle}
-
- \noindent
- whereby @{text "\<beta>"} is a generic type for the object @{text
- x}.\footnote{We will write @{text "((op \<bullet>) \<pi>)
- x"} for this operation in the few cases where we need to indicate
- that it is a function applied with two arguments.} The definition
- of this operation will be given by in terms of `induction' over this
- generic type. The type-class mechanism of Isabelle/HOL
- \cite{Wenzel04} allows us to give a definition for `base' types,
- such as atoms, permutations, booleans and natural numbers:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l@ {\hspace{4mm}}l@ {}}
- atoms: & @{thm permute_atom_def[where p="\<pi>",no_vars, THEN eq_reflection]}\\
- permutations: & @{thm permute_perm_def[where p="\<pi>" and q="\<pi>'", THEN eq_reflection]}\\
- booleans: & @{thm permute_bool_def[where p="\<pi>", no_vars, THEN eq_reflection]}\\
- nats: & @{thm permute_nat_def[where p="\<pi>", no_vars, THEN eq_reflection]}\\
- \end{tabular}\hfill\numbered{permdefsbase}
- \end{isabelle}
-
- \noindent
- and for type-constructors, such as functions, sets, lists and products:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l@ {\hspace{4mm}}l@ {}}
- functions: & @{text "\<pi> \<bullet> f \<equiv> \<lambda>x. \<pi> \<bullet> (f ((-\<pi>) \<bullet> x))"}\\
- sets: & @{thm permute_set_eq[where p="\<pi>", no_vars, THEN eq_reflection]}\\
- lists: & @{thm permute_list.simps(1)[where p="\<pi>", no_vars, THEN eq_reflection]}\\
- & @{thm permute_list.simps(2)[where p="\<pi>", no_vars, THEN eq_reflection]}\\
- products: & @{thm permute_prod.simps[where p="\<pi>", no_vars, THEN eq_reflection]}\\
- \end{tabular}\hfill\numbered{permdefsconstrs}
- \end{isabelle}
-
- \noindent
- The type classes also allow us to reason abstractly about the permutation operation.
- For this we state the following two
- \emph{permutation properties}:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}r@ {\hspace{4mm}}p{10cm}}
- i) & @{thm permute_zero[no_vars]}\\
- ii) & @{thm permute_plus[where p="\<pi>\<^isub>1" and q="\<pi>\<^isub>2",no_vars]}
- \end{tabular}\hfill\numbered{newpermprops}
- \end{isabelle}
-
- \noindent
- From these properties and law (\ref{grouplaws}.{\it iv}) about groups
- follows that a permutation and its inverse cancel each other. That is
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm permute_minus_cancel(1)[where p="\<pi>", no_vars]}\hspace{10mm}
- @{thm permute_minus_cancel(2)[where p="\<pi>", no_vars]}
- \end{tabular}\hfill\numbered{cancel}
- \end{isabelle}
-
- \noindent
- Consequently, the permutation operation @{text "\<pi> \<bullet> _"}~~is bijective,
- which in turn implies the property
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm (lhs) permute_eq_iff[where p="\<pi>", no_vars]}
- $\;$if and only if$\;$
- @{thm (rhs) permute_eq_iff[where p="\<pi>", no_vars]}.
- \end{tabular}\hfill\numbered{permuteequ}
- \end{isabelle}
-
- \noindent
- We can also show that the following property holds for the permutation
- operation.
-
- \begin{lemma}\label{permutecompose}
- @{text "\<pi>\<^isub>1 \<bullet> (\<pi>\<^isub>2 \<bullet> x) = (\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2) \<bullet> (\<pi>\<^isub>1 \<bullet> x)"}.
- \end{lemma}
-
- \begin{proof} The proof is as follows:
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}[b]{@ {}c@ {\hspace{2mm}}l@ {\hspace{8mm}}l}
- & @{text "\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2 \<bullet> x"}\\
- @{text "="} & @{text "\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2 \<bullet> (-\<pi>\<^isub>1) \<bullet> \<pi>\<^isub>1 \<bullet> x"} & by \eqref{cancel}\\
- @{text "="} & @{text "(\<pi>\<^isub>1 + \<pi>\<^isub>2 - \<pi>\<^isub>1) \<bullet> (\<pi>\<^isub>1 \<bullet> x)"} & by {\rm(\ref{newpermprops}.@{text "ii"})}\\
- @{text "\<equiv>"} & @{text "(\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2) \<bullet> (\<pi>\<^isub>1 \<bullet> x)"}\\
- \end{tabular}\hfill\qed
- \end{isabelle}
- \end{proof}
-
- \noindent
- Note that the permutation operation for functions is defined so that
- we have for applications the equation
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "\<pi> \<bullet> (f x) ="}
- @{thm (rhs) permute_fun_app_eq[where p="\<pi>", no_vars]}
- \hfill\numbered{permutefunapp}
- \end{isabelle}
-
- \noindent
- provided the permutation properties hold for @{text x}. This equation can
- be easily shown by unfolding the permutation operation for functions on
- the right-hand side of the equation, simplifying the resulting beta-redex
- and eliminating the permutations in front of @{text x} using \eqref{cancel}.
-
- The main benefit of the use of type classes is that it allows us to delegate
- much of the routine resoning involved in determining whether the permutation properties
- are satisfied to Isabelle/HOL's type system: we only have to
- establish that base types satisfy them and that type-constructors
- preserve them. Then Isabelle/HOL will use this information and determine
- whether an object @{text x} with a compound type, like @{typ "atom \<Rightarrow> (atom set * nat)"}, satisfies the
- permutation properties. For this we define the notion of a
- \emph{permutation type}:
-
- \begin{definition}[Permutation Type]
- A type @{text "\<beta>"} is a \emph{permutation type} if the permutation
- properties in \eqref{newpermprops} are satisfied for every @{text
- "x"} of type @{text "\<beta>"}.
- \end{definition}
-
- \noindent
- and establish:
-
- \begin{theorem}
- The types @{type atom}, @{type perm}, @{type bool} and @{type nat}
- are permutation types, and if @{text \<beta>}, @{text "\<beta>\<^isub>1"} and @{text
- "\<beta>\<^isub>2"} are permutation types, then so are \mbox{@{text "\<beta>\<^isub>1 \<Rightarrow> \<beta>\<^isub>2"}},
- @{text "\<beta> set"}, @{text "\<beta> list"} and @{text "\<beta>\<^isub>1 \<times> \<beta>\<^isub>2"}.
- \end{theorem}
-
- \begin{proof}
- All statements are by unfolding the definitions of the permutation
- operations and simple calculations involving addition and
- minus. In case of permutations for example we have
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}[b]{@ {}rcl}
- @{text "0 \<bullet> \<pi>'"} & @{text "\<equiv>"} & @{text "0 + \<pi>' - 0 = \<pi>'"}\smallskip\\
- @{text "(\<pi>\<^isub>1 + \<pi>\<^isub>2) \<bullet> \<pi>'"} & @{text "\<equiv>"} & @{text "(\<pi>\<^isub>1 + \<pi>\<^isub>2) + \<pi>' - (\<pi>\<^isub>1 + \<pi>\<^isub>2)"}\\
- & @{text "="} & @{text "(\<pi>\<^isub>1 + \<pi>\<^isub>2) + \<pi>' - \<pi>\<^isub>2 - \<pi>\<^isub>1"}\\
- & @{text "="} & @{text "\<pi>\<^isub>1 + (\<pi>\<^isub>2 + \<pi>' - \<pi>\<^isub>2) - \<pi>\<^isub>1"}\\
- & @{text "\<equiv>"} & @{text "\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2 \<bullet> \<pi>'"}
- \end{tabular}\hfill\qed
- \end{isabelle}
- \end{proof}
-*}
-
-section {* Equivariance *}
-
-text {*
- (mention alpha-structural paper by Andy)
-
- Two important notions in the nominal logic work are what Pitts calls
- \emph{equivariance} and the \emph{equivariance principle}. These
- notions allows us to characterise how permutations act upon compound
- statements in HOL by analysing how these statements are constructed.
- The notion of equivariance means that an object is invariant under
- any permutations. This can be defined as follows:
-
- \begin{definition}[Equivariance]\label{equivariance}
- An object @{text "x"} of permutation type is \emph{equivariant} provided
- for all permutations @{text "\<pi>"}, \mbox{@{term "\<pi> \<bullet> x = x"}} holds.
- \end{definition}
-
- \noindent
- In what follows we will primarily be interested in the cases where
- @{text x} is a constant, but of course there is no way in
- Isabelle/HOL to restrict this definition to just these cases.
-
- There are a number of equivalent formulations for equivariance.
- For example, assuming @{text f} is a function of permutation
- type @{text "\<alpha> \<Rightarrow> \<beta>"}, then equivariance of @{text f} can also be stated as
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{text "\<forall>\<pi> x. \<pi> \<bullet> (f x) = f (\<pi> \<bullet> x)"}
- \end{tabular}\hfill\numbered{altequivariance}
- \end{isabelle}
-
- \noindent
- We will say this formulation of equivariance is in \emph{fully applied form}.
- To see that this formulation implies the definition, we just unfold
- the definition of the permutation operation for functions and
- simplify with the equation and the cancellation property shown in
- \eqref{cancel}. To see the other direction, we use
- \eqref{permutefunapp}. Similarly for functions that take more than
- one argument. The point to note is that equivariance and equivariance in fully
- applied form are always interderivable (for permutation types).
-
- Both formulations of equivariance have their advantages and
- disadvantages: \eqref{altequivariance} is usually more convenient to
- establish, since statements in HOL are commonly given in a
- form where functions are fully applied. For example we can easily
- show that equality is equivariant
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm eq_eqvt[where p="\<pi>", no_vars]}
- \end{tabular}\hfill\numbered{eqeqvt}
- \end{isabelle}
-
- \noindent
- using the permutation operation on booleans and property
- \eqref{permuteequ}.
- Lemma~\ref{permutecompose} establishes that the
- permutation operation is equivariant. The permutation operation for
- lists and products, shown in \eqref{permdefsconstrs}, state that the
- constructors for products, @{text "Nil"} and @{text Cons} are
- equivariant. Furthermore a simple calculation will show that our
- swapping functions are equivariant, that is
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm swap_eqvt[where p="\<pi>", no_vars]}
- \end{tabular}\hfill\numbered{swapeqvt}
- \end{isabelle}
-
- \noindent
- for all @{text a}, @{text b} and @{text \<pi>}. Also the booleans
- @{const True} and @{const False} are equivariant by the definition
- of the permutation operation for booleans. Given this definition, it
- is also easy to see that the boolean operators, like @{text "\<and>"},
- @{text "\<or>"}, @{text "\<longrightarrow>"} and @{text "\<not>"} are equivariant:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}lcl}
- @{text "\<pi> \<bullet> (A \<and> B) = (\<pi> \<bullet> A) \<and> (\<pi> \<bullet> B)"}\\
- @{text "\<pi> \<bullet> (A \<or> B) = (\<pi> \<bullet> A) \<or> (\<pi> \<bullet> B)"}\\
- @{text "\<pi> \<bullet> (A \<longrightarrow> B) = (\<pi> \<bullet> A) \<longrightarrow> (\<pi> \<bullet> B)"}\\
- @{text "\<pi> \<bullet> (\<not>A) = \<not>(\<pi> \<bullet> A)"}\\
- \end{tabular}
- \end{isabelle}
-
- In contrast, the advantage of Definition \ref{equivariance} is that
- it allows us to state a general principle how permutations act on
- statements in HOL. For this we will define a rewrite system that
- `pushes' a permutation towards the leaves of statements (i.e.~constants
- and variables). Then the permutations disappear in cases where the
- constants are equivariant. To do so, let us first define
- \emph{HOL-terms}, which are the building blocks of statements in HOL.
- They are given by the grammar
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "t ::= c | x | t\<^isub>1 t\<^isub>2 | \<lambda>x. t"}
- \hfill\numbered{holterms}
- \end{isabelle}
-
- \noindent
- where @{text c} stands for constants and @{text x} for variables.
- We assume HOL-terms are fully typed, but for the sake of better
- legibility we leave the typing information implicit. We also assume
- the usual notions for free and bound variables of a HOL-term.
- Furthermore, HOL-terms are regarded as equal modulo alpha-, beta-
- and eta-equivalence. The equivariance principle can now
- be stated formally as follows:
-
- \begin{theorem}[Equivariance Principle]\label{eqvtprin}
- Suppose a HOL-term @{text t} whose constants are all equivariant. For any
- permutation @{text \<pi>}, let @{text t'} be @{text t} except every
- free variable @{text x} in @{term t} is replaced by @{text "\<pi> \<bullet> x"}, then
- @{text "\<pi> \<bullet> t = t'"}.
- \end{theorem}
-
- \noindent
- The significance of this principle is that we can automatically establish
- the equivariance of a constant for which equivariance is not yet
- known. For this we only have to establish that the definiens of this
- constant is a HOL-term whose constants are all equivariant.
- This meshes well with how HOL is designed: except for a few axioms, every constant
- is defined in terms of existing constants. For example an alternative way
- to deduce that @{term True} is equivariant is to look at its
- definition
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{thm True_def}
- \end{isabelle}
-
- \noindent
- and observing that the only constant in the definiens, namely @{text "="}, is
- equivariant. Similarly, the universal quantifier @{text "\<forall>"} is definied in HOL as
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "\<forall>x. P x \<equiv> "}~@{thm (rhs) All_def[no_vars]}
- \end{isabelle}
-
- \noindent
- The constants in the definiens @{thm (rhs) All_def[no_vars]}, namely @{text "="}
- and @{text "True"}, are equivariant (we shown this above). Therefore
- the equivariance principle gives us
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{r@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
- @{text "\<pi> \<bullet> (\<forall>x. P x)"} & @{text "\<equiv>"} & @{text "\<pi> \<bullet> (P = (\<lambda>x. True))"}\\
- & @{text "="} & @{text "(\<pi> \<bullet> P) = (\<lambda>x. True)"}\\
- & @{text "\<equiv>"} & @{text "\<forall>x. (\<pi> \<bullet> P) x"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- which means the constant @{text "\<forall>"} must be equivariant. From this
- we can deduce that the existential quantifier @{text "\<exists>"} is equivariant.
- Its definition in HOL is
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "\<exists>x. P x \<equiv> "}~@{thm (rhs) Ex_def[no_vars]}
- \end{isabelle}
-
- \noindent
- where again the HOL-term on the right-hand side only contains equivariant constants
- (namely @{text "\<forall>"} and @{text "\<longrightarrow>"}). Taking both facts together, we can deduce that
- the unique existential quantifier @{text "\<exists>!"} is equivariant. Its definition
- is
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "\<exists>!x. P x \<equiv> "}~@{thm (rhs) Ex1_def[no_vars]}
- \end{isabelle}
-
- \noindent
- with all constants on the right-hand side being equivariant. With this kind
- of reasoning we can build up a database of equivariant constants, which will
- be handy for more complex calculations later on.
-
- Before we proceed, let us give a justification for the equivariance principle.
- This justification cannot be given directly inside Isabelle/HOL since we cannot
- prove any statement about HOL-terms. Instead, we will use a rewrite
- system consisting of a series of equalities
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "\<pi> \<cdot> t = ... = t'"}
- \end{isabelle}
-
- \noindent
- that establish the equality between @{term "\<pi> \<bullet> t"} and
- @{text "t'"}. The idea of the rewrite system is to push the
- permutation inside the term @{text t}. We have implemented this as a
- conversion tactic on the ML-level of Isabelle/HOL. In what follows,
- we will show that this tactic produces only finitely many equations
- and also show that it is correct (in the sense of pushing a permutation
- @{text "\<pi>"} inside a term and the only remaining instances of @{text
- "\<pi>"} are in front of the term's free variables).
-
- The tactic applies four `oriented' equations.
- We will first give a naive version of
- our tactic, which however in some corner cases produces incorrect
- results or does not terminate. We then give a modification in order
- to obtain the desired properties.
- Consider the following for oriented equations
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}lr@ {\hspace{3mm}}c@ {\hspace{3mm}}l}
- i) & @{text "\<pi> \<bullet> (t\<^isub>1 t\<^isub>2)"} & \rrh & @{term "(\<pi> \<bullet> t\<^isub>1) (\<pi> \<bullet> t\<^isub>2)"}\\
- ii) & @{text "\<pi> \<bullet> (\<lambda>x. t)"} & \rrh & @{text "\<lambda>x. \<pi> \<bullet> (t[x := (-\<pi>) \<bullet> x])"}\\
- iii) & @{term "\<pi> \<bullet> (- \<pi>) \<bullet> x"} & \rrh & @{term "x"}\\
- iv) & @{term "\<pi> \<bullet> c"} & \rrh &
- {\rm @{term "c"}\hspace{6mm}provided @{text c} is equivariant}\\
- \end{tabular}\hfill\numbered{rewriteapplam}
- \end{isabelle}
-
- \noindent
- These equation are oriented in the sense of being applied in the left-to-right
- direction. The first equation we established in \eqref{permutefunapp};
- the second follows from the definition of permutations acting on functions
- and the fact that HOL-terms are equal modulo beta-equivalence.
- The third is a consequence of \eqref{cancel} and the fourth from
- Definition~\ref{equivariance}. Unfortunately, we have to be careful with
- the rules {\it i)} and {\it iv}) since they can lead to loops whenever
- \mbox{@{text "t\<^isub>1 t\<^isub>2"}} is of the form @{text "((op \<bullet>) \<pi>') t"}.
- Recall that we established in Lemma \ref{permutecompose} that the
- constant @{text "(op \<bullet>)"} is equivariant and consider the infinite
- reduction sequence
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{text "\<pi> \<bullet> (\<pi>' \<bullet> t)"}
- $\;\;\stackrel{\text{\it i)}}{\rrh}\stackrel{\text{\it i)}}{\rrh}\stackrel{\text{\it iv)}}{\rrh}\;\;$
- @{text "(\<pi> \<bullet> \<pi>') \<bullet> (\<pi> \<bullet> t)"}
- $\;\;\stackrel{\text{\it i)}}{\rrh}\stackrel{\text{\it i)}}{\rrh}\stackrel{\text{\it iv)}}{\rrh}\;\;$
- @{text "((\<pi> \<bullet> \<pi>') \<bullet> \<pi>) \<bullet> ((\<pi> \<bullet> \<pi>') \<bullet> t)"}~~\ldots%
-
- \end{tabular}
- \end{isabelle}
-
- \noindent
- where the last term is again an instance of rewrite rule {\it i}), but larger.
- To avoid this loop we will apply the rewrite rule
- using an `outside to inside' strategy. This strategy is sufficient
- since we are only interested of rewriting terms of the form @{term
- "\<pi> \<bullet> t"}, where an outermost permutation needs to pushed inside a term.
-
- Another problem we have to avoid is that the rules {\it i)} and {\it
- iii)} can `overlap'. For this note that the term @{term "\<pi>
- \<bullet>(\<lambda>x. x)"} reduces by {\it ii)} to @{term "\<lambda>x. \<pi> \<bullet> (- \<pi>) \<bullet> x"}, to
- which we can apply rule {\it iii)} in order to obtain @{term
- "\<lambda>x. x"}, as is desired: since there is no free variable in the original
- term, the permutation should completely vanish. However, the
- subterm @{text "(- \<pi>) \<bullet> x"} is also an application. Consequently,
- the term @{term "\<lambda>x. \<pi> \<bullet> (- \<pi>) \<bullet>x"} can also reduce to @{text "\<lambda>x. (- (\<pi>
- \<bullet> \<pi>)) \<bullet> (\<pi> \<bullet> x)"} using {\it i)}. Given our strategy, we cannot
- apply rule {\it iii)} anymore in order to eliminate the permutation.
- In contrast, if we start
- with the term @{text "\<pi> \<bullet> ((- \<pi>) \<bullet> x)"} where @{text \<pi>} and @{text
- x} are free variables, then we \emph{do} want to apply rule {\it i)}
- in order to obtain @{text "(\<pi> \<bullet> (- \<pi>)) \<bullet> (\<pi> \<bullet> x)"}
- and not rule {\it iii)}. The latter would eliminate @{text \<pi>}
- completely and thus violating our correctness property. The problem is that
- rule {\it iii)} should only apply to
- instances where the corresponding variable is to bound; for free variables we want
- to use {\it ii)}. In order to distinguish these cases we have to
- maintain the information which variable is bound when inductively
- taking a term `apart'. This, unfortunately, does not mesh well with
- the way how conversion tactics are implemented in Isabelle/HOL.
-
- Our remedy is to use a standard trick in HOL: we introduce a
- separate definition for terms of the form @{text "(- \<pi>) \<bullet> x"},
- namely as
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{term "unpermute \<pi> x \<equiv> (- \<pi>) \<bullet> x"}
- \end{isabelle}
-
- \noindent
- The point is that now we can re-formulate the rewrite rules as follows
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}lrcl}
- i') & @{text "\<pi> \<bullet> (t\<^isub>1 t\<^isub>2)"} & \rrh &
- @{term "(\<pi> \<bullet> t\<^isub>1) (\<pi> \<bullet> t\<^isub>2)"}\hspace{45mm}\mbox{}\\
- \multicolumn{4}{r}{\rm provided @{text "t\<^isub>1 t\<^isub>2"} is not of the form @{text "unpermute \<pi> x"}}\smallskip\\
- ii') & @{text "\<pi> \<bullet> (\<lambda>x. t)"} & \rrh & @{text "\<lambda>x. \<pi> \<bullet> (t[x := unpermute \<pi> x])"}\\
- iii') & @{text "\<pi> \<bullet> (unpermute \<pi> x)"} & \rrh & @{term x}\\
- iv') & @{term "\<pi> \<bullet> c"} & \rrh & @{term "c"}
- \hspace{6mm}{\rm provided @{text c} is equivariant}\\
- \end{tabular}
- \end{isabelle}
-
- \noindent
- where @{text unpermutes} are only generated in case of bound variables.
- Clearly none of these rules overlap. Moreover, given our
- outside-to-inside strategy, applying them repeatedly must terminate.
-To see this, notice that
- the permutation on the right-hand side of the rewrite rules is
- always applied to a smaller term, provided we take the measure consisting
- of lexicographically ordered pairs whose first component is the size
- of a term (counting terms of the form @{text "unpermute \<pi> x"} as
- leaves) and the second is the number of occurences of @{text
- "unpermute \<pi> x"} and @{text "\<pi> \<bullet> c"}.
-
- With the rewrite rules of the conversions tactic in place, we can
- establish its correctness. The property we are after is that
- for a HOL-term @{text t} whose constants are all equivariant the
- term \mbox{@{text "\<pi> \<bullet> t"}} is equal to @{text "t'"} with @{text "t'"}
- being equal to @{text t} except that every free variable @{text x}
- in @{text t} is replaced by \mbox{@{text "\<pi> \<bullet> x"}}. Let us call
- a variable @{text x} \emph{really free}, if it is free and not occuring
- in an @{term unpermute}, such as @{text "unpermute _ x"} and @{text "unpermute x _"}.
- We need the following technical notion characterising \emph{@{text "\<pi>"}-proper} HOL-terms
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}ll}
- $\bullet$ & variables and constants are @{text \<pi>}-proper,\\
- $\bullet$ & @{term "unpermute \<pi> x"} is @{text \<pi>}-proper,\\
- $\bullet$ & @{term "\<lambda>x. t"} is @{text \<pi>}-proper, if @{text t} is @{text \<pi>}-proper and @{text x} is
- really free in @{text t}, and\\
- $\bullet$ & @{term "t\<^isub>1 t\<^isub>2"} is @{text \<pi>}-proper, if @{text "t\<^isub>1"} and @{text "t\<^isub>2"} are
- @{text \<pi>}-proper.
- \end{tabular}
- \end{isabelle}
-
- \begin{proof}[Theorem~\ref{eqvtprin}] We establish the property if @{text t}
- is @{text \<pi>}-proper and only contains equivaraint constants, then
- @{text "\<pi> \<bullet> t = t'"} where @{text "t'"} is equal to @{text "t"} except all really
- free variables @{text x} are replaced by @{text "\<pi> \<bullet> x"}, and all semi-free variables
- @{text "unpermute \<pi> x"} by @{text "x"}. We establish this property by induction
- on the size of HOL-terms, counting terms like @{text "unpermuting \<pi> x"} as leafes,
- like variables and constants. The cases for variables, constants and @{text unpermutes}
- are routine. In the case of abstractions we have by induction hypothesis that
- @{text "\<pi> \<bullet> (t[x := unpermute \<pi> x]) = t'"} with @{text "t'"} satisfying our
- correctness property. This implies that @{text "\<lambda>x. \<pi> \<bullet> (t[x := unpermute \<pi> x]) = \<lambda>x. t'"}
- and hence @{text "\<pi> \<bullet> (\<lambda>x. t) = \<lambda>x. t'"} as needed.\hfill\qed
- \end{proof}
-
- Pitts calls this property \emph{equivariance principle} (book ref ???).
-
- Problems with @{text undefined}
-
- Lines of code
-*}
-
-
-section {* Support and Freshness *}
-
-text {*
- The most original aspect of the nominal logic work of Pitts is a general
- definition for `the set of free variables, or free atoms, of an object @{text "x"}'. This
- definition is general in the sense that it applies not only to lambda terms,
- but to any type for which a permutation operation is defined
- (like lists, sets, functions and so on).
-
- \begin{definition}[Support] \label{support}
- Given @{text x} is of permutation type, then
-
- @{thm [display,indent=10] supp_def[no_vars, THEN eq_reflection]}
- \end{definition}
-
- \noindent
- (Note that due to the definition of swapping in \eqref{swapdef}, we do not
- need to explicitly restrict @{text a} and @{text b} to have the same sort.)
- There is also the derived notion for when an atom @{text a} is \emph{fresh}
- for an @{text x} of permutation type, defined as
-
- @{thm [display,indent=10] fresh_def[no_vars]}
-
- \noindent
- We also use the notation @{thm (lhs) fresh_star_def[no_vars]} for sets ot atoms
- defined as follows
-
- @{thm [display,indent=10] fresh_star_def[no_vars]}
-
- \noindent
- Using the equivariance principle, it can be easily checked that all three notions
- are equivariant. A simple consequence of the definition of support and equivariance
- is that if @{text x} is equivariant then we have
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm (concl) supp_fun_eqvt[where f="x", no_vars]}
- \end{tabular}\hfill\numbered{suppeqvtfun}
- \end{isabelle}
-
- \noindent
- For function applications we can establish the following two properties.
-
- \begin{lemma}\label{suppfunapp} Let @{text f} and @{text x} be of permutation type, then
- \begin{isabelle}
- \begin{tabular}{r@ {\hspace{4mm}}p{10cm}}
- {\it i)} & @{thm[mode=IfThen] fresh_fun_app[no_vars]}\\
- {\it ii)} & @{thm supp_fun_app[no_vars]}\\
- \end{tabular}
- \end{isabelle}
- \end{lemma}
-
- \begin{proof}
- For the first property, we know from the assumption that @{term
- "finite {b. (a \<rightleftharpoons> b) \<bullet> f \<noteq> f}"} and @{term "finite {b . (a \<rightleftharpoons> b) \<bullet> x \<noteq>
- x}"} hold. That means for all, but finitely many @{text b} we have
- @{term "(a \<rightleftharpoons> b) \<bullet> f = f"} and @{term "(a \<rightleftharpoons> b) \<bullet> x = x"}. Similarly,
- we have to show that for but, but finitely @{text b} the equation
- @{term "(a \<rightleftharpoons> b) \<bullet> f x = f x"} holds. The left-hand side of this
- equation is equal to @{term "((a \<rightleftharpoons> b) \<bullet> f) ((a \<rightleftharpoons> b) \<bullet> x)"} by
- \eqref{permutefunapp}, which we know by the previous two facts for
- @{text f} and @{text x} is equal to the right-hand side for all,
- but finitely many @{text b}. This establishes the first
- property. The second is a simple corollary of {\it i)} by
- unfolding the definition of freshness.\qed
- \end{proof}
-
- A striking consequence of the definitions for support and freshness
- is that we can prove without knowing anything about the structure of
- @{term x} that swapping two fresh atoms, say @{text a} and @{text
- b}, leave @{text x} unchanged. For the proof we use the following
- lemma about swappings applied to an @{text x}:
-
- \begin{lemma}\label{swaptriple}
- Assuming @{text x} is of permutation type, and @{text a}, @{text b} and @{text c}
- have the same sort, then \mbox{@{thm (prem 3) swap_rel_trans[no_vars]}} and
- @{thm (prem 4) swap_rel_trans[no_vars]} imply @{thm (concl) swap_rel_trans[no_vars]}.
- \end{lemma}
-
- \begin{proof}
- The cases where @{text "a = c"} and @{text "b = c"} are immediate.
- For the remaining case it is, given our assumptions, easy to calculate
- that the permutations
-
- @{thm [display,indent=10] (concl) swap_triple[no_vars]}
-
- \noindent
- are equal. The lemma is then by application of the second permutation
- property shown in~\eqref{newpermprops}.\qed
- \end{proof}
-
- \begin{theorem}\label{swapfreshfresh}
- Let @{text x} be of permutation type.
- @{thm [mode=IfThen] swap_fresh_fresh[no_vars]}
- \end{theorem}
-
- \begin{proof}
- If @{text a} and @{text b} have different sort, then the swapping is the identity.
- If they have the same sort, we know by definition of support that both
- @{term "finite {c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x}"} and @{term "finite {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x}"}
- hold. So the union of these sets is finite too, and we know by Proposition~\ref{choosefresh}
- that there is an atom @{term c}, with the same sort as @{term a} and @{term b},
- that satisfies \mbox{@{term "(a \<rightleftharpoons> c) \<bullet> x = x"}} and @{term "(b \<rightleftharpoons> c) \<bullet> x = x"}.
- Now the theorem follows from Lemma~\ref{swaptriple}.\hfill\qed
- \end{proof}
-
- While the abstract properties of support and freshness, particularly
- Theorem~\ref{swapfreshfresh}, are useful for developing Nominal Isabelle,
- one often has to calculate the support of concrete objects.
- For booleans, nats, products and lists it is easy to check that
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l@ {\hspace{4mm}}l@ {}}
- @{text "booleans"}: & @{term "supp b = {}"}\\
- @{text "nats"}: & @{term "supp n = {}"}\\
- @{text "products"}: & @{thm supp_Pair[no_vars]}\\
- @{text "lists:"} & @{thm supp_Nil[no_vars]}\\
- & @{thm supp_Cons[no_vars]}\\
- \end{tabular}
- \end{isabelle}
-
- \noindent
- hold. Establishing the support of atoms and permutations is a bit
- trickier. To do so we will use the following notion about a \emph{supporting set}.
-
- \begin{definition}[Supporting Set]
- A set @{text S} \emph{supports} @{text x} if for all atoms @{text a} and @{text b}
- not in @{text S} we have @{term "(a \<rightleftharpoons> b) \<bullet> x = x"}.
- \end{definition}
-
- \noindent
- The main motivation for this notion is that we can characterise @{text "supp x"}
- as the smallest finite set that supports @{text "x"}. For this we prove:
-
- \begin{lemma}\label{supports} Let @{text x} be of permutation type.
- \begin{isabelle}
- \begin{tabular}{r@ {\hspace{4mm}}p{10cm}}
- i) & @{thm[mode=IfThen] supp_is_subset[no_vars]}\\
- ii) & @{thm[mode=IfThen] supp_supports[no_vars]}\\
- iii) & @{thm (concl) supp_is_least_supports[no_vars]}
- provided @{thm (prem 1) supp_is_least_supports[no_vars]},
- @{thm (prem 2) supp_is_least_supports[no_vars]}
- and @{text "S"} is the least such set, that means formally,
- for all @{text "S'"}, if @{term "finite S'"} and
- @{term "S' supports x"} then @{text "S \<subseteq> S'"}.
- \end{tabular}
- \end{isabelle}
- \end{lemma}
-
- \begin{proof}
- For {\it i)} we derive a contradiction by assuming there is an atom @{text a}
- with @{term "a \<in> supp x"} and @{text "a \<notin> S"}. Using the second fact, the
- assumption that @{term "S supports x"} gives us that @{text S} is a superset of
- @{term "{b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}"}, which is finite by the assumption of @{text S}
- being finite. But this means @{term "a \<notin> supp x"}, contradicting our assumption.
- Property {\it ii)} is by a direct application of
- Theorem~\ref{swapfreshfresh}. For the last property, part {\it i)} proves
- one ``half'' of the claimed equation. The other ``half'' is by property
- {\it ii)} and the fact that @{term "supp x"} is finite by {\it i)}.\hfill\qed
- \end{proof}
-
- \noindent
- These are all relatively straightforward proofs adapted from the existing
- nominal logic work. However for establishing the support of atoms and
- permutations we found the following `optimised' variant of {\it iii)}
- more useful:
-
- \begin{lemma}\label{optimised} Let @{text x} be of permutation type.
- We have that @{thm (concl) finite_supp_unique[no_vars]}
- provided @{thm (prem 1) finite_supp_unique[no_vars]},
- @{thm (prem 2) finite_supp_unique[no_vars]}, and for
- all @{text "a \<in> S"} and all @{text "b \<notin> S"}, with @{text a}
- and @{text b} having the same sort, \mbox{@{term "(a \<rightleftharpoons> b) \<bullet> x \<noteq> x"}}
- \end{lemma}
-
- \begin{proof}
- By Lemma \ref{supports}@{text ".iii)"} we have to show that for every finite
- set @{text S'} that supports @{text x}, \mbox{@{text "S \<subseteq> S'"}} holds. We will
- assume that there is an atom @{text "a"} that is element of @{text S}, but
- not @{text "S'"} and derive a contradiction. Since both @{text S} and
- @{text S'} are finite, we can by Proposition \ref{choosefresh} obtain an atom
- @{text b}, which has the same sort as @{text "a"} and for which we know
- @{text "b \<notin> S"} and @{text "b \<notin> S'"}. Since we assumed @{text "a \<notin> S'"} and
- we have that @{text "S' supports x"}, we have on one hand @{term "(a \<rightleftharpoons> b) \<bullet> x
- = x"}. On the other hand, the fact @{text "a \<in> S"} and @{text "b \<notin> S"} imply
- @{term "(a \<rightleftharpoons> b) \<bullet> x \<noteq> x"} using the assumed implication. This gives us the
- contradiction.\hfill\qed
- \end{proof}
-
- \noindent
- Using this lemma we only have to show the following three proof-obligations
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}r@ {\hspace{4mm}}l}
- i) & @{term "{c} supports c"}\\
- ii) & @{term "finite {c}"}\\
- iii) & @{text "\<forall>a \<in> {c} b \<notin> {c}. sort a = sort b \<longrightarrow> (a b) \<bullet> c \<noteq> c"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- in order to establish that @{thm supp_atom[where a="c", no_vars]} holds. In
- Isabelle/HOL these proof-obligations can be discharged by easy
- simplifications. Similar proof-obligations arise for the support of
- permutations, which is
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm supp_perm[where p="\<pi>", no_vars]}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- The only proof-obligation that is
- interesting is the one where we have to show that
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{text "If \<pi> \<bullet> a \<noteq> a, \<pi> \<bullet> b = b and sort a = sort b, then (a b) \<bullet> \<pi> \<noteq> \<pi>"}.
- \end{tabular}
- \end{isabelle}
-
- \noindent
- For this we observe that
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}rcl}
- @{thm (lhs) perm_swap_eq[where p="\<pi>", no_vars]} &
- if and only if &
- @{thm (rhs) perm_swap_eq[where p="\<pi>", no_vars]}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- holds by a simple calculation using the group properties of permutations.
- The proof-obligation can then be discharged by analysing the inequality
- between the permutations @{term "(\<pi> \<bullet> a \<rightleftharpoons> b)"} and @{term "(a \<rightleftharpoons> b)"}.
-
- The main point about support is that whenever an object @{text x} has finite
- support, then Proposition~\ref{choosefresh} allows us to choose for @{text x} a
- fresh atom with arbitrary sort. This is a crucial operation in Nominal
- Isabelle in situations where, for example, a bound variable needs to be
- renamed. To allow such a choice, we only have to assume that
- @{text "finite (supp x)"} holds. For more convenience we
- can define a type class in Isabelle/HOL corresponding to the
- property:
-
- \begin{definition}[Finitely Supported Type]
- A type @{text "\<beta>"} is \emph{finitely supported} if @{term "finite (supp x)"}
- holds for all @{text x} of type @{text "\<beta>"}.
- \end{definition}
-
- \noindent
- By the calculations above we can easily establish
-
- \begin{theorem}\label{finsuptype}
- The types @{type atom}, @{type perm}, @{type bool} and @{type nat}
- are fintitely supported, and assuming @{text \<beta>}, @{text "\<beta>\<^isub>1"} and
- @{text "\<beta>\<^isub>2"} are finitely supported types, then @{text "\<beta> list"} and
- @{text "\<beta>\<^isub>1 \<times> \<beta>\<^isub>2"} are finitely supported.
- \end{theorem}
-
- \noindent
- The main benefit of using the finite support property for choosing a
- fresh atom is that the reasoning is `compositional'. To see this,
- assume we have a list of atoms and a method of choosing a fresh atom
- that is not a member in this list---for example the maximum plus
- one. Then if we enlarge this list \emph{after} the choice, then
- obviously the fresh atom might not be fresh anymore. In contrast, by
- the classical reasoning of Proposition~\ref{choosefresh} we know a
- fresh atom exists for every list of atoms and no matter how we
- extend this list of atoms, we always preserve the property of being
- finitely supported. Consequently the existence of a fresh atom is
- still guarantied by Proposition~\ref{choosefresh}. Using the method
- of `maximum plus one' we might have to adapt the choice of a fresh
- atom.
-
- Unfortunately, Theorem~\ref{finsuptype} does not work in general for the
- types of sets and functions. There are functions definable in HOL
- for which the finite support property does not hold. A simple
- example of a function with infinite support is @{const nat_of} shown
- in \eqref{sortnatof}. This function's support is the set of
- \emph{all} atoms @{term "UNIV::atom set"}. To establish this we
- show @{term "\<not> a \<sharp> nat_of"}. This is equivalent to assuming the set
- @{term "{b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of}"} is finite and deriving a
- contradiction. From the assumption we also know that @{term "{a} \<union>
- {b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of}"} is finite. Then we can use
- Proposition~\ref{choosefresh} to choose an atom @{text c} such that
- @{term "c \<noteq> a"}, @{term "sort_of c = sort_of a"} and @{term "(a \<rightleftharpoons> c)
- \<bullet> nat_of = nat_of"}. Now we can reason as follows:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}[b]{@ {}rcl@ {\hspace{5mm}}l}
- @{text "nat_of a"} & @{text "="} & @{text "(a \<rightleftharpoons> c) \<bullet> (nat_of a)"} & by def.~of permutations on nats\\
- & @{text "="} & @{term "((a \<rightleftharpoons> c) \<bullet> nat_of) ((a \<rightleftharpoons> c) \<bullet> a)"} & by \eqref{permutefunapp}\\
- & @{text "="} & @{term "nat_of c"} & by assumptions on @{text c}\\
- \end{tabular}
- \end{isabelle}
-
- \noindent
- But this means we have that @{term "nat_of a = nat_of c"} and @{term "sort_of a = sort_of c"}.
- This implies that atoms @{term a} and @{term c} must be equal, which clashes with our
- assumption @{term "c \<noteq> a"} about how we chose @{text c}.\footnote{Cheney \cite{Cheney06}
- gives similar examples for constructions that have infinite support.}
-*}
-
-section {* Support of Finite Sets *}
-
-text {*
- Also the set type is an instance whose elements are not generally finitely
- supported (we will give an example in Section~\ref{concrete}).
- However, we can easily show that finite sets and co-finite sets of atoms are finitely
- supported. Their support can be characterised as:
-
- \begin{lemma}\label{finatomsets}
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}[b]{@ {}rl}
- {\it i)} & If @{text S} is a finite set of atoms, then @{thm (concl) supp_finite_atom_set[no_vars]}.\\
- {\it ii)} & If @{term "UNIV - (S::atom set)"} is a finite set of atoms, then
- @{thm (concl) supp_cofinite_atom_set[no_vars]}.
- \end{tabular}
- \end{isabelle}
- \end{lemma}
-
- \begin{proof}
- Both parts can be easily shown by Lemma~\ref{optimised}. We only have to observe
- that a swapping @{text "(a b)"} leaves a set @{text S} unchanged provided both
- @{text a} and @{text b} are elements in @{text S} or both are not in @{text S}.
- However if the sorts of a @{text a} and @{text b} agree, then the swapping will
- change @{text S} if either of them is an element in @{text S} and the other is
- not.\hfill\qed
- \end{proof}
-
- \noindent
- Note that a consequence of the second part of this lemma is that
- @{term "supp (UNIV::atom set) = {}"}.
- More difficult, however, is it to establish that finite sets of finitely
- supported objects are finitely supported. For this we first show that
- the union of the supports of finitely many and finitely supported objects
- is finite, namely
-
- \begin{lemma}\label{unionsupp}
- If @{text S} is a finite set whose elements are all finitely supported, then
- %
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}[b]{@ {}rl}
- {\it i)} & @{thm (concl) Union_of_finite_supp_sets[no_vars]} and\\
- {\it ii)} & @{thm (concl) Union_included_in_supp[no_vars]}.
- \end{tabular}
- \end{isabelle}
- \end{lemma}
-
- \begin{proof}
- The first part is by a straightforward induction on the finiteness
- of @{text S}. For the second part, we know that @{term "\<Union>x\<in>S. supp
- x"} is a set of atoms, which by the first part is finite. Therefore
- we know by Lemma~\ref{finatomsets}.{\it i)} that @{term "(\<Union>x\<in>S. supp
- x) = supp (\<Union>x\<in>S. supp x)"}. Taking @{text "f"} to be the function
- \mbox{@{text "\<lambda>S. \<Union> (supp ` S)"}}, we can write the right-hand side
- as @{text "supp (f S)"}. Since @{text "f"} is an equivariant
- function (can be easily checked by the equivariance principle), we
- have that @{text "supp (f S) \<subseteq> supp S"} by
- Lemma~\ref{suppfunapp}.{\it ii)}. This completes the second
- part.\hfill\qed
- \end{proof}
-
- \noindent
- With this lemma in place we can establish that
-
- \begin{lemma}
- @{thm[mode=IfThen] supp_of_finite_sets[no_vars]}
- \end{lemma}
-
- \begin{proof}
- The right-to-left inclusion is proved in Lemma~\ref{unionsupp}.{\it ii)}. To show the inclusion
- in the other direction we can use Lemma~\ref{supports}.{\it i)}. This means
- for all @{text a} and @{text b} that are not in @{text S} we have to show that
- @{term "(a \<rightleftharpoons> b) \<bullet> (\<Union>x \<in> S. supp x) = (\<Union>x \<in> S. supp x)"} holds. By the equivariance
- principle, the left-hand side is equal to @{term "\<Union>x \<in> ((a \<rightleftharpoons> b) \<bullet> S). supp x"}. Now
- the swapping in front of @{text S} disappears, since @{term "a \<sharp> S"} and @{term "b \<sharp> S"}
- whenever @{text "a, b \<notin> S"}. Thus we are done.\hfill\qed
- \end{proof}
-
- \noindent
- To sum up, every finite set of finitely supported elements has
- finite support. Unfortunately, we cannot use
- Theorem~\ref{finsuptype} to let Isabelle/HOL find this out
- automatically. This would require to introduce a separate type of
- finite sets, which however is not so convenient to reason about as
- Isabelle's standard set type.
-*}
-
-
-section {* Induction Principles for Permutations *}
-
-text {*
- While the use of functions as permutation provides us with a unique
- representation for permutations (for example @{term "(a \<rightleftharpoons> b)"} and
- @{term "(b \<rightleftharpoons> a)"} are equal permutations), this representation does
- not come automatically with an induction principle. Such an
- induction principle is however useful for generalising
- Lemma~\ref{swapfreshfresh} from swappings to permutations, namely
-
- \begin{lemma}
- @{thm [mode=IfThen] perm_supp_eq[where p="\<pi>", no_vars]}
- \end{lemma}
-
- \noindent
- In this section we will establish an induction principle for permutations
- with which this lemma can be easily proved. It is not too difficult to derive
- an induction principle for permutations, given the fact that we allow only
- permutations having a finite support.
-
- Using a the property from \cite{???}
-
- \begin{lemma}\label{smallersupp}
- @{thm [mode=IfThen] smaller_supp[where p="\<pi>", no_vars]}
- \end{lemma}
-*}
-
-
-section {* An Abstraction Type *}
-
-text {*
- To that end, we will consider
- first pairs @{text "(as, x)"} of type @{text "(atom set) \<times> \<beta>"}. These pairs
- are intended to represent the abstraction, or binding, of the set of atoms @{text
- "as"} in the body @{text "x"}.
-
- The first question we have to answer is when two pairs @{text "(as, x)"} and
- @{text "(bs, y)"} are $\alpha$-equivalent? (For the moment we are interested in
- the notion of $\alpha$-equivalence that is \emph{not} preserved by adding
- vacuous binders.) To answer this question, we identify four conditions: {\it (i)}
- given a free-atom function @{text "fa"} of type \mbox{@{text "\<beta> \<Rightarrow> atom
- set"}}, then @{text x} and @{text y} need to have the same set of free
- atoms; moreover there must be a permutation @{text p} such that {\it
- (ii)} @{text p} leaves the free atoms of @{text x} and @{text y} unchanged, but
- {\it (iii)} ``moves'' their bound names so that we obtain modulo a relation,
- say \mbox{@{text "_ R _"}}, two equivalent terms. We also require that {\it (iv)}
- @{text p} makes the sets of abstracted atoms @{text as} and @{text bs} equal. The
- requirements {\it (i)} to {\it (iv)} can be stated formally as follows:
- %
- \begin{equation}\label{alphaset}
- \begin{array}{@ {\hspace{10mm}}r@ {\hspace{2mm}}l@ {\hspace{4mm}}r}
- \multicolumn{3}{l}{@{text "(as, x) \<approx>set R fa p (bs, y)"}\hspace{2mm}@{text "\<equiv>"}}\\[1mm]
- & @{term "fa(x) - as = fa(y) - bs"} & \mbox{\it (i)}\\
- @{text "\<and>"} & @{term "(fa(x) - as) \<sharp>* p"} & \mbox{\it (ii)}\\
- @{text "\<and>"} & @{text "(p \<bullet> x) R y"} & \mbox{\it (iii)}\\
- @{text "\<and>"} & @{term "(p \<bullet> as) = bs"} & \mbox{\it (iv)}\\
- \end{array}
- \end{equation}
-
- \noindent
- Note that this relation depends on the permutation @{text
- "p"}; $\alpha$-equivalence between two pairs is then the relation where we
- existentially quantify over this @{text "p"}. Also note that the relation is
- dependent on a free-atom function @{text "fa"} and a relation @{text
- "R"}. The reason for this extra generality is that we will use
- $\approx_{\,\textit{set}}$ for both ``raw'' terms and $\alpha$-equated terms. In
- the latter case, @{text R} will be replaced by equality @{text "="} and we
- will prove that @{text "fa"} is equal to @{text "supp"}.
-
- It might be useful to consider first some examples about how these definitions
- of $\alpha$-equivalence pan out in practice. For this consider the case of
- abstracting a set of atoms over types (as in type-schemes). We set
- @{text R} to be the usual equality @{text "="} and for @{text "fa(T)"} we
- define
-
- \begin{center}
- @{text "fa(x) = {x}"} \hspace{5mm} @{text "fa(T\<^isub>1 \<rightarrow> T\<^isub>2) = fa(T\<^isub>1) \<union> fa(T\<^isub>2)"}
- \end{center}
-
- \noindent
- Now recall the examples shown in \eqref{ex1}, \eqref{ex2} and
- \eqref{ex3}. It can be easily checked that @{text "({x, y}, x \<rightarrow> y)"} and
- @{text "({y, x}, y \<rightarrow> x)"} are $\alpha$-equivalent according to
- $\approx_{\,\textit{set}}$ and $\approx_{\,\textit{res}}$ by taking @{text p} to
- be the swapping @{term "(x \<rightleftharpoons> y)"}. In case of @{text "x \<noteq> y"}, then @{text
- "([x, y], x \<rightarrow> y)"} $\not\approx_{\,\textit{list}}$ @{text "([y, x], x \<rightarrow> y)"}
- since there is no permutation that makes the lists @{text "[x, y]"} and
- @{text "[y, x]"} equal, and also leaves the type \mbox{@{text "x \<rightarrow> y"}}
- unchanged. Another example is @{text "({x}, x)"} $\approx_{\,\textit{res}}$
- @{text "({x, y}, x)"} which holds by taking @{text p} to be the identity
- permutation. However, if @{text "x \<noteq> y"}, then @{text "({x}, x)"}
- $\not\approx_{\,\textit{set}}$ @{text "({x, y}, x)"} since there is no
- permutation that makes the sets @{text "{x}"} and @{text "{x, y}"} equal
- (similarly for $\approx_{\,\textit{list}}$). It can also relatively easily be
- shown that all three notions of $\alpha$-equivalence coincide, if we only
- abstract a single atom.
-
- In the rest of this section we are going to introduce three abstraction
- types. For this we define
- %
- \begin{equation}
- @{term "abs_set (as, x) (bs, x) \<equiv> \<exists>p. alpha_set (as, x) equal supp p (bs, x)"}
- \end{equation}
-
- \noindent
- (similarly for $\approx_{\,\textit{abs\_res}}$
- and $\approx_{\,\textit{abs\_list}}$). We can show that these relations are equivalence
- relations and equivariant.
-
- \begin{lemma}\label{alphaeq}
- The relations $\approx_{\,\textit{abs\_set}}$, $\approx_{\,\textit{abs\_list}}$
- and $\approx_{\,\textit{abs\_res}}$ are equivalence relations, and if @{term
- "abs_set (as, x) (bs, y)"} then also @{term "abs_set (p \<bullet> as, p \<bullet> x) (p \<bullet>
- bs, p \<bullet> y)"} (similarly for the other two relations).
- \end{lemma}
-
- \begin{proof}
- Reflexivity is by taking @{text "p"} to be @{text "0"}. For symmetry we have
- a permutation @{text p} and for the proof obligation take @{term "-p"}. In case
- of transitivity, we have two permutations @{text p} and @{text q}, and for the
- proof obligation use @{text "q + p"}. All conditions are then by simple
- calculations.
- \end{proof}
-
- \noindent
- This lemma allows us to use our quotient package for introducing
- new types @{text "\<beta> abs_set"}, @{text "\<beta> abs_res"} and @{text "\<beta> abs_list"}
- representing $\alpha$-equivalence classes of pairs of type
- @{text "(atom set) \<times> \<beta>"} (in the first two cases) and of type @{text "(atom list) \<times> \<beta>"}
- (in the third case).
- The elements in these types will be, respectively, written as:
-
- \begin{center}
- @{term "Abs_set as x"} \hspace{5mm}
- @{term "Abs_res as x"} \hspace{5mm}
- @{term "Abs_lst as x"}
- \end{center}
-
- \noindent
- indicating that a set (or list) of atoms @{text as} is abstracted in @{text x}. We will
- call the types \emph{abstraction types} and their elements
- \emph{abstractions}. The important property we need to derive is the support of
- abstractions, namely:
-
- \begin{theorem}[Support of Abstractions]\label{suppabs}
- Assuming @{text x} has finite support, then\\[-6mm]
- \begin{center}
- \begin{tabular}{l@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
- %@ {thm (lhs) supp_abs(1)[no_vars]} & $=$ & @ {thm (rhs) supp_abs(1)[no_vars]}\\
- %@ {thm (lhs) supp_abs(2)[no_vars]} & $=$ & @ {thm (rhs) supp_abs(2)[no_vars]}\\
- %@ {thm (lhs) supp_abs(3)[where bs="as", no_vars]} & $=$ & @ {thm (rhs) supp_abs(3)[where bs="as", no_vars]}
- \end{tabular}
- \end{center}
- \end{theorem}
-
- \noindent
- Below we will show the first equation. The others
- follow by similar arguments. By definition of the abstraction type @{text "abs_set"}
- we have
- %
- \begin{equation}\label{abseqiff}
- %@ {thm (lhs) abs_eq_iff(1)[where bs="as" and cs="bs", no_vars]} \;\;\text{if and only if}\;\;
- %@ {thm (rhs) abs_eq_iff(1)[where bs="as" and cs="bs", no_vars]}
- \end{equation}
-
- \noindent
- and also
- %
- \begin{equation}\label{absperm}
- @{thm permute_Abs[no_vars]}
- \end{equation}
-
- \noindent
- The second fact derives from the definition of permutations acting on pairs
- \eqref{permute} and $\alpha$-equivalence being equivariant
- (see Lemma~\ref{alphaeq}). With these two facts at our disposal, we can show
- the following lemma about swapping two atoms in an abstraction.
-
- \begin{lemma}
- %@ {thm[mode=IfThen] abs_swap1(1)[where bs="as", no_vars]}
- \end{lemma}
-
- \begin{proof}
- This lemma is straightforward using \eqref{abseqiff} and observing that
- the assumptions give us @{term "(a \<rightleftharpoons> b) \<bullet> (supp x - as) = (supp x - as)"}.
- Moreover @{text supp} and set difference are equivariant (see \cite{HuffmanUrban10}).
- \end{proof}
-
- \noindent
- Assuming that @{text "x"} has finite support, this lemma together
- with \eqref{absperm} allows us to show
- %
- \begin{equation}\label{halfone}
- %@ {thm abs_supports(1)[no_vars]}
- \end{equation}
-
- \noindent
- which by Property~\ref{supportsprop} gives us ``one half'' of
- Theorem~\ref{suppabs}. The ``other half'' is a bit more involved. To establish
- it, we use a trick from \cite{Pitts04} and first define an auxiliary
- function @{text aux}, taking an abstraction as argument:
- %
- \begin{center}
- @{thm supp_set.simps[THEN eq_reflection, no_vars]}
- \end{center}
-
- \noindent
- Using the second equation in \eqref{equivariance}, we can show that
- @{text "aux"} is equivariant (since @{term "p \<bullet> (supp x - as) =
- (supp (p \<bullet> x)) - (p \<bullet> as)"}) and therefore has empty support.
- This in turn means
- %
- \begin{center}
- @{term "supp (supp_gen (Abs_set as x)) \<subseteq> supp (Abs_set as x)"}
- \end{center}
-
- \noindent
- using \eqref{suppfun}. Assuming @{term "supp x - as"} is a finite set,
- we further obtain
- %
- \begin{equation}\label{halftwo}
- %@ {thm (concl) supp_abs_subset1(1)[no_vars]}
- \end{equation}
-
- \noindent
- since for finite sets of atoms, @{text "bs"}, we have
- @{thm (concl) supp_finite_atom_set[where S="bs", no_vars]}.
- Finally, taking \eqref{halfone} and \eqref{halftwo} together establishes
- Theorem~\ref{suppabs}.
-
- The method of first considering abstractions of the
- form @{term "Abs_set as x"} etc is motivated by the fact that
- we can conveniently establish at the Isabelle/HOL level
- properties about them. It would be
- laborious to write custom ML-code that derives automatically such properties
- for every term-constructor that binds some atoms. Also the generality of
- the definitions for $\alpha$-equivalence will help us in the next section.
-*}
-
-
-section {* Concrete Atom Types\label{concrete} *}
-
-text {*
-
- So far, we have presented a system that uses only a single multi-sorted atom
- type. This design gives us the flexibility to define operations and prove
- theorems that are generic with respect to atom sorts. For example, as
- illustrated above the @{term supp} function returns a set that includes the
- free atoms of \emph{all} sorts together.
-
- However, the single multi-sorted atom type does not make an ideal interface
- for end-users of Nominal Isabelle. If sorts are not distinguished by
- Isabelle's type system, users must reason about atom sorts manually. That
- means for example that subgoals involving sorts must be discharged explicitly within proof
- scripts, instead of being inferred automatically. In other
- cases, lemmas might require additional side conditions about sorts to be true.
- For example, swapping @{text a} and @{text b} in the pair \mbox{@{term "(a,
- b)"}} will only produce the expected result if we state the lemma in
- Isabelle/HOL as:
-*}
-
- lemma
- fixes a b :: "atom"
- assumes asm: "sort a = sort b"
- shows "(a \<rightleftharpoons> b) \<bullet> (a, b) = (b, a)"
- using asm by simp
-
-text {*
- \noindent
- Fortunately, it is possible to regain most of the type-checking automation
- that is lost by moving to a single atom type. We accomplish this by defining
- \emph{subtypes} of the generic atom type that only include atoms of a single
- specific sort. We call such subtypes \emph{concrete atom types}.
-
- The following Isabelle/HOL command defines a concrete atom type called
- \emph{name}, which consists of atoms whose sort equals the string @{term
- "''name''"}.
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \isacommand{typedef}\ \ @{typ name} = @{term "{a. sort\<iota> a = ''name''}"}
- \end{isabelle}
-
- \noindent
- This command automatically generates injective functions that map from the
- concrete atom type into the generic atom type and back, called
- representation and abstraction functions, respectively. We will write these
- functions as follows:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l@ {\hspace{10mm}}l}
- @{text "\<lfloor>_\<rfloor> :: name \<Rightarrow> atom"} &
- @{text "\<lceil>_\<rceil> :: atom \<Rightarrow> name"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- With the definition @{thm permute_name_def [where p="\<pi>", THEN
- eq_reflection, no_vars]}, it is straightforward to verify that the type
- @{typ name} is a permutation type.
-
- In order to reason uniformly about arbitrary concrete atom types, we define a
- type class that characterises type @{typ name} and other similarly-defined
- types. The definition of the concrete atom type class is as follows: First,
- every concrete atom type must be a permutation type. In addition, the class
- defines an overloaded function that maps from the concrete type into the
- generic atom type, which we will write @{text "|_|"}. For each class
- instance, this function must be injective and equivariant, and its outputs
- must all have the same sort, that is
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}r@ {\hspace{3mm}}l}
- i) & if @{thm (lhs) atom_eq_iff [no_vars]} then @{thm (rhs) atom_eq_iff [no_vars]}\\
- ii) & @{thm atom_eqvt[where p="\<pi>", no_vars]}\\
- iii) & @{thm sort_of_atom_eq [no_vars]}
- \end{tabular}\hfill\numbered{atomprops}
- \end{isabelle}
-
- \noindent
- With the definition @{thm atom_name_def [THEN eq_reflection, no_vars]} we can
- show that @{typ name} satisfies all the above requirements of a concrete atom
- type.
-
- The whole point of defining the concrete atom type class is to let users
- avoid explicit reasoning about sorts. This benefit is realised by defining a
- special swapping operation of type @{text "\<alpha> \<Rightarrow> \<alpha>
- \<Rightarrow> perm"}, where @{text "\<alpha>"} is a concrete atom type:
-
- @{thm [display,indent=10] flip_def [THEN eq_reflection, no_vars]}
-
- \noindent
- As a consequence of its type, the @{text "\<leftrightarrow>"}-swapping
- operation works just like the generic swapping operation, but it does not
- require any sort-checking side conditions---the sort-correctness is ensured by
- the types! For @{text "\<leftrightarrow>"} we can establish the following
- simplification rule:
-
- @{thm [display,indent=10] permute_flip_at[no_vars]}
-
- \noindent
- If we now want to swap the \emph{concrete} atoms @{text a} and @{text b}
- in the pair @{term "(a, b)"} we can establish the lemma as follows:
-*}
-
- lemma
- fixes a b :: "name"
- shows "(a \<leftrightarrow> b) \<bullet> (a, b) = (b, a)"
- by simp
-
-text {*
- \noindent
- There is no need to state an explicit premise involving sorts.
-
- We can automate the process of creating concrete atom types, so that users
- can define a new one simply by issuing the command
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- \isacommand{atom\_decl}~~@{text "name"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- This command can be implemented using less than 100 lines of custom ML-code.
-
-*}
-
-
-
-section {* Related Work\label{related} *}
-
-text {*
- Coq-tries, but failed
-
- Add here comparison with old work.
-
- In comparison, the old version of Nominal Isabelle included more than 1000
- lines of ML-code for creating concrete atom types, and for defining various
- type classes and instantiating generic lemmas for them. In addition to
- simplifying the ML-code, the setup here also offers user-visible improvements:
- Now concrete atoms can be declared at any point of a formalisation, and
- theories that separately declare different atom types can be merged
- together---it is no longer required to collect all atom declarations in one
- place.
-
- Using a single atom type to represent atoms of different sorts and
- representing permutations as functions are not new ideas; see
- \cite{GunterOsbornPopescu09} \footnote{function rep.} The main contribution
- of this paper is to show an example of how to make better theorem proving
- tools by choosing the right level of abstraction for the underlying
- theory---our design choices take advantage of Isabelle's type system, type
- classes and reasoning infrastructure. The novel technical contribution is a
- mechanism for dealing with ``Church-style'' lambda-terms \cite{Church40} and
- HOL-based languages \cite{PittsHOL4} where variables and variable binding
- depend on type annotations.
-
- The paper is organised as follows\ldots
-
-
- The main point is that the above reasoning blends smoothly with the reasoning
- infrastructure of Isabelle/HOL; no custom ML-code is necessary and a single
- type class suffices.
-
- With this
- design one can represent permutations as lists of pairs of atoms and the
- operation of applying a permutation to an object as the function
-
-
- @{text [display,indent=10] "_ \<bullet> _ :: (\<alpha> \<times> \<alpha>) list \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
-
- \noindent
- where @{text "\<alpha>"} stands for a type of atoms and @{text "\<beta>"} for the type
- of the objects on which the permutation acts. For atoms
- the permutation operation is defined over the length of lists as follows
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
- @{text "[] \<bullet> c"} & @{text "="} & @{text c}\\
- @{text "(a b)::\<pi> \<bullet> c"} & @{text "="} &
- $\begin{cases} @{text a} & \textrm{if}~@{text "\<pi> \<bullet> c = b"}\\
- @{text b} & \textrm{if}~@{text "\<pi> \<bullet> c = a"}\\
- @{text "\<pi> \<bullet> c"} & \textrm{otherwise}\end{cases}$
- \end{tabular}\hfill\numbered{atomperm}
- \end{isabelle}
-
- \noindent
- where we write @{text "(a b)"} for a swapping of atoms @{text "a"} and
- @{text "b"}. For atoms with different type than the permutation, we
- define @{text "\<pi> \<bullet> c \<equiv> c"}.
-
- With the separate atom types and the list representation of permutations it
- is impossible in systems like Isabelle/HOL to state an ``ill-sorted''
- permutation, since the type system excludes lists containing atoms of
- different type. However, a disadvantage is that whenever we need to
- generalise induction hypotheses by quantifying over permutations, we have to
- build quantifications like
-
- @{text [display,indent=10] "\<forall>\<pi>\<^isub>1 \<dots> \<forall>\<pi>\<^isub>n. \<dots>"}
-
- \noindent
- where the @{text "\<pi>\<^isub>i"} are of type @{text "(\<alpha>\<^isub>i \<times> \<alpha>\<^isub>i) list"}.
- The reason is that the permutation operation behaves differently for
- every @{text "\<alpha>\<^isub>i"} and the type system does not allow use to have a
- single quantification to stand for all permutations. Similarly, the
- notion of support
-
- @{text [display,indent=10] "supp _ :: \<beta> \<Rightarrow> \<alpha> set"}
-
- \noindent
- which we will define later, cannot be
- used to express the support of an object over \emph{all} atoms. The reason
- is that support can behave differently for each @{text
- "\<alpha>\<^isub>i"}. This problem is annoying, because if we need to know in
- a statement that an object, say @{text "x"}, is finitely supported we end up
- with having to state premises of the form
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{text "finite ((supp x) :: \<alpha>\<^isub>1 set) , \<dots>, finite ((supp x) :: \<alpha>\<^isub>n set)"}
- \end{tabular}\hfill\numbered{fssequence}
- \end{isabelle}
-
- \noindent
- Because of these disadvantages, we will use in this paper a single unified atom type to
- represent atoms of different sorts. Consequently, we have to deal with the
- case that a swapping of two atoms is ill-sorted: we cannot rely anymore on
- the type systems to exclude them.
-
- We also will not represent permutations as lists of pairs of atoms (as done in
- \cite{Urban08}). Although an
- advantage of this representation is that the basic operations on
- permutations are already defined in Isabelle's list library: composition of
- two permutations (written @{text "_ @ _"}) is just list append, and
- inversion of a permutation (written @{text "_\<^sup>-\<^sup>1"}) is just
- list reversal, and another advantage is that there is a well-understood
- induction principle for lists, a disadvantage is that permutations
- do not have unique representations as lists. We have to explicitly identify
- them according to the relation
-
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{text "\<pi>\<^isub>1 \<sim> \<pi>\<^isub>2 \<equiv> \<forall>a. \<pi>\<^isub>1 \<bullet> a = \<pi>\<^isub>2 \<bullet> a"}
- \end{tabular}\hfill\numbered{permequ}
- \end{isabelle}
-
- \noindent
- This is a problem when lifting the permutation operation to other types, for
- example sets, functions and so on. For this we need to ensure that every definition
- is well-behaved in the sense that it satisfies some
- \emph{permutation properties}. In the list representation we need
- to state these properties as follows:
-
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}r@ {\hspace{4mm}}p{10cm}}
- i) & @{text "[] \<bullet> x = x"}\\
- ii) & @{text "(\<pi>\<^isub>1 @ \<pi>\<^isub>2) \<bullet> x = \<pi>\<^isub>1 \<bullet> (\<pi>\<^isub>2 \<bullet> x)"}\\
- iii) & if @{text "\<pi>\<^isub>1 \<sim> \<pi>\<^isub>2"} then @{text "\<pi>\<^isub>1 \<bullet> x = \<pi>\<^isub>2 \<bullet> x"}
- \end{tabular}\hfill\numbered{permprops}
- \end{isabelle}
-
- \noindent
- where the last clause explicitly states that the permutation operation has
- to produce the same result for related permutations. Moreover,
- ``permutations-as-lists'' do not satisfy the group properties. This means by
- using this representation we will not be able to reuse the extensive
- reasoning infrastructure in Isabelle about groups. Because of this, we will represent
- in this paper permutations as functions from atoms to atoms. This representation
- is unique and satisfies the laws of non-commutative groups.
-*}
-
-
-section {* Conclusion *}
-
-text {*
- This proof pearl describes a new formalisation of the nominal logic work by
- Pitts et al. With the definitions we presented here, the formal reasoning blends
- smoothly with the infrastructure of the Isabelle/HOL theorem prover.
- Therefore the formalisation will be the underlying theory for a
- new version of Nominal Isabelle.
-
- The main difference of this paper with respect to existing work on Nominal
- Isabelle is the representation of atoms and permutations. First, we used a
- single type for sorted atoms. This design choice means for a term @{term t},
- say, that its support is completely characterised by @{term "supp t"}, even
- if the term contains different kinds of atoms. Also, whenever we have to
- generalise an induction so that a property @{text P} is not just established
- for all @{text t}, but for all @{text t} \emph{and} under all permutations
- @{text \<pi>}, then we only have to state @{term "\<forall>\<pi>. P (\<pi> \<bullet> t)"}. The reason is
- that permutations can now consist of multiple swapping each of which can
- swap different kinds of atoms. This simplifies considerably the reasoning
- involved in building Nominal Isabelle.
-
- Second, we represented permutations as functions so that the associated
- permutation operation has only a single type parameter. This is very convenient
- because the abstract reasoning about permutations fits cleanly
- with Isabelle/HOL's type classes. No custom ML-code is required to work
- around rough edges. Moreover, by establishing that our permutations-as-functions
- representation satisfy the group properties, we were able to use extensively
- Isabelle/HOL's reasoning infrastructure for groups. This often reduced proofs
- to simple calculations over @{text "+"}, @{text "-"} and @{text "0"}.
- An interesting point is that we defined the swapping operation so that a
- swapping of two atoms with different sorts is \emph{not} excluded, like
- in our older work on Nominal Isabelle, but there is no ``effect'' of such
- a swapping (it is defined as the identity). This is a crucial insight
- in order to make the approach based on a single type of sorted atoms to work.
- But of course it is analogous to the well-known trick of defining division by
- zero to return zero.
-
- We noticed only one disadvantage of the permutations-as-functions: Over
- lists we can easily perform inductions. For permutations made up from
- functions, we have to manually derive an appropriate induction principle. We
- can establish such a principle, but we have no real experience yet whether ours
- is the most useful principle: such an induction principle was not needed in
- any of the reasoning we ported from the old Nominal Isabelle, except
- when showing that if @{term "\<forall>a \<in> supp x. a \<sharp> p"} implies @{term "p \<bullet> x = x"}.
-
- Finally, our implementation of sorted atoms turned out powerful enough to
- use it for representing variables that carry on additional information, for
- example typing annotations. This information is encoded into the sorts. With
- this we can represent conveniently binding in ``Church-style'' lambda-terms
- and HOL-based languages. While dealing with such additional information in
- dependent type-theories, such as LF or Coq, is straightforward, we are not
- aware of any other approach in a non-dependent HOL-setting that can deal
- conveniently with such binders.
-
- The formalisation presented here will eventually become part of the Isabelle
- distribution, but for the moment it can be downloaded from the
- Mercurial repository linked at
- \href{http://isabelle.in.tum.de/nominal/download}
- {http://isabelle.in.tum.de/nominal/download}.\smallskip
-
- \noindent
- {\bf Acknowledgements:} We are very grateful to Jesper Bengtson, Stefan
- Berghofer and Cezary Kaliszyk for their comments on earlier versions
- of this paper. We are also grateful to the anonymous referee who helped us to
- put the work into the right context.
-*}
-
-
-(*<*)
-end
-(*>*)
\ No newline at end of file
--- a/Pearl-jv/ROOT.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,5 +0,0 @@
-
-use_thys ["../Nominal/Nominal2_Base",
- "../Nominal/Atoms",
- "../Nominal/Nominal2_Abs",
- "~~/src/HOL/Library/LaTeXsugar"];
--- a/Pearl-jv/ROOT2.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-use_thys ["Paper"];
\ No newline at end of file
--- a/Pearl-jv/document/llncs.cls Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1189 +0,0 @@
-% LLNCS DOCUMENT CLASS -- version 2.13 (28-Jan-2002)
-% Springer Verlag LaTeX2e support for Lecture Notes in Computer Science
-%
-%%
-%% \CharacterTable
-%% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
-%% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
-%% Digits \0\1\2\3\4\5\6\7\8\9
-%% Exclamation \! Double quote \" Hash (number) \#
-%% Dollar \$ Percent \% Ampersand \&
-%% Acute accent \' Left paren \( Right paren \)
-%% Asterisk \* Plus \+ Comma \,
-%% Minus \- Point \. Solidus \/
-%% Colon \: Semicolon \; Less than \<
-%% Equals \= Greater than \> Question mark \?
-%% Commercial at \@ Left bracket \[ Backslash \\
-%% Right bracket \] Circumflex \^ Underscore \_
-%% Grave accent \` Left brace \{ Vertical bar \|
-%% Right brace \} Tilde \~}
-%%
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-\ProvidesClass{llncs}[2002/01/28 v2.13
-^^J LaTeX document class for Lecture Notes in Computer Science]
-% Options
-\let\if@envcntreset\iffalse
-\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue}
-\DeclareOption{citeauthoryear}{\let\citeauthoryear=Y}
-\DeclareOption{oribibl}{\let\oribibl=Y}
-\let\if@custvec\iftrue
-\DeclareOption{orivec}{\let\if@custvec\iffalse}
-\let\if@envcntsame\iffalse
-\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue}
-\let\if@envcntsect\iffalse
-\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue}
-\let\if@runhead\iffalse
-\DeclareOption{runningheads}{\let\if@runhead\iftrue}
-
-\let\if@openbib\iffalse
-\DeclareOption{openbib}{\let\if@openbib\iftrue}
-
-% languages
-\let\switcht@@therlang\relax
-\def\ds@deutsch{\def\switcht@@therlang{\switcht@deutsch}}
-\def\ds@francais{\def\switcht@@therlang{\switcht@francais}}
-
-\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}}
-
-\ProcessOptions
-
-\LoadClass[twoside]{article}
-\RequirePackage{multicol} % needed for the list of participants, index
-
-\setlength{\textwidth}{12.2cm}
-\setlength{\textheight}{19.3cm}
-\renewcommand\@pnumwidth{2em}
-\renewcommand\@tocrmarg{3.5em}
-%
-\def\@dottedtocline#1#2#3#4#5{%
- \ifnum #1>\c@tocdepth \else
- \vskip \z@ \@plus.2\p@
- {\leftskip #2\relax \rightskip \@tocrmarg \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \parindent #2\relax\@afterindenttrue
- \interlinepenalty\@M
- \leavevmode
- \@tempdima #3\relax
- \advance\leftskip \@tempdima \null\nobreak\hskip -\leftskip
- {#4}\nobreak
- \leaders\hbox{$\m@th
- \mkern \@dotsep mu\hbox{.}\mkern \@dotsep
- mu$}\hfill
- \nobreak
- \hb@xt@\@pnumwidth{\hfil\normalfont \normalcolor #5}%
- \par}%
- \fi}
-%
-\def\switcht@albion{%
-\def\abstractname{Abstract.}
-\def\ackname{Acknowledgement.}
-\def\andname{and}
-\def\lastandname{\unskip, and}
-\def\appendixname{Appendix}
-\def\chaptername{Chapter}
-\def\claimname{Claim}
-\def\conjecturename{Conjecture}
-\def\contentsname{Table of Contents}
-\def\corollaryname{Corollary}
-\def\definitionname{Definition}
-\def\examplename{Example}
-\def\exercisename{Exercise}
-\def\figurename{Fig.}
-\def\keywordname{{\bf Key words:}}
-\def\indexname{Index}
-\def\lemmaname{Lemma}
-\def\contriblistname{List of Contributors}
-\def\listfigurename{List of Figures}
-\def\listtablename{List of Tables}
-\def\mailname{{\it Correspondence to\/}:}
-\def\noteaddname{Note added in proof}
-\def\notename{Note}
-\def\partname{Part}
-\def\problemname{Problem}
-\def\proofname{Proof}
-\def\propertyname{Property}
-\def\propositionname{Proposition}
-\def\questionname{Question}
-\def\remarkname{Remark}
-\def\seename{see}
-\def\solutionname{Solution}
-\def\subclassname{{\it Subject Classifications\/}:}
-\def\tablename{Table}
-\def\theoremname{Theorem}}
-\switcht@albion
-% Names of theorem like environments are already defined
-% but must be translated if another language is chosen
-%
-% French section
-\def\switcht@francais{%\typeout{On parle francais.}%
- \def\abstractname{R\'esum\'e.}%
- \def\ackname{Remerciements.}%
- \def\andname{et}%
- \def\lastandname{ et}%
- \def\appendixname{Appendice}
- \def\chaptername{Chapitre}%
- \def\claimname{Pr\'etention}%
- \def\conjecturename{Hypoth\`ese}%
- \def\contentsname{Table des mati\`eres}%
- \def\corollaryname{Corollaire}%
- \def\definitionname{D\'efinition}%
- \def\examplename{Exemple}%
- \def\exercisename{Exercice}%
- \def\figurename{Fig.}%
- \def\keywordname{{\bf Mots-cl\'e:}}
- \def\indexname{Index}
- \def\lemmaname{Lemme}%
- \def\contriblistname{Liste des contributeurs}
- \def\listfigurename{Liste des figures}%
- \def\listtablename{Liste des tables}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}%
- \def\notename{Remarque}%
- \def\partname{Partie}%
- \def\problemname{Probl\`eme}%
- \def\proofname{Preuve}%
- \def\propertyname{Caract\'eristique}%
-%\def\propositionname{Proposition}%
- \def\questionname{Question}%
- \def\remarkname{Remarque}%
- \def\seename{voir}
- \def\solutionname{Solution}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tableau}%
- \def\theoremname{Th\'eor\`eme}%
-}
-%
-% German section
-\def\switcht@deutsch{%\typeout{Man spricht deutsch.}%
- \def\abstractname{Zusammenfassung.}%
- \def\ackname{Danksagung.}%
- \def\andname{und}%
- \def\lastandname{ und}%
- \def\appendixname{Anhang}%
- \def\chaptername{Kapitel}%
- \def\claimname{Behauptung}%
- \def\conjecturename{Hypothese}%
- \def\contentsname{Inhaltsverzeichnis}%
- \def\corollaryname{Korollar}%
-%\def\definitionname{Definition}%
- \def\examplename{Beispiel}%
- \def\exercisename{\"Ubung}%
- \def\figurename{Abb.}%
- \def\keywordname{{\bf Schl\"usselw\"orter:}}
- \def\indexname{Index}
-%\def\lemmaname{Lemma}%
- \def\contriblistname{Mitarbeiter}
- \def\listfigurename{Abbildungsverzeichnis}%
- \def\listtablename{Tabellenverzeichnis}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Nachtrag}%
- \def\notename{Anmerkung}%
- \def\partname{Teil}%
-%\def\problemname{Problem}%
- \def\proofname{Beweis}%
- \def\propertyname{Eigenschaft}%
-%\def\propositionname{Proposition}%
- \def\questionname{Frage}%
- \def\remarkname{Anmerkung}%
- \def\seename{siehe}
- \def\solutionname{L\"osung}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tabelle}%
-%\def\theoremname{Theorem}%
-}
-
-% Ragged bottom for the actual page
-\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil
-\global\let\@textbottom\relax}}
-
-\renewcommand\small{%
- \@setfontsize\small\@ixpt{11}%
- \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
- \abovedisplayshortskip \z@ \@plus2\p@
- \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
- \def\@listi{\leftmargin\leftmargini
- \parsep 0\p@ \@plus1\p@ \@minus\p@
- \topsep 8\p@ \@plus2\p@ \@minus4\p@
- \itemsep0\p@}%
- \belowdisplayskip \abovedisplayskip
-}
-
-\frenchspacing
-\widowpenalty=10000
-\clubpenalty=10000
-
-\setlength\oddsidemargin {63\p@}
-\setlength\evensidemargin {63\p@}
-\setlength\marginparwidth {90\p@}
-
-\setlength\headsep {16\p@}
-
-\setlength\footnotesep{7.7\p@}
-\setlength\textfloatsep{8mm\@plus 2\p@ \@minus 4\p@}
-\setlength\intextsep {8mm\@plus 2\p@ \@minus 2\p@}
-
-\setcounter{secnumdepth}{2}
-
-\newcounter {chapter}
-\renewcommand\thechapter {\@arabic\c@chapter}
-
-\newif\if@mainmatter \@mainmattertrue
-\newcommand\frontmatter{\cleardoublepage
- \@mainmatterfalse\pagenumbering{Roman}}
-\newcommand\mainmatter{\cleardoublepage
- \@mainmattertrue\pagenumbering{arabic}}
-\newcommand\backmatter{\if@openright\cleardoublepage\else\clearpage\fi
- \@mainmatterfalse}
-
-\renewcommand\part{\cleardoublepage
- \thispagestyle{empty}%
- \if@twocolumn
- \onecolumn
- \@tempswatrue
- \else
- \@tempswafalse
- \fi
- \null\vfil
- \secdef\@part\@spart}
-
-\def\@part[#1]#2{%
- \ifnum \c@secnumdepth >-2\relax
- \refstepcounter{part}%
- \addcontentsline{toc}{part}{\thepart\hspace{1em}#1}%
- \else
- \addcontentsline{toc}{part}{#1}%
- \fi
- \markboth{}{}%
- {\centering
- \interlinepenalty \@M
- \normalfont
- \ifnum \c@secnumdepth >-2\relax
- \huge\bfseries \partname~\thepart
- \par
- \vskip 20\p@
- \fi
- \Huge \bfseries #2\par}%
- \@endpart}
-\def\@spart#1{%
- {\centering
- \interlinepenalty \@M
- \normalfont
- \Huge \bfseries #1\par}%
- \@endpart}
-\def\@endpart{\vfil\newpage
- \if@twoside
- \null
- \thispagestyle{empty}%
- \newpage
- \fi
- \if@tempswa
- \twocolumn
- \fi}
-
-\newcommand\chapter{\clearpage
- \thispagestyle{empty}%
- \global\@topnum\z@
- \@afterindentfalse
- \secdef\@chapter\@schapter}
-\def\@chapter[#1]#2{\ifnum \c@secnumdepth >\m@ne
- \if@mainmatter
- \refstepcounter{chapter}%
- \typeout{\@chapapp\space\thechapter.}%
- \addcontentsline{toc}{chapter}%
- {\protect\numberline{\thechapter}#1}%
- \else
- \addcontentsline{toc}{chapter}{#1}%
- \fi
- \else
- \addcontentsline{toc}{chapter}{#1}%
- \fi
- \chaptermark{#1}%
- \addtocontents{lof}{\protect\addvspace{10\p@}}%
- \addtocontents{lot}{\protect\addvspace{10\p@}}%
- \if@twocolumn
- \@topnewpage[\@makechapterhead{#2}]%
- \else
- \@makechapterhead{#2}%
- \@afterheading
- \fi}
-\def\@makechapterhead#1{%
-% \vspace*{50\p@}%
- {\centering
- \ifnum \c@secnumdepth >\m@ne
- \if@mainmatter
- \large\bfseries \@chapapp{} \thechapter
- \par\nobreak
- \vskip 20\p@
- \fi
- \fi
- \interlinepenalty\@M
- \Large \bfseries #1\par\nobreak
- \vskip 40\p@
- }}
-\def\@schapter#1{\if@twocolumn
- \@topnewpage[\@makeschapterhead{#1}]%
- \else
- \@makeschapterhead{#1}%
- \@afterheading
- \fi}
-\def\@makeschapterhead#1{%
-% \vspace*{50\p@}%
- {\centering
- \normalfont
- \interlinepenalty\@M
- \Large \bfseries #1\par\nobreak
- \vskip 40\p@
- }}
-
-\renewcommand\section{\@startsection{section}{1}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {12\p@ \@plus 4\p@ \@minus 4\p@}%
- {\normalfont\large\bfseries\boldmath
- \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsection{\@startsection{subsection}{2}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {8\p@ \@plus 4\p@ \@minus 4\p@}%
- {\normalfont\normalsize\bfseries\boldmath
- \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {-0.5em \@plus -0.22em \@minus -0.1em}%
- {\normalfont\normalsize\bfseries\boldmath}}
-\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}%
- {-12\p@ \@plus -4\p@ \@minus -4\p@}%
- {-0.5em \@plus -0.22em \@minus -0.1em}%
- {\normalfont\normalsize\itshape}}
-\renewcommand\subparagraph[1]{\typeout{LLNCS warning: You should not use
- \string\subparagraph\space with this class}\vskip0.5cm
-You should not use \verb|\subparagraph| with this class.\vskip0.5cm}
-
-\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00}
-\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01}
-\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02}
-\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03}
-\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04}
-\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05}
-\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06}
-\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07}
-\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08}
-\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09}
-\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A}
-
-\let\footnotesize\small
-
-\if@custvec
-\def\vec#1{\mathchoice{\mbox{\boldmath$\displaystyle#1$}}
-{\mbox{\boldmath$\textstyle#1$}}
-{\mbox{\boldmath$\scriptstyle#1$}}
-{\mbox{\boldmath$\scriptscriptstyle#1$}}}
-\fi
-
-\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}}
-\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil
-\penalty50\hskip1em\null\nobreak\hfil\squareforqed
-\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi}
-
-\def\getsto{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr\gets\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-\gets\cr\to\cr}}}}}
-\def\lid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr<\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr<\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr<\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-<\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\gid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr>\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr>\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\grole{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr
->\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.8pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.3pt}<\cr}}}}}
-\def\bbbr{{\rm I\!R}} %reelle Zahlen
-\def\bbbm{{\rm I\!M}}
-\def\bbbn{{\rm I\!N}} %natuerliche Zahlen
-\def\bbbf{{\rm I\!F}}
-\def\bbbh{{\rm I\!H}}
-\def\bbbk{{\rm I\!K}}
-\def\bbbp{{\rm I\!P}}
-\def\bbbone{{\mathchoice {\rm 1\mskip-4mu l} {\rm 1\mskip-4mu l}
-{\rm 1\mskip-4.5mu l} {\rm 1\mskip-5mu l}}}
-\def\bbbc{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbq{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}}}
-\def\bbbt{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-T$}\hbox{\hbox to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbs{{\mathchoice
-{\setbox0=\hbox{$\displaystyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.5\wd0\vrule height0.45\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.4\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.55\wd0\vrule height0.45\ht0\hss}\box0}}}}
-\def\bbbz{{\mathchoice {\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\scriptstyle Z\kern-0.3em Z$}}
-{\hbox{$\mathsf\scriptscriptstyle Z\kern-0.2em Z$}}}}
-
-\let\ts\,
-
-\setlength\leftmargini {17\p@}
-\setlength\leftmargin {\leftmargini}
-\setlength\leftmarginii {\leftmargini}
-\setlength\leftmarginiii {\leftmargini}
-\setlength\leftmarginiv {\leftmargini}
-\setlength \labelsep {.5em}
-\setlength \labelwidth{\leftmargini}
-\addtolength\labelwidth{-\labelsep}
-
-\def\@listI{\leftmargin\leftmargini
- \parsep 0\p@ \@plus1\p@ \@minus\p@
- \topsep 8\p@ \@plus2\p@ \@minus4\p@
- \itemsep0\p@}
-\let\@listi\@listI
-\@listi
-\def\@listii {\leftmargin\leftmarginii
- \labelwidth\leftmarginii
- \advance\labelwidth-\labelsep
- \topsep 0\p@ \@plus2\p@ \@minus\p@}
-\def\@listiii{\leftmargin\leftmarginiii
- \labelwidth\leftmarginiii
- \advance\labelwidth-\labelsep
- \topsep 0\p@ \@plus\p@\@minus\p@
- \parsep \z@
- \partopsep \p@ \@plus\z@ \@minus\p@}
-
-\renewcommand\labelitemi{\normalfont\bfseries --}
-\renewcommand\labelitemii{$\m@th\bullet$}
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\def\tableofcontents{\chapter*{\contentsname\@mkboth{{\contentsname}}%
- {{\contentsname}}}
- \def\authcount##1{\setcounter{auco}{##1}\setcounter{@auth}{1}}
- \def\lastand{\ifnum\value{auco}=2\relax
- \unskip{} \andname\
- \else
- \unskip \lastandname\
- \fi}%
- \def\and{\stepcounter{@auth}\relax
- \ifnum\value{@auth}=\value{auco}%
- \lastand
- \else
- \unskip,
- \fi}%
- \@starttoc{toc}\if@restonecol\twocolumn\fi}
-
-\def\l@part#1#2{\addpenalty{\@secpenalty}%
- \addvspace{2em plus\p@}% % space above part line
- \begingroup
- \parindent \z@
- \rightskip \z@ plus 5em
- \hrule\vskip5pt
- \large % same size as for a contribution heading
- \bfseries\boldmath % set line in boldface
- \leavevmode % TeX command to enter horizontal mode.
- #1\par
- \vskip5pt
- \hrule
- \vskip1pt
- \nobreak % Never break after part entry
- \endgroup}
-
-\def\@dotsep{2}
-
-\def\hyperhrefextend{\ifx\hyper@anchor\@undefined\else
-{chapter.\thechapter}\fi}
-
-\def\addnumcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{\protect\numberline
- {\thechapter}#3}{\thepage}\hyperhrefextend}}
-\def\addcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{\thepage}\hyperhrefextend}}
-\def\addcontentsmarkwop#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{0}\hyperhrefextend}}
-
-\def\@adcmk[#1]{\ifcase #1 \or
-\def\@gtempa{\addnumcontentsmark}%
- \or \def\@gtempa{\addcontentsmark}%
- \or \def\@gtempa{\addcontentsmarkwop}%
- \fi\@gtempa{toc}{chapter}}
-\def\addtocmark{\@ifnextchar[{\@adcmk}{\@adcmk[3]}}
-
-\def\l@chapter#1#2{\addpenalty{-\@highpenalty}
- \vskip 1.0em plus 1pt \@tempdima 1.5em \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- {\large\bfseries\boldmath#1}\ifx0#2\hfil\null
- \else
- \nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}%
- \fi\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@title#1#2{\addpenalty{-\@highpenalty}
- \addvspace{8pt plus 1pt}
- \@tempdima \z@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- #1\nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@author#1#2{\addpenalty{\@highpenalty}
- \@tempdima=\z@ %15\p@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima %\hskip -\leftskip
- \textit{#1}\par
- \penalty\@highpenalty \endgroup}
-
-%\setcounter{tocdepth}{0}
-\newdimen\tocchpnum
-\newdimen\tocsecnum
-\newdimen\tocsectotal
-\newdimen\tocsubsecnum
-\newdimen\tocsubsectotal
-\newdimen\tocsubsubsecnum
-\newdimen\tocsubsubsectotal
-\newdimen\tocparanum
-\newdimen\tocparatotal
-\newdimen\tocsubparanum
-\tocchpnum=\z@ % no chapter numbers
-\tocsecnum=15\p@ % section 88. plus 2.222pt
-\tocsubsecnum=23\p@ % subsection 88.8 plus 2.222pt
-\tocsubsubsecnum=27\p@ % subsubsection 88.8.8 plus 1.444pt
-\tocparanum=35\p@ % paragraph 88.8.8.8 plus 1.666pt
-\tocsubparanum=43\p@ % subparagraph 88.8.8.8.8 plus 1.888pt
-\def\calctocindent{%
-\tocsectotal=\tocchpnum
-\advance\tocsectotal by\tocsecnum
-\tocsubsectotal=\tocsectotal
-\advance\tocsubsectotal by\tocsubsecnum
-\tocsubsubsectotal=\tocsubsectotal
-\advance\tocsubsubsectotal by\tocsubsubsecnum
-\tocparatotal=\tocsubsubsectotal
-\advance\tocparatotal by\tocparanum}
-\calctocindent
-
-\def\l@section{\@dottedtocline{1}{\tocchpnum}{\tocsecnum}}
-\def\l@subsection{\@dottedtocline{2}{\tocsectotal}{\tocsubsecnum}}
-\def\l@subsubsection{\@dottedtocline{3}{\tocsubsectotal}{\tocsubsubsecnum}}
-\def\l@paragraph{\@dottedtocline{4}{\tocsubsubsectotal}{\tocparanum}}
-\def\l@subparagraph{\@dottedtocline{5}{\tocparatotal}{\tocsubparanum}}
-
-\def\listoffigures{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}}
- \@starttoc{lof}\if@restonecol\twocolumn\fi}
-\def\l@figure{\@dottedtocline{1}{0em}{1.5em}}
-
-\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}}
- \@starttoc{lot}\if@restonecol\twocolumn\fi}
-\let\l@table\l@figure
-
-\renewcommand\listoffigures{%
- \section*{\listfigurename
- \@mkboth{\listfigurename}{\listfigurename}}%
- \@starttoc{lof}%
- }
-
-\renewcommand\listoftables{%
- \section*{\listtablename
- \@mkboth{\listtablename}{\listtablename}}%
- \@starttoc{lot}%
- }
-
-\ifx\oribibl\undefined
-\ifx\citeauthoryear\undefined
-\renewenvironment{thebibliography}[1]
- {\section*{\refname}
- \def\@biblabel##1{##1.}
- \small
- \list{\@biblabel{\@arabic\c@enumiv}}%
- {\settowidth\labelwidth{\@biblabel{#1}}%
- \leftmargin\labelwidth
- \advance\leftmargin\labelsep
- \if@openbib
- \advance\leftmargin\bibindent
- \itemindent -\bibindent
- \listparindent \itemindent
- \parsep \z@
- \fi
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{\@arabic\c@enumiv}}%
- \if@openbib
- \renewcommand\newblock{\par}%
- \else
- \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
- \fi
- \sloppy\clubpenalty4000\widowpenalty4000%
- \sfcode`\.=\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
-\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw
- {\let\protect\noexpand\immediate
- \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
-\newcount\@tempcntc
-\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi
- \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do
- {\@ifundefined
- {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries
- ?}\@warning
- {Citation `\@citeb' on page \thepage \space undefined}}%
- {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}%
- \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne
- \@citea\def\@citea{,}\hbox{\csname b@\@citeb\endcsname}%
- \else
- \advance\@tempcntb\@ne
- \ifnum\@tempcntb=\@tempcntc
- \else\advance\@tempcntb\m@ne\@citeo
- \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}}
-\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else
- \@citea\def\@citea{,\,\hskip\z@skip}%
- \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else
- {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else
- \def\@citea{--}\fi
- \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi}
-\else
-\renewenvironment{thebibliography}[1]
- {\section*{\refname}
- \small
- \list{}%
- {\settowidth\labelwidth{}%
- \leftmargin\parindent
- \itemindent=-\parindent
- \labelsep=\z@
- \if@openbib
- \advance\leftmargin\bibindent
- \itemindent -\bibindent
- \listparindent \itemindent
- \parsep \z@
- \fi
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{}}%
- \if@openbib
- \renewcommand\newblock{\par}%
- \else
- \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
- \fi
- \sloppy\clubpenalty4000\widowpenalty4000%
- \sfcode`\.=\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
- \def\@cite#1{#1}%
- \def\@lbibitem[#1]#2{\item[]\if@filesw
- {\def\protect##1{\string ##1\space}\immediate
- \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
- \fi
-\else
-\@cons\@openbib@code{\noexpand\small}
-\fi
-
-\def\idxquad{\hskip 10\p@}% space that divides entry from number
-
-\def\@idxitem{\par\hangindent 10\p@}
-
-\def\subitem{\par\setbox0=\hbox{--\enspace}% second order
- \noindent\hangindent\wd0\box0}% index entry
-
-\def\subsubitem{\par\setbox0=\hbox{--\,--\enspace}% third
- \noindent\hangindent\wd0\box0}% order index entry
-
-\def\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax}
-
-\renewenvironment{theindex}
- {\@mkboth{\indexname}{\indexname}%
- \thispagestyle{empty}\parindent\z@
- \parskip\z@ \@plus .3\p@\relax
- \let\item\par
- \def\,{\relax\ifmmode\mskip\thinmuskip
- \else\hskip0.2em\ignorespaces\fi}%
- \normalfont\small
- \begin{multicols}{2}[\@makeschapterhead{\indexname}]%
- }
- {\end{multicols}}
-
-\renewcommand\footnoterule{%
- \kern-3\p@
- \hrule\@width 2truecm
- \kern2.6\p@}
- \newdimen\fnindent
- \fnindent1em
-\long\def\@makefntext#1{%
- \parindent \fnindent%
- \leftskip \fnindent%
- \noindent
- \llap{\hb@xt@1em{\hss\@makefnmark\ }}\ignorespaces#1}
-
-\long\def\@makecaption#1#2{%
- \vskip\abovecaptionskip
- \sbox\@tempboxa{{\bfseries #1.} #2}%
- \ifdim \wd\@tempboxa >\hsize
- {\bfseries #1.} #2\par
- \else
- \global \@minipagefalse
- \hb@xt@\hsize{\hfil\box\@tempboxa\hfil}%
- \fi
- \vskip\belowcaptionskip}
-
-\def\fps@figure{htbp}
-\def\fnum@figure{\figurename\thinspace\thefigure}
-\def \@floatboxreset {%
- \reset@font
- \small
- \@setnobreak
- \@setminipage
-}
-\def\fps@table{htbp}
-\def\fnum@table{\tablename~\thetable}
-\renewenvironment{table}
- {\setlength\abovecaptionskip{0\p@}%
- \setlength\belowcaptionskip{10\p@}%
- \@float{table}}
- {\end@float}
-\renewenvironment{table*}
- {\setlength\abovecaptionskip{0\p@}%
- \setlength\belowcaptionskip{10\p@}%
- \@dblfloat{table}}
- {\end@dblfloat}
-
-\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname
- ext@#1\endcsname}{#1}{\protect\numberline{\csname
- the#1\endcsname}{\ignorespaces #2}}\begingroup
- \@parboxrestore
- \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
- \endgroup}
-
-% LaTeX does not provide a command to enter the authors institute
-% addresses. The \institute command is defined here.
-
-\newcounter{@inst}
-\newcounter{@auth}
-\newcounter{auco}
-\newdimen\instindent
-\newbox\authrun
-\newtoks\authorrunning
-\newtoks\tocauthor
-\newbox\titrun
-\newtoks\titlerunning
-\newtoks\toctitle
-
-\def\clearheadinfo{\gdef\@author{No Author Given}%
- \gdef\@title{No Title Given}%
- \gdef\@subtitle{}%
- \gdef\@institute{No Institute Given}%
- \gdef\@thanks{}%
- \global\titlerunning={}\global\authorrunning={}%
- \global\toctitle={}\global\tocauthor={}}
-
-\def\institute#1{\gdef\@institute{#1}}
-
-\def\institutename{\par
- \begingroup
- \parskip=\z@
- \parindent=\z@
- \setcounter{@inst}{1}%
- \def\and{\par\stepcounter{@inst}%
- \noindent$^{\the@inst}$\enspace\ignorespaces}%
- \setbox0=\vbox{\def\thanks##1{}\@institute}%
- \ifnum\c@@inst=1\relax
- \gdef\fnnstart{0}%
- \else
- \xdef\fnnstart{\c@@inst}%
- \setcounter{@inst}{1}%
- \noindent$^{\the@inst}$\enspace
- \fi
- \ignorespaces
- \@institute\par
- \endgroup}
-
-\def\@fnsymbol#1{\ensuremath{\ifcase#1\or\star\or{\star\star}\or
- {\star\star\star}\or \dagger\or \ddagger\or
- \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger
- \or \ddagger\ddagger \else\@ctrerr\fi}}
-
-\def\inst#1{\unskip$^{#1}$}
-\def\fnmsep{\unskip$^,$}
-\def\email#1{{\tt#1}}
-\AtBeginDocument{\@ifundefined{url}{\def\url#1{#1}}{}%
-\@ifpackageloaded{babel}{%
-\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}%
-\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}%
-\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}%
-}{\switcht@@therlang}%
-}
-\def\homedir{\~{ }}
-
-\def\subtitle#1{\gdef\@subtitle{#1}}
-\clearheadinfo
-
-\renewcommand\maketitle{\newpage
- \refstepcounter{chapter}%
- \stepcounter{section}%
- \setcounter{section}{0}%
- \setcounter{subsection}{0}%
- \setcounter{figure}{0}
- \setcounter{table}{0}
- \setcounter{equation}{0}
- \setcounter{footnote}{0}%
- \begingroup
- \parindent=\z@
- \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
- \if@twocolumn
- \ifnum \col@number=\@ne
- \@maketitle
- \else
- \twocolumn[\@maketitle]%
- \fi
- \else
- \newpage
- \global\@topnum\z@ % Prevents figures from going at top of page.
- \@maketitle
- \fi
- \thispagestyle{empty}\@thanks
-%
- \def\\{\unskip\ \ignorespaces}\def\inst##1{\unskip{}}%
- \def\thanks##1{\unskip{}}\def\fnmsep{\unskip}%
- \instindent=\hsize
- \advance\instindent by-\headlineindent
-% \if!\the\toctitle!\addcontentsline{toc}{title}{\@title}\else
-% \addcontentsline{toc}{title}{\the\toctitle}\fi
- \if@runhead
- \if!\the\titlerunning!\else
- \edef\@title{\the\titlerunning}%
- \fi
- \global\setbox\titrun=\hbox{\small\rm\unboldmath\ignorespaces\@title}%
- \ifdim\wd\titrun>\instindent
- \typeout{Title too long for running head. Please supply}%
- \typeout{a shorter form with \string\titlerunning\space prior to
- \string\maketitle}%
- \global\setbox\titrun=\hbox{\small\rm
- Title Suppressed Due to Excessive Length}%
- \fi
- \xdef\@title{\copy\titrun}%
- \fi
-%
- \if!\the\tocauthor!\relax
- {\def\and{\noexpand\protect\noexpand\and}%
- \protected@xdef\toc@uthor{\@author}}%
- \else
- \def\\{\noexpand\protect\noexpand\newline}%
- \protected@xdef\scratch{\the\tocauthor}%
- \protected@xdef\toc@uthor{\scratch}%
- \fi
-% \addcontentsline{toc}{author}{\toc@uthor}%
- \if@runhead
- \if!\the\authorrunning!
- \value{@inst}=\value{@auth}%
- \setcounter{@auth}{1}%
- \else
- \edef\@author{\the\authorrunning}%
- \fi
- \global\setbox\authrun=\hbox{\small\unboldmath\@author\unskip}%
- \ifdim\wd\authrun>\instindent
- \typeout{Names of authors too long for running head. Please supply}%
- \typeout{a shorter form with \string\authorrunning\space prior to
- \string\maketitle}%
- \global\setbox\authrun=\hbox{\small\rm
- Authors Suppressed Due to Excessive Length}%
- \fi
- \xdef\@author{\copy\authrun}%
- \markboth{\@author}{\@title}%
- \fi
- \endgroup
- \setcounter{footnote}{\fnnstart}%
- \clearheadinfo}
-%
-\def\@maketitle{\newpage
- \markboth{}{}%
- \def\lastand{\ifnum\value{@inst}=2\relax
- \unskip{} \andname\
- \else
- \unskip \lastandname\
- \fi}%
- \def\and{\stepcounter{@auth}\relax
- \ifnum\value{@auth}=\value{@inst}%
- \lastand
- \else
- \unskip,
- \fi}%
- \begin{center}%
- \let\newline\\
- {\Large \bfseries\boldmath
- \pretolerance=10000
- \@title \par}\vskip .8cm
-\if!\@subtitle!\else {\large \bfseries\boldmath
- \vskip -.65cm
- \pretolerance=10000
- \@subtitle \par}\vskip .8cm\fi
- \setbox0=\vbox{\setcounter{@auth}{1}\def\and{\stepcounter{@auth}}%
- \def\thanks##1{}\@author}%
- \global\value{@inst}=\value{@auth}%
- \global\value{auco}=\value{@auth}%
- \setcounter{@auth}{1}%
-{\lineskip .5em
-\noindent\ignorespaces
-\@author\vskip.35cm}
- {\small\institutename}
- \end{center}%
- }
-
-% definition of the "\spnewtheorem" command.
-%
-% Usage:
-%
-% \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font}
-% or \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font}
-% or \spnewtheorem*{env_nam}{caption}{cap_font}{body_font}
-%
-% New is "cap_font" and "body_font". It stands for
-% fontdefinition of the caption and the text itself.
-%
-% "\spnewtheorem*" gives a theorem without number.
-%
-% A defined spnewthoerem environment is used as described
-% by Lamport.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\def\@thmcountersep{}
-\def\@thmcounterend{.}
-
-\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}}
-
-% definition of \spnewtheorem with number
-
-\def\@spnthm#1#2{%
- \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}}
-\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}}
-
-\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}\@addtoreset{#1}{#3}%
- \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand
- \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}%
- \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spothm#1[#2]#3#4#5{%
- \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}%
- {\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{the#1}{\@nameuse{the#2}}%
- \expandafter\xdef\csname #1name\endcsname{#3}%
- \global\@namedef{#1}{\@spthm{#2}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}}
-
-\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\refstepcounter{#1}%
-\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}}
-
-\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}%
- \ignorespaces}
-
-\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname
- the#1\endcsname}{#5}{#3}{#4}\ignorespaces}
-
-\def\@spbegintheorem#1#2#3#4{\trivlist
- \item[\hskip\labelsep{#3#1\ #2\@thmcounterend}]#4}
-
-\def\@spopargbegintheorem#1#2#3#4#5{\trivlist
- \item[\hskip\labelsep{#4#1\ #2}]{#4(#3)\@thmcounterend\ }#5}
-
-% definition of \spnewtheorem* without number
-
-\def\@sthm#1#2{\@Ynthm{#1}{#2}}
-
-\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}}
-
-\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces}
-
-\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1}
- {#4}{#2}{#3}\ignorespaces}
-
-\def\@Begintheorem#1#2#3{#3\trivlist
- \item[\hskip\labelsep{#2#1\@thmcounterend}]}
-
-\def\@Opargbegintheorem#1#2#3#4{#4\trivlist
- \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }}
-
-\if@envcntsect
- \def\@thmcountersep{.}
- \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape}
-\else
- \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape}
- \if@envcntreset
- \@addtoreset{theorem}{section}
- \else
- \@addtoreset{theorem}{chapter}
- \fi
-\fi
-
-%definition of divers theorem environments
-\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily}
-\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily}
-\if@envcntsame % alle Umgebungen wie Theorem.
- \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}}
-\else % alle Umgebungen mit eigenem Zaehler
- \if@envcntsect % mit section numeriert
- \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}}
- \else % nicht mit section numeriert
- \if@envcntreset
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{section}}
- \else
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{chapter}}%
- \fi
- \fi
-\fi
-\spn@wtheorem{case}{Case}{\itshape}{\rmfamily}
-\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily}
-\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape}
-\spn@wtheorem{definition}{Definition}{\bfseries}{\itshape}
-\spn@wtheorem{example}{Example}{\itshape}{\rmfamily}
-\spn@wtheorem{exercise}{Exercise}{\itshape}{\rmfamily}
-\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape}
-\spn@wtheorem{note}{Note}{\itshape}{\rmfamily}
-\spn@wtheorem{problem}{Problem}{\itshape}{\rmfamily}
-\spn@wtheorem{property}{Property}{\itshape}{\rmfamily}
-\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape}
-\spn@wtheorem{question}{Question}{\itshape}{\rmfamily}
-\spn@wtheorem{solution}{Solution}{\itshape}{\rmfamily}
-\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily}
-
-\def\@takefromreset#1#2{%
- \def\@tempa{#1}%
- \let\@tempd\@elt
- \def\@elt##1{%
- \def\@tempb{##1}%
- \ifx\@tempa\@tempb\else
- \@addtoreset{##1}{#2}%
- \fi}%
- \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
- \expandafter\def\csname cl@#2\endcsname{}%
- \@tempc
- \let\@elt\@tempd}
-
-\def\theopargself{\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5}
- \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
- \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }}
- }
-
-\renewenvironment{abstract}{%
- \list{}{\advance\topsep by0.35cm\relax\small
- \leftmargin=1cm
- \labelwidth=\z@
- \listparindent=\z@
- \itemindent\listparindent
- \rightmargin\leftmargin}\item[\hskip\labelsep
- \bfseries\abstractname]}
- {\endlist}
-
-\newdimen\headlineindent % dimension for space between
-\headlineindent=1.166cm % number and text of headings.
-
-\def\ps@headings{\let\@mkboth\@gobbletwo
- \let\@oddfoot\@empty\let\@evenfoot\@empty
- \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
- \leftmark\hfil}
- \def\@oddhead{\normalfont\small\hfil\rightmark\hspace{\headlineindent}%
- \llap{\thepage}}
- \def\chaptermark##1{}%
- \def\sectionmark##1{}%
- \def\subsectionmark##1{}}
-
-\def\ps@titlepage{\let\@mkboth\@gobbletwo
- \let\@oddfoot\@empty\let\@evenfoot\@empty
- \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
- \hfil}
- \def\@oddhead{\normalfont\small\hfil\hspace{\headlineindent}%
- \llap{\thepage}}
- \def\chaptermark##1{}%
- \def\sectionmark##1{}%
- \def\subsectionmark##1{}}
-
-\if@runhead\ps@headings\else
-\ps@empty\fi
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\endinput
-%end of file llncs.cls
--- a/Pearl-jv/document/root.bib Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,171 +0,0 @@
-@Article{ Urban08,
- author = "C. Urban",
- title = "{N}ominal {T}echniques in {I}sabelle/{HOL}",
- journal = "Journal of Automated Reasoning",
- volume = "40",
- number = "4",
- pages = "327--356",
- year = "2008"
-}
-
-@InProceedings{norrish04,
- author = {M.~Norrish},
- title = {{R}ecursive {F}unction {D}efinition for {T}ypes with {B}inders},
- booktitle = {Proc.~of the 17th International Conference Theorem Proving in
- Higher Order Logics (TPHOLs)},
- pages = {241--256},
- year = {2004},
- volume = {3223},
- series = {LNCS}
-}
-
-@InProceedings{GunterOsbornPopescu09,
- author = {E.L.~Gunter and C.J.~Osborn and A.~Popescu},
- title = {{T}heory {S}upport for {W}eak {H}igher {O}rder {A}bstract {S}yntax in
- {I}sabelle/{HOL}},
- booktitle = {Proc.~of the 4th International Workshop on Logical Frameworks and Meta-Languages:
- Theory and Practice (LFMTP)},
- pages = {12--20},
- year = {2009},
- series = {ENTCS}
-}
-
-@InProceedings{AydemirBohannonWeirich07,
- author = {B.~Aydemir and A.~Bohannon and S.~Weihrich},
- title = {{N}ominal {R}easoning {T}echniques in {C}oq ({E}xtended {A}bstract)},
- booktitle = {Proc.~of the 2st International Workshop on Logical Frameworks and Meta-Languages:
- Theory and Practice (LFMTP)},
- pages = {69--77},
- year = {2007},
- series = {ENTCS}
-}
-
-@Unpublished{SatoPollack10,
- author = {M.~Sato and R.~Pollack},
- title = {{E}xternal and {I}nternal {S}yntax of the {L}ambda-{C}alculus},
- note = {To appear in {\it Journal of Symbolic Computation}}
-}
-
-@article{GabbayPitts02,
- author = {M.~J.~Gabbay and A.~M.~Pitts},
- title = {{A} {N}ew {A}pproach to {A}bstract {S}yntax with {V}ariable
- {B}inding},
- journal = {Formal Aspects of Computing},
- volume = {13},
- year = 2002,
- pages = {341--363}
-}
-
-@article{Pitts03,
- author = {A.~M.~Pitts},
- title = {{N}ominal {L}ogic, {A} {F}irst {O}rder {T}heory of {N}ames and
- {B}inding},
- journal = {Information and Computation},
- year = {2003},
- volume = {183},
- pages = {165--193}
-}
-
-@InProceedings{BengtsonParrow07,
- author = {J.~Bengtson and J.~Parrow},
- title = {Formalising the pi-{C}alculus using {N}ominal {L}ogic},
- booktitle = {Proc.~of the 10th International Conference on
- Foundations of Software Science and Computation Structures (FOSSACS)},
- year = 2007,
- pages = {63--77},
- series = {LNCS},
- volume = {4423}
-}
-
-@inproceedings{BengtsonParow09,
- author = {J.~Bengtson and J.~Parrow},
- title = {{P}si-{C}alculi in {I}sabelle},
- booktitle = {Proc of the 22nd Conference on Theorem Proving in
- Higher Order Logics (TPHOLs)},
- year = 2009,
- pages = {99--114},
- series = {LNCS},
- volume = {5674}
-}
-
-@inproceedings{TobinHochstadtFelleisen08,
- author = {S.~Tobin-Hochstadt and M.~Felleisen},
- booktitle = {Proc.~of the 35rd Symposium on
- Principles of Programming Languages (POPL)},
- title = {{T}he {D}esign and {I}mplementation of {T}yped {S}cheme},
- publisher = {ACM},
- year = {2008},
- pages = {395--406}
-}
-
-@InProceedings{UrbanCheneyBerghofer08,
- author = "C.~Urban and J.~Cheney and S.~Berghofer",
- title = "{M}echanizing the {M}etatheory of {LF}",
- pages = "45--56",
- year = 2008,
- booktitle = "Proc.~of the 23rd IEEE Symposium on Logic in Computer Science (LICS)"
-}
-
-@InProceedings{UrbanZhu08,
- title = "{R}evisiting {C}ut-{E}limination: {O}ne {D}ifficult {P}roof is {R}eally a {P}roof",
- author = "C.~Urban and B.~Zhu",
- booktitle = "Proc.~of the 9th International Conference on Rewriting Techniques
- and Applications (RTA)",
- year = "2008",
- pages = "409--424",
- series = "LNCS",
- volume = 5117
-}
-
-@Article{UrbanPittsGabbay04,
- title = "{N}ominal {U}nification",
- author = "C.~Urban and A.M.~Pitts and M.J.~Gabbay",
- journal = "Theoretical Computer Science",
- pages = "473--497",
- volume = "323",
- number = "1-3",
- year = "2004"
-}
-
-@Article{Church40,
- author = {A.~Church},
- title = {{A} {F}ormulation of the {S}imple {T}heory of {T}ypes},
- journal = {Journal of Symbolic Logic},
- year = {1940},
- volume = {5},
- number = {2},
- pages = {56--68}
-}
-
-
-@Manual{PittsHOL4,
- title = {{S}yntax and {S}emantics},
- author = {A.~M.~Pitts},
- note = {Part of the documentation for the HOL4 system.}
-}
-
-
-@book{PaulsonBenzmueller,
- year={2009},
- author={Benzm{\"u}ller, Christoph and Paulson, Lawrence C.},
- title={Quantified Multimodal Logics in Simple Type Theory},
- note={{http://arxiv.org/abs/0905.2435}},
- series={{SEKI Report SR--2009--02 (ISSN 1437-4447)}},
- publisher={{SEKI Publications}}
-}
-
-@Article{Cheney06,
- author = {J.~Cheney},
- title = {{C}ompleteness and {H}erbrand {T}heorems for {N}ominal {L}ogic},
- journal = {Journal of Symbolic Logic},
- year = {2006},
- volume = {71},
- number = {1},
- pages = {299--320}
-}
-
-@Manual{Wenzel04,
- title = {{U}sing {A}xiomatic {T}ype {C}lasses in {I}sabelle},
- author = {M.~Wenzel},
- note = {Manual in the Isabelle distribution.}
-}
\ No newline at end of file
--- a/Pearl-jv/document/root.tex Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,68 +0,0 @@
-\documentclass{svjour3}
-\usepackage{times}
-\usepackage{isabelle}
-\usepackage{isabellesym}
-\usepackage{amsmath}
-\usepackage{amssymb}
-\usepackage{mathabx}
-\usepackage{proof}
-\usepackage{longtable}
-\usepackage{graphics}
-\usepackage{pdfsetup}
-
-\urlstyle{rm}
-\isabellestyle{it}
-\renewcommand{\isastyle}{\isastyleminor}
-\renewcommand{\isacharunderscore}{\mbox{$\_\!\_$}}
-\renewcommand{\isasymbullet}{{\raisebox{-0.4mm}{\Large$\boldsymbol{\cdot}$}}}
-\def\dn{\,\stackrel{\mbox{\scriptsize def}}{=}\,}
-\renewcommand{\isasymequiv}{$\dn$}
-\renewcommand{\isasymiota}{}
-\renewcommand{\isasymrightleftharpoons}{}
-\renewcommand{\isasymemptyset}{$\varnothing$}
-\newcommand{\isasymallatoms}{\ensuremath{\mathbb{A}}}
-\newcommand{\rrh}{\mbox{\footnotesize$\rightrightharpoons$}}
-
-\newcommand{\numbered}[1]{\refstepcounter{equation}{\rm(\arabic{equation})}\label{#1}}
-\newcommand\new[0]{\reflectbox{\ensuremath{\mathsf{N}}}}
-
-\changenotsign
-
-\begin{document}
-
-\title{Implementing the Nominal Logic Work in Isabelle/HOL}
-\author{Christian Urban \and Brian Huffman}
-\institute{C.~Urban \at Technical University of Munich
- \and B.~Huffman \at Portland State University}
-\date{Received: date / Accepted: date}
-
-\maketitle
-
-\begin{abstract}
-In his nominal logic work, Pitts introduced a beautiful theory about names and
-binding based on the notions of atoms, permutations and support. The
-engineering challenge is to smoothly adapt this theory to a theorem prover
-environment, in our case Isabelle/HOL. For this we have to formulate the
-theory so that it is compatible with Higher-Order Logic, which the original formulation by
-Pitts is not. We achieve this by not requiring that every construction has
-to have finite support. We present a formalisation that is based on a
-unified atom type and that represents permutations by bijective functions from
-atoms to atoms. Interestingly, we allow swappings, which are permutations
-build from two atoms, to be ill-sorted. We also describe a reasoning infrastructure
-that automates properties about equivariance, and present a formalisation of
-two abstraction operators that bind sets of names.
-\end{abstract}
-
-% generated text of all theories
-\input{session}
-
-% optional bibliography
-\bibliographystyle{abbrv}
-\bibliography{root}
-
-\end{document}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
--- a/Pearl-jv/document/svglov3.clo Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,101 +0,0 @@
-% SVJour3 DOCUMENT CLASS OPTION SVGLOV3 -- for standardised journals
-%
-% This is an enhancement for the LaTeX
-% SVJour3 document class for Springer journals
-%
-%%
-%%
-%% \CharacterTable
-%% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
-%% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
-%% Digits \0\1\2\3\4\5\6\7\8\9
-%% Exclamation \! Double quote \" Hash (number) \#
-%% Dollar \$ Percent \% Ampersand \&
-%% Acute accent \' Left paren \( Right paren \)
-%% Asterisk \* Plus \+ Comma \,
-%% Minus \- Point \. Solidus \/
-%% Colon \: Semicolon \; Less than \<
-%% Equals \= Greater than \> Question mark \?
-%% Commercial at \@ Left bracket \[ Backslash \\
-%% Right bracket \] Circumflex \^ Underscore \_
-%% Grave accent \` Left brace \{ Vertical bar \|
-%% Right brace \} Tilde \~}
-\ProvidesFile{svglov3.clo}
- [2006/02/03 v3.1
- style option for standardised journals]
-\typeout{SVJour Class option: svglov3.clo for standardised journals}
-\def\validfor{svjour3}
-\ExecuteOptions{final,10pt,runningheads}
-% No size changing allowed, hence a "copy" of size10.clo is included
-\renewcommand\normalsize{%
-\if@twocolumn
- \@setfontsize\normalsize\@xpt{12.5pt}%
-\else
- \if@smallext
- \@setfontsize\normalsize\@xpt\@xiipt
- \else
- \@setfontsize\normalsize{9.5pt}{11.5pt}%
- \fi
-\fi
- \abovedisplayskip=3 mm plus6pt minus 4pt
- \belowdisplayskip=3 mm plus6pt minus 4pt
- \abovedisplayshortskip=0.0 mm plus6pt
- \belowdisplayshortskip=2 mm plus4pt minus 4pt
- \let\@listi\@listI}
-\normalsize
-\newcommand\small{%
-\if@twocolumn
- \@setfontsize\small{8.5pt}\@xpt
-\else
- \if@smallext
- \@setfontsize\small\@viiipt{9.5pt}%
- \else
- \@setfontsize\small\@viiipt{9.25pt}%
- \fi
-\fi
- \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
- \abovedisplayshortskip \z@ \@plus2\p@
- \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
- \def\@listi{\leftmargin\leftmargini
- \parsep 0\p@ \@plus1\p@ \@minus\p@
- \topsep 4\p@ \@plus2\p@ \@minus4\p@
- \itemsep0\p@}%
- \belowdisplayskip \abovedisplayskip
-}
-\let\footnotesize\small
-\newcommand\scriptsize{\@setfontsize\scriptsize\@viipt\@viiipt}
-\newcommand\tiny{\@setfontsize\tiny\@vpt\@vipt}
-\if@twocolumn
- \newcommand\large{\@setfontsize\large\@xiipt\@xivpt}
- \newcommand\LARGE{\@setfontsize\LARGE{16pt}{18pt}}
-\else
- \newcommand\large{\@setfontsize\large\@xipt\@xiipt}
- \newcommand\LARGE{\@setfontsize\LARGE{13pt}{15pt}}
-\fi
-\newcommand\Large{\@setfontsize\Large\@xivpt{16dd}}
-\newcommand\huge{\@setfontsize\huge\@xxpt{25}}
-\newcommand\Huge{\@setfontsize\Huge\@xxvpt{30}}
-%
-\def\runheadhook{\rlap{\smash{\lower6.5pt\hbox to\textwidth{\hrulefill}}}}
-\if@twocolumn
-\setlength{\textwidth}{17.4cm}
-\setlength{\textheight}{234mm}
-\AtEndOfClass{\setlength\columnsep{6mm}}
-\else
- \if@smallext
- \setlength{\textwidth}{11.9cm}
- \setlength{\textheight}{19.4cm}
- \else
- \setlength{\textwidth}{12.2cm}
- \setlength{\textheight}{19.8cm}
- \fi
-\fi
-%
-\AtBeginDocument{%
-\@ifundefined{@journalname}
- {\typeout{Unknown journal: specify \string\journalname\string{%
-<name of your journal>\string} in preambel^^J}}{}}
-%
-\endinput
-%%
-%% End of file `svglov3.clo'.
--- a/Pearl-jv/document/svjour3.cls Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1431 +0,0 @@
-% SVJour3 DOCUMENT CLASS -- version 3.2 for LaTeX2e
-%
-% LaTeX document class for Springer journals
-%
-%%
-%%
-%% \CharacterTable
-%% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
-%% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
-%% Digits \0\1\2\3\4\5\6\7\8\9
-%% Exclamation \! Double quote \" Hash (number) \#
-%% Dollar \$ Percent \% Ampersand \&
-%% Acute accent \' Left paren \( Right paren \)
-%% Asterisk \* Plus \+ Comma \,
-%% Minus \- Point \. Solidus \/
-%% Colon \: Semicolon \; Less than \<
-%% Equals \= Greater than \> Question mark \?
-%% Commercial at \@ Left bracket \[ Backslash \\
-%% Right bracket \] Circumflex \^ Underscore \_
-%% Grave accent \` Left brace \{ Vertical bar \|
-%% Right brace \} Tilde \~}
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-\ProvidesClass{svjour3}[2007/05/08 v3.2
-^^JLaTeX document class for Springer journals]
-\newcommand\@ptsize{}
-\newif\if@restonecol
-\newif\if@titlepage
-\@titlepagefalse
-\DeclareOption{a4paper}
- {\setlength\paperheight {297mm}%
- \setlength\paperwidth {210mm}}
-\DeclareOption{10pt}{\renewcommand\@ptsize{0}}
-\DeclareOption{twoside}{\@twosidetrue \@mparswitchtrue}
-\DeclareOption{draft}{\setlength\overfullrule{5pt}}
-\DeclareOption{final}{\setlength\overfullrule{0pt}}
-\DeclareOption{fleqn}{\input{fleqn.clo}\AtBeginDocument{\mathindent\z@}%
-\AtBeginDocument{\@ifpackageloaded{amsmath}{\@mathmargin\z@}{}}%
-\PassOptionsToPackage{fleqn}{amsmath}}
-%%%
-\DeclareOption{onecolumn}{}
-\DeclareOption{smallcondensed}{}
-\DeclareOption{twocolumn}{\@twocolumntrue\ExecuteOptions{fleqn}}
-\newif\if@smallext\@smallextfalse
-\DeclareOption{smallextended}{\@smallexttrue}
-\let\if@mathematic\iftrue
-\let\if@numbook\iffalse
-\DeclareOption{numbook}{\let\if@envcntsect\iftrue
- \AtEndOfPackage{%
- \renewcommand\thefigure{\thesection.\@arabic\c@figure}%
- \renewcommand\thetable{\thesection.\@arabic\c@table}%
- \renewcommand\theequation{\thesection.\@arabic\c@equation}%
- \@addtoreset{figure}{section}%
- \@addtoreset{table}{section}%
- \@addtoreset{equation}{section}%
- }%
-}
-\DeclareOption{openbib}{%
- \AtEndOfPackage{%
- \renewcommand\@openbib@code{%
- \advance\leftmargin\bibindent
- \itemindent -\bibindent
- \listparindent \itemindent
- \parsep \z@
- }%
- \renewcommand\newblock{\par}}%
-}
-\DeclareOption{natbib}{%
-\AtEndOfClass{\RequirePackage{natbib}%
-% Changing some parameters of NATBIB
-\setlength{\bibhang}{\parindent}%
-%\setlength{\bibsep}{0mm}%
-\let\bibfont=\small
-\def\@biblabel#1{#1.}%
-\newcommand{\etal}{et al.}%
-\bibpunct{(}{)}{;}{a}{}{,}}}
-%
-\let\if@runhead\iffalse
-\DeclareOption{runningheads}{\let\if@runhead\iftrue}
-\let\if@smartrunh\iffalse
-\DeclareOption{smartrunhead}{\let\if@smartrunh\iftrue}
-\DeclareOption{nosmartrunhead}{\let\if@smartrunh\iffalse}
-\let\if@envcntreset\iffalse
-\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue}
-\let\if@envcntsame\iffalse
-\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue}
-\let\if@envcntsect\iffalse
-\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue}
-\let\if@referee\iffalse
-\DeclareOption{referee}{\let\if@referee\iftrue}
-\def\makereferee{\def\baselinestretch{2}}
-\let\if@instindent\iffalse
-\DeclareOption{instindent}{\let\if@instindent\iftrue}
-\let\if@smartand\iffalse
-\DeclareOption{smartand}{\let\if@smartand\iftrue}
-\let\if@spthms\iftrue
-\DeclareOption{nospthms}{\let\if@spthms\iffalse}
-%
-% language and babel dependencies
-\DeclareOption{deutsch}{\def\switcht@@therlang{\switcht@deutsch}%
-\gdef\svlanginfo{\typeout{Man spricht deutsch.}\global\let\svlanginfo\relax}}
-\DeclareOption{francais}{\def\switcht@@therlang{\switcht@francais}%
-\gdef\svlanginfo{\typeout{On parle francais.}\global\let\svlanginfo\relax}}
-\let\switcht@@therlang\relax
-\let\svlanginfo\relax
-%
-\AtBeginDocument{\@ifpackageloaded{babel}{%
-\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}%
-\@ifundefined{extrasUKenglish}{}{\addto\extrasUKenglish{\switcht@albion}}%
-\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}%
-\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}%
-\@ifundefined{extrasngerman}{}{\addto\extrasngerman{\switcht@deutsch}}%
-}{\switcht@@therlang}%
-}
-%
-\def\ClassInfoNoLine#1#2{%
- \ClassInfo{#1}{#2\@gobble}%
-}
-\let\journalopt\@empty
-\DeclareOption*{%
-\InputIfFileExists{sv\CurrentOption.clo}{%
-\global\let\journalopt\CurrentOption}{%
-\ClassWarning{Springer-SVJour3}{Specified option or subpackage
-"\CurrentOption" not found -}\OptionNotUsed}}
-\ExecuteOptions{a4paper,twoside,10pt,instindent}
-\ProcessOptions
-%
-\ifx\journalopt\@empty\relax
-\ClassInfoNoLine{Springer-SVJour3}{extra/valid Springer sub-package (-> *.clo)
-\MessageBreak not found in option list of \string\documentclass
-\MessageBreak - autoactivating "global" style}{}
-\input{svglov3.clo}
-\else
-\@ifundefined{validfor}{%
-\ClassError{Springer-SVJour3}{Possible option clash for sub-package
-\MessageBreak "sv\journalopt.clo" - option file not valid
-\MessageBreak for this class}{Perhaps you used an option of the old
-Springer class SVJour!}
-}{}
-\fi
-%
-\if@smartrunh\AtEndDocument{\islastpageeven\getlastpagenumber}\fi
-%
-\newcommand{\twocoltest}[2]{\if@twocolumn\def\@gtempa{#2}\else\def\@gtempa{#1}\fi
-\@gtempa\makeatother}
-\newcommand{\columncase}{\makeatletter\twocoltest}
-%
-\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00}
-\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01}
-\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02}
-\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03}
-\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04}
-\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05}
-\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06}
-\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07}
-\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08}
-\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09}
-\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A}
-%
-\setlength\parindent{15\p@}
-\setlength\smallskipamount{3\p@ \@plus 1\p@ \@minus 1\p@}
-\setlength\medskipamount{6\p@ \@plus 2\p@ \@minus 2\p@}
-\setlength\bigskipamount{12\p@ \@plus 4\p@ \@minus 4\p@}
-\setlength\headheight{12\p@}
-\setlength\headsep {14.50dd}
-\setlength\topskip {10\p@}
-\setlength\footskip{30\p@}
-\setlength\maxdepth{.5\topskip}
-%
-\@settopoint\textwidth
-\setlength\marginparsep {10\p@}
-\setlength\marginparpush{5\p@}
-\setlength\topmargin{-10pt}
-\if@twocolumn
- \setlength\oddsidemargin {-30\p@}
- \setlength\evensidemargin{-30\p@}
-\else
- \setlength\oddsidemargin {\z@}
- \setlength\evensidemargin{\z@}
-\fi
-\setlength\marginparwidth {48\p@}
-\setlength\footnotesep{8\p@}
-\setlength{\skip\footins}{9\p@ \@plus 4\p@ \@minus 2\p@}
-\setlength\floatsep {12\p@ \@plus 2\p@ \@minus 2\p@}
-\setlength\textfloatsep{20\p@ \@plus 2\p@ \@minus 4\p@}
-\setlength\intextsep {20\p@ \@plus 2\p@ \@minus 2\p@}
-\setlength\dblfloatsep {12\p@ \@plus 2\p@ \@minus 2\p@}
-\setlength\dbltextfloatsep{20\p@ \@plus 2\p@ \@minus 4\p@}
-\setlength\@fptop{0\p@}
-\setlength\@fpsep{12\p@ \@plus 2\p@ \@minus 2\p@}
-\setlength\@fpbot{0\p@ \@plus 1fil}
-\setlength\@dblfptop{0\p@}
-\setlength\@dblfpsep{12\p@ \@plus 2\p@ \@minus 2\p@}
-\setlength\@dblfpbot{0\p@ \@plus 1fil}
-\setlength\partopsep{2\p@ \@plus 1\p@ \@minus 1\p@}
-\def\@listi{\leftmargin\leftmargini
- \parsep \z@
- \topsep 6\p@ \@plus2\p@ \@minus4\p@
- \itemsep\parsep}
-\let\@listI\@listi
-\@listi
-\def\@listii {\leftmargin\leftmarginii
- \labelwidth\leftmarginii
- \advance\labelwidth-\labelsep
- \topsep \z@
- \parsep \topsep
- \itemsep \parsep}
-\def\@listiii{\leftmargin\leftmarginiii
- \labelwidth\leftmarginiii
- \advance\labelwidth-\labelsep
- \topsep \z@
- \parsep \topsep
- \itemsep \parsep}
-\def\@listiv {\leftmargin\leftmarginiv
- \labelwidth\leftmarginiv
- \advance\labelwidth-\labelsep}
-\def\@listv {\leftmargin\leftmarginv
- \labelwidth\leftmarginv
- \advance\labelwidth-\labelsep}
-\def\@listvi {\leftmargin\leftmarginvi
- \labelwidth\leftmarginvi
- \advance\labelwidth-\labelsep}
-%
-\setlength\lineskip{1\p@}
-\setlength\normallineskip{1\p@}
-\renewcommand\baselinestretch{}
-\setlength\parskip{0\p@ \@plus \p@}
-\@lowpenalty 51
-\@medpenalty 151
-\@highpenalty 301
-\setcounter{topnumber}{4}
-\renewcommand\topfraction{.9}
-\setcounter{bottomnumber}{2}
-\renewcommand\bottomfraction{.7}
-\setcounter{totalnumber}{6}
-\renewcommand\textfraction{.1}
-\renewcommand\floatpagefraction{.85}
-\setcounter{dbltopnumber}{3}
-\renewcommand\dbltopfraction{.85}
-\renewcommand\dblfloatpagefraction{.85}
-\def\ps@headings{%
- \let\@oddfoot\@empty\let\@evenfoot\@empty
- \def\@evenhead{\small\csname runheadhook\endcsname
- \rlap{\thepage}\hfil\leftmark\unskip}%
- \def\@oddhead{\small\csname runheadhook\endcsname
- \ignorespaces\rightmark\hfil\llap{\thepage}}%
- \let\@mkboth\@gobbletwo
- \let\sectionmark\@gobble
- \let\subsectionmark\@gobble
- }
-% make indentations changeable
-\def\setitemindent#1{\settowidth{\labelwidth}{#1}%
- \leftmargini\labelwidth
- \advance\leftmargini\labelsep
- \def\@listi{\leftmargin\leftmargini
- \labelwidth\leftmargini\advance\labelwidth by -\labelsep
- \parsep=\parskip
- \topsep=\medskipamount
- \itemsep=\parskip \advance\itemsep by -\parsep}}
-\def\setitemitemindent#1{\settowidth{\labelwidth}{#1}%
- \leftmarginii\labelwidth
- \advance\leftmarginii\labelsep
-\def\@listii{\leftmargin\leftmarginii
- \labelwidth\leftmarginii\advance\labelwidth by -\labelsep
- \parsep=\parskip
- \topsep=\z@
- \itemsep=\parskip \advance\itemsep by -\parsep}}
-% labels of description
-\def\descriptionlabel#1{\hspace\labelsep #1\hfil}
-% adjusted environment "description"
-% if an optional parameter (at the first two levels of lists)
-% is present, its width is considered to be the widest mark
-% throughout the current list.
-\def\description{\@ifnextchar[{\@describe}{\list{}{\labelwidth\z@
- \itemindent-\leftmargin \let\makelabel\descriptionlabel}}}
-\let\enddescription\endlist
-%
-\def\describelabel#1{#1\hfil}
-\def\@describe[#1]{\relax\ifnum\@listdepth=0
-\setitemindent{#1}\else\ifnum\@listdepth=1
-\setitemitemindent{#1}\fi\fi
-\list{--}{\let\makelabel\describelabel}}
-%
-\newdimen\logodepth
-\logodepth=1.2cm
-\newdimen\headerboxheight
-\headerboxheight=163pt % 18 10.5dd-lines - 2\baselineskip
-\if@twocolumn\else\advance\headerboxheight by-14.5mm\fi
-\newdimen\betweenumberspace % dimension for space between
-\betweenumberspace=3.33pt % number and text of titles.
-\newdimen\aftertext % dimension for space after
-\aftertext=5pt % text of title.
-\newdimen\headlineindent % dimension for space between
-\headlineindent=1.166cm % number and text of headings.
-\if@mathematic
- \def\runinend{} % \enspace}
- \def\floatcounterend{\enspace}
- \def\sectcounterend{}
-\else
- \def\runinend{.}
- \def\floatcounterend{.\ }
- \def\sectcounterend{.}
-\fi
-\def\email#1{\emailname: #1}
-\def\keywords#1{\par\addvspace\medskipamount{\rightskip=0pt plus1cm
-\def\and{\ifhmode\unskip\nobreak\fi\ $\cdot$
-}\noindent\keywordname\enspace\ignorespaces#1\par}}
-%
-\def\subclassname{{\bfseries Mathematics Subject Classification
-(2000)}\enspace}
-\def\subclass#1{\par\addvspace\medskipamount{\rightskip=0pt plus1cm
-\def\and{\ifhmode\unskip\nobreak\fi\ $\cdot$
-}\noindent\subclassname\ignorespaces#1\par}}
-%
-\def\PACSname{\textbf{PACS}\enspace}
-\def\PACS#1{\par\addvspace\medskipamount{\rightskip=0pt plus1cm
-\def\and{\ifhmode\unskip\nobreak\fi\ $\cdot$
-}\noindent\PACSname\ignorespaces#1\par}}
-%
-\def\CRclassname{{\bfseries CR Subject Classification}\enspace}
-\def\CRclass#1{\par\addvspace\medskipamount{\rightskip=0pt plus1cm
-\def\and{\ifhmode\unskip\nobreak\fi\ $\cdot$
-}\noindent\CRclassname\ignorespaces#1\par}}
-%
-\def\ESMname{\textbf{Electronic Supplementary Material}\enspace}
-\def\ESM#1{\par\addvspace\medskipamount
-\noindent\ESMname\ignorespaces#1\par}
-%
-\newcounter{inst}
-\newcounter{auth}
-\def\authdepth{2}
-\newdimen\instindent
-\newbox\authrun
-\newtoks\authorrunning
-\newbox\titrun
-\newtoks\titlerunning
-\def\authorfont{\bfseries}
-
-\def\combirunning#1{\gdef\@combi{#1}}
-\def\@combi{}
-\newbox\combirun
-%
-\def\ps@last{\def\@evenhead{\small\rlap{\thepage}\hfil
-\lastevenhead}}
-\newcounter{lastpage}
-\def\islastpageeven{\@ifundefined{lastpagenumber}
-{\setcounter{lastpage}{0}}{\setcounter{lastpage}{\lastpagenumber}}
-\ifnum\value{lastpage}>0
- \ifodd\value{lastpage}%
- \else
- \if@smartrunh
- \thispagestyle{last}%
- \fi
- \fi
-\fi}
-\def\getlastpagenumber{\clearpage
-\addtocounter{page}{-1}%
- \immediate\write\@auxout{\string\gdef\string\lastpagenumber{\thepage}}%
- \immediate\write\@auxout{\string\newlabel{LastPage}{{}{\thepage}}}%
- \addtocounter{page}{1}}
-
-\def\journalname#1{\gdef\@journalname{#1}}
-
-\def\dedication#1{\gdef\@dedic{#1}}
-\def\@dedic{}
-
-\let\@date\undefined
-\def\notused{~}
-
-\def\institute#1{\gdef\@institute{#1}}
-
-\def\offprints#1{\begingroup
-\def\protect{\noexpand\protect\noexpand}\xdef\@thanks{\@thanks
-\protect\footnotetext[0]{\unskip\hskip-15pt{\itshape Send offprint requests
-to\/}: \ignorespaces#1}}\endgroup\ignorespaces}
-
-%\def\mail#1{\gdef\@mail{#1}}
-%\def\@mail{}
-
-\def\@thanks{}
-
-\def\@fnsymbol#1{\ifcase#1\or\star\or{\star\star}\or{\star\star\star}%
- \or \dagger\or \ddagger\or
- \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger
- \or \ddagger\ddagger \else\@ctrerr\fi\relax}
-%
-%\def\invthanks#1{\footnotetext[0]{\kern-\bibindent#1}}
-%
-\def\nothanksmarks{\def\thanks##1{\protected@xdef\@thanks{\@thanks
- \protect\footnotetext[0]{\kern-\bibindent##1}}}}
-%
-\def\subtitle#1{\gdef\@subtitle{#1}}
-\def\@subtitle{}
-
-\def\headnote#1{\gdef\@headnote{#1}}
-\def\@headnote{}
-
-\def\papertype#1{\gdef\paper@type{\MakeUppercase{#1}}}
-\def\paper@type{}
-
-\def\ch@ckobl#1#2{\@ifundefined{@#1}
- {\typeout{SVJour3 warning: Missing
-\expandafter\string\csname#1\endcsname}%
- \csname #1\endcsname{#2}}
- {}}
-%
-\def\ProcessRunnHead{%
- \def\\{\unskip\ \ignorespaces}%
- \def\thanks##1{\unskip{}}%
- \instindent=\textwidth
- \advance\instindent by-\headlineindent
- \if!\the\titlerunning!\else
- \edef\@title{\the\titlerunning}%
- \fi
- \global\setbox\titrun=\hbox{\small\rmfamily\unboldmath\ignorespaces\@title
- \unskip}%
- \ifdim\wd\titrun>\instindent
- \typeout{^^JSVJour3 Warning: Title too long for running head.}%
- \typeout{Please supply a shorter form with \string\titlerunning
- \space prior to \string\maketitle}%
- \global\setbox\titrun=\hbox{\small\rmfamily
- Title Suppressed Due to Excessive Length}%
- \fi
- \xdef\@title{\copy\titrun}%
-%
- \if!\the\authorrunning!
- \else
- \setcounter{auth}{1}%
- \edef\@author{\the\authorrunning}%
- \fi
- \ifnum\value{inst}>\authdepth
- \def\stripauthor##1\and##2\endauthor{%
- \protected@xdef\@author{##1\unskip\unskip\if!##2!\else\ et al.\fi}}%
- \expandafter\stripauthor\@author\and\endauthor
- \else
- \gdef\and{\unskip, \ignorespaces}%
- {\def\and{\noexpand\protect\noexpand\and}%
- \protected@xdef\@author{\@author}}
- \fi
- \global\setbox\authrun=\hbox{\small\rmfamily\unboldmath\ignorespaces
- \@author\unskip}%
- \ifdim\wd\authrun>\instindent
- \typeout{^^JSVJour3 Warning: Author name(s) too long for running head.
- ^^JPlease supply a shorter form with \string\authorrunning
- \space prior to \string\maketitle}%
- \global\setbox\authrun=\hbox{\small\rmfamily Please give a shorter version
- with: {\tt\string\authorrunning\space and
- \string\titlerunning\space prior to \string\maketitle}}%
- \fi
- \xdef\@author{\copy\authrun}%
- \markboth{\@author}{\@title}%
-}
-%
-\let\orithanks=\thanks
-\def\thanks#1{\ClassWarning{SVJour3}{\string\thanks\space may only be
-used inside of \string\title, \string\author,\MessageBreak
-and \string\date\space prior to \string\maketitle}}
-%
-\def\maketitle{\par\let\thanks=\orithanks
-\ch@ckobl{journalname}{Noname}
-\ch@ckobl{date}{the date of receipt and acceptance should be inserted
-later}
-\ch@ckobl{title}{A title should be given}
-\ch@ckobl{author}{Name(s) and initial(s) of author(s) should be given}
-\ch@ckobl{institute}{Address(es) of author(s) should be given}
-\begingroup
-%
- \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
- \def\@makefnmark{$^{\@thefnmark}$}%
- \renewcommand\@makefntext[1]{%
- \noindent
- \hb@xt@\bibindent{\hss\@makefnmark\enspace}##1\vrule height0pt
- width0pt depth8pt}
-%
- \def\lastand{\ifnum\value{inst}=2\relax
- \unskip{} \andname\
- \else
- \unskip, \andname\
- \fi}%
- \def\and{\stepcounter{auth}\relax
- \if@smartand
- \ifnum\value{auth}=\value{inst}%
- \lastand
- \else
- \unskip,
- \fi
- \else
- \unskip,
- \fi}%
- \thispagestyle{empty}
- \ifnum \col@number=\@ne
- \@maketitle
- \else
- \twocolumn[\@maketitle]%
- \fi
-%
- \global\@topnum\z@
- \if!\@thanks!\else
- \@thanks
-\insert\footins{\vskip-3pt\hrule\@width\if@twocolumn\columnwidth
-\else 38mm\fi\vskip3pt}%
- \fi
- {\def\thanks##1{\unskip{}}%
- \def\iand{\\[5pt]\let\and=\nand}%
- \def\nand{\ifhmode\unskip\nobreak\fi\ $\cdot$ }%
- \let\and=\nand
- \def\at{\\\let\and=\iand}%
- \footnotetext[0]{\kern-\bibindent
- \ignorespaces\@institute}\vspace{5dd}}%
-%\if!\@mail!\else
-% \footnotetext[0]{\kern-\bibindent\mailname\
-% \ignorespaces\@mail}%
-%\fi
-%
- \if@runhead
- \ProcessRunnHead
- \fi
-%
- \endgroup
- \setcounter{footnote}{0}
- \global\let\thanks\relax
- \global\let\maketitle\relax
- \global\let\@maketitle\relax
- \global\let\@thanks\@empty
- \global\let\@author\@empty
- \global\let\@date\@empty
- \global\let\@title\@empty
- \global\let\@subtitle\@empty
- \global\let\title\relax
- \global\let\author\relax
- \global\let\date\relax
- \global\let\and\relax}
-
-\def\makeheadbox{{%
-\hbox to0pt{\vbox{\baselineskip=10dd\hrule\hbox
-to\hsize{\vrule\kern3pt\vbox{\kern3pt
-\hbox{\bfseries\@journalname\ manuscript No.}
-\hbox{(will be inserted by the editor)}
-\kern3pt}\hfil\kern3pt\vrule}\hrule}%
-\hss}}}
-%
-\def\rubric{\setbox0=\hbox{\small\strut}\@tempdima=\ht0\advance
-\@tempdima\dp0\advance\@tempdima2\fboxsep\vrule\@height\@tempdima
-\@width\z@}
-\newdimen\rubricwidth
-%
-\def\@maketitle{\newpage
-\normalfont
-\vbox to0pt{\if@twocolumn\vskip-39pt\else\vskip-49pt\fi
-\nointerlineskip
-\makeheadbox\vss}\nointerlineskip
-\vbox to 0pt{\offinterlineskip\rubricwidth=\columnwidth
-%%%%\vskip-12.5pt % -12.5pt
-\if@twocolumn\else % one column journal
- \divide\rubricwidth by144\multiply\rubricwidth by89 % perform golden section
- \vskip-\topskip
-\fi
-\hrule\@height0.35mm\noindent
-\advance\fboxsep by.25mm
-\global\advance\rubricwidth by0pt
-\rubric
-\vss}\vskip19.5pt % war 9pt
-%
-\if@twocolumn\else
- \gdef\footnoterule{%
- \kern-3\p@
- \hrule\@width38mm % \columnwidth \rubricwidth
- \kern2.6\p@}
-\fi
-%
- \setbox\authrun=\vbox\bgroup
- \if@twocolumn
- \hrule\@height10.5mm\@width0\p@
- \else
- \hrule\@height 2mm\@width0\p@
- \fi
- \pretolerance=10000
- \rightskip=0pt plus 4cm
- \nothanksmarks
-% \if!\@headnote!\else
-% \noindent
-% {\LARGE\normalfont\itshape\ignorespaces\@headnote\par}\vskip 3.5mm
-% \fi
- {\LARGE\bfseries
- \noindent\ignorespaces
- \@title \par}\vskip 17pt\relax
- \if!\@subtitle!\else
- {\large\bfseries
- \pretolerance=10000
- \rightskip=0pt plus 3cm
- \vskip-12pt
-% \noindent\ignorespaces\@subtitle \par}\vskip 11.24pt\relax
- \noindent\ignorespaces\@subtitle \par}\vskip 17pt\relax
- \fi
- {\authorfont
- \setbox0=\vbox{\setcounter{auth}{1}\def\and{\stepcounter{auth} }%
- \hfuzz=2\textwidth\def\thanks##1{}\@author}%
- \setcounter{footnote}{0}%
- \global\value{inst}=\value{auth}%
- \setcounter{auth}{1}%
- \if@twocolumn
- \rightskip43mm plus 4cm minus 3mm
- \else % one column journal
- \rightskip=\linewidth
- \advance\rightskip by-\rubricwidth
- \advance\rightskip by0pt plus 4cm minus 3mm
- \fi
-%
-\def\and{\unskip\nobreak\enskip{\boldmath$\cdot$}\enskip\ignorespaces}%
- \noindent\ignorespaces\@author\vskip7.23pt}
-%
- \small
- \if!\@dedic!\else
- \par
- \normalsize\it
- \addvspace\baselineskip
- \noindent\@dedic
- \fi
- \egroup % end of header box
- \@tempdima=\headerboxheight
- \advance\@tempdima by-\ht\authrun
- \unvbox\authrun
- \ifdim\@tempdima>0pt
- \vrule width0pt height\@tempdima\par
- \fi
- \noindent{\small\@date\if@twocolumn\vskip 7.2mm\else\vskip 5.2mm\fi}
- \global\@minipagetrue
- \global\everypar{\global\@minipagefalse\global\everypar{}}%
-%\vskip22.47pt
-}
-%
-\if@mathematic
- \def\vec#1{\ensuremath{\mathchoice
- {\mbox{\boldmath$\displaystyle\mathbf{#1}$}}
- {\mbox{\boldmath$\textstyle\mathbf{#1}$}}
- {\mbox{\boldmath$\scriptstyle\mathbf{#1}$}}
- {\mbox{\boldmath$\scriptscriptstyle\mathbf{#1}$}}}}
-\else
- \def\vec#1{\ensuremath{\mathchoice
- {\mbox{\boldmath$\displaystyle#1$}}
- {\mbox{\boldmath$\textstyle#1$}}
- {\mbox{\boldmath$\scriptstyle#1$}}
- {\mbox{\boldmath$\scriptscriptstyle#1$}}}}
-\fi
-%
-\def\tens#1{\ensuremath{\mathsf{#1}}}
-%
-\setcounter{secnumdepth}{3}
-\newcounter {section}
-\newcounter {subsection}[section]
-\newcounter {subsubsection}[subsection]
-\newcounter {paragraph}[subsubsection]
-\newcounter {subparagraph}[paragraph]
-\renewcommand\thesection {\@arabic\c@section}
-\renewcommand\thesubsection {\thesection.\@arabic\c@subsection}
-\renewcommand\thesubsubsection{\thesubsection.\@arabic\c@subsubsection}
-\renewcommand\theparagraph {\thesubsubsection.\@arabic\c@paragraph}
-\renewcommand\thesubparagraph {\theparagraph.\@arabic\c@subparagraph}
-%
-\def\@hangfrom#1{\setbox\@tempboxa\hbox{#1}%
- \hangindent \z@\noindent\box\@tempboxa}
-%
-\def\@seccntformat#1{\csname the#1\endcsname\sectcounterend
-\hskip\betweenumberspace}
-%
-% \newif\if@sectrule
-% \if@twocolumn\else\let\@sectruletrue=\relax\fi
-% \if@avier\let\@sectruletrue=\relax\fi
-% \def\makesectrule{\if@sectrule\global\@sectrulefalse\null\vglue-\topskip
-% \hrule\nobreak\parskip=5pt\relax\fi}
-% %
-% \let\makesectruleori=\makesectrule
-% \def\restoresectrule{\global\let\makesectrule=\makesectruleori\global\@sectrulefalse}
-% \def\nosectrule{\let\makesectrule=\restoresectrule}
-%
-\def\@startsection#1#2#3#4#5#6{%
- \if@noskipsec \leavevmode \fi
- \par
- \@tempskipa #4\relax
- \@afterindenttrue
- \ifdim \@tempskipa <\z@
- \@tempskipa -\@tempskipa \@afterindentfalse
- \fi
- \if@nobreak
- \everypar{}%
- \else
- \addpenalty\@secpenalty\addvspace\@tempskipa
- \fi
-% \ifnum#2=1\relax\@sectruletrue\fi
- \@ifstar
- {\@ssect{#3}{#4}{#5}{#6}}%
- {\@dblarg{\@sect{#1}{#2}{#3}{#4}{#5}{#6}}}}
-%
-\def\@sect#1#2#3#4#5#6[#7]#8{%
- \ifnum #2>\c@secnumdepth
- \let\@svsec\@empty
- \else
- \refstepcounter{#1}%
- \protected@edef\@svsec{\@seccntformat{#1}\relax}%
- \fi
- \@tempskipa #5\relax
- \ifdim \@tempskipa>\z@
- \begingroup
- #6{% \makesectrule
- \@hangfrom{\hskip #3\relax\@svsec}%
- \raggedright
- \hyphenpenalty \@M%
- \interlinepenalty \@M #8\@@par}%
- \endgroup
- \csname #1mark\endcsname{#7}%
- \addcontentsline{toc}{#1}{%
- \ifnum #2>\c@secnumdepth \else
- \protect\numberline{\csname the#1\endcsname\sectcounterend}%
- \fi
- #7}%
- \else
- \def\@svsechd{%
- #6{\hskip #3\relax
- \@svsec #8\/\hskip\aftertext}%
- \csname #1mark\endcsname{#7}%
- \addcontentsline{toc}{#1}{%
- \ifnum #2>\c@secnumdepth \else
- \protect\numberline{\csname the#1\endcsname}%
- \fi
- #7}}%
- \fi
- \@xsect{#5}}
-%
-\def\@ssect#1#2#3#4#5{%
- \@tempskipa #3\relax
- \ifdim \@tempskipa>\z@
- \begingroup
- #4{% \makesectrule
- \@hangfrom{\hskip #1}%
- \interlinepenalty \@M #5\@@par}%
- \endgroup
- \else
- \def\@svsechd{#4{\hskip #1\relax #5}}%
- \fi
- \@xsect{#3}}
-
-%
-% measures and setting of sections
-%
-\def\section{\@startsection{section}{1}{\z@}%
- {-21dd plus-8pt minus-4pt}{10.5dd}
- {\normalsize\bfseries\boldmath}}
-\def\subsection{\@startsection{subsection}{2}{\z@}%
- {-21dd plus-8pt minus-4pt}{10.5dd}
- {\normalsize\upshape}}
-\def\subsubsection{\@startsection{subsubsection}{3}{\z@}%
- {-13dd plus-8pt minus-4pt}{10.5dd}
- {\normalsize\itshape}}
-\def\paragraph{\@startsection{paragraph}{4}{\z@}%
- {-13pt plus-8pt minus-4pt}{\z@}{\normalsize\itshape}}
-
-\setlength\leftmargini {\parindent}
-\leftmargin \leftmargini
-\setlength\leftmarginii {\parindent}
-\setlength\leftmarginiii {1.87em}
-\setlength\leftmarginiv {1.7em}
-\setlength\leftmarginv {.5em}
-\setlength\leftmarginvi {.5em}
-\setlength \labelsep {.5em}
-\setlength \labelwidth{\leftmargini}
-\addtolength\labelwidth{-\labelsep}
-\@beginparpenalty -\@lowpenalty
-\@endparpenalty -\@lowpenalty
-\@itempenalty -\@lowpenalty
-\renewcommand\theenumi{\@arabic\c@enumi}
-\renewcommand\theenumii{\@alph\c@enumii}
-\renewcommand\theenumiii{\@roman\c@enumiii}
-\renewcommand\theenumiv{\@Alph\c@enumiv}
-\newcommand\labelenumi{\theenumi.}
-\newcommand\labelenumii{(\theenumii)}
-\newcommand\labelenumiii{\theenumiii.}
-\newcommand\labelenumiv{\theenumiv.}
-\renewcommand\p@enumii{\theenumi}
-\renewcommand\p@enumiii{\theenumi(\theenumii)}
-\renewcommand\p@enumiv{\p@enumiii\theenumiii}
-\newcommand\labelitemi{\normalfont\bfseries --}
-\newcommand\labelitemii{\normalfont\bfseries --}
-\newcommand\labelitemiii{$\m@th\bullet$}
-\newcommand\labelitemiv{$\m@th\cdot$}
-
-\if@spthms
-% definition of the "\spnewtheorem" command.
-%
-% Usage:
-%
-% \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font}
-% or \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font}
-% or \spnewtheorem*{env_nam}{caption}{cap_font}{body_font}
-%
-% New is "cap_font" and "body_font". It stands for
-% fontdefinition of the caption and the text itself.
-%
-% "\spnewtheorem*" gives a theorem without number.
-%
-% A defined spnewthoerem environment is used as described
-% by Lamport.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\def\@thmcountersep{}
-\def\@thmcounterend{}
-\newcommand\nocaption{\noexpand\@gobble}
-\newdimen\spthmsep \spthmsep=5pt
-
-\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}}
-
-% definition of \spnewtheorem with number
-
-\def\@spnthm#1#2{%
- \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}}
-\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}}
-
-\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}\@addtoreset{#1}{#3}%
- \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand
- \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}%
- \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spothm#1[#2]#3#4#5{%
- \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}%
- {\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{the#1}{\@nameuse{the#2}}%
- \expandafter\xdef\csname #1name\endcsname{#3}%
- \global\@namedef{#1}{\@spthm{#2}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}}
-
-\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\labelsep=\spthmsep\refstepcounter{#1}%
-\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}}
-
-\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}%
- \ignorespaces}
-
-\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname
- the#1\endcsname}{#5}{#3}{#4}\ignorespaces}
-
-\def\normalthmheadings{\def\@spbegintheorem##1##2##3##4{\trivlist\normalfont
- \item[\hskip\labelsep{##3##1\ ##2\@thmcounterend}]##4}
-\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##1\ ##2}]{##4(##3)\@thmcounterend\ }##5}}
-\normalthmheadings
-
-\def\reversethmheadings{\def\@spbegintheorem##1##2##3##4{\trivlist\normalfont
- \item[\hskip\labelsep{##3##2\ ##1\@thmcounterend}]##4}
-\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##2\ ##1}]{##4(##3)\@thmcounterend\ }##5}}
-
-% definition of \spnewtheorem* without number
-
-\def\@sthm#1#2{\@Ynthm{#1}{#2}}
-
-\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}}
-
-\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces}
-
-\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1}
- {#4}{#2}{#3}\ignorespaces}
-
-\def\@Begintheorem#1#2#3{#3\trivlist
- \item[\hskip\labelsep{#2#1\@thmcounterend}]}
-
-\def\@Opargbegintheorem#1#2#3#4{#4\trivlist
- \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }}
-
-% initialize theorem environment
-
-\if@envcntsect
- \def\@thmcountersep{.}
- \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape}
-\else
- \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape}
- \if@envcntreset
- \@addtoreset{theorem}{section}
- \else
- \@addtoreset{theorem}{chapter}
- \fi
-\fi
-
-%definition of divers theorem environments
-\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily}
-\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily}
-\if@envcntsame % all environments like "Theorem" - using its counter
- \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}}
-\else % all environments with their own counter
- \if@envcntsect % show section counter
- \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}}
- \else % not numbered with section
- \if@envcntreset
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{section}}
- \else
- \let\spn@wtheorem=\@spynthm
- \fi
- \fi
-\fi
-%
-\let\spdefaulttheorem=\spn@wtheorem
-%
-\spn@wtheorem{case}{Case}{\itshape}{\rmfamily}
-\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily}
-\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape}
-\spn@wtheorem{definition}{Definition}{\bfseries}{\rmfamily}
-\spn@wtheorem{example}{Example}{\itshape}{\rmfamily}
-\spn@wtheorem{exercise}{Exercise}{\bfseries}{\rmfamily}
-\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape}
-\spn@wtheorem{note}{Note}{\itshape}{\rmfamily}
-\spn@wtheorem{problem}{Problem}{\bfseries}{\rmfamily}
-\spn@wtheorem{property}{Property}{\itshape}{\rmfamily}
-\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape}
-\spn@wtheorem{question}{Question}{\itshape}{\rmfamily}
-\spn@wtheorem{solution}{Solution}{\bfseries}{\rmfamily}
-\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily}
-%
-\newenvironment{theopargself}
- {\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5}
- \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
- \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }}}{}
-\newenvironment{theopargself*}
- {\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##1\ ##2}]{\hspace*{-\labelsep}##4##3\@thmcounterend}##5}
- \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
- \item[\hskip\labelsep{##3##1}]{\hspace*{-\labelsep}##3##2\@thmcounterend}}}{}
-%
-\fi
-
-\def\@takefromreset#1#2{%
- \def\@tempa{#1}%
- \let\@tempd\@elt
- \def\@elt##1{%
- \def\@tempb{##1}%
- \ifx\@tempa\@tempb\else
- \@addtoreset{##1}{#2}%
- \fi}%
- \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
- \expandafter\def\csname cl@#2\endcsname{}%
- \@tempc
- \let\@elt\@tempd}
-
-\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}}
-\def\qed{\ifmmode\else\unskip\quad\fi\squareforqed}
-\def\smartqed{\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil
-\penalty50\hskip1em\null\nobreak\hfil\squareforqed
-\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi}}
-
-% Define `abstract' environment
-\def\abstract{\topsep=0pt\partopsep=0pt\parsep=0pt\itemsep=0pt\relax
-\trivlist\item[\hskip\labelsep
-{\bfseries\abstractname}]\if!\abstractname!\hskip-\labelsep\fi}
-\if@twocolumn
-% \if@avier
-% \def\endabstract{\endtrivlist\addvspace{5mm}\strich}
-% \def\strich{\hrule\vskip1ptplus12pt}
-% \else
- \def\endabstract{\endtrivlist\addvspace{3mm}}
-% \fi
-\else
-\fi
-%
-\newenvironment{verse}
- {\let\\\@centercr
- \list{}{\itemsep \z@
- \itemindent -1.5em%
- \listparindent\itemindent
- \rightmargin \leftmargin
- \advance\leftmargin 1.5em}%
- \item\relax}
- {\endlist}
-\newenvironment{quotation}
- {\list{}{\listparindent 1.5em%
- \itemindent \listparindent
- \rightmargin \leftmargin
- \parsep \z@ \@plus\p@}%
- \item\relax}
- {\endlist}
-\newenvironment{quote}
- {\list{}{\rightmargin\leftmargin}%
- \item\relax}
- {\endlist}
-\newcommand\appendix{\par\small
- \setcounter{section}{0}%
- \setcounter{subsection}{0}%
- \renewcommand\thesection{\@Alph\c@section}}
-\setlength\arraycolsep{1.5\p@}
-\setlength\tabcolsep{6\p@}
-\setlength\arrayrulewidth{.4\p@}
-\setlength\doublerulesep{2\p@}
-\setlength\tabbingsep{\labelsep}
-\skip\@mpfootins = \skip\footins
-\setlength\fboxsep{3\p@}
-\setlength\fboxrule{.4\p@}
-\renewcommand\theequation{\@arabic\c@equation}
-\newcounter{figure}
-\renewcommand\thefigure{\@arabic\c@figure}
-\def\fps@figure{tbp}
-\def\ftype@figure{1}
-\def\ext@figure{lof}
-\def\fnum@figure{\figurename~\thefigure}
-\newenvironment{figure}
- {\@float{figure}}
- {\end@float}
-\newenvironment{figure*}
- {\@dblfloat{figure}}
- {\end@dblfloat}
-\newcounter{table}
-\renewcommand\thetable{\@arabic\c@table}
-\def\fps@table{tbp}
-\def\ftype@table{2}
-\def\ext@table{lot}
-\def\fnum@table{\tablename~\thetable}
-\newenvironment{table}
- {\@float{table}}
- {\end@float}
-\newenvironment{table*}
- {\@dblfloat{table}}
- {\end@dblfloat}
-%
-\def \@floatboxreset {%
- \reset@font
- \small
- \@setnobreak
- \@setminipage
-}
-%
-\newcommand{\tableheadseprule}{\noalign{\hrule height.375mm}}
-%
-\newlength\abovecaptionskip
-\newlength\belowcaptionskip
-\setlength\abovecaptionskip{10\p@}
-\setlength\belowcaptionskip{0\p@}
-\newcommand\leftlegendglue{}
-
-\def\fig@type{figure}
-
-\newdimen\figcapgap\figcapgap=3pt
-\newdimen\tabcapgap\tabcapgap=5.5pt
-
-\@ifundefined{floatlegendstyle}{\def\floatlegendstyle{\bfseries}}{}
-
-\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname
- ext@#1\endcsname}{#1}{\protect\numberline{\csname
- the#1\endcsname}{\ignorespaces #2}}\begingroup
- \@parboxrestore
- \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
- \endgroup}
-
-\def\capstrut{\vrule\@width\z@\@height\topskip}
-
-\@ifundefined{captionstyle}{\def\captionstyle{\normalfont\small}}{}
-
-\long\def\@makecaption#1#2{%
- \captionstyle
- \ifx\@captype\fig@type
- \vskip\figcapgap
- \fi
- \setbox\@tempboxa\hbox{{\floatlegendstyle #1\floatcounterend}%
- \capstrut #2}%
- \ifdim \wd\@tempboxa >\hsize
- {\floatlegendstyle #1\floatcounterend}\capstrut #2\par
- \else
- \hbox to\hsize{\leftlegendglue\unhbox\@tempboxa\hfil}%
- \fi
- \ifx\@captype\fig@type\else
- \vskip\tabcapgap
- \fi}
-
-\newdimen\figgap\figgap=1cc
-\long\def\@makesidecaption#1#2{%
- \parbox[b]{\@tempdimb}{\captionstyle{\floatlegendstyle
- #1\floatcounterend}#2}}
-\def\sidecaption#1\caption{%
-\setbox\@tempboxa=\hbox{#1\unskip}%
-\if@twocolumn
- \ifdim\hsize<\textwidth\else
- \ifdim\wd\@tempboxa<\columnwidth
- \typeout{Double column float fits into single column -
- ^^Jyou'd better switch the environment. }%
- \fi
- \fi
-\fi
-\@tempdimb=\hsize
-\advance\@tempdimb by-\figgap
-\advance\@tempdimb by-\wd\@tempboxa
-\ifdim\@tempdimb<3cm
- \typeout{\string\sidecaption: No sufficient room for the legend;
- using normal \string\caption. }%
- \unhbox\@tempboxa
- \let\@capcommand=\@caption
-\else
- \let\@capcommand=\@sidecaption
- \leavevmode
- \unhbox\@tempboxa
- \hfill
-\fi
-\refstepcounter\@captype
-\@dblarg{\@capcommand\@captype}}
-
-\long\def\@sidecaption#1[#2]#3{\addcontentsline{\csname
- ext@#1\endcsname}{#1}{\protect\numberline{\csname
- the#1\endcsname}{\ignorespaces #2}}\begingroup
- \@parboxrestore
- \@makesidecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
- \endgroup}
-
-% Define `acknowledgement' environment
-\def\acknowledgement{\par\addvspace{17pt}\small\rmfamily
-\trivlist\if!\ackname!\item[]\else
-\item[\hskip\labelsep
-{\bfseries\ackname}]\fi}
-\def\endacknowledgement{\endtrivlist\addvspace{6pt}}
-\newenvironment{acknowledgements}{\begin{acknowledgement}}
-{\end{acknowledgement}}
-% Define `noteadd' environment
-\def\noteadd{\par\addvspace{17pt}\small\rmfamily
-\trivlist\item[\hskip\labelsep
-{\itshape\noteaddname}]}
-\def\endnoteadd{\endtrivlist\addvspace{6pt}}
-
-\DeclareOldFontCommand{\rm}{\normalfont\rmfamily}{\mathrm}
-\DeclareOldFontCommand{\sf}{\normalfont\sffamily}{\mathsf}
-\DeclareOldFontCommand{\tt}{\normalfont\ttfamily}{\mathtt}
-\DeclareOldFontCommand{\bf}{\normalfont\bfseries}{\mathbf}
-\DeclareOldFontCommand{\it}{\normalfont\itshape}{\mathit}
-\DeclareOldFontCommand{\sl}{\normalfont\slshape}{\@nomath\sl}
-\DeclareOldFontCommand{\sc}{\normalfont\scshape}{\@nomath\sc}
-\DeclareRobustCommand*\cal{\@fontswitch\relax\mathcal}
-\DeclareRobustCommand*\mit{\@fontswitch\relax\mathnormal}
-\newcommand\@pnumwidth{1.55em}
-\newcommand\@tocrmarg{2.55em}
-\newcommand\@dotsep{4.5}
-\setcounter{tocdepth}{1}
-\newcommand\tableofcontents{%
- \section*{\contentsname}%
- \@starttoc{toc}%
- \addtocontents{toc}{\begingroup\protect\small}%
- \AtEndDocument{\addtocontents{toc}{\endgroup}}%
- }
-\newcommand*\l@part[2]{%
- \ifnum \c@tocdepth >-2\relax
- \addpenalty\@secpenalty
- \addvspace{2.25em \@plus\p@}%
- \begingroup
- \setlength\@tempdima{3em}%
- \parindent \z@ \rightskip \@pnumwidth
- \parfillskip -\@pnumwidth
- {\leavevmode
- \large \bfseries #1\hfil \hb@xt@\@pnumwidth{\hss #2}}\par
- \nobreak
- \if@compatibility
- \global\@nobreaktrue
- \everypar{\global\@nobreakfalse\everypar{}}%
- \fi
- \endgroup
- \fi}
-\newcommand*\l@section{\@dottedtocline{1}{0pt}{1.5em}}
-\newcommand*\l@subsection{\@dottedtocline{2}{1.5em}{2.3em}}
-\newcommand*\l@subsubsection{\@dottedtocline{3}{3.8em}{3.2em}}
-\newcommand*\l@paragraph{\@dottedtocline{4}{7.0em}{4.1em}}
-\newcommand*\l@subparagraph{\@dottedtocline{5}{10em}{5em}}
-\newcommand\listoffigures{%
- \section*{\listfigurename
- \@mkboth{\listfigurename}%
- {\listfigurename}}%
- \@starttoc{lof}%
- }
-\newcommand*\l@figure{\@dottedtocline{1}{1.5em}{2.3em}}
-\newcommand\listoftables{%
- \section*{\listtablename
- \@mkboth{\listtablename}{\listtablename}}%
- \@starttoc{lot}%
- }
-\let\l@table\l@figure
-\newdimen\bibindent
-\setlength\bibindent{\parindent}
-\def\@biblabel#1{#1.}
-\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw
- {\let\protect\noexpand
- \immediate
- \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
-\newenvironment{thebibliography}[1]
- {\section*{\refname
- \@mkboth{\refname}{\refname}}\small
- \list{\@biblabel{\@arabic\c@enumiv}}%
- {\settowidth\labelwidth{\@biblabel{#1}}%
- \leftmargin\labelwidth
- \advance\leftmargin\labelsep
- \@openbib@code
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{\@arabic\c@enumiv}}%
- \sloppy\clubpenalty4000\widowpenalty4000%
- \sfcode`\.\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
-%
-\newcount\@tempcntc
-\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi
- \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do
- {\@ifundefined
- {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries
- ?}\@warning
- {Citation `\@citeb' on page \thepage \space undefined}}%
- {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}%
- \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne
- \@citea\def\@citea{,\hskip0.1em\ignorespaces}\hbox{\csname b@\@citeb\endcsname}%
- \else
- \advance\@tempcntb\@ne
- \ifnum\@tempcntb=\@tempcntc
- \else\advance\@tempcntb\m@ne\@citeo
- \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}}
-\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else
- \@citea\def\@citea{,\hskip0.1em\ignorespaces}%
- \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else
- {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else \def\@citea{--}\fi
- \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi}
-%
-\newcommand\newblock{\hskip .11em\@plus.33em\@minus.07em}
-\let\@openbib@code\@empty
-\newenvironment{theindex}
- {\if@twocolumn
- \@restonecolfalse
- \else
- \@restonecoltrue
- \fi
- \columnseprule \z@
- \columnsep 35\p@
- \twocolumn[\section*{\indexname}]%
- \@mkboth{\indexname}{\indexname}%
- \thispagestyle{plain}\parindent\z@
- \parskip\z@ \@plus .3\p@\relax
- \let\item\@idxitem}
- {\if@restonecol\onecolumn\else\clearpage\fi}
-\newcommand\@idxitem{\par\hangindent 40\p@}
-\newcommand\subitem{\@idxitem \hspace*{20\p@}}
-\newcommand\subsubitem{\@idxitem \hspace*{30\p@}}
-\newcommand\indexspace{\par \vskip 10\p@ \@plus5\p@ \@minus3\p@\relax}
-
-\if@twocolumn
- \renewcommand\footnoterule{%
- \kern-3\p@
- \hrule\@width\columnwidth
- \kern2.6\p@}
-\else
- \renewcommand\footnoterule{%
- \kern-3\p@
- \hrule\@width.382\columnwidth
- \kern2.6\p@}
-\fi
-\newcommand\@makefntext[1]{%
- \noindent
- \hb@xt@\bibindent{\hss\@makefnmark\enspace}#1}
-%
-\def\trans@english{\switcht@albion}
-\def\trans@french{\switcht@francais}
-\def\trans@german{\switcht@deutsch}
-\newenvironment{translation}[1]{\if!#1!\else
-\@ifundefined{selectlanguage}{\csname trans@#1\endcsname}{\selectlanguage{#1}}%
-\fi}{}
-% languages
-% English section
-\def\switcht@albion{%\typeout{English spoken.}%
- \def\abstractname{Abstract}%
- \def\ackname{Acknowledgements}%
- \def\andname{and}%
- \def\lastandname{, and}%
- \def\appendixname{Appendix}%
- \def\chaptername{Chapter}%
- \def\claimname{Claim}%
- \def\conjecturename{Conjecture}%
- \def\contentsname{Contents}%
- \def\corollaryname{Corollary}%
- \def\definitionname{Definition}%
- \def\emailname{E-mail}%
- \def\examplename{Example}%
- \def\exercisename{Exercise}%
- \def\figurename{Fig.}%
- \def\keywordname{{\bfseries Keywords}}%
- \def\indexname{Index}%
- \def\lemmaname{Lemma}%
- \def\contriblistname{List of Contributors}%
- \def\listfigurename{List of Figures}%
- \def\listtablename{List of Tables}%
- \def\mailname{{\itshape Correspondence to\/}:}%
- \def\noteaddname{Note added in proof}%
- \def\notename{Note}%
- \def\partname{Part}%
- \def\problemname{Problem}%
- \def\proofname{Proof}%
- \def\propertyname{Property}%
- \def\questionname{Question}%
- \def\refname{References}%
- \def\remarkname{Remark}%
- \def\seename{see}%
- \def\solutionname{Solution}%
- \def\tablename{Table}%
- \def\theoremname{Theorem}%
-}\switcht@albion % make English default
-%
-% French section
-\def\switcht@francais{\svlanginfo
-%\typeout{On parle francais.}%
- \def\abstractname{R\'esum\'e\runinend}%
- \def\ackname{Remerciements\runinend}%
- \def\andname{et}%
- \def\lastandname{ et}%
- \def\appendixname{Appendice}%
- \def\chaptername{Chapitre}%
- \def\claimname{Pr\'etention}%
- \def\conjecturename{Hypoth\`ese}%
- \def\contentsname{Table des mati\`eres}%
- \def\corollaryname{Corollaire}%
- \def\definitionname{D\'efinition}%
- \def\emailname{E-mail}%
- \def\examplename{Exemple}%
- \def\exercisename{Exercice}%
- \def\figurename{Fig.}%
- \def\keywordname{{\bfseries Mots-cl\'e\runinend}}%
- \def\indexname{Index}%
- \def\lemmaname{Lemme}%
- \def\contriblistname{Liste des contributeurs}%
- \def\listfigurename{Liste des figures}%
- \def\listtablename{Liste des tables}%
- \def\mailname{{\itshape Correspondence to\/}:}%
- \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}%
- \def\notename{Remarque}%
- \def\partname{Partie}%
- \def\problemname{Probl\`eme}%
- \def\proofname{Preuve}%
- \def\propertyname{Caract\'eristique}%
-%\def\propositionname{Proposition}%
- \def\questionname{Question}%
- \def\refname{Bibliographie}%
- \def\remarkname{Remarque}%
- \def\seename{voyez}%
- \def\solutionname{Solution}%
-%\def\subclassname{{\it Subject Classifications\/}:}%
- \def\tablename{Tableau}%
- \def\theoremname{Th\'eor\`eme}%
-}
-%
-% German section
-\def\switcht@deutsch{\svlanginfo
-%\typeout{Man spricht deutsch.}%
- \def\abstractname{Zusammenfassung\runinend}%
- \def\ackname{Danksagung\runinend}%
- \def\andname{und}%
- \def\lastandname{ und}%
- \def\appendixname{Anhang}%
- \def\chaptername{Kapitel}%
- \def\claimname{Behauptung}%
- \def\conjecturename{Hypothese}%
- \def\contentsname{Inhaltsverzeichnis}%
- \def\corollaryname{Korollar}%
-%\def\definitionname{Definition}%
- \def\emailname{E-Mail}%
- \def\examplename{Beispiel}%
- \def\exercisename{\"Ubung}%
- \def\figurename{Abb.}%
- \def\keywordname{{\bfseries Schl\"usselw\"orter\runinend}}%
- \def\indexname{Index}%
-%\def\lemmaname{Lemma}%
- \def\contriblistname{Mitarbeiter}%
- \def\listfigurename{Abbildungsverzeichnis}%
- \def\listtablename{Tabellenverzeichnis}%
- \def\mailname{{\itshape Correspondence to\/}:}%
- \def\noteaddname{Nachtrag}%
- \def\notename{Anmerkung}%
- \def\partname{Teil}%
-%\def\problemname{Problem}%
- \def\proofname{Beweis}%
- \def\propertyname{Eigenschaft}%
-%\def\propositionname{Proposition}%
- \def\questionname{Frage}%
- \def\refname{Literatur}%
- \def\remarkname{Anmerkung}%
- \def\seename{siehe}%
- \def\solutionname{L\"osung}%
-%\def\subclassname{{\it Subject Classifications\/}:}%
- \def\tablename{Tabelle}%
-%\def\theoremname{Theorem}%
-}
-\newcommand\today{}
-\edef\today{\ifcase\month\or
- January\or February\or March\or April\or May\or June\or
- July\or August\or September\or October\or November\or December\fi
- \space\number\day, \number\year}
-\setlength\columnsep{1.5cc}
-\setlength\columnseprule{0\p@}
-%
-\frenchspacing
-\clubpenalty=10000
-\widowpenalty=10000
-\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil
-\global\let\@textbottom\relax}}
-\pagestyle{headings}
-\pagenumbering{arabic}
-\if@twocolumn
- \twocolumn
-\fi
-%\if@avier
-% \onecolumn
-% \setlength{\textwidth}{156mm}
-% \setlength{\textheight}{226mm}
-%\fi
-\if@referee
- \makereferee
-\fi
-\flushbottom
-\endinput
-%%
-%% End of file `svjour3.cls'.
--- a/Pearl/Paper.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1231 +0,0 @@
-(*<*)
-theory Paper
-imports "../Nominal/Nominal2_Base"
- "../Nominal/Atoms"
- "~~/src/HOL/Library/LaTeXsugar"
-begin
-
-notation (latex output)
- sort_of ("sort _" [1000] 100) and
- Abs_perm ("_") and
- Rep_perm ("_") and
- swap ("'(_ _')" [1000, 1000] 1000) and
- fresh ("_ # _" [51, 51] 50) and
- Cons ("_::_" [78,77] 73) and
- supp ("supp _" [78] 73) and
- uminus ("-_" [78] 73) and
- atom ("|_|") and
- If ("if _ then _ else _" 10) and
- Rep_name ("\<lfloor>_\<rfloor>") and
- Abs_name ("\<lceil>_\<rceil>") and
- Rep_var ("\<lfloor>_\<rfloor>") and
- Abs_var ("\<lceil>_\<rceil>") and
- sort_of_ty ("sort'_ty _")
-
-(* BH: uncomment if you really prefer the dot notation
-syntax (latex output)
- "_Collect" :: "pttrn => bool => 'a set" ("(1{_ . _})")
-*)
-
-(* sort is used in Lists for sorting *)
-hide_const sort
-
-abbreviation
- "sort \<equiv> sort_of"
-
-abbreviation
- "sort_ty \<equiv> sort_of_ty"
-
-(*>*)
-
-section {* Introduction *}
-
-text {*
- Nominal Isabelle is a definitional extension of the Isabelle/HOL theorem
- prover providing a proving infrastructure for convenient reasoning about
- programming languages. It has been used to formalise an equivalence checking
- algorithm for LF \cite{UrbanCheneyBerghofer08},
- Typed Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for concurrency
- \cite{BengtsonParrow07} and a strong normalisation result for
- cut-elimination in classical logic \cite{UrbanZhu08}. It has also been used
- by Pollack for formalisations in the locally-nameless approach to binding
- \cite{SatoPollack10}.
-
- At its core Nominal Isabelle is based on the nominal logic work of Pitts et
- al \cite{GabbayPitts02,Pitts03}. The most basic notion in this work is a
- sort-respecting permutation operation defined over a countably infinite
- collection of sorted atoms. The atoms are used for representing variables
- that might be bound. Multiple sorts are necessary for being
- able to represent different kinds of variables. For example, in the language
- Mini-ML there are bound term variables and bound type variables; each kind
- needs to be represented by a different sort of atoms.
-
- Unfortunately, the type system of Isabelle/HOL is not a good fit for the way
- atoms and sorts are used in the original formulation of the nominal logic work.
- Therefore it was decided in earlier versions of Nominal Isabelle to use a
- separate type for each sort of atoms and let the type system enforce the
- sort-respecting property of permutations. Inspired by the work on nominal
- unification \cite{UrbanPittsGabbay04}, it seemed best at the time to also
- implement permutations concretely as lists of pairs of atoms. Thus Nominal
- Isabelle used the two-place permutation operation with the generic type
-
- @{text [display,indent=10] "_ \<bullet> _ :: (\<alpha> \<times> \<alpha>) list \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
-
- \noindent
- where @{text "\<alpha>"} stands for the type of atoms and @{text "\<beta>"} for the type
- of the objects on which the permutation acts. For atoms of type @{text "\<alpha>"}
- the permutation operation is defined over the length of lists as follows
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
- @{text "[] \<bullet> c"} & @{text "="} & @{text c}\\
- \end{tabular}\hspace{12mm}
- \begin{tabular}{@ {}r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
- @{text "(a b)::\<pi> \<bullet> c"} & @{text "="} &
- $\begin{cases} @{text a} & \textrm{if}~@{text "\<pi> \<bullet> c = b"}\\
- @{text b} & \textrm{if}~@{text "\<pi> \<bullet> c = a"}\\
- @{text "\<pi> \<bullet> c"} & \textrm{otherwise}\end{cases}$
- \end{tabular}\hfill\numbered{atomperm}
- \end{isabelle}
-
- \noindent
- where we write @{text "(a b)"} for a swapping of atoms @{text "a"} and
- @{text "b"}. For atoms of different type, the permutation operation
- is defined as @{text "\<pi> \<bullet> c \<equiv> c"}.
-
- With the list representation of permutations it is impossible to state an
- ``ill-sorted'' permutation, since the type system excludes lists containing
- atoms of different type. Another advantage of the list representation is that
- the basic operations on permutations are already defined in the list library:
- composition of two permutations (written @{text "_ @ _"}) is just list append,
- and inversion of a permutation (written @{text "_\<^sup>-\<^sup>1"}) is just
- list reversal. A disadvantage is that permutations do not have unique
- representations as lists; we had to explicitly identify permutations according
- to the relation
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{text "\<pi>\<^isub>1 \<sim> \<pi>\<^isub>2 \<equiv> \<forall>a. \<pi>\<^isub>1 \<bullet> a = \<pi>\<^isub>2 \<bullet> a"}
- \end{tabular}\hfill\numbered{permequ}
- \end{isabelle}
-
- When lifting the permutation operation to other types, for example sets,
- functions and so on, we needed to ensure that every definition is
- well-behaved in the sense that it satisfies the following three
- \emph{permutation properties}:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}r@ {\hspace{4mm}}p{10cm}}
- i) & @{text "[] \<bullet> x = x"}\\
- ii) & @{text "(\<pi>\<^isub>1 @ \<pi>\<^isub>2) \<bullet> x = \<pi>\<^isub>1 \<bullet> (\<pi>\<^isub>2 \<bullet> x)"}\\
- iii) & if @{text "\<pi>\<^isub>1 \<sim> \<pi>\<^isub>2"} then @{text "\<pi>\<^isub>1 \<bullet> x = \<pi>\<^isub>2 \<bullet> x"}
- \end{tabular}\hfill\numbered{permprops}
- \end{isabelle}
-
- \noindent
- From these properties we were able to derive most facts about permutations, and
- the type classes of Isabelle/HOL allowed us to reason abstractly about these
- three properties, and then let the type system automatically enforce these
- properties for each type.
-
- The major problem with Isabelle/HOL's type classes, however, is that they
- support operations with only a single type parameter and the permutation
- operations @{text "_ \<bullet> _"} used above in the permutation properties
- contain two! To work around this obstacle, Nominal Isabelle
- required the user to
- declare up-front the collection of \emph{all} atom types, say @{text
- "\<alpha>\<^isub>1,\<dots>,\<alpha>\<^isub>n"}. From this collection it used custom ML-code to
- generate @{text n} type classes corresponding to the permutation properties,
- whereby in these type classes the permutation operation is restricted to
-
- @{text [display,indent=10] "_ \<bullet> _ :: (\<alpha>\<^isub>i \<times> \<alpha>\<^isub>i) list \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
-
- \noindent
- This operation has only a single type parameter @{text "\<beta>"} (the @{text "\<alpha>\<^isub>i"} are the
- atom types given by the user).
-
- While the representation of permutations-as-lists solved the
- ``sort-respecting'' requirement and the declaration of all atom types
- up-front solved the problem with Isabelle/HOL's type classes, this setup
- caused several problems for formalising the nominal logic work: First,
- Nominal Isabelle had to generate @{text "n\<^sup>2"} definitions for the
- permutation operation over @{text "n"} types of atoms. Second, whenever we
- need to generalise induction hypotheses by quantifying over permutations, we
- have to build cumbersome quantifications like
-
- @{text [display,indent=10] "\<forall>\<pi>\<^isub>1 \<dots> \<forall>\<pi>\<^isub>n. \<dots>"}
-
- \noindent
- where the @{text "\<pi>\<^isub>i"} are of type @{text "(\<alpha>\<^isub>i \<times> \<alpha>\<^isub>i) list"}.
- The reason is that the permutation operation behaves differently for
- every @{text "\<alpha>\<^isub>i"}. Third, although the notion of support
-
- @{text [display,indent=10] "supp _ :: \<beta> \<Rightarrow> \<alpha> set"}
-
- \noindent
- which we will define later, has a generic type @{text "\<alpha> set"}, it cannot be
- used to express the support of an object over \emph{all} atoms. The reason
- is again that support can behave differently for each @{text
- "\<alpha>\<^isub>i"}. This problem is annoying, because if we need to know in
- a statement that an object, say @{text "x"}, is finitely supported we end up
- with having to state premises of the form
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{text "finite ((supp x) :: \<alpha>\<^isub>1 set) , \<dots>, finite ((supp x) :: \<alpha>\<^isub>n set)"}
- \end{tabular}\hfill\numbered{fssequence}
- \end{isabelle}
-
- \noindent
- Sometimes we can avoid such premises completely, if @{text x} is a member of a
- \emph{finitely supported type}. However, keeping track of finitely supported
- types requires another @{text n} type classes, and for technical reasons not
- all types can be shown to be finitely supported.
-
- The real pain of having a separate type for each atom sort arises, however,
- from another permutation property
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}r@ {\hspace{4mm}}p{10cm}}
- iv) & @{text "\<pi>\<^isub>1 \<bullet> (\<pi>\<^isub>2 \<bullet> x) = (\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2) \<bullet> (\<pi>\<^isub>1 \<bullet> x)"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- where permutation @{text "\<pi>\<^isub>1"} has type @{text "(\<alpha> \<times> \<alpha>) list"},
- @{text "\<pi>\<^isub>2"} type @{text "(\<alpha>' \<times> \<alpha>') list"} and @{text x} type @{text
- "\<beta>"}. This property is needed in order to derive facts about how
- permutations of different types interact, which is not covered by the
- permutation properties @{text "i"}-@{text "iii"} shown in
- \eqref{permprops}. The problem is that this property involves three type
- parameters. In order to use again Isabelle/HOL's type class mechanism with
- only permitting a single type parameter, we have to instantiate the atom
- types. Consequently we end up with an additional @{text "n\<^sup>2"}
- slightly different type classes for this permutation property.
-
- While the problems and pain can be almost completely hidden from the user in
- the existing implementation of Nominal Isabelle, the work is \emph{not}
- pretty. It requires a large amount of custom ML-code and also forces the
- user to declare up-front all atom-types that are ever going to be used in a
- formalisation. In this paper we set out to solve the problems with multiple
- type parameters in the permutation operation, and in this way can dispense
- with the large amounts of custom ML-code for generating multiple variants
- for some basic definitions. The result is that we can implement a pleasingly
- simple formalisation of the nominal logic work.\smallskip
-
- \noindent
- {\bf Contributions of the paper:} Using a single atom type to represent
- atoms of different sorts and representing permutations as functions are not
- new ideas. The main contribution of this paper is to show an example of how
- to make better theorem proving tools by choosing the right level of
- abstraction for the underlying theory---our design choices take advantage of
- Isabelle's type system, type classes, and reasoning infrastructure.
- The novel
- technical contribution is a mechanism for dealing with
- ``Church-style'' lambda-terms \cite{Church40} and HOL-based languages
- \cite{PittsHOL4} where variables and variable binding depend on type
- annotations.
-*}
-
-section {* Sorted Atoms and Sort-Respecting Permutations *}
-
-text {*
- In the nominal logic work of Pitts, binders and bound variables are
- represented by \emph{atoms}. As stated above, we need to have different
- \emph{sorts} of atoms to be able to bind different kinds of variables. A
- basic requirement is that there must be a countably infinite number of atoms
- of each sort. Unlike in our earlier work, where we identified each sort with
- a separate type, we implement here atoms to be
-*}
-
- datatype atom\<iota> = Atom\<iota> string nat
-
-text {*
- \noindent
- whereby the string argument specifies the sort of the atom.\footnote{A similar
- design choice was made by Gunter et al \cite{GunterOsbornPopescu09}
- for their variables.} (The use type
- \emph{string} is merely for convenience; any countably infinite type would work
- as well.)
- We have an auxiliary function @{text sort} that is defined as @{thm
- sort_of.simps[no_vars]}, and we clearly have for every finite set @{text X} of
- atoms and every sort @{text s} the property:
-
- \begin{proposition}\label{choosefresh}
- @{text "If finite X then there exists an atom a such that
- sort a = s and a \<notin> X"}.
- \end{proposition}
-
- For implementing sort-respecting permutations, we use functions of type @{typ
- "atom => atom"} that @{text "i)"} are bijective; @{text "ii)"} are the
- identity on all atoms, except a finite number of them; and @{text "iii)"} map
- each atom to one of the same sort. These properties can be conveniently stated
- for a function @{text \<pi>} as follows:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- i)~~~@{term "bij \<pi>"}\hspace{5mm}
- ii)~~~@{term "finite {a. \<pi> a \<noteq> a}"}\hspace{5mm}
- iii)~~~@{term "\<forall>a. sort (\<pi> a) = sort a"}\hfill\numbered{permtype}
- \end{isabelle}
-
- \noindent
- Like all HOL-based theorem provers, Isabelle/HOL allows us to
- introduce a new type @{typ perm} that includes just those functions
- satisfying all three properties. For example the identity function,
- written @{term id}, is included in @{typ perm}. Also function composition,
- written \mbox{@{text "_ \<circ> _"}}, and function inversion, given by Isabelle/HOL's
- inverse operator and written \mbox{@{text "inv _"}}, preserve the properties
- @{text "i"}-@{text "iii"}.
-
- However, a moment of thought is needed about how to construct non-trivial
- permutations. In the nominal logic work it turned out to be most convenient
- to work with swappings, written @{text "(a b)"}. In our setting the
- type of swappings must be
-
- @{text [display,indent=10] "(_ _) :: atom \<Rightarrow> atom \<Rightarrow> perm"}
-
- \noindent
- but since permutations are required to respect sorts, we must carefully
- consider what happens if a user states a swapping of atoms with different
- sorts. In earlier versions of Nominal Isabelle, we avoided this problem by
- using different types for different sorts; the type system prevented users
- from stating ill-sorted swappings. Here, however, definitions such
- as\footnote{To increase legibility, we omit here and in what follows the
- @{term Rep_perm} and @{term "Abs_perm"} wrappers that are needed in our
- implementation since we defined permutation not to be the full function space,
- but only those functions of type @{typ perm} satisfying properties @{text
- i}-@{text "iii"}.}
-
- @{text [display,indent=10] "(a b) \<equiv> \<lambda>c. if a = c then b else (if b = c then a else c)"}
-
- \noindent
- do not work in general, because the type system does not prevent @{text a}
- and @{text b} from having different sorts---in which case the function would
- violate property @{text iii}. We could make the definition of swappings
- partial by adding the precondition @{term "sort a = sort b"},
- which would mean that in case @{text a} and @{text b} have different sorts,
- the value of @{text "(a b)"} is unspecified. However, this looked like a
- cumbersome solution, since sort-related side conditions would be required
- everywhere, even to unfold the definition. It turned out to be more
- convenient to actually allow the user to state ``ill-sorted'' swappings but
- limit their ``damage'' by defaulting to the identity permutation in the
- ill-sorted case:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}rl}
- @{text "(a b) \<equiv>"} & @{text "if (sort a = sort b)"}\\
- & \hspace{3mm}@{text "then \<lambda>c. if a = c then b else (if b = c then a else c)"}\\
- & \hspace{3mm}@{text "else id"}
- \end{tabular}\hfill\numbered{swapdef}
- \end{isabelle}
-
- \noindent
- This function is bijective, the identity on all atoms except
- @{text a} and @{text b}, and sort respecting. Therefore it is
- a function in @{typ perm}.
-
- One advantage of using functions instead of lists as a representation for
- permutations is that for example the swappings
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm swap_commute[no_vars]}\hspace{10mm}
- @{text "(a a) = id"}
- \end{tabular}\hfill\numbered{swapeqs}
- \end{isabelle}
-
- \noindent
- are \emph{equal}. We do not have to use the equivalence relation shown
- in~\eqref{permequ} to identify them, as we would if they had been represented
- as lists of pairs. Another advantage of the function representation is that
- they form a (non-commutative) group, provided we define
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm zero_perm_def[no_vars, THEN eq_reflection]} \hspace{4mm}
- @{thm plus_perm_def[where p="\<pi>\<^isub>1" and q="\<pi>\<^isub>2", THEN eq_reflection]} \hspace{4mm}
- @{thm uminus_perm_def[where p="\<pi>", THEN eq_reflection]} \hspace{4mm}
- @{thm minus_perm_def[where ?p1.0="\<pi>\<^isub>1" and ?p2.0="\<pi>\<^isub>2"]}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- and verify the simple properties
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm add_assoc[where a="\<pi>\<^isub>1" and b="\<pi>\<^isub>2" and c="\<pi>\<^isub>3"]} \hspace{3mm}
- @{thm monoid_add_class.add_0_left[where a="\<pi>::perm"]} \hspace{3mm}
- @{thm monoid_add_class.add_0_right[where a="\<pi>::perm"]} \hspace{3mm}
- @{thm group_add_class.left_minus[where a="\<pi>::perm"]}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- Again this is in contrast to the list-of-pairs representation which does not
- form a group. The technical importance of this fact is that we can rely on
- Isabelle/HOL's existing simplification infrastructure for groups, which will
- come in handy when we have to do calculations with permutations.
- Note that Isabelle/HOL defies standard conventions of mathematical notation
- by using additive syntax even for non-commutative groups. Obviously,
- composition of permutations is not commutative in general---@{text
- "\<pi>\<^sub>1 + \<pi>\<^sub>2 \<noteq> \<pi>\<^sub>2 + \<pi>\<^sub>1"}. But since the point of this paper is to implement the
- nominal theory as smoothly as possible in Isabelle/HOL, we tolerate
- the non-standard notation in order to reuse the existing libraries.
-
- By formalising permutations abstractly as functions, and using a single type
- for all atoms, we can now restate the \emph{permutation properties} from
- \eqref{permprops} as just the two equations
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}r@ {\hspace{4mm}}p{10cm}}
- i) & @{thm permute_zero[no_vars]}\\
- ii) & @{thm permute_plus[where p="\<pi>\<^isub>1" and q="\<pi>\<^isub>2",no_vars]}
- \end{tabular}\hfill\numbered{newpermprops}
- \end{isabelle}
-
- \noindent
- in which the permutation operations are of type @{text "perm \<Rightarrow> \<beta> \<Rightarrow> \<beta>"} and so
- have only a single type parameter. Consequently, these properties are
- compatible with the one-parameter restriction of Isabelle/HOL's type classes.
- There is no need to introduce a separate type class instantiated for each
- sort, like in the old approach.
-
- The next notion allows us to establish generic lemmas involving the
- permutation operation.
-
- \begin{definition}
- A type @{text "\<beta>"} is a \emph{permutation type} if the permutation
- properties in \eqref{newpermprops} are satisfied for every @{text "x"} of type
- @{text "\<beta>"}.
- \end{definition}
-
- \noindent
- First, it follows from the laws governing
- groups that a permutation and its inverse cancel each other. That is, for any
- @{text "x"} of a permutation type:
-
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm permute_minus_cancel(1)[where p="\<pi>", no_vars]}\hspace{10mm}
- @{thm permute_minus_cancel(2)[where p="\<pi>", no_vars]}
- \end{tabular}\hfill\numbered{cancel}
- \end{isabelle}
-
- \noindent
- Consequently, in a permutation type the permutation operation @{text "\<pi> \<bullet> _"} is bijective,
- which in turn implies the property
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm (lhs) permute_eq_iff[where p="\<pi>", no_vars]}
- $\;$if and only if$\;$
- @{thm (rhs) permute_eq_iff[where p="\<pi>", no_vars]}.
- \end{tabular}\hfill\numbered{permuteequ}
- \end{isabelle}
-
- \noindent
- In order to lift the permutation operation to other types, we can define for:
-
- \begin{isabelle}
- \begin{tabular}{@ {}c@ {\hspace{-1mm}}c@ {}}
- \begin{tabular}{@ {}r@ {\hspace{2mm}}l@ {}}
- atoms: & @{thm permute_atom_def[where p="\<pi>",no_vars, THEN eq_reflection]}\\
- functions: & @{text "\<pi> \<bullet> f \<equiv> \<lambda>x. \<pi> \<bullet> (f ((-\<pi>) \<bullet> x))"}\\
- permutations: & @{thm permute_perm_def[where p="\<pi>" and q="\<pi>'", THEN eq_reflection]}\\
- sets: & @{thm permute_set_eq[where p="\<pi>", no_vars, THEN eq_reflection]}\\
- booleans: & @{thm permute_bool_def[where p="\<pi>", no_vars, THEN eq_reflection]}\\
- \end{tabular} &
- \begin{tabular}{@ {}r@ {\hspace{2mm}}l@ {}}
- lists: & @{thm permute_list.simps(1)[where p="\<pi>", no_vars, THEN eq_reflection]}\\
- & @{thm permute_list.simps(2)[where p="\<pi>", no_vars, THEN eq_reflection]}\\[2mm]
- products: & @{thm permute_prod.simps[where p="\<pi>", no_vars, THEN eq_reflection]}\\
- nats: & @{thm permute_nat_def[where p="\<pi>", no_vars, THEN eq_reflection]}\\
- \end{tabular}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- and then establish:
-
- \begin{theorem}
- If @{text \<beta>}, @{text "\<beta>\<^isub>1"} and @{text "\<beta>\<^isub>2"} are permutation types,
- then so are @{text "atom"}, @{text "\<beta>\<^isub>1 \<Rightarrow> \<beta>\<^isub>2"},
- @{text perm}, @{term "\<beta> set"}, @{term "\<beta> list"}, @{term "\<beta>\<^isub>1 \<times> \<beta>\<^isub>2"},
- @{text bool} and @{text "nat"}.
- \end{theorem}
-
- \begin{proof}
- All statements are by unfolding the definitions of the permutation operations and simple
- calculations involving addition and minus. With permutations for example we
- have
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}[b]{@ {}rcl}
- @{text "0 \<bullet> \<pi>'"} & @{text "\<equiv>"} & @{text "0 + \<pi>' - 0 = \<pi>'"}\\
- @{text "(\<pi>\<^isub>1 + \<pi>\<^isub>2) \<bullet> \<pi>'"} & @{text "\<equiv>"} & @{text "(\<pi>\<^isub>1 + \<pi>\<^isub>2) + \<pi>' - (\<pi>\<^isub>1 + \<pi>\<^isub>2)"}\\
- & @{text "="} & @{text "(\<pi>\<^isub>1 + \<pi>\<^isub>2) + \<pi>' - \<pi>\<^isub>2 - \<pi>\<^isub>1"}\\
- & @{text "="} & @{text "\<pi>\<^isub>1 + (\<pi>\<^isub>2 + \<pi>' - \<pi>\<^isub>2) - \<pi>\<^isub>1"} @{text "\<equiv>"} @{text "\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2 \<bullet> \<pi>'"}
- \end{tabular}\hfill\qed
- \end{isabelle}
- \end{proof}
-
- \noindent
- The main point is that the above reasoning blends smoothly with the reasoning
- infrastructure of Isabelle/HOL; no custom ML-code is necessary and a single
- type class suffices. We can also show once and for all that the following
- property---which caused so many headaches in our earlier setup---holds for any
- permutation type.
-
- \begin{lemma}\label{permutecompose}
- Given @{term x} is of permutation type, then
- @{text "\<pi>\<^isub>1 \<bullet> (\<pi>\<^isub>2 \<bullet> x) = (\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2) \<bullet> (\<pi>\<^isub>1 \<bullet> x)"}.
- \end{lemma}
-
- \begin{proof} The proof is as follows:
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}[b]{@ {}rcl@ {\hspace{8mm}}l}
- @{text "\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2 \<bullet> x"}
- & @{text "="} & @{text "\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2 \<bullet> (-\<pi>\<^isub>1) \<bullet> \<pi>\<^isub>1 \<bullet> x"} & by \eqref{cancel}\\
- & @{text "="} & @{text "(\<pi>\<^isub>1 + \<pi>\<^isub>2 - \<pi>\<^isub>1) \<bullet> \<pi>\<^isub>1 \<bullet> x"} & by {\rm(\ref{newpermprops}.@{text "ii"})}\\
- & @{text "\<equiv>"} & @{text "(\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2) \<bullet> (\<pi>\<^isub>1 \<bullet> x)"}\\
- \end{tabular}\hfill\qed
- \end{isabelle}
- \end{proof}
-
-%* }
-%
-%section { * Equivariance * }
-%
-%text { *
-
- An \emph{equivariant} function or predicate is one that is invariant under
- the swapping of atoms. Having a notion of equivariance with nice logical
- properties is a major advantage of bijective permutations over traditional
- renaming substitutions \cite[\S2]{Pitts03}. Equivariance can be defined
- uniformly for all permutation types, and it is satisfied by most HOL
- functions and constants.
-
- \begin{definition}\label{equivariance}
- A function @{text f} is \emph{equivariant} if @{term "\<forall>\<pi>. \<pi> \<bullet> f = f"}.
- \end{definition}
-
- \noindent
- There are a number of equivalent formulations for the equivariance property.
- For example, assuming @{text f} is of type @{text "\<alpha> \<Rightarrow> \<beta>"}, then equivariance
- can also be stated as
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{text "\<forall>\<pi> x. \<pi> \<bullet> (f x) = f (\<pi> \<bullet> x)"}
- \end{tabular}\hfill\numbered{altequivariance}
- \end{isabelle}
-
- \noindent
- To see that this formulation implies the definition, we just unfold the
- definition of the permutation operation for functions and simplify with the equation
- and the cancellation property shown in \eqref{cancel}. To see the other direction, we use
- the fact
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{text "\<pi> \<bullet> (f x) = (\<pi> \<bullet> f) (\<pi> \<bullet> x)"}
- \end{tabular}\hfill\numbered{permutefunapp}
- \end{isabelle}
-
- \noindent
- which follows again directly
- from the definition of the permutation operation for functions and the cancellation
- property. Similarly for functions with more than one argument.
-
- Both formulations of equivariance have their advantages and disadvantages:
- \eqref{altequivariance} is often easier to establish. For example we
- can easily show that equality is equivariant
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm eq_eqvt[where p="\<pi>", no_vars]}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- using the permutation operation on booleans and property \eqref{permuteequ}.
- Lemma~\ref{permutecompose} establishes that the permutation operation is
- equivariant. It is also easy to see that the boolean operators, like
- @{text "\<and>"}, @{text "\<or>"} and @{text "\<longrightarrow>"} are all equivariant. Furthermore
- a simple calculation will show that our swapping functions are equivariant, that is
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm swap_eqvt[where p="\<pi>", no_vars]}
- \end{tabular}\hfill\numbered{swapeqvt}
- \end{isabelle}
-
- \noindent
- for all @{text a}, @{text b} and @{text \<pi>}. These equivariance properties
- are tremendously helpful later on when we have to push permutations inside
- terms.
-*}
-
-
-section {* Support and Freshness *}
-
-text {*
- The most original aspect of the nominal logic work of Pitts et al is a general
- definition for ``the set of free variables of an object @{text "x"}''. This
- definition is general in the sense that it applies not only to lambda-terms,
- but also to lists, products, sets and even functions. The definition depends
- only on the permutation operation and on the notion of equality defined for
- the type of @{text x}, namely:
-
- @{thm [display,indent=10] supp_def[no_vars, THEN eq_reflection]}
-
- \noindent
- (Note that due to the definition of swapping in \eqref{swapdef}, we do not
- need to explicitly restrict @{text a} and @{text b} to have the same sort.)
- There is also the derived notion for when an atom @{text a} is \emph{fresh}
- for an @{text x}, defined as
-
- @{thm [display,indent=10] fresh_def[no_vars]}
-
- \noindent
- A striking consequence of these definitions is that we can prove
- without knowing anything about the structure of @{term x} that
- swapping two fresh atoms, say @{text a} and @{text b}, leave
- @{text x} unchanged. For the proof we use the following lemma
- about swappings applied to an @{text x}:
-
- \begin{lemma}\label{swaptriple}
- Assuming @{text x} is of permutation type, and @{text a}, @{text b} and @{text c}
- have the same sort, then @{thm (prem 3) swap_rel_trans[no_vars]} and
- @{thm (prem 4) swap_rel_trans[no_vars]} imply @{thm (concl) swap_rel_trans[no_vars]}.
- \end{lemma}
-
- \begin{proof}
- The cases where @{text "a = c"} and @{text "b = c"} are immediate.
- For the remaining case it is, given our assumptions, easy to calculate
- that the permutations
-
- @{thm [display,indent=10] (concl) swap_triple[no_vars]}
-
- \noindent
- are equal. The lemma is then by application of the second permutation
- property shown in \eqref{newpermprops}.\hfill\qed
- \end{proof}
-
- \begin{theorem}\label{swapfreshfresh}
- Let @{text x} be of permutation type.
- @{thm [mode=IfThen] swap_fresh_fresh[no_vars]}
- \end{theorem}
-
- \begin{proof}
- If @{text a} and @{text b} have different sort, then the swapping is the identity.
- If they have the same sort, we know by definition of support that both
- @{term "finite {c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x}"} and @{term "finite {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x}"}
- hold. So the union of these sets is finite too, and we know by Proposition~\ref{choosefresh}
- that there is an atom @{term c}, with the same sort as @{term a} and @{term b},
- that satisfies \mbox{@{term "(a \<rightleftharpoons> c) \<bullet> x = x"}} and @{term "(b \<rightleftharpoons> c) \<bullet> x = x"}.
- Now the theorem follows from Lemma~\ref{swaptriple}.\hfill\qed
- \end{proof}
-
- \noindent
- Two important properties that need to be established for later calculations is
- that @{text "supp"} and freshness are equivariant. For this we first show that:
-
- \begin{lemma}\label{half}
- If @{term x} is a permutation type, then @{thm (lhs) fresh_permute_iff[where p="\<pi>",no_vars]}
- if and only if @{thm (rhs) fresh_permute_iff[where p="\<pi>",no_vars]}.
- \end{lemma}
-
- \begin{proof}
- \begin{isabelle}
- \begin{tabular}[t]{c@ {\hspace{2mm}}l@ {\hspace{5mm}}l}
- & \multicolumn{2}{@ {}l}{@{thm (lhs) fresh_permute_iff[where p="\<pi>",no_vars]} @{text "\<equiv>"}
- @{term "finite {b. (\<pi> \<bullet> a \<rightleftharpoons> b) \<bullet> \<pi> \<bullet> x \<noteq> \<pi> \<bullet> x}"}}\\
- @{text "\<Leftrightarrow>"}
- & @{term "finite {b. (\<pi> \<bullet> a \<rightleftharpoons> \<pi> \<bullet> b) \<bullet> \<pi> \<bullet> x \<noteq> \<pi> \<bullet> x}"}
- & since @{text "\<pi> \<bullet> _"} is bijective\\
- @{text "\<Leftrightarrow>"}
- & @{term "finite {b. \<pi> \<bullet> (a \<rightleftharpoons> b) \<bullet> x \<noteq> \<pi> \<bullet> x}"}
- & by \eqref{permutecompose} and \eqref{swapeqvt}\\
- @{text "\<Leftrightarrow>"}
- & @{term "finite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}"} @{text "\<equiv>"}
- @{thm (rhs) fresh_permute_iff[where p="\<pi>",no_vars]}
- & by \eqref{permuteequ}\\
- \end{tabular}
- \end{isabelle}\hfill\qed
- \end{proof}
-
- \noindent
- Together with the definition of the permutation operation on booleans,
- we can immediately infer equivariance of freshness:
-
- @{thm [display,indent=10] fresh_eqvt[where p="\<pi>",no_vars]}
-
- \noindent
- Now equivariance of @{text "supp"}, namely
-
- @{thm [display,indent=10] supp_eqvt[where p="\<pi>",no_vars]}
-
- \noindent
- is by noting that @{thm supp_conv_fresh[no_vars]} and that freshness and
- the logical connectives are equivariant.
-
- While the abstract properties of support and freshness, particularly
- Theorem~\ref{swapfreshfresh}, are useful for developing Nominal Isabelle,
- one often has to calculate the support of some concrete object. This is
- straightforward for example for booleans, nats, products and lists:
-
- \begin{center}
- \begin{tabular}{@ {}c@ {\hspace{2mm}}c@ {}}
- \begin{tabular}{@ {}r@ {\hspace{2mm}}l}
- @{text "booleans"}: & @{term "supp b = {}"}\\
- @{text "nats"}: & @{term "supp n = {}"}\\
- @{text "products"}: & @{thm supp_Pair[no_vars]}\\
- \end{tabular} &
- \begin{tabular}{r@ {\hspace{2mm}}l@ {}}
- @{text "lists:"} & @{thm supp_Nil[no_vars]}\\
- & @{thm supp_Cons[no_vars]}\\
- \end{tabular}
- \end{tabular}
- \end{center}
-
- \noindent
- But establishing the support of atoms and permutations in our setup here is a bit
- trickier. To do so we will use the following notion about a \emph{supporting set}.
-
- \begin{definition}
- A set @{text S} \emph{supports} @{text x} if for all atoms @{text a} and @{text b}
- not in @{text S} we have @{term "(a \<rightleftharpoons> b) \<bullet> x = x"}.
- \end{definition}
-
- \noindent
- The main motivation for this notion is that we can characterise @{text "supp x"}
- as the smallest finite set that supports @{text "x"}. For this we prove:
-
- \begin{lemma}\label{supports} Let @{text x} be of permutation type.
- \begin{isabelle}
- \begin{tabular}{r@ {\hspace{4mm}}p{10cm}}
- i) & @{thm[mode=IfThen] supp_is_subset[no_vars]}\\
- ii) & @{thm[mode=IfThen] supp_supports[no_vars]}\\
- iii) & @{thm (concl) supp_is_least_supports[no_vars]}
- provided @{thm (prem 1) supp_is_least_supports[no_vars]},
- @{thm (prem 2) supp_is_least_supports[no_vars]}
- and @{text "S"} is the least such set, that means formally,
- for all @{text "S'"}, if @{term "finite S'"} and
- @{term "S' supports x"} then @{text "S \<subseteq> S'"}.
- \end{tabular}
- \end{isabelle}
- \end{lemma}
-
- \begin{proof}
- For @{text "i)"} we derive a contradiction by assuming there is an atom @{text a}
- with @{term "a \<in> supp x"} and @{text "a \<notin> S"}. Using the second fact, the
- assumption that @{term "S supports x"} gives us that @{text S} is a superset of
- @{term "{b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}"}, which is finite by the assumption of @{text S}
- being finite. But this means @{term "a \<notin> supp x"}, contradicting our assumption.
- Property @{text "ii)"} is by a direct application of
- Theorem~\ref{swapfreshfresh}. For the last property, part @{text "i)"} proves
- one ``half'' of the claimed equation. The other ``half'' is by property
- @{text "ii)"} and the fact that @{term "supp x"} is finite by @{text "i)"}.\hfill\qed
- \end{proof}
-
- \noindent
- These are all relatively straightforward proofs adapted from the existing
- nominal logic work. However for establishing the support of atoms and
- permutations we found the following ``optimised'' variant of @{text "iii)"}
- more useful:
-
- \begin{lemma}\label{optimised} Let @{text x} be of permutation type.
- We have that @{thm (concl) finite_supp_unique[no_vars]}
- provided @{thm (prem 1) finite_supp_unique[no_vars]},
- @{thm (prem 2) finite_supp_unique[no_vars]}, and for
- all @{text "a \<in> S"} and all @{text "b \<notin> S"}, with @{text a}
- and @{text b} having the same sort, \mbox{@{term "(a \<rightleftharpoons> b) \<bullet> x \<noteq> x"}}
- \end{lemma}
-
- \begin{proof}
- By Lemma \ref{supports}@{text ".iii)"} we have to show that for every finite
- set @{text S'} that supports @{text x}, \mbox{@{text "S \<subseteq> S'"}} holds. We will
- assume that there is an atom @{text "a"} that is element of @{text S}, but
- not @{text "S'"} and derive a contradiction. Since both @{text S} and
- @{text S'} are finite, we can by Proposition \ref{choosefresh} obtain an atom
- @{text b}, which has the same sort as @{text "a"} and for which we know
- @{text "b \<notin> S"} and @{text "b \<notin> S'"}. Since we assumed @{text "a \<notin> S'"} and
- we have that @{text "S' supports x"}, we have on one hand @{term "(a \<rightleftharpoons> b) \<bullet> x
- = x"}. On the other hand, the fact @{text "a \<in> S"} and @{text "b \<notin> S"} imply
- @{term "(a \<rightleftharpoons> b) \<bullet> x \<noteq> x"} using the assumed implication. This gives us the
- contradiction.\hfill\qed
- \end{proof}
-
- \noindent
- Using this lemma we only have to show the following three proof-obligations
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}r@ {\hspace{4mm}}l}
- i) & @{term "{c} supports c"}\\
- ii) & @{term "finite {c}"}\\
- iii) & @{text "\<forall>a \<in> {c} b \<notin> {c}. sort a = sort b \<longrightarrow> (a b) \<bullet> c \<noteq> c"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- in order to establish that @{thm supp_atom[where a="c", no_vars]} holds. In
- Isabelle/HOL these proof-obligations can be discharged by easy
- simplifications. Similar proof-obligations arise for the support of
- permutations, which is
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{thm supp_perm[where p="\<pi>", no_vars]}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- The only proof-obligation that is
- interesting is the one where we have to show that
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{text "If \<pi> \<bullet> a \<noteq> a, \<pi> \<bullet> b = b and sort a = sort b, then (a b) \<bullet> \<pi> \<noteq> \<pi>"}.
- \end{tabular}
- \end{isabelle}
-
- \noindent
- For this we observe that
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}rcl}
- @{thm (lhs) perm_swap_eq[where p="\<pi>", no_vars]} &
- if and only if &
- @{thm (rhs) perm_swap_eq[where p="\<pi>", no_vars]}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- holds by a simple calculation using the group properties of permutations.
- The proof-obligation can then be discharged by analysing the inequality
- between the permutations @{term "(\<pi> \<bullet> a \<rightleftharpoons> b)"} and @{term "(a \<rightleftharpoons> b)"}.
-
- The main point about support is that whenever an object @{text x} has finite
- support, then Proposition~\ref{choosefresh} allows us to choose for @{text x} a
- fresh atom with arbitrary sort. This is an important operation in Nominal
- Isabelle in situations where, for example, a bound variable needs to be
- renamed. To allow such a choice, we only have to assume \emph{one} premise
- of the form @{text "finite (supp x)"}
- for each @{text x}. Compare that with the sequence of premises in our earlier
- version of Nominal Isabelle (see~\eqref{fssequence}). For more convenience we
- can define a type class for types where every element has finite support, and
- prove that the types @{term "atom"}, @{term "perm"}, lists, products and
- booleans are instances of this type class. Then \emph{no} premise is needed,
- as the type system of Isabelle/HOL can figure out automatically when an object
- is finitely supported.
-
- Unfortunately, this does not work for sets or Isabelle/HOL's function type.
- There are functions and sets definable in Isabelle/HOL for which the finite
- support property does not hold. A simple example of a function with
- infinite support is the function that returns the natural number of an atom
-
- @{text [display, indent=10] "nat_of (Atom s i) \<equiv> i"}
-
- \noindent
- This function's support is the set of \emph{all} atoms. To establish this we show @{term "\<not> a \<sharp> nat_of"}.
- This is equivalent to assuming the set @{term "{b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of}"} is finite
- and deriving a contradiction. From the assumption we also know that
- @{term "{a} \<union> {b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of}"} is finite. Then we can use
- Proposition~\ref{choosefresh} to choose an atom @{text c} such that
- @{term "c \<noteq> a"}, @{term "sort_of c = sort_of a"} and @{term "(a \<rightleftharpoons> c) \<bullet> nat_of = nat_of"}.
- Now we can reason as follows:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}[b]{@ {}rcl@ {\hspace{5mm}}l}
- @{text "nat_of a"} & @{text "="} & @{text "(a \<rightleftharpoons> c) \<bullet> (nat_of a)"} & by def.~of permutations on nats\\
- & @{text "="} & @{term "((a \<rightleftharpoons> c) \<bullet> nat_of) ((a \<rightleftharpoons> c) \<bullet> a)"} & by \eqref{permutefunapp}\\
- & @{text "="} & @{term "nat_of c"} & by assumptions on @{text c}\\
- \end{tabular}
- \end{isabelle}
-
-
- \noindent
- But this means we have that @{term "nat_of a = nat_of c"} and @{term "sort_of a = sort_of c"}.
- This implies that atoms @{term a} and @{term c} must be equal, which clashes with our
- assumption @{term "c \<noteq> a"} about how we chose @{text c}.
- Cheney \cite{Cheney06} gives similar examples for constructions that have infinite support.
-*}
-
-section {* Concrete Atom Types *}
-
-text {*
-
- So far, we have presented a system that uses only a single multi-sorted atom
- type. This design gives us the flexibility to define operations and prove
- theorems that are generic with respect to atom sorts. For example, as
- illustrated above the @{term supp} function returns a set that includes the
- free atoms of \emph{all} sorts together; the flexibility offered by the new
- atom type makes this possible.
-
- However, the single multi-sorted atom type does not make an ideal interface
- for end-users of Nominal Isabelle. If sorts are not distinguished by
- Isabelle's type system, users must reason about atom sorts manually. That
- means subgoals involving sorts must be discharged explicitly within proof
- scripts, instead of being inferred by Isabelle/HOL's type checker. In other
- cases, lemmas might require additional side conditions about sorts to be true.
- For example, swapping @{text a} and @{text b} in the pair \mbox{@{term "(a,
- b)"}} will only produce the expected result if we state the lemma in
- Isabelle/HOL as:
-*}
-
- lemma
- fixes a b :: "atom"
- assumes asm: "sort a = sort b"
- shows "(a \<rightleftharpoons> b) \<bullet> (a, b) = (b, a)"
- using asm by simp
-
-text {*
- \noindent
- Fortunately, it is possible to regain most of the type-checking automation
- that is lost by moving to a single atom type. We accomplish this by defining
- \emph{subtypes} of the generic atom type that only include atoms of a single
- specific sort. We call such subtypes \emph{concrete atom types}.
-
- The following Isabelle/HOL command defines a concrete atom type called
- \emph{name}, which consists of atoms whose sort equals the string @{term
- "''name''"}.
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \isacommand{typedef}\ \ @{typ name} = @{term "{a. sort\<iota> a = ''name''}"}
- \end{isabelle}
-
- \noindent
- This command automatically generates injective functions that map from the
- concrete atom type into the generic atom type and back, called
- representation and abstraction functions, respectively. We will write these
- functions as follows:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l@ {\hspace{10mm}}l}
- @{text "\<lfloor>_\<rfloor> :: name \<Rightarrow> atom"} &
- @{text "\<lceil>_\<rceil> :: atom \<Rightarrow> name"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- With the definition @{thm permute_name_def [where p="\<pi>", THEN
- eq_reflection, no_vars]}, it is straightforward to verify that the type
- @{typ name} is a permutation type.
-
- In order to reason uniformly about arbitrary concrete atom types, we define a
- type class that characterises type @{typ name} and other similarly-defined
- types. The definition of the concrete atom type class is as follows: First,
- every concrete atom type must be a permutation type. In addition, the class
- defines an overloaded function that maps from the concrete type into the
- generic atom type, which we will write @{text "|_|"}. For each class
- instance, this function must be injective and equivariant, and its outputs
- must all have the same sort, that is
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- i) \hspace{1mm}if @{thm (lhs) atom_eq_iff [no_vars]} then @{thm (rhs) atom_eq_iff [no_vars]}\hspace{4mm}
- ii) \hspace{1mm}@{thm atom_eqvt[where p="\<pi>", no_vars]}\hspace{4mm}
- iii) \hspace{1mm}@{thm sort_of_atom_eq [no_vars]}
- \hfill\numbered{atomprops}
- \end{isabelle}
-
- \noindent
- With the definition @{thm atom_name_def [THEN eq_reflection, no_vars]} we can
- show that @{typ name} satisfies all the above requirements of a concrete atom
- type.
-
- The whole point of defining the concrete atom type class was to let users
- avoid explicit reasoning about sorts. This benefit is realised by defining a
- special swapping operation of type @{text "\<alpha> \<Rightarrow> \<alpha>
- \<Rightarrow> perm"}, where @{text "\<alpha>"} is a concrete atom type:
-
- @{thm [display,indent=10] flip_def [THEN eq_reflection, no_vars]}
-
- \noindent
- As a consequence of its type, the @{text "\<leftrightarrow>"}-swapping
- operation works just like the generic swapping operation, but it does not
- require any sort-checking side conditions---the sort-correctness is ensured by
- the types! For @{text "\<leftrightarrow>"} we can establish the following
- simplification rule:
-
- @{thm [display,indent=10] permute_flip_at[no_vars]}
-
- \noindent
- If we now want to swap the \emph{concrete} atoms @{text a} and @{text b}
- in the pair @{term "(a, b)"} we can establish the lemma as follows:
-*}
-
- lemma
- fixes a b :: "name"
- shows "(a \<leftrightarrow> b) \<bullet> (a, b) = (b, a)"
- by simp
-
-text {*
- \noindent
- There is no need to state an explicit premise involving sorts.
-
- We can automate the process of creating concrete atom types, so that users
- can define a new one simply by issuing the command
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- \isacommand{atom\_decl}~~@{text "name"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- This command can be implemented using less than 100 lines of custom ML-code.
- In comparison, the old version of Nominal Isabelle included more than 1000
- lines of ML-code for creating concrete atom types, and for defining various
- type classes and instantiating generic lemmas for them. In addition to
- simplifying the ML-code, the setup here also offers user-visible improvements:
- Now concrete atoms can be declared at any point of a formalisation, and
- theories that separately declare different atom types can be merged
- together---it is no longer required to collect all atom declarations in one
- place.
-*}
-
-
-section {* Multi-Sorted Concrete Atoms *}
-
-(*<*)
-datatype ty = TVar string | Fun ty ty ("_ \<rightarrow> _")
-(*>*)
-
-text {*
- The formalisation presented so far allows us to streamline proofs and reduce
- the amount of custom ML-code in the existing implementation of Nominal
- Isabelle. In this section we describe a mechanism that extends the
- capabilities of Nominal Isabelle. This mechanism is about variables with
- additional information, for example typing constraints.
- While we leave a detailed treatment of binders and binding of variables for a
- later paper, we will have a look here at how such variables can be
- represented by concrete atoms.
-
- In the previous section we considered concrete atoms that can be used in
- simple binders like \emph{@{text "\<lambda>x. x"}}. Such concrete atoms do
- not carry any information beyond their identities---comparing for equality
- is really the only way to analyse ordinary concrete atoms.
- However, in ``Church-style'' lambda-terms \cite{Church40} and in the terms
- underlying HOL-systems \cite{PittsHOL4} binders and bound variables have a
- more complicated structure. For example in the ``Church-style'' lambda-term
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{text "\<lambda>x\<^isub>\<alpha>. x\<^isub>\<alpha> x\<^isub>\<beta>"}
- \end{tabular}\hfill\numbered{church}
- \end{isabelle}
-
- \noindent
- both variables and binders include typing information indicated by @{text \<alpha>}
- and @{text \<beta>}. In this setting, we treat @{text "x\<^isub>\<alpha>"} and @{text
- "x\<^isub>\<beta>"} as distinct variables (assuming @{term "\<alpha>\<noteq>\<beta>"}) so that the
- variable @{text "x\<^isub>\<alpha>"} is bound by the lambda-abstraction, but not
- @{text "x\<^isub>\<beta>"}.
-
- To illustrate how we can deal with this phenomenon, let us represent object
- types like @{text \<alpha>} and @{text \<beta>} by the datatype
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- \isacommand{datatype}~~@{text "ty = TVar string | ty \<rightarrow> ty"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- If we attempt to encode a variable naively as a pair of a @{text name} and a @{text ty}, we have the
- problem that a swapping, say @{term "(x \<leftrightarrow> y)"}, applied to the pair @{text "((x, \<alpha>), (x, \<beta>))"}
- will always permute \emph{both} occurrences of @{text x}, even if the types
- @{text "\<alpha>"} and @{text "\<beta>"} are different. This is unwanted, because it will
- eventually mean that both occurrences of @{text x} will become bound by a
- corresponding binder.
-
- Another attempt might be to define variables as an instance of the concrete
- atom type class, where a @{text ty} is somehow encoded within each variable.
- Remember we defined atoms as the datatype:
-*}
-
- datatype atom\<iota>\<iota> = Atom\<iota>\<iota> string nat
-
-text {*
- \noindent
- Considering our method of defining concrete atom types, the usage of a string
- for the sort of atoms seems a natural choice. However, none of the results so
- far depend on this choice and we are free to change it.
- One possibility is to encode types or any other information by making the sort
- argument parametric as follows:
-*}
-
- datatype 'a \<iota>atom\<iota>\<iota>\<iota> = \<iota>Atom\<iota>\<iota> 'a nat
-
-text {*
- \noindent
- The problem with this possibility is that we are then back in the old
- situation where our permutation operation is parametric in two types and
- this would require to work around Isabelle/HOL's restriction on type
- classes. Fortunately, encoding the types in a separate parameter is not
- necessary for what we want to achieve, as we only have to know when two
- types are equal or not. The solution is to use a different sort for each
- object type. Then we can use the fact that permutations respect \emph{sorts} to
- ensure that permutations also respect \emph{object types}. In order to do
- this, we must define an injective function @{text "sort_ty"} mapping from
- object types to sorts. For defining functions like @{text "sort_ty"}, it is
- more convenient to use a tree datatype for sorts. Therefore we define
-*}
-
- datatype sort = Sort string "(sort list)"
- datatype atom\<iota>\<iota>\<iota>\<iota> = Atom\<iota>\<iota>\<iota>\<iota> sort nat
-
-text {*
- \noindent
- With this definition,
- the sorts we considered so far can be encoded just as \mbox{@{text "Sort s []"}}.
- The point, however, is that we can now define the function @{text sort_ty} simply as
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- @{text "sort_ty (TVar s) = Sort ''TVar'' [Sort s []]"}\\
- @{text "sort_ty (\<tau>\<^isub>1 \<rightarrow> \<tau>\<^isub>2) = Sort ''Fun'' [sort_ty \<tau>\<^isub>1, sort_ty \<tau>\<^isub>2]"}
- \end{tabular}\hfill\numbered{sortty}
- \end{isabelle}
-
- \noindent
- which can easily be shown to be injective.
-
- Having settled on what the sorts should be for ``Church-like'' atoms, we have to
- give a subtype definition for concrete atoms. Previously we identified a subtype consisting
- of atoms of only one specified sort. This must be generalised to all sorts the
- function @{text "sort_ty"} might produce, i.e.~the
- range of @{text "sort_ty"}. Therefore we define
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \isacommand{typedef}\ \ @{text var} = @{term "{a. sort a : range sort_ty}"}
- \end{isabelle}
-
- \noindent
- This command gives us again injective representation and abstraction
- functions. We will write them also as \mbox{@{text "\<lfloor>_\<rfloor> :: var \<Rightarrow> atom"}} and
- @{text "\<lceil>_\<rceil> :: atom \<Rightarrow> var"}, respectively.
-
- We can define the permutation operation for @{text var} as @{thm
- permute_var_def[where p="\<pi>", THEN eq_reflection, no_vars]} and the
- injective function to type @{typ atom} as @{thm atom_var_def[THEN
- eq_reflection, no_vars]}. Finally, we can define a constructor function that
- makes a @{text var} from a variable name and an object type:
-
- @{thm [display,indent=10] Var_def[where t="\<alpha>", THEN eq_reflection, no_vars]}
-
- \noindent
- With these definitions we can verify all the properties for concrete atom
- types except Property \ref{atomprops}@{text ".iii)"}, which requires every
- atom to have the same sort. This last property is clearly not true for type
- @{text "var"}.
- This fact is slightly unfortunate since this
- property allowed us to use the type-checker in order to shield the user from
- all sort-constraints. But this failure is expected here, because we cannot
- burden the type-system of Isabelle/HOL with the task of deciding when two
- object types are equal. This means we sometimes need to explicitly state sort
- constraints or explicitly discharge them, but as we will see in the lemma
- below this seems a natural price to pay in these circumstances.
-
- To sum up this section, the encoding of type-information into atoms allows us
- to form the swapping @{term "(Var x \<alpha> \<leftrightarrow> Var y \<alpha>)"} and to prove the following
- lemma
-*}
-
- lemma
- assumes asm: "\<alpha> \<noteq> \<beta>"
- shows "(Var x \<alpha> \<leftrightarrow> Var y \<alpha>) \<bullet> (Var x \<alpha>, Var x \<beta>) = (Var y \<alpha>, Var x \<beta>)"
- using asm by simp
-
-text {*
- \noindent
- As we expect, the atom @{term "Var x \<beta>"} is left unchanged by the
- swapping. With this we can faithfully represent bindings in languages
- involving ``Church-style'' terms and bindings as shown in \eqref{church}. We
- expect that the creation of such atoms can be easily automated so that the
- user just needs to specify \isacommand{atom\_decl}~~@{text "var (ty)"}
- where the argument, or arguments, are datatypes for which we can automatically
- define an injective function like @{text "sort_ty"} (see \eqref{sortty}).
- Our hope is that with this approach Benzmueller and Paulson can make
- headway with formalising their results
- about simple type theory \cite{PaulsonBenzmueller}.
- Because of its limitations, they did not attempt this with the old version
- of Nominal Isabelle. We also hope we can make progress with formalisations of
- HOL-based languages.
-*}
-
-
-section {* Conclusion *}
-
-text {*
- This proof pearl describes a new formalisation of the nominal logic work by
- Pitts et al. With the definitions we presented here, the formal reasoning blends
- smoothly with the infrastructure of the Isabelle/HOL theorem prover.
- Therefore the formalisation will be the underlying theory for a
- new version of Nominal Isabelle.
-
- The main difference of this paper with respect to existing work on Nominal
- Isabelle is the representation of atoms and permutations. First, we used a
- single type for sorted atoms. This design choice means for a term @{term t},
- say, that its support is completely characterised by @{term "supp t"}, even
- if the term contains different kinds of atoms. Also, whenever we have to
- generalise an induction so that a property @{text P} is not just established
- for all @{text t}, but for all @{text t} \emph{and} under all permutations
- @{text \<pi>}, then we only have to state @{term "\<forall>\<pi>. P (\<pi> \<bullet> t)"}. The reason is
- that permutations can now consist of multiple swapping each of which can
- swap different kinds of atoms. This simplifies considerably the reasoning
- involved in building Nominal Isabelle.
-
- Second, we represented permutations as functions so that the associated
- permutation operation has only a single type parameter. This is very convenient
- because the abstract reasoning about permutations fits cleanly
- with Isabelle/HOL's type classes. No custom ML-code is required to work
- around rough edges. Moreover, by establishing that our permutations-as-functions
- representation satisfy the group properties, we were able to use extensively
- Isabelle/HOL's reasoning infrastructure for groups. This often reduced proofs
- to simple calculations over @{text "+"}, @{text "-"} and @{text "0"}.
- An interesting point is that we defined the swapping operation so that a
- swapping of two atoms with different sorts is \emph{not} excluded, like
- in our older work on Nominal Isabelle, but there is no ``effect'' of such
- a swapping (it is defined as the identity). This is a crucial insight
- in order to make the approach based on a single type of sorted atoms to work.
- But of course it is analogous to the well-known trick of defining division by
- zero to return zero.
-
- We noticed only one disadvantage of the permutations-as-functions: Over
- lists we can easily perform inductions. For permutations made up from
- functions, we have to manually derive an appropriate induction principle. We
- can establish such a principle, but we have no real experience yet whether ours
- is the most useful principle: such an induction principle was not needed in
- any of the reasoning we ported from the old Nominal Isabelle, except
- when showing that if @{term "\<forall>a \<in> supp x. a \<sharp> p"} implies @{term "p \<bullet> x = x"}.
-
- Finally, our implementation of sorted atoms turned out powerful enough to
- use it for representing variables that carry on additional information, for
- example typing annotations. This information is encoded into the sorts. With
- this we can represent conveniently binding in ``Church-style'' lambda-terms
- and HOL-based languages. While dealing with such additional information in
- dependent type-theories, such as LF or Coq, is straightforward, we are not
- aware of any other approach in a non-dependent HOL-setting that can deal
- conveniently with such binders.
-
- The formalisation presented here will eventually become part of the Isabelle
- distribution, but for the moment it can be downloaded from the
- Mercurial repository linked at
- \href{http://isabelle.in.tum.de/nominal/download}
- {http://isabelle.in.tum.de/nominal/download}.\smallskip
-
- \noindent
- {\bf Acknowledgements:} We are very grateful to Jesper Bengtson, Stefan
- Berghofer and Cezary Kaliszyk for their comments on earlier versions
- of this paper. We are also grateful to the anonymous referee who helped us to
- put the work into the right context.
-*}
-
-
-(*<*)
-end
-(*>*)
\ No newline at end of file
--- a/Pearl/ROOT.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,5 +0,0 @@
-no_document use_thys ["../Nominal/Nominal2_Base",
- "../Nominal/Atoms",
- "~~/src/HOL/Library/LaTeXsugar"];
-
-use_thys ["Paper"];
\ No newline at end of file
--- a/Pearl/document/llncs.cls Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1189 +0,0 @@
-% LLNCS DOCUMENT CLASS -- version 2.13 (28-Jan-2002)
-% Springer Verlag LaTeX2e support for Lecture Notes in Computer Science
-%
-%%
-%% \CharacterTable
-%% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
-%% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
-%% Digits \0\1\2\3\4\5\6\7\8\9
-%% Exclamation \! Double quote \" Hash (number) \#
-%% Dollar \$ Percent \% Ampersand \&
-%% Acute accent \' Left paren \( Right paren \)
-%% Asterisk \* Plus \+ Comma \,
-%% Minus \- Point \. Solidus \/
-%% Colon \: Semicolon \; Less than \<
-%% Equals \= Greater than \> Question mark \?
-%% Commercial at \@ Left bracket \[ Backslash \\
-%% Right bracket \] Circumflex \^ Underscore \_
-%% Grave accent \` Left brace \{ Vertical bar \|
-%% Right brace \} Tilde \~}
-%%
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-\ProvidesClass{llncs}[2002/01/28 v2.13
-^^J LaTeX document class for Lecture Notes in Computer Science]
-% Options
-\let\if@envcntreset\iffalse
-\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue}
-\DeclareOption{citeauthoryear}{\let\citeauthoryear=Y}
-\DeclareOption{oribibl}{\let\oribibl=Y}
-\let\if@custvec\iftrue
-\DeclareOption{orivec}{\let\if@custvec\iffalse}
-\let\if@envcntsame\iffalse
-\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue}
-\let\if@envcntsect\iffalse
-\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue}
-\let\if@runhead\iffalse
-\DeclareOption{runningheads}{\let\if@runhead\iftrue}
-
-\let\if@openbib\iffalse
-\DeclareOption{openbib}{\let\if@openbib\iftrue}
-
-% languages
-\let\switcht@@therlang\relax
-\def\ds@deutsch{\def\switcht@@therlang{\switcht@deutsch}}
-\def\ds@francais{\def\switcht@@therlang{\switcht@francais}}
-
-\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}}
-
-\ProcessOptions
-
-\LoadClass[twoside]{article}
-\RequirePackage{multicol} % needed for the list of participants, index
-
-\setlength{\textwidth}{12.2cm}
-\setlength{\textheight}{19.3cm}
-\renewcommand\@pnumwidth{2em}
-\renewcommand\@tocrmarg{3.5em}
-%
-\def\@dottedtocline#1#2#3#4#5{%
- \ifnum #1>\c@tocdepth \else
- \vskip \z@ \@plus.2\p@
- {\leftskip #2\relax \rightskip \@tocrmarg \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \parindent #2\relax\@afterindenttrue
- \interlinepenalty\@M
- \leavevmode
- \@tempdima #3\relax
- \advance\leftskip \@tempdima \null\nobreak\hskip -\leftskip
- {#4}\nobreak
- \leaders\hbox{$\m@th
- \mkern \@dotsep mu\hbox{.}\mkern \@dotsep
- mu$}\hfill
- \nobreak
- \hb@xt@\@pnumwidth{\hfil\normalfont \normalcolor #5}%
- \par}%
- \fi}
-%
-\def\switcht@albion{%
-\def\abstractname{Abstract.}
-\def\ackname{Acknowledgement.}
-\def\andname{and}
-\def\lastandname{\unskip, and}
-\def\appendixname{Appendix}
-\def\chaptername{Chapter}
-\def\claimname{Claim}
-\def\conjecturename{Conjecture}
-\def\contentsname{Table of Contents}
-\def\corollaryname{Corollary}
-\def\definitionname{Definition}
-\def\examplename{Example}
-\def\exercisename{Exercise}
-\def\figurename{Fig.}
-\def\keywordname{{\bf Key words:}}
-\def\indexname{Index}
-\def\lemmaname{Lemma}
-\def\contriblistname{List of Contributors}
-\def\listfigurename{List of Figures}
-\def\listtablename{List of Tables}
-\def\mailname{{\it Correspondence to\/}:}
-\def\noteaddname{Note added in proof}
-\def\notename{Note}
-\def\partname{Part}
-\def\problemname{Problem}
-\def\proofname{Proof}
-\def\propertyname{Property}
-\def\propositionname{Proposition}
-\def\questionname{Question}
-\def\remarkname{Remark}
-\def\seename{see}
-\def\solutionname{Solution}
-\def\subclassname{{\it Subject Classifications\/}:}
-\def\tablename{Table}
-\def\theoremname{Theorem}}
-\switcht@albion
-% Names of theorem like environments are already defined
-% but must be translated if another language is chosen
-%
-% French section
-\def\switcht@francais{%\typeout{On parle francais.}%
- \def\abstractname{R\'esum\'e.}%
- \def\ackname{Remerciements.}%
- \def\andname{et}%
- \def\lastandname{ et}%
- \def\appendixname{Appendice}
- \def\chaptername{Chapitre}%
- \def\claimname{Pr\'etention}%
- \def\conjecturename{Hypoth\`ese}%
- \def\contentsname{Table des mati\`eres}%
- \def\corollaryname{Corollaire}%
- \def\definitionname{D\'efinition}%
- \def\examplename{Exemple}%
- \def\exercisename{Exercice}%
- \def\figurename{Fig.}%
- \def\keywordname{{\bf Mots-cl\'e:}}
- \def\indexname{Index}
- \def\lemmaname{Lemme}%
- \def\contriblistname{Liste des contributeurs}
- \def\listfigurename{Liste des figures}%
- \def\listtablename{Liste des tables}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}%
- \def\notename{Remarque}%
- \def\partname{Partie}%
- \def\problemname{Probl\`eme}%
- \def\proofname{Preuve}%
- \def\propertyname{Caract\'eristique}%
-%\def\propositionname{Proposition}%
- \def\questionname{Question}%
- \def\remarkname{Remarque}%
- \def\seename{voir}
- \def\solutionname{Solution}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tableau}%
- \def\theoremname{Th\'eor\`eme}%
-}
-%
-% German section
-\def\switcht@deutsch{%\typeout{Man spricht deutsch.}%
- \def\abstractname{Zusammenfassung.}%
- \def\ackname{Danksagung.}%
- \def\andname{und}%
- \def\lastandname{ und}%
- \def\appendixname{Anhang}%
- \def\chaptername{Kapitel}%
- \def\claimname{Behauptung}%
- \def\conjecturename{Hypothese}%
- \def\contentsname{Inhaltsverzeichnis}%
- \def\corollaryname{Korollar}%
-%\def\definitionname{Definition}%
- \def\examplename{Beispiel}%
- \def\exercisename{\"Ubung}%
- \def\figurename{Abb.}%
- \def\keywordname{{\bf Schl\"usselw\"orter:}}
- \def\indexname{Index}
-%\def\lemmaname{Lemma}%
- \def\contriblistname{Mitarbeiter}
- \def\listfigurename{Abbildungsverzeichnis}%
- \def\listtablename{Tabellenverzeichnis}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Nachtrag}%
- \def\notename{Anmerkung}%
- \def\partname{Teil}%
-%\def\problemname{Problem}%
- \def\proofname{Beweis}%
- \def\propertyname{Eigenschaft}%
-%\def\propositionname{Proposition}%
- \def\questionname{Frage}%
- \def\remarkname{Anmerkung}%
- \def\seename{siehe}
- \def\solutionname{L\"osung}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tabelle}%
-%\def\theoremname{Theorem}%
-}
-
-% Ragged bottom for the actual page
-\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil
-\global\let\@textbottom\relax}}
-
-\renewcommand\small{%
- \@setfontsize\small\@ixpt{11}%
- \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
- \abovedisplayshortskip \z@ \@plus2\p@
- \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
- \def\@listi{\leftmargin\leftmargini
- \parsep 0\p@ \@plus1\p@ \@minus\p@
- \topsep 8\p@ \@plus2\p@ \@minus4\p@
- \itemsep0\p@}%
- \belowdisplayskip \abovedisplayskip
-}
-
-\frenchspacing
-\widowpenalty=10000
-\clubpenalty=10000
-
-\setlength\oddsidemargin {63\p@}
-\setlength\evensidemargin {63\p@}
-\setlength\marginparwidth {90\p@}
-
-\setlength\headsep {16\p@}
-
-\setlength\footnotesep{7.7\p@}
-\setlength\textfloatsep{8mm\@plus 2\p@ \@minus 4\p@}
-\setlength\intextsep {8mm\@plus 2\p@ \@minus 2\p@}
-
-\setcounter{secnumdepth}{2}
-
-\newcounter {chapter}
-\renewcommand\thechapter {\@arabic\c@chapter}
-
-\newif\if@mainmatter \@mainmattertrue
-\newcommand\frontmatter{\cleardoublepage
- \@mainmatterfalse\pagenumbering{Roman}}
-\newcommand\mainmatter{\cleardoublepage
- \@mainmattertrue\pagenumbering{arabic}}
-\newcommand\backmatter{\if@openright\cleardoublepage\else\clearpage\fi
- \@mainmatterfalse}
-
-\renewcommand\part{\cleardoublepage
- \thispagestyle{empty}%
- \if@twocolumn
- \onecolumn
- \@tempswatrue
- \else
- \@tempswafalse
- \fi
- \null\vfil
- \secdef\@part\@spart}
-
-\def\@part[#1]#2{%
- \ifnum \c@secnumdepth >-2\relax
- \refstepcounter{part}%
- \addcontentsline{toc}{part}{\thepart\hspace{1em}#1}%
- \else
- \addcontentsline{toc}{part}{#1}%
- \fi
- \markboth{}{}%
- {\centering
- \interlinepenalty \@M
- \normalfont
- \ifnum \c@secnumdepth >-2\relax
- \huge\bfseries \partname~\thepart
- \par
- \vskip 20\p@
- \fi
- \Huge \bfseries #2\par}%
- \@endpart}
-\def\@spart#1{%
- {\centering
- \interlinepenalty \@M
- \normalfont
- \Huge \bfseries #1\par}%
- \@endpart}
-\def\@endpart{\vfil\newpage
- \if@twoside
- \null
- \thispagestyle{empty}%
- \newpage
- \fi
- \if@tempswa
- \twocolumn
- \fi}
-
-\newcommand\chapter{\clearpage
- \thispagestyle{empty}%
- \global\@topnum\z@
- \@afterindentfalse
- \secdef\@chapter\@schapter}
-\def\@chapter[#1]#2{\ifnum \c@secnumdepth >\m@ne
- \if@mainmatter
- \refstepcounter{chapter}%
- \typeout{\@chapapp\space\thechapter.}%
- \addcontentsline{toc}{chapter}%
- {\protect\numberline{\thechapter}#1}%
- \else
- \addcontentsline{toc}{chapter}{#1}%
- \fi
- \else
- \addcontentsline{toc}{chapter}{#1}%
- \fi
- \chaptermark{#1}%
- \addtocontents{lof}{\protect\addvspace{10\p@}}%
- \addtocontents{lot}{\protect\addvspace{10\p@}}%
- \if@twocolumn
- \@topnewpage[\@makechapterhead{#2}]%
- \else
- \@makechapterhead{#2}%
- \@afterheading
- \fi}
-\def\@makechapterhead#1{%
-% \vspace*{50\p@}%
- {\centering
- \ifnum \c@secnumdepth >\m@ne
- \if@mainmatter
- \large\bfseries \@chapapp{} \thechapter
- \par\nobreak
- \vskip 20\p@
- \fi
- \fi
- \interlinepenalty\@M
- \Large \bfseries #1\par\nobreak
- \vskip 40\p@
- }}
-\def\@schapter#1{\if@twocolumn
- \@topnewpage[\@makeschapterhead{#1}]%
- \else
- \@makeschapterhead{#1}%
- \@afterheading
- \fi}
-\def\@makeschapterhead#1{%
-% \vspace*{50\p@}%
- {\centering
- \normalfont
- \interlinepenalty\@M
- \Large \bfseries #1\par\nobreak
- \vskip 40\p@
- }}
-
-\renewcommand\section{\@startsection{section}{1}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {12\p@ \@plus 4\p@ \@minus 4\p@}%
- {\normalfont\large\bfseries\boldmath
- \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsection{\@startsection{subsection}{2}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {8\p@ \@plus 4\p@ \@minus 4\p@}%
- {\normalfont\normalsize\bfseries\boldmath
- \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {-0.5em \@plus -0.22em \@minus -0.1em}%
- {\normalfont\normalsize\bfseries\boldmath}}
-\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}%
- {-12\p@ \@plus -4\p@ \@minus -4\p@}%
- {-0.5em \@plus -0.22em \@minus -0.1em}%
- {\normalfont\normalsize\itshape}}
-\renewcommand\subparagraph[1]{\typeout{LLNCS warning: You should not use
- \string\subparagraph\space with this class}\vskip0.5cm
-You should not use \verb|\subparagraph| with this class.\vskip0.5cm}
-
-\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00}
-\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01}
-\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02}
-\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03}
-\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04}
-\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05}
-\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06}
-\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07}
-\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08}
-\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09}
-\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A}
-
-\let\footnotesize\small
-
-\if@custvec
-\def\vec#1{\mathchoice{\mbox{\boldmath$\displaystyle#1$}}
-{\mbox{\boldmath$\textstyle#1$}}
-{\mbox{\boldmath$\scriptstyle#1$}}
-{\mbox{\boldmath$\scriptscriptstyle#1$}}}
-\fi
-
-\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}}
-\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil
-\penalty50\hskip1em\null\nobreak\hfil\squareforqed
-\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi}
-
-\def\getsto{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr\gets\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-\gets\cr\to\cr}}}}}
-\def\lid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr<\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr<\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr<\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-<\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\gid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr>\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr>\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\grole{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr
->\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.8pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.3pt}<\cr}}}}}
-\def\bbbr{{\rm I\!R}} %reelle Zahlen
-\def\bbbm{{\rm I\!M}}
-\def\bbbn{{\rm I\!N}} %natuerliche Zahlen
-\def\bbbf{{\rm I\!F}}
-\def\bbbh{{\rm I\!H}}
-\def\bbbk{{\rm I\!K}}
-\def\bbbp{{\rm I\!P}}
-\def\bbbone{{\mathchoice {\rm 1\mskip-4mu l} {\rm 1\mskip-4mu l}
-{\rm 1\mskip-4.5mu l} {\rm 1\mskip-5mu l}}}
-\def\bbbc{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbq{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}}}
-\def\bbbt{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-T$}\hbox{\hbox to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbs{{\mathchoice
-{\setbox0=\hbox{$\displaystyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.5\wd0\vrule height0.45\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.4\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.55\wd0\vrule height0.45\ht0\hss}\box0}}}}
-\def\bbbz{{\mathchoice {\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\scriptstyle Z\kern-0.3em Z$}}
-{\hbox{$\mathsf\scriptscriptstyle Z\kern-0.2em Z$}}}}
-
-\let\ts\,
-
-\setlength\leftmargini {17\p@}
-\setlength\leftmargin {\leftmargini}
-\setlength\leftmarginii {\leftmargini}
-\setlength\leftmarginiii {\leftmargini}
-\setlength\leftmarginiv {\leftmargini}
-\setlength \labelsep {.5em}
-\setlength \labelwidth{\leftmargini}
-\addtolength\labelwidth{-\labelsep}
-
-\def\@listI{\leftmargin\leftmargini
- \parsep 0\p@ \@plus1\p@ \@minus\p@
- \topsep 8\p@ \@plus2\p@ \@minus4\p@
- \itemsep0\p@}
-\let\@listi\@listI
-\@listi
-\def\@listii {\leftmargin\leftmarginii
- \labelwidth\leftmarginii
- \advance\labelwidth-\labelsep
- \topsep 0\p@ \@plus2\p@ \@minus\p@}
-\def\@listiii{\leftmargin\leftmarginiii
- \labelwidth\leftmarginiii
- \advance\labelwidth-\labelsep
- \topsep 0\p@ \@plus\p@\@minus\p@
- \parsep \z@
- \partopsep \p@ \@plus\z@ \@minus\p@}
-
-\renewcommand\labelitemi{\normalfont\bfseries --}
-\renewcommand\labelitemii{$\m@th\bullet$}
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\def\tableofcontents{\chapter*{\contentsname\@mkboth{{\contentsname}}%
- {{\contentsname}}}
- \def\authcount##1{\setcounter{auco}{##1}\setcounter{@auth}{1}}
- \def\lastand{\ifnum\value{auco}=2\relax
- \unskip{} \andname\
- \else
- \unskip \lastandname\
- \fi}%
- \def\and{\stepcounter{@auth}\relax
- \ifnum\value{@auth}=\value{auco}%
- \lastand
- \else
- \unskip,
- \fi}%
- \@starttoc{toc}\if@restonecol\twocolumn\fi}
-
-\def\l@part#1#2{\addpenalty{\@secpenalty}%
- \addvspace{2em plus\p@}% % space above part line
- \begingroup
- \parindent \z@
- \rightskip \z@ plus 5em
- \hrule\vskip5pt
- \large % same size as for a contribution heading
- \bfseries\boldmath % set line in boldface
- \leavevmode % TeX command to enter horizontal mode.
- #1\par
- \vskip5pt
- \hrule
- \vskip1pt
- \nobreak % Never break after part entry
- \endgroup}
-
-\def\@dotsep{2}
-
-\def\hyperhrefextend{\ifx\hyper@anchor\@undefined\else
-{chapter.\thechapter}\fi}
-
-\def\addnumcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{\protect\numberline
- {\thechapter}#3}{\thepage}\hyperhrefextend}}
-\def\addcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{\thepage}\hyperhrefextend}}
-\def\addcontentsmarkwop#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{0}\hyperhrefextend}}
-
-\def\@adcmk[#1]{\ifcase #1 \or
-\def\@gtempa{\addnumcontentsmark}%
- \or \def\@gtempa{\addcontentsmark}%
- \or \def\@gtempa{\addcontentsmarkwop}%
- \fi\@gtempa{toc}{chapter}}
-\def\addtocmark{\@ifnextchar[{\@adcmk}{\@adcmk[3]}}
-
-\def\l@chapter#1#2{\addpenalty{-\@highpenalty}
- \vskip 1.0em plus 1pt \@tempdima 1.5em \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- {\large\bfseries\boldmath#1}\ifx0#2\hfil\null
- \else
- \nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}%
- \fi\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@title#1#2{\addpenalty{-\@highpenalty}
- \addvspace{8pt plus 1pt}
- \@tempdima \z@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- #1\nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@author#1#2{\addpenalty{\@highpenalty}
- \@tempdima=\z@ %15\p@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima %\hskip -\leftskip
- \textit{#1}\par
- \penalty\@highpenalty \endgroup}
-
-%\setcounter{tocdepth}{0}
-\newdimen\tocchpnum
-\newdimen\tocsecnum
-\newdimen\tocsectotal
-\newdimen\tocsubsecnum
-\newdimen\tocsubsectotal
-\newdimen\tocsubsubsecnum
-\newdimen\tocsubsubsectotal
-\newdimen\tocparanum
-\newdimen\tocparatotal
-\newdimen\tocsubparanum
-\tocchpnum=\z@ % no chapter numbers
-\tocsecnum=15\p@ % section 88. plus 2.222pt
-\tocsubsecnum=23\p@ % subsection 88.8 plus 2.222pt
-\tocsubsubsecnum=27\p@ % subsubsection 88.8.8 plus 1.444pt
-\tocparanum=35\p@ % paragraph 88.8.8.8 plus 1.666pt
-\tocsubparanum=43\p@ % subparagraph 88.8.8.8.8 plus 1.888pt
-\def\calctocindent{%
-\tocsectotal=\tocchpnum
-\advance\tocsectotal by\tocsecnum
-\tocsubsectotal=\tocsectotal
-\advance\tocsubsectotal by\tocsubsecnum
-\tocsubsubsectotal=\tocsubsectotal
-\advance\tocsubsubsectotal by\tocsubsubsecnum
-\tocparatotal=\tocsubsubsectotal
-\advance\tocparatotal by\tocparanum}
-\calctocindent
-
-\def\l@section{\@dottedtocline{1}{\tocchpnum}{\tocsecnum}}
-\def\l@subsection{\@dottedtocline{2}{\tocsectotal}{\tocsubsecnum}}
-\def\l@subsubsection{\@dottedtocline{3}{\tocsubsectotal}{\tocsubsubsecnum}}
-\def\l@paragraph{\@dottedtocline{4}{\tocsubsubsectotal}{\tocparanum}}
-\def\l@subparagraph{\@dottedtocline{5}{\tocparatotal}{\tocsubparanum}}
-
-\def\listoffigures{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}}
- \@starttoc{lof}\if@restonecol\twocolumn\fi}
-\def\l@figure{\@dottedtocline{1}{0em}{1.5em}}
-
-\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}}
- \@starttoc{lot}\if@restonecol\twocolumn\fi}
-\let\l@table\l@figure
-
-\renewcommand\listoffigures{%
- \section*{\listfigurename
- \@mkboth{\listfigurename}{\listfigurename}}%
- \@starttoc{lof}%
- }
-
-\renewcommand\listoftables{%
- \section*{\listtablename
- \@mkboth{\listtablename}{\listtablename}}%
- \@starttoc{lot}%
- }
-
-\ifx\oribibl\undefined
-\ifx\citeauthoryear\undefined
-\renewenvironment{thebibliography}[1]
- {\section*{\refname}
- \def\@biblabel##1{##1.}
- \small
- \list{\@biblabel{\@arabic\c@enumiv}}%
- {\settowidth\labelwidth{\@biblabel{#1}}%
- \leftmargin\labelwidth
- \advance\leftmargin\labelsep
- \if@openbib
- \advance\leftmargin\bibindent
- \itemindent -\bibindent
- \listparindent \itemindent
- \parsep \z@
- \fi
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{\@arabic\c@enumiv}}%
- \if@openbib
- \renewcommand\newblock{\par}%
- \else
- \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
- \fi
- \sloppy\clubpenalty4000\widowpenalty4000%
- \sfcode`\.=\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
-\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw
- {\let\protect\noexpand\immediate
- \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
-\newcount\@tempcntc
-\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi
- \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do
- {\@ifundefined
- {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries
- ?}\@warning
- {Citation `\@citeb' on page \thepage \space undefined}}%
- {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}%
- \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne
- \@citea\def\@citea{,}\hbox{\csname b@\@citeb\endcsname}%
- \else
- \advance\@tempcntb\@ne
- \ifnum\@tempcntb=\@tempcntc
- \else\advance\@tempcntb\m@ne\@citeo
- \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}}
-\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else
- \@citea\def\@citea{,\,\hskip\z@skip}%
- \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else
- {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else
- \def\@citea{--}\fi
- \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi}
-\else
-\renewenvironment{thebibliography}[1]
- {\section*{\refname}
- \small
- \list{}%
- {\settowidth\labelwidth{}%
- \leftmargin\parindent
- \itemindent=-\parindent
- \labelsep=\z@
- \if@openbib
- \advance\leftmargin\bibindent
- \itemindent -\bibindent
- \listparindent \itemindent
- \parsep \z@
- \fi
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{}}%
- \if@openbib
- \renewcommand\newblock{\par}%
- \else
- \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
- \fi
- \sloppy\clubpenalty4000\widowpenalty4000%
- \sfcode`\.=\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
- \def\@cite#1{#1}%
- \def\@lbibitem[#1]#2{\item[]\if@filesw
- {\def\protect##1{\string ##1\space}\immediate
- \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
- \fi
-\else
-\@cons\@openbib@code{\noexpand\small}
-\fi
-
-\def\idxquad{\hskip 10\p@}% space that divides entry from number
-
-\def\@idxitem{\par\hangindent 10\p@}
-
-\def\subitem{\par\setbox0=\hbox{--\enspace}% second order
- \noindent\hangindent\wd0\box0}% index entry
-
-\def\subsubitem{\par\setbox0=\hbox{--\,--\enspace}% third
- \noindent\hangindent\wd0\box0}% order index entry
-
-\def\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax}
-
-\renewenvironment{theindex}
- {\@mkboth{\indexname}{\indexname}%
- \thispagestyle{empty}\parindent\z@
- \parskip\z@ \@plus .3\p@\relax
- \let\item\par
- \def\,{\relax\ifmmode\mskip\thinmuskip
- \else\hskip0.2em\ignorespaces\fi}%
- \normalfont\small
- \begin{multicols}{2}[\@makeschapterhead{\indexname}]%
- }
- {\end{multicols}}
-
-\renewcommand\footnoterule{%
- \kern-3\p@
- \hrule\@width 2truecm
- \kern2.6\p@}
- \newdimen\fnindent
- \fnindent1em
-\long\def\@makefntext#1{%
- \parindent \fnindent%
- \leftskip \fnindent%
- \noindent
- \llap{\hb@xt@1em{\hss\@makefnmark\ }}\ignorespaces#1}
-
-\long\def\@makecaption#1#2{%
- \vskip\abovecaptionskip
- \sbox\@tempboxa{{\bfseries #1.} #2}%
- \ifdim \wd\@tempboxa >\hsize
- {\bfseries #1.} #2\par
- \else
- \global \@minipagefalse
- \hb@xt@\hsize{\hfil\box\@tempboxa\hfil}%
- \fi
- \vskip\belowcaptionskip}
-
-\def\fps@figure{htbp}
-\def\fnum@figure{\figurename\thinspace\thefigure}
-\def \@floatboxreset {%
- \reset@font
- \small
- \@setnobreak
- \@setminipage
-}
-\def\fps@table{htbp}
-\def\fnum@table{\tablename~\thetable}
-\renewenvironment{table}
- {\setlength\abovecaptionskip{0\p@}%
- \setlength\belowcaptionskip{10\p@}%
- \@float{table}}
- {\end@float}
-\renewenvironment{table*}
- {\setlength\abovecaptionskip{0\p@}%
- \setlength\belowcaptionskip{10\p@}%
- \@dblfloat{table}}
- {\end@dblfloat}
-
-\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname
- ext@#1\endcsname}{#1}{\protect\numberline{\csname
- the#1\endcsname}{\ignorespaces #2}}\begingroup
- \@parboxrestore
- \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
- \endgroup}
-
-% LaTeX does not provide a command to enter the authors institute
-% addresses. The \institute command is defined here.
-
-\newcounter{@inst}
-\newcounter{@auth}
-\newcounter{auco}
-\newdimen\instindent
-\newbox\authrun
-\newtoks\authorrunning
-\newtoks\tocauthor
-\newbox\titrun
-\newtoks\titlerunning
-\newtoks\toctitle
-
-\def\clearheadinfo{\gdef\@author{No Author Given}%
- \gdef\@title{No Title Given}%
- \gdef\@subtitle{}%
- \gdef\@institute{No Institute Given}%
- \gdef\@thanks{}%
- \global\titlerunning={}\global\authorrunning={}%
- \global\toctitle={}\global\tocauthor={}}
-
-\def\institute#1{\gdef\@institute{#1}}
-
-\def\institutename{\par
- \begingroup
- \parskip=\z@
- \parindent=\z@
- \setcounter{@inst}{1}%
- \def\and{\par\stepcounter{@inst}%
- \noindent$^{\the@inst}$\enspace\ignorespaces}%
- \setbox0=\vbox{\def\thanks##1{}\@institute}%
- \ifnum\c@@inst=1\relax
- \gdef\fnnstart{0}%
- \else
- \xdef\fnnstart{\c@@inst}%
- \setcounter{@inst}{1}%
- \noindent$^{\the@inst}$\enspace
- \fi
- \ignorespaces
- \@institute\par
- \endgroup}
-
-\def\@fnsymbol#1{\ensuremath{\ifcase#1\or\star\or{\star\star}\or
- {\star\star\star}\or \dagger\or \ddagger\or
- \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger
- \or \ddagger\ddagger \else\@ctrerr\fi}}
-
-\def\inst#1{\unskip$^{#1}$}
-\def\fnmsep{\unskip$^,$}
-\def\email#1{{\tt#1}}
-\AtBeginDocument{\@ifundefined{url}{\def\url#1{#1}}{}%
-\@ifpackageloaded{babel}{%
-\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}%
-\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}%
-\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}%
-}{\switcht@@therlang}%
-}
-\def\homedir{\~{ }}
-
-\def\subtitle#1{\gdef\@subtitle{#1}}
-\clearheadinfo
-
-\renewcommand\maketitle{\newpage
- \refstepcounter{chapter}%
- \stepcounter{section}%
- \setcounter{section}{0}%
- \setcounter{subsection}{0}%
- \setcounter{figure}{0}
- \setcounter{table}{0}
- \setcounter{equation}{0}
- \setcounter{footnote}{0}%
- \begingroup
- \parindent=\z@
- \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
- \if@twocolumn
- \ifnum \col@number=\@ne
- \@maketitle
- \else
- \twocolumn[\@maketitle]%
- \fi
- \else
- \newpage
- \global\@topnum\z@ % Prevents figures from going at top of page.
- \@maketitle
- \fi
- \thispagestyle{empty}\@thanks
-%
- \def\\{\unskip\ \ignorespaces}\def\inst##1{\unskip{}}%
- \def\thanks##1{\unskip{}}\def\fnmsep{\unskip}%
- \instindent=\hsize
- \advance\instindent by-\headlineindent
-% \if!\the\toctitle!\addcontentsline{toc}{title}{\@title}\else
-% \addcontentsline{toc}{title}{\the\toctitle}\fi
- \if@runhead
- \if!\the\titlerunning!\else
- \edef\@title{\the\titlerunning}%
- \fi
- \global\setbox\titrun=\hbox{\small\rm\unboldmath\ignorespaces\@title}%
- \ifdim\wd\titrun>\instindent
- \typeout{Title too long for running head. Please supply}%
- \typeout{a shorter form with \string\titlerunning\space prior to
- \string\maketitle}%
- \global\setbox\titrun=\hbox{\small\rm
- Title Suppressed Due to Excessive Length}%
- \fi
- \xdef\@title{\copy\titrun}%
- \fi
-%
- \if!\the\tocauthor!\relax
- {\def\and{\noexpand\protect\noexpand\and}%
- \protected@xdef\toc@uthor{\@author}}%
- \else
- \def\\{\noexpand\protect\noexpand\newline}%
- \protected@xdef\scratch{\the\tocauthor}%
- \protected@xdef\toc@uthor{\scratch}%
- \fi
-% \addcontentsline{toc}{author}{\toc@uthor}%
- \if@runhead
- \if!\the\authorrunning!
- \value{@inst}=\value{@auth}%
- \setcounter{@auth}{1}%
- \else
- \edef\@author{\the\authorrunning}%
- \fi
- \global\setbox\authrun=\hbox{\small\unboldmath\@author\unskip}%
- \ifdim\wd\authrun>\instindent
- \typeout{Names of authors too long for running head. Please supply}%
- \typeout{a shorter form with \string\authorrunning\space prior to
- \string\maketitle}%
- \global\setbox\authrun=\hbox{\small\rm
- Authors Suppressed Due to Excessive Length}%
- \fi
- \xdef\@author{\copy\authrun}%
- \markboth{\@author}{\@title}%
- \fi
- \endgroup
- \setcounter{footnote}{\fnnstart}%
- \clearheadinfo}
-%
-\def\@maketitle{\newpage
- \markboth{}{}%
- \def\lastand{\ifnum\value{@inst}=2\relax
- \unskip{} \andname\
- \else
- \unskip \lastandname\
- \fi}%
- \def\and{\stepcounter{@auth}\relax
- \ifnum\value{@auth}=\value{@inst}%
- \lastand
- \else
- \unskip,
- \fi}%
- \begin{center}%
- \let\newline\\
- {\Large \bfseries\boldmath
- \pretolerance=10000
- \@title \par}\vskip .8cm
-\if!\@subtitle!\else {\large \bfseries\boldmath
- \vskip -.65cm
- \pretolerance=10000
- \@subtitle \par}\vskip .8cm\fi
- \setbox0=\vbox{\setcounter{@auth}{1}\def\and{\stepcounter{@auth}}%
- \def\thanks##1{}\@author}%
- \global\value{@inst}=\value{@auth}%
- \global\value{auco}=\value{@auth}%
- \setcounter{@auth}{1}%
-{\lineskip .5em
-\noindent\ignorespaces
-\@author\vskip.35cm}
- {\small\institutename}
- \end{center}%
- }
-
-% definition of the "\spnewtheorem" command.
-%
-% Usage:
-%
-% \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font}
-% or \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font}
-% or \spnewtheorem*{env_nam}{caption}{cap_font}{body_font}
-%
-% New is "cap_font" and "body_font". It stands for
-% fontdefinition of the caption and the text itself.
-%
-% "\spnewtheorem*" gives a theorem without number.
-%
-% A defined spnewthoerem environment is used as described
-% by Lamport.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\def\@thmcountersep{}
-\def\@thmcounterend{.}
-
-\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}}
-
-% definition of \spnewtheorem with number
-
-\def\@spnthm#1#2{%
- \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}}
-\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}}
-
-\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}\@addtoreset{#1}{#3}%
- \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand
- \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}%
- \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spothm#1[#2]#3#4#5{%
- \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}%
- {\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{the#1}{\@nameuse{the#2}}%
- \expandafter\xdef\csname #1name\endcsname{#3}%
- \global\@namedef{#1}{\@spthm{#2}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}}
-
-\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\refstepcounter{#1}%
-\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}}
-
-\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}%
- \ignorespaces}
-
-\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname
- the#1\endcsname}{#5}{#3}{#4}\ignorespaces}
-
-\def\@spbegintheorem#1#2#3#4{\trivlist
- \item[\hskip\labelsep{#3#1\ #2\@thmcounterend}]#4}
-
-\def\@spopargbegintheorem#1#2#3#4#5{\trivlist
- \item[\hskip\labelsep{#4#1\ #2}]{#4(#3)\@thmcounterend\ }#5}
-
-% definition of \spnewtheorem* without number
-
-\def\@sthm#1#2{\@Ynthm{#1}{#2}}
-
-\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}}
-
-\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces}
-
-\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1}
- {#4}{#2}{#3}\ignorespaces}
-
-\def\@Begintheorem#1#2#3{#3\trivlist
- \item[\hskip\labelsep{#2#1\@thmcounterend}]}
-
-\def\@Opargbegintheorem#1#2#3#4{#4\trivlist
- \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }}
-
-\if@envcntsect
- \def\@thmcountersep{.}
- \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape}
-\else
- \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape}
- \if@envcntreset
- \@addtoreset{theorem}{section}
- \else
- \@addtoreset{theorem}{chapter}
- \fi
-\fi
-
-%definition of divers theorem environments
-\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily}
-\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily}
-\if@envcntsame % alle Umgebungen wie Theorem.
- \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}}
-\else % alle Umgebungen mit eigenem Zaehler
- \if@envcntsect % mit section numeriert
- \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}}
- \else % nicht mit section numeriert
- \if@envcntreset
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{section}}
- \else
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{chapter}}%
- \fi
- \fi
-\fi
-\spn@wtheorem{case}{Case}{\itshape}{\rmfamily}
-\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily}
-\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape}
-\spn@wtheorem{definition}{Definition}{\bfseries}{\itshape}
-\spn@wtheorem{example}{Example}{\itshape}{\rmfamily}
-\spn@wtheorem{exercise}{Exercise}{\itshape}{\rmfamily}
-\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape}
-\spn@wtheorem{note}{Note}{\itshape}{\rmfamily}
-\spn@wtheorem{problem}{Problem}{\itshape}{\rmfamily}
-\spn@wtheorem{property}{Property}{\itshape}{\rmfamily}
-\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape}
-\spn@wtheorem{question}{Question}{\itshape}{\rmfamily}
-\spn@wtheorem{solution}{Solution}{\itshape}{\rmfamily}
-\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily}
-
-\def\@takefromreset#1#2{%
- \def\@tempa{#1}%
- \let\@tempd\@elt
- \def\@elt##1{%
- \def\@tempb{##1}%
- \ifx\@tempa\@tempb\else
- \@addtoreset{##1}{#2}%
- \fi}%
- \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
- \expandafter\def\csname cl@#2\endcsname{}%
- \@tempc
- \let\@elt\@tempd}
-
-\def\theopargself{\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5}
- \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
- \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }}
- }
-
-\renewenvironment{abstract}{%
- \list{}{\advance\topsep by0.35cm\relax\small
- \leftmargin=1cm
- \labelwidth=\z@
- \listparindent=\z@
- \itemindent\listparindent
- \rightmargin\leftmargin}\item[\hskip\labelsep
- \bfseries\abstractname]}
- {\endlist}
-
-\newdimen\headlineindent % dimension for space between
-\headlineindent=1.166cm % number and text of headings.
-
-\def\ps@headings{\let\@mkboth\@gobbletwo
- \let\@oddfoot\@empty\let\@evenfoot\@empty
- \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
- \leftmark\hfil}
- \def\@oddhead{\normalfont\small\hfil\rightmark\hspace{\headlineindent}%
- \llap{\thepage}}
- \def\chaptermark##1{}%
- \def\sectionmark##1{}%
- \def\subsectionmark##1{}}
-
-\def\ps@titlepage{\let\@mkboth\@gobbletwo
- \let\@oddfoot\@empty\let\@evenfoot\@empty
- \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
- \hfil}
- \def\@oddhead{\normalfont\small\hfil\hspace{\headlineindent}%
- \llap{\thepage}}
- \def\chaptermark##1{}%
- \def\sectionmark##1{}%
- \def\subsectionmark##1{}}
-
-\if@runhead\ps@headings\else
-\ps@empty\fi
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\endinput
-%end of file llncs.cls
--- a/Pearl/document/root.bib Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,135 +0,0 @@
-@InProceedings{GunterOsbornPopescu09,
- author = {E.L.~Gunter and C.J.~Osborn and A.~Popescu},
- title = {{T}heory {S}upport for {W}eak {H}igher {O}rder {A}bstract {S}yntax in
- {I}sabelle/{HOL}},
- booktitle = {Proc.~of the 4th International Workshop on Logical Frameworks and Meta-Languages:
- Theory and Practice (LFMTP)},
- pages = {12--20},
- year = {2009},
- series = {ENTCS}
-}
-
-@Unpublished{SatoPollack10,
- author = {M.~Sato and R.~Pollack},
- title = {{E}xternal and {I}nternal {S}yntax of the {L}ambda-{C}alculus},
- note = {To appear in {\it Journal of Symbolic Computation}}
-}
-
-@article{GabbayPitts02,
- author = {M.~J.~Gabbay and A.~M.~Pitts},
- title = {{A} {N}ew {A}pproach to {A}bstract {S}yntax with {V}ariable
- {B}inding},
- journal = {Formal Aspects of Computing},
- volume = {13},
- year = 2002,
- pages = {341--363}
-}
-
-@article{Pitts03,
- author = {A.~M.~Pitts},
- title = {{N}ominal {L}ogic, {A} {F}irst {O}rder {T}heory of {N}ames and
- {B}inding},
- journal = {Information and Computation},
- year = {2003},
- volume = {183},
- pages = {165--193}
-}
-
-@InProceedings{BengtsonParrow07,
- author = {J.~Bengtson and J.~Parrow},
- title = {Formalising the pi-{C}alculus using {N}ominal {L}ogic},
- booktitle = {Proc.~of the 10th International Conference on
- Foundations of Software Science and Computation Structures (FOSSACS)},
- year = 2007,
- pages = {63--77},
- series = {LNCS},
- volume = {4423}
-}
-
-@inproceedings{BengtsonParow09,
- author = {J.~Bengtson and J.~Parrow},
- title = {{P}si-{C}alculi in {I}sabelle},
- booktitle = {Proc of the 22nd Conference on Theorem Proving in
- Higher Order Logics (TPHOLs)},
- year = 2009,
- pages = {99--114},
- series = {LNCS},
- volume = {5674}
-}
-
-@inproceedings{TobinHochstadtFelleisen08,
- author = {S.~Tobin-Hochstadt and M.~Felleisen},
- booktitle = {Proc.~of the 35rd Symposium on
- Principles of Programming Languages (POPL)},
- title = {{T}he {D}esign and {I}mplementation of {T}yped {S}cheme},
- publisher = {ACM},
- year = {2008},
- pages = {395--406}
-}
-
-@InProceedings{UrbanCheneyBerghofer08,
- author = "C.~Urban and J.~Cheney and S.~Berghofer",
- title = "{M}echanizing the {M}etatheory of {LF}",
- pages = "45--56",
- year = 2008,
- booktitle = "Proc.~of the 23rd IEEE Symposium on Logic in Computer Science (LICS)"
-}
-
-@InProceedings{UrbanZhu08,
- title = "{R}evisiting {C}ut-{E}limination: {O}ne {D}ifficult {P}roof is {R}eally a {P}roof",
- author = "C.~Urban and B.~Zhu",
- booktitle = "Proc.~of the 9th International Conference on Rewriting Techniques
- and Applications (RTA)",
- year = "2008",
- pages = "409--424",
- series = "LNCS",
- volume = 5117
-}
-
-@Article{UrbanPittsGabbay04,
- title = "{N}ominal {U}nification",
- author = "C.~Urban and A.M.~Pitts and M.J.~Gabbay",
- journal = "Theoretical Computer Science",
- pages = "473--497",
- volume = "323",
- number = "1-3",
- year = "2004"
-}
-
-@Article{Church40,
- author = {A.~Church},
- title = {{A} {F}ormulation of the {S}imple {T}heory of {T}ypes},
- journal = {Journal of Symbolic Logic},
- year = {1940},
- volume = {5},
- number = {2},
- pages = {56--68}
-}
-
-
-@Manual{PittsHOL4,
- title = {{S}yntax and {S}emantics},
- author = {A.~M.~Pitts},
- note = {Part of the documentation for the HOL4 system.}
-}
-
-
-@book{PaulsonBenzmueller,
- year={2009},
- author={Benzm{\"u}ller, Christoph and Paulson, Lawrence C.},
- title={Quantified Multimodal Logics in Simple Type Theory},
- note={{http://arxiv.org/abs/0905.2435}},
- series={{SEKI Report SR--2009--02 (ISSN 1437-4447)}},
- publisher={{SEKI Publications}}
-}
-
-@Article{Cheney06,
- author = {J.~Cheney},
- title = {{C}ompleteness and {H}erbrand {T}heorems for {N}ominal {L}ogic},
- journal = {Journal of Symbolic Logic},
- year = {2006},
- volume = {71},
- number = {1},
- pages = {299--320}
-}
-
--- a/Pearl/document/root.tex Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,62 +0,0 @@
-\documentclass{llncs}
-\usepackage{times}
-\usepackage{isabelle}
-\usepackage{isabellesym}
-\usepackage{amsmath}
-\usepackage{amssymb}
-
-
-\usepackage{pdfsetup}
-\urlstyle{rm}
-\isabellestyle{it}
-\renewcommand{\isastyle}{\isastyleminor}
-\renewcommand{\isacharunderscore}{\mbox{$\_\!\_$}}
-\renewcommand{\isasymbullet}{{\raisebox{-0.4mm}{\Large$\boldsymbol{\cdot}$}}}
-\def\dn{\,\stackrel{\mbox{\scriptsize def}}{=}\,}
-\renewcommand{\isasymequiv}{$\dn$}
-\renewcommand{\isasymiota}{}
-\renewcommand{\isasymrightleftharpoons}{}
-\renewcommand{\isasymemptyset}{$\varnothing$}
-
-\newcommand{\numbered}[1]{\refstepcounter{equation}{\rm(\arabic{equation})}\label{#1}}
-
-
-\begin{document}
-
-\title{Proof Pearl: A New Foundation for Nominal Isabelle}
-\author{Brian Huffman\inst{1} and Christian Urban\inst{2}}
-\institute{Portland State University \and Technical University of Munich}
-\maketitle
-
-\begin{abstract}
-Pitts et al introduced a beautiful theory about names and binding based on the
-notions of permutation and support. The engineering challenge is to smoothly
-adapt this theory to a theorem prover environment, in our case Isabelle/HOL.
-We present a formalisation of this work that differs from our earlier approach
-in two important respects: First, instead of representing permutations as
-lists of pairs of atoms, we now use a more abstract representation based on
-functions. Second, whereas the earlier work modeled different sorts of atoms
-using different types, we now introduce a unified atom type that includes all
-sorts of atoms. Interestingly, we allow swappings, that is permutations build from
-two atoms, to be ill-sorted. As a result of these design changes, we can iron
-out inconveniences for the user, considerably simplify proofs and also
-drastically reduce the amount of custom ML-code. Furthermore we can extend the
-capabilities of Nominal Isabelle to deal with variables that carry additional
-information. We end up with a pleasing and formalised theory of permutations
-and support, on which we can build an improved and more powerful version of
-Nominal Isabelle.
-\end{abstract}
-
-% generated text of all theories
-\input{session}
-
-% optional bibliography
-\bibliographystyle{abbrv}
-\bibliography{root}
-
-\end{document}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
--- a/Quotient-Paper/Paper-old.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1239 +0,0 @@
-(*<*)
-theory Paper
-imports "Quotient"
- "LaTeXsugar"
- "../Nominal/FSet"
-begin
-
-(****
-
-** things to do for the next version
-*
-* - what are quot_thms?
-* - what do all preservation theorems look like,
- in particular preservation for quotient
- compositions
- - explain how Quotient R Abs Rep is proved (j-version)
- - give an example where precise specification helps (core Haskell in nominal?)
-
- - Quote from Peter:
-
- One might think quotient have been studied to death, but
-
- - Mention Andreas Lochbiler in Acknowledgements and 'desceding'.
-
-*)
-
-notation (latex output)
- rel_conj ("_ \<circ>\<circ>\<circ> _" [53, 53] 52) and
- pred_comp ("_ \<circ>\<circ> _" [1, 1] 30) and
- "op -->" (infix "\<longrightarrow>" 100) and
- "==>" (infix "\<Longrightarrow>" 100) and
- fun_map ("_ \<^raw:\mbox{\singlearr}> _" 51) and
- fun_rel ("_ \<^raw:\mbox{\doublearr}> _" 51) and
- list_eq (infix "\<approx>" 50) and (* Not sure if we want this notation...? *)
- fempty ("\<emptyset>") and
- funion ("_ \<union> _") and
- finsert ("{_} \<union> _") and
- Cons ("_::_") and
- concat ("flat") and
- fconcat ("\<Union>")
-
-
-
-ML {*
-fun nth_conj n (_, r) = nth (HOLogic.dest_conj r) n;
-
-fun style_lhs_rhs proj = Scan.succeed (fn ctxt => fn t =>
- let
- val concl =
- Object_Logic.drop_judgment (ProofContext.theory_of ctxt) (Logic.strip_imp_concl t)
- in
- case concl of (_ $ l $ r) => proj (l, r)
- | _ => error ("Binary operator expected in term: " ^ Syntax.string_of_term ctxt concl)
- end);
-*}
-
-setup {*
- Term_Style.setup "rhs1" (style_lhs_rhs (nth_conj 0)) #>
- Term_Style.setup "rhs2" (style_lhs_rhs (nth_conj 1)) #>
- Term_Style.setup "rhs3" (style_lhs_rhs (nth_conj 2))
-*}
-
-(*>*)
-
-
-section {* Introduction *}
-
-text {*
- \begin{flushright}
- {\em ``Not using a [quotient] package has its advantages: we do not have to\\
- collect all the theorems we shall ever want into one giant list;''}\\
- Larry Paulson \cite{Paulson06}
- \end{flushright}
-
- \noindent
- Isabelle is a popular generic theorem prover in which many logics can be
- implemented. The most widely used one, however, is Higher-Order Logic
- (HOL). This logic consists of a small number of axioms and inference rules
- over a simply-typed term-language. Safe reasoning in HOL is ensured by two
- very restricted mechanisms for extending the logic: one is the definition of
- new constants in terms of existing ones; the other is the introduction of
- new types by identifying non-empty subsets in existing types. It is well
- understood how to use both mechanisms for dealing with quotient
- constructions in HOL (see \cite{Homeier05,Paulson06}). For example the
- integers in Isabelle/HOL are constructed by a quotient construction over the
- type @{typ "nat \<times> nat"} and the equivalence relation
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "(n\<^isub>1, n\<^isub>2) \<approx> (m\<^isub>1, m\<^isub>2) \<equiv> n\<^isub>1 + m\<^isub>2 = m\<^isub>1 + n\<^isub>2"}\hfill\numbered{natpairequiv}
- \end{isabelle}
-
- \noindent
- This constructions yields the new type @{typ int} and definitions for @{text
- "0"} and @{text "1"} of type @{typ int} can be given in terms of pairs of
- natural numbers (namely @{text "(0, 0)"} and @{text "(1, 0)"}). Operations
- such as @{text "add"} with type @{typ "int \<Rightarrow> int \<Rightarrow> int"} can be defined in
- terms of operations on pairs of natural numbers (namely @{text
- "add_pair (n\<^isub>1, m\<^isub>1) (n\<^isub>2,
- m\<^isub>2) \<equiv> (n\<^isub>1 + n\<^isub>2, m\<^isub>1 + m\<^isub>2)"}).
- Similarly one can construct the type of finite sets, written @{term "\<alpha> fset"},
- by quotienting the type @{text "\<alpha> list"} according to the equivalence relation
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "xs \<approx> ys \<equiv> (\<forall>x. memb x xs \<longleftrightarrow> memb x ys)"}\hfill\numbered{listequiv}
- \end{isabelle}
-
- \noindent
- which states that two lists are equivalent if every element in one list is
- also member in the other. The empty finite set, written @{term "{||}"}, can
- then be defined as the empty list and the union of two finite sets, written
- @{text "\<union>"}, as list append.
-
- Quotients are important in a variety of areas, but they are really ubiquitous in
- the area of reasoning about programming language calculi. A simple example
- is the lambda-calculus, whose raw terms are defined as
-
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "t ::= x | t t | \<lambda>x.t"}\hfill\numbered{lambda}
- \end{isabelle}
-
- \noindent
- The problem with this definition arises, for instance, when one attempts to
- prove formally the substitution lemma \cite{Barendregt81} by induction
- over the structure of terms. This can be fiendishly complicated (see
- \cite[Pages 94--104]{CurryFeys58} for some ``rough'' sketches of a proof
- about raw lambda-terms). In contrast, if we reason about
- $\alpha$-equated lambda-terms, that means terms quotient according to
- $\alpha$-equivalence, then the reasoning infrastructure provided,
- for example, by Nominal Isabelle \cite{UrbanKaliszyk11} makes the formal
- proof of the substitution lemma almost trivial.
-
- The difficulty is that in order to be able to reason about integers, finite
- sets or $\alpha$-equated lambda-terms one needs to establish a reasoning
- infrastructure by transferring, or \emph{lifting}, definitions and theorems
- from the raw type @{typ "nat \<times> nat"} to the quotient type @{typ int}
- (similarly for finite sets and $\alpha$-equated lambda-terms). This lifting
- usually requires a \emph{lot} of tedious reasoning effort \cite{Paulson06}.
- It is feasible to do this work manually, if one has only a few quotient
- constructions at hand. But if they have to be done over and over again, as in
- Nominal Isabelle, then manual reasoning is not an option.
-
- The purpose of a \emph{quotient package} is to ease the lifting of theorems
- and automate the reasoning as much as possible. In the
- context of HOL, there have been a few quotient packages already
- \cite{harrison-thesis,Slotosch97}. The most notable one is by Homeier
- \cite{Homeier05} implemented in HOL4. The fundamental construction these
- quotient packages perform can be illustrated by the following picture:
-
-%%% FIXME: Referee 1 says:
-%%% Diagram is unclear. Firstly, isn't an existing type a "set (not sets) of raw elements"?
-%%% Secondly, isn't the _set of_ equivalence classes mapped to and from the new type?
-%%% Thirdly, what do the words "non-empty subset" refer to ?
-
-%%% Cezary: I like the diagram, maybe 'new type' could be outside, but otherwise
-%%% I wouldn't change it.
-
- \begin{center}
- \mbox{}\hspace{20mm}\begin{tikzpicture}
- %%\draw[step=2mm] (-4,-1) grid (4,1);
-
- \draw[very thick] (0.7,0.3) circle (4.85mm);
- \draw[rounded corners=1mm, very thick] ( 0.0,-0.9) rectangle ( 1.8, 0.9);
- \draw[rounded corners=1mm, very thick] (-1.95,0.8) rectangle (-2.9,-0.195);
-
- \draw (-2.0, 0.8) -- (0.7,0.8);
- \draw (-2.0,-0.195) -- (0.7,-0.195);
-
- \draw ( 0.7, 0.23) node {\begin{tabular}{@ {}c@ {}}equiv-\\[-1mm]clas.\end{tabular}};
- \draw (-2.45, 0.35) node {\begin{tabular}{@ {}c@ {}}new\\[-1mm]type\end{tabular}};
- \draw (1.8, 0.35) node[right=-0.1mm]
- {\begin{tabular}{@ {}l@ {}}existing\\[-1mm] type\\ (sets of raw elements)\end{tabular}};
- \draw (0.9, -0.55) node {\begin{tabular}{@ {}l@ {}}non-empty\\[-1mm]subset\end{tabular}};
-
- \draw[->, very thick] (-1.8, 0.36) -- (-0.1,0.36);
- \draw[<-, very thick] (-1.8, 0.16) -- (-0.1,0.16);
- \draw (-0.95, 0.26) node[above=0.4mm] {@{text Rep}};
- \draw (-0.95, 0.26) node[below=0.4mm] {@{text Abs}};
-
- \end{tikzpicture}
- \end{center}
-
- \noindent
- The starting point is an existing type, to which we refer as the
- \emph{raw type} and over which an equivalence relation given by the user is
- defined. With this input the package introduces a new type, to which we
- refer as the \emph{quotient type}. This type comes with an
- \emph{abstraction} and a \emph{representation} function, written @{text Abs}
- and @{text Rep}.\footnote{Actually slightly more basic functions are given;
- the functions @{text Abs} and @{text Rep} need to be derived from them. We
- will show the details later. } They relate elements in the
- existing type to elements in the new type and vice versa, and can be uniquely
- identified by their quotient type. For example for the integer quotient construction
- the types of @{text Abs} and @{text Rep} are
-
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "Abs :: nat \<times> nat \<Rightarrow> int"}\hspace{10mm}@{text "Rep :: int \<Rightarrow> nat \<times> nat"}
- \end{isabelle}
-
- \noindent
- We therefore often write @{text Abs_int} and @{text Rep_int} if the
- typing information is important.
-
- Every abstraction and representation function stands for an isomorphism
- between the non-empty subset and elements in the new type. They are
- necessary for making definitions involving the new type. For example @{text
- "0"} and @{text "1"} of type @{typ int} can be defined as
-
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "0 \<equiv> Abs_int (0, 0)"}\hspace{10mm}@{text "1 \<equiv> Abs_int (1, 0)"}
- \end{isabelle}
-
- \noindent
- Slightly more complicated is the definition of @{text "add"} having type
- @{typ "int \<Rightarrow> int \<Rightarrow> int"}. Its definition is as follows
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "add n m \<equiv> Abs_int (add_pair (Rep_int n) (Rep_int m))"}
- \hfill\numbered{adddef}
- \end{isabelle}
-
- \noindent
- where we take the representation of the arguments @{text n} and @{text m},
- add them according to the function @{text "add_pair"} and then take the
- abstraction of the result. This is all straightforward and the existing
- quotient packages can deal with such definitions. But what is surprising is
- that none of them can deal with slightly more complicated definitions involving
- \emph{compositions} of quotients. Such compositions are needed for example
- in case of quotienting lists to yield finite sets and the operator that
- flattens lists of lists, defined as follows
-
- @{thm [display, indent=10] concat.simps(1) concat.simps(2)[no_vars]}
-
- \noindent
- We expect that the corresponding operator on finite sets, written @{term "fconcat"},
- builds finite unions of finite sets:
-
- @{thm [display, indent=10] fconcat_empty[no_vars] fconcat_insert[no_vars]}
-
- \noindent
- The quotient package should automatically provide us with a definition for @{text "\<Union>"} in
- terms of @{text flat}, @{text Rep_fset} and @{text Abs_fset}. The problem is
- that the method used in the existing quotient
- packages of just taking the representation of the arguments and then taking
- the abstraction of the result is \emph{not} enough. The reason is that in case
- of @{text "\<Union>"} we obtain the incorrect definition
-
- @{text [display, indent=10] "\<Union> S \<equiv> Abs_fset (flat (Rep_fset S))"}
-
- \noindent
- where the right-hand side is not even typable! This problem can be remedied in the
- existing quotient packages by introducing an intermediate step and reasoning
- about flattening of lists of finite sets. However, this remedy is rather
- cumbersome and inelegant in light of our work, which can deal with such
- definitions directly. The solution is that we need to build aggregate
- representation and abstraction functions, which in case of @{text "\<Union>"}
- generate the following definition
-
- @{text [display, indent=10] "\<Union> S \<equiv> Abs_fset (flat ((map_list Rep_fset \<circ> Rep_fset) S))"}
-
- \noindent
- where @{term map_list} is the usual mapping function for lists. In this paper we
- will present a formal definition of our aggregate abstraction and
- representation functions (this definition was omitted in \cite{Homeier05}).
- They generate definitions, like the one above for @{text "\<Union>"},
- according to the type of the raw constant and the type
- of the quotient constant. This means we also have to extend the notions
- of \emph{aggregate equivalence relation}, \emph{respectfulness} and \emph{preservation}
- from Homeier \cite{Homeier05}.
-
- In addition we are able to address the criticism by Paulson \cite{Paulson06} cited
- at the beginning of this section about having to collect theorems that are
- lifted from the raw level to the quotient level into one giant list. Homeier's and
- also our quotient package are modular so that they allow lifting
- theorems separately. This has the advantage for the user of being able to develop a
- formal theory interactively as a natural progression. A pleasing side-result of
- the modularity is that we are able to clearly specify what is involved
- in the lifting process (this was only hinted at in \cite{Homeier05} and
- implemented as a ``rough recipe'' in ML-code).
-
-
- The paper is organised as follows: Section \ref{sec:prelims} presents briefly
- some necessary preliminaries; Section \ref{sec:type} describes the definitions
- of quotient types and shows how definitions of constants can be made over
- quotient types. Section \ref{sec:resp} introduces the notions of respectfulness
- and preservation; Section \ref{sec:lift} describes the lifting of theorems;
- Section \ref{sec:examples} presents some examples
- and Section \ref{sec:conc} concludes and compares our results to existing
- work.
-*}
-
-section {* Preliminaries and General Quotients\label{sec:prelims} *}
-
-text {*
- We give in this section a crude overview of HOL and describe the main
- definitions given by Homeier for quotients \cite{Homeier05}.
-
- At its core, HOL is based on a simply-typed term language, where types are
- recorded in Church-style fashion (that means, we can always infer the type of
- a term and its subterms without any additional information). The grammars
- for types and terms are as follows
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}rl@ {\hspace{3mm}}l@ {}}
- @{text "\<sigma>, \<tau> ::="} & @{text "\<alpha> | (\<sigma>,\<dots>, \<sigma>) \<kappa>"} & (type variables and type constructors)\\
- @{text "t, s ::="} & @{text "x\<^isup>\<sigma> | c\<^isup>\<sigma> | t t | \<lambda>x\<^isup>\<sigma>. t"} &
- (variables, constants, applications and abstractions)\\
- \end{tabular}
- \end{isabelle}
-
- \noindent
- We often write just @{text \<kappa>} for @{text "() \<kappa>"}, and use @{text "\<alpha>s"} and
- @{text "\<sigma>s"} to stand for collections of type variables and types,
- respectively. The type of a term is often made explicit by writing @{text
- "t :: \<sigma>"}. HOL includes a type @{typ bool} for booleans and the function
- type, written @{text "\<sigma> \<Rightarrow> \<tau>"}. HOL also contains many primitive and defined
- constants; for example, a primitive constant is equality, with type @{text "= :: \<sigma> \<Rightarrow> \<sigma> \<Rightarrow>
- bool"}, and the identity function with type @{text "id :: \<sigma> \<Rightarrow> \<sigma>"} is
- defined as @{text "\<lambda>x\<^sup>\<sigma>. x\<^sup>\<sigma>"}.
-
- An important point to note is that theorems in HOL can be seen as a subset
- of terms that are constructed specially (namely through axioms and proof
- rules). As a result we are able to define automatic proof
- procedures showing that one theorem implies another by decomposing the term
- underlying the first theorem.
-
- Like Homeier's, our work relies on map-functions defined for every type
- constructor taking some arguments, for example @{text map_list} for lists. Homeier
- describes in \cite{Homeier05} map-functions for products, sums, options and
- also the following map for function types
-
- @{thm [display, indent=10] fun_map_def[no_vars, THEN eq_reflection]}
-
- \noindent
- Using this map-function, we can give the following, equivalent, but more
- uniform definition for @{text add} shown in \eqref{adddef}:
-
- @{text [display, indent=10] "add \<equiv> (Rep_int \<singlearr> Rep_int \<singlearr> Abs_int) add_pair"}
-
- \noindent
- Using extensionality and unfolding the definition of @{text "\<singlearr>"},
- we can get back to \eqref{adddef}.
- In what follows we shall use the convention to write @{text "map_\<kappa>"} for a map-function
- of the type-constructor @{text \<kappa>}. For a type @{text \<kappa>} with arguments @{text "\<alpha>\<^isub>1\<^isub>\<dots>\<^isub>n"} the
- type of @{text "map_\<kappa>"} has to be @{text "\<alpha>\<^isub>1\<Rightarrow>\<dots>\<Rightarrow>\<alpha>\<^isub>n\<Rightarrow>\<alpha>\<^isub>1\<dots>\<alpha>\<^isub>n \<kappa>"}. For example @{text "map_list"}
- has to have the type @{text "\<alpha>\<Rightarrow>\<alpha> list"}.
- In our implementation we maintain
- a database of these map-functions that can be dynamically extended.
-
- It will also be necessary to have operators, referred to as @{text "rel_\<kappa>"},
- which define equivalence relations in terms of constituent equivalence
- relations. For example given two equivalence relations @{text "R\<^isub>1"}
- and @{text "R\<^isub>2"}, we can define an equivalence relations over
- products as follows
- %
- @{text [display, indent=10] "(R\<^isub>1 \<tripple> R\<^isub>2) (x\<^isub>1, x\<^isub>2) (y\<^isub>1, y\<^isub>2) \<equiv> R\<^isub>1 x\<^isub>1 y\<^isub>1 \<and> R\<^isub>2 x\<^isub>2 y\<^isub>2"}
-
- \noindent
- Homeier gives also the following operator for defining equivalence
- relations over function types
- %
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{thm fun_rel_def[of "R\<^isub>1" "R\<^isub>2", no_vars, THEN eq_reflection]}
- \hfill\numbered{relfun}
- \end{isabelle}
-
- \noindent
- In the context of quotients, the following two notions from \cite{Homeier05}
- are needed later on.
-
- \begin{definition}[Respects]\label{def:respects}
- An element @{text "x"} respects a relation @{text "R"} provided @{text "R x x"}.
- \end{definition}
-
- \begin{definition}[Bounded Quantification and Bounded Abstractions]\label{def:babs}
- @{text "\<forall>x \<in> S. P x"} holds if for all @{text x}, @{text "x \<in> S"} implies @{text "P x"};
- and @{text "(\<lambda>x \<in> S. f x) = f x"} provided @{text "x \<in> S"}.
- \end{definition}
-
- The central definition in Homeier's work \cite{Homeier05} relates equivalence
- relations, abstraction and representation functions:
-
- \begin{definition}[Quotient Types]
- Given a relation $R$, an abstraction function $Abs$
- and a representation function $Rep$, the predicate @{term "Quotient R Abs Rep"}
- holds if and only if
- \begin{enumerate}
- \item @{thm (rhs1) Quotient_def[of "R", no_vars]}
- \item @{thm (rhs2) Quotient_def[of "R", no_vars]}
- \item @{thm (rhs3) Quotient_def[of "R", no_vars]}
- \end{enumerate}
- \end{definition}
-
- \noindent
- The value of this definition lies in the fact that validity of @{text "Quotient R Abs Rep"} can
- often be proved in terms of the validity of @{text "Quotient"} over the constituent
- types of @{text "R"}, @{text Abs} and @{text Rep}.
- For example Homeier proves the following property for higher-order quotient
- types:
-
- \begin{proposition}\label{funquot}
- @{thm[mode=IfThen] fun_quotient[where ?R1.0="R\<^isub>1" and ?R2.0="R\<^isub>2"
- and ?abs1.0="Abs\<^isub>1" and ?abs2.0="Abs\<^isub>2" and ?rep1.0="Rep\<^isub>1" and ?rep2.0="Rep\<^isub>2"]}
- \end{proposition}
-
- \noindent
- As a result, Homeier is able to build an automatic prover that can nearly
- always discharge a proof obligation involving @{text "Quotient"}. Our quotient
- package makes heavy
- use of this part of Homeier's work including an extension
- for dealing with compositions of equivalence relations defined as follows:
-
-%%% FIXME Referee 2 claims that composition-of-relations means OO, and this is also
-%%% what wikipedia says. Any idea for a different name? Conjugation of Relations?
-
- \begin{definition}[Composition of Relations]
- @{abbrev "rel_conj R\<^isub>1 R\<^isub>2"} where @{text "\<circ>\<circ>"} is the predicate
- composition defined by
- @{thm (concl) pred_compI[of "R\<^isub>1" "x" "y" "R\<^isub>2" "z"]}
- holds if and only if there exists a @{text y} such that @{thm (prem 1) pred_compI[of "R\<^isub>1" "x" "y" "R\<^isub>2" "z"]} and
- @{thm (prem 2) pred_compI[of "R\<^isub>1" "x" "y" "R\<^isub>2" "z"]}.
- \end{definition}
-
- \noindent
- Unfortunately a general quotient theorem for @{text "\<circ>\<circ>\<circ>"}, analogous to the one
- for @{text "\<singlearr>"} given in Proposition \ref{funquot}, would not be true
- in general. It cannot even be stated inside HOL, because of restrictions on types.
- However, we can prove specific instances of a
- quotient theorem for composing particular quotient relations.
- For example, to lift theorems involving @{term flat} the quotient theorem for
- composing @{text "\<approx>\<^bsub>list\<^esub>"} will be necessary: given @{term "Quotient R Abs Rep"}
- with @{text R} being an equivalence relation, then
-
- @{text [display, indent=2] "Quotient (rel_list R \<circ>\<circ>\<circ> \<approx>\<^bsub>list\<^esub>) (Abs_fset \<circ> map_list Abs) (map_list Rep \<circ> Rep_fset)"}
-
- \vspace{-.5mm}
-*}
-
-section {* Quotient Types and Quotient Definitions\label{sec:type} *}
-
-text {*
- The first step in a quotient construction is to take a name for the new
- type, say @{text "\<kappa>\<^isub>q"}, and an equivalence relation, say @{text R},
- defined over a raw type, say @{text "\<sigma>"}. The type of the equivalence
- relation must be @{text "\<sigma> \<Rightarrow> \<sigma> \<Rightarrow> bool"}. The user-visible part of
- the quotient type declaration is therefore
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \isacommand{quotient\_type}~~@{text "\<alpha>s \<kappa>\<^isub>q = \<sigma> / R"}\hfill\numbered{typedecl}
- \end{isabelle}
-
- \noindent
- and a proof that @{text "R"} is indeed an equivalence relation. Two concrete
- examples are
-
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- \isacommand{quotient\_type}~~@{text "int = nat \<times> nat / \<approx>\<^bsub>nat \<times> nat\<^esub>"}\\
- \isacommand{quotient\_type}~~@{text "\<alpha> fset = \<alpha> list / \<approx>\<^bsub>list\<^esub>"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- which introduce the type of integers and of finite sets using the
- equivalence relations @{text "\<approx>\<^bsub>nat \<times> nat\<^esub>"} and @{text
- "\<approx>\<^bsub>list\<^esub>"} defined in \eqref{natpairequiv} and
- \eqref{listequiv}, respectively (the proofs about being equivalence
- relations is omitted). Given this data, we define for declarations shown in
- \eqref{typedecl} the quotient types internally as
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \isacommand{typedef}~~@{text "\<alpha>s \<kappa>\<^isub>q = {c. \<exists>x. c = R x}"}
- \end{isabelle}
-
- \noindent
- where the right-hand side is the (non-empty) set of equivalence classes of
- @{text "R"}. The constraint in this declaration is that the type variables
- in the raw type @{text "\<sigma>"} must be included in the type variables @{text
- "\<alpha>s"} declared for @{text "\<kappa>\<^isub>q"}. HOL will then provide us with the following
- abstraction and representation functions
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "abs_\<kappa>\<^isub>q :: \<sigma> set \<Rightarrow> \<alpha>s \<kappa>\<^isub>q"}\hspace{10mm}@{text "rep_\<kappa>\<^isub>q :: \<alpha>s \<kappa>\<^isub>q \<Rightarrow> \<sigma> set"}
- \end{isabelle}
-
- \noindent
- As can be seen from the type, they relate the new quotient type and equivalence classes of the raw
- type. However, as Homeier \cite{Homeier05} noted, it is much more convenient
- to work with the following derived abstraction and representation functions
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "Abs_\<kappa>\<^isub>q x \<equiv> abs_\<kappa>\<^isub>q (R x)"}\hspace{10mm}@{text "Rep_\<kappa>\<^isub>q x \<equiv> \<epsilon> (rep_\<kappa>\<^isub>q x)"}
- \end{isabelle}
-
- \noindent
- on the expense of having to use Hilbert's choice operator @{text "\<epsilon>"} in the
- definition of @{text "Rep_\<kappa>\<^isub>q"}. These derived notions relate the
- quotient type and the raw type directly, as can be seen from their type,
- namely @{text "\<sigma> \<Rightarrow> \<alpha>s \<kappa>\<^isub>q"} and @{text "\<alpha>s \<kappa>\<^isub>q \<Rightarrow> \<sigma>"},
- respectively. Given that @{text "R"} is an equivalence relation, the
- following property holds for every quotient type
- (for the proof see \cite{Homeier05}).
-
- \begin{proposition}
- @{text "Quotient R Abs_\<kappa>\<^isub>q Rep_\<kappa>\<^isub>q"}.
- \end{proposition}
-
- The next step in a quotient construction is to introduce definitions of new constants
- involving the quotient type. These definitions need to be given in terms of concepts
- of the raw type (remember this is the only way how to extend HOL
- with new definitions). For the user the visible part of such definitions is the declaration
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \isacommand{quotient\_definition}~~@{text "c :: \<tau>"}~~\isacommand{is}~~@{text "t :: \<sigma>"}
- \end{isabelle}
-
- \noindent
- where @{text t} is the definiens (its type @{text \<sigma>} can always be inferred)
- and @{text "c"} is the name of definiendum, whose type @{text "\<tau>"} needs to be
- given explicitly (the point is that @{text "\<tau>"} and @{text "\<sigma>"} can only differ
- in places where a quotient and raw type is involved). Two concrete examples are
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- \isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0::nat, 0::nat)"}\\
- \isacommand{quotient\_definition}~~@{text "\<Union> :: (\<alpha> fset) fset \<Rightarrow> \<alpha> fset"}~~%
- \isacommand{is}~~@{text "flat"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- The first one declares zero for integers and the second the operator for
- building unions of finite sets (@{text "flat"} having the type
- @{text "(\<alpha> list) list \<Rightarrow> \<alpha> list"}).
-
- From such declarations given by the user, the quotient package needs to derive proper
- definitions using @{text "Abs"} and @{text "Rep"}. The data we rely on is the given quotient type
- @{text "\<tau>"} and the raw type @{text "\<sigma>"}. They allow us to define \emph{aggregate
- abstraction} and \emph{representation functions} using the functions @{text "ABS (\<sigma>,
- \<tau>)"} and @{text "REP (\<sigma>, \<tau>)"} whose clauses we shall give below. The idea behind
- these two functions is to simultaneously descend into the raw types @{text \<sigma>} and
- quotient types @{text \<tau>}, and generate the appropriate
- @{text "Abs"} and @{text "Rep"} in places where the types differ. Therefore
- we generate just the identity whenever the types are equal. On the ``way'' down,
- however we might have to use map-functions to let @{text Abs} and @{text Rep} act
- over the appropriate types. In what follows we use the short-hand notation
- @{text "ABS (\<sigma>s, \<tau>s)"} to mean @{text "ABS (\<sigma>\<^isub>1, \<tau>\<^isub>1)\<dots>ABS (\<sigma>\<^isub>n, \<tau>\<^isub>n)"}; similarly
- for @{text REP}.
- %
- \begin{center}
- \hfill
- \begin{tabular}{rcl}
- \multicolumn{3}{@ {\hspace{-4mm}}l}{equal types:}\\
- @{text "ABS (\<sigma>, \<sigma>)"} & $\dn$ & @{text "id :: \<sigma> \<Rightarrow> \<sigma>"}\\
- @{text "REP (\<sigma>, \<sigma>)"} & $\dn$ & @{text "id :: \<sigma> \<Rightarrow> \<sigma>"}\smallskip\\
- \multicolumn{3}{@ {\hspace{-4mm}}l}{function types:}\\
- @{text "ABS (\<sigma>\<^isub>1 \<Rightarrow> \<sigma>\<^isub>2, \<tau>\<^isub>1 \<Rightarrow> \<tau>\<^isub>2)"} & $\dn$ & @{text "REP (\<sigma>\<^isub>1, \<tau>\<^isub>1) \<singlearr> ABS (\<sigma>\<^isub>2, \<tau>\<^isub>2)"}\\
- @{text "REP (\<sigma>\<^isub>1 \<Rightarrow> \<sigma>\<^isub>2, \<tau>\<^isub>1 \<Rightarrow> \<tau>\<^isub>2)"} & $\dn$ & @{text "ABS (\<sigma>\<^isub>1, \<tau>\<^isub>1) \<singlearr> REP (\<sigma>\<^isub>2, \<tau>\<^isub>2)"}\smallskip\\
- \multicolumn{3}{@ {\hspace{-4mm}}l}{equal type constructors:}\\
- @{text "ABS (\<sigma>s \<kappa>, \<tau>s \<kappa>)"} & $\dn$ & @{text "map_\<kappa> (ABS (\<sigma>s, \<tau>s))"}\\
- @{text "REP (\<sigma>s \<kappa>, \<tau>s \<kappa>)"} & $\dn$ & @{text "map_\<kappa> (REP (\<sigma>s, \<tau>s))"}\smallskip\\
- \multicolumn{3}{@ {\hspace{-4mm}}l}{unequal type constructors with @{text "\<alpha>s
- \<kappa>\<^isub>q"} being the quotient of the raw type @{text "\<rho>s \<kappa>"}:}\\
- @{text "ABS (\<sigma>s \<kappa>, \<tau>s \<kappa>\<^isub>q)"} & $\dn$ & @{text "Abs_\<kappa>\<^isub>q \<circ> (MAP(\<rho>s \<kappa>) (ABS (\<sigma>s', \<tau>s)))"}\\
- @{text "REP (\<sigma>s \<kappa>, \<tau>s \<kappa>\<^isub>q)"} & $\dn$ & @{text "(MAP(\<rho>s \<kappa>) (REP (\<sigma>s', \<tau>s))) \<circ> Rep_\<kappa>\<^isub>q"}
- \end{tabular}\hfill\numbered{ABSREP}
- \end{center}
- %
- \noindent
- In the last two clauses we rely on the fact that the type @{text "\<alpha>s
- \<kappa>\<^isub>q"} is the quotient of the raw type @{text "\<rho>s \<kappa>"} (for example
- @{text "int"} and @{text "nat \<times> nat"}, or @{text "\<alpha> fset"} and @{text "\<alpha>
- list"}). The quotient construction ensures that the type variables in @{text
- "\<rho>s \<kappa>"} must be among the @{text "\<alpha>s"}. The @{text "\<sigma>s'"} are given by the
- substitutions for the @{text "\<alpha>s"} when matching @{text "\<sigma>s \<kappa>"} against
- @{text "\<rho>s \<kappa>"}. The
- function @{text "MAP"} calculates an \emph{aggregate map-function} for a raw
- type as follows:
- %
- \begin{center}
- \begin{tabular}{rcl}
- @{text "MAP' (\<alpha>)"} & $\dn$ & @{text "a\<^sup>\<alpha>"}\\
- @{text "MAP' (\<kappa>)"} & $\dn$ & @{text "id :: \<kappa> \<Rightarrow> \<kappa>"}\\
- @{text "MAP' (\<sigma>s \<kappa>)"} & $\dn$ & @{text "map_\<kappa> (MAP'(\<sigma>s))"}\smallskip\\
- @{text "MAP (\<sigma>)"} & $\dn$ & @{text "\<lambda>as. MAP'(\<sigma>)"}
- \end{tabular}
- \end{center}
- %
- \noindent
- In this definition we rely on the fact that in the first clause we can interpret type-variables @{text \<alpha>} as
- term variables @{text a}. In the last clause we build an abstraction over all
- term-variables of the map-function generated by the auxiliary function
- @{text "MAP'"}.
- The need for aggregate map-functions can be seen in cases where we build quotients,
- say @{text "(\<alpha>, \<beta>) \<kappa>\<^isub>q"}, out of compound raw types, say @{text "(\<alpha> list) \<times> \<beta>"}.
- In this case @{text MAP} generates the
- aggregate map-function:
-
-%%% FIXME: Reviewer 2 asks: last two lines defining ABS and REP for
-%%% unequal type constructors: How are the $\varrho$s defined? The
-%%% following paragraph mentions them, but this paragraph is unclear,
-%%% since it then mentions $\alpha$s, which do not seem to be defined
-%%% either. As a result, I do not understand the first two sentences
-%%% in this paragraph. I can imagine roughly what the following
-%%% sentence `The $\sigma$s' are given by the matchers for the
-%%% $\alpha$s$ when matching $\varrho$s $\kappa$ against $\sigma$s
-%%% $\kappa$.' means, but also think that it is too vague.
-
- @{text [display, indent=10] "\<lambda>a b. map_prod (map_list a) b"}
-
- \noindent
- which is essential in order to define the corresponding aggregate
- abstraction and representation functions.
-
- To see how these definitions pan out in practise, let us return to our
- example about @{term "concat"} and @{term "fconcat"}, where we have the raw type
- @{text "(\<alpha> list) list \<Rightarrow> \<alpha> list"} and the quotient type @{text "(\<alpha> fset) fset \<Rightarrow> \<alpha>
- fset"}. Feeding these types into @{text ABS} gives us (after some @{text "\<beta>"}-simplifications)
- the abstraction function
-
- @{text [display, indent=10] "(map_list (map_list id \<circ> Rep_fset) \<circ> Rep_fset) \<singlearr> Abs_fset \<circ> map_list id"}
-
- \noindent
- In our implementation we further
- simplify this function by rewriting with the usual laws about @{text
- "map"}s and @{text "id"}, for example @{term "map_list id = id"} and @{text "f \<circ> id =
- id \<circ> f = f"}. This gives us the simpler abstraction function
-
- @{text [display, indent=10] "(map_list Rep_fset \<circ> Rep_fset) \<singlearr> Abs_fset"}
-
- \noindent
- which we can use for defining @{term "fconcat"} as follows
-
- @{text [display, indent=10] "\<Union> \<equiv> ((map_list Rep_fset \<circ> Rep_fset) \<singlearr> Abs_fset) flat"}
-
- \noindent
- Note that by using the operator @{text "\<singlearr>"} and special clauses
- for function types in \eqref{ABSREP}, we do not have to
- distinguish between arguments and results, but can deal with them uniformly.
- Consequently, all definitions in the quotient package
- are of the general form
-
- @{text [display, indent=10] "c \<equiv> ABS (\<sigma>, \<tau>) t"}
-
- \noindent
- where @{text \<sigma>} is the type of the definiens @{text "t"} and @{text "\<tau>"} the
- type of the defined quotient constant @{text "c"}. This data can be easily
- generated from the declaration given by the user.
- To increase the confidence in this way of making definitions, we can prove
- that the terms involved are all typable.
-
- \begin{lemma}
- If @{text "ABS (\<sigma>, \<tau>)"} returns some abstraction function @{text "Abs"}
- and @{text "REP (\<sigma>, \<tau>)"} some representation function @{text "Rep"},
- then @{text "Abs"} is of type @{text "\<sigma> \<Rightarrow> \<tau>"} and @{text "Rep"} of type
- @{text "\<tau> \<Rightarrow> \<sigma>"}.
- \end{lemma}
-
- \begin{proof}
- By mutual induction and analysing the definitions of @{text "ABS"} and @{text "REP"}.
- The cases of equal types and function types are
- straightforward (the latter follows from @{text "\<singlearr>"} having the
- type @{text "(\<alpha> \<Rightarrow> \<beta>) \<Rightarrow> (\<gamma> \<Rightarrow> \<delta>) \<Rightarrow> (\<beta> \<Rightarrow> \<gamma>) \<Rightarrow> (\<alpha> \<Rightarrow> \<delta>)"}). In case of equal type
- constructors we can observe that a map-function after applying the functions
- @{text "ABS (\<sigma>s, \<tau>s)"} produces a term of type @{text "\<sigma>s \<kappa> \<Rightarrow> \<tau>s \<kappa>"}. The
- interesting case is the one with unequal type constructors. Since we know
- the quotient is between @{text "\<alpha>s \<kappa>\<^isub>q"} and @{text "\<rho>s \<kappa>"}, we have
- that @{text "Abs_\<kappa>\<^isub>q"} is of type @{text "\<rho>s \<kappa> \<Rightarrow> \<alpha>s
- \<kappa>\<^isub>q"}. This type can be more specialised to @{text "\<rho>s[\<tau>s] \<kappa> \<Rightarrow> \<tau>s
- \<kappa>\<^isub>q"} where the type variables @{text "\<alpha>s"} are instantiated with the
- @{text "\<tau>s"}. The complete type can be calculated by observing that @{text
- "MAP (\<rho>s \<kappa>)"}, after applying the functions @{text "ABS (\<sigma>s', \<tau>s)"} to it,
- returns a term of type @{text "\<rho>s[\<sigma>s'] \<kappa> \<Rightarrow> \<rho>s[\<tau>s] \<kappa>"}. This type is
- equivalent to @{text "\<sigma>s \<kappa> \<Rightarrow> \<rho>s[\<tau>s] \<kappa>"}, which we just have to compose with
- @{text "\<rho>s[\<tau>s] \<kappa> \<Rightarrow> \<tau>s \<kappa>\<^isub>q"} according to the type of @{text "\<circ>"}.\qed
- \end{proof}
-*}
-
-section {* Respectfulness and Preservation \label{sec:resp} *}
-
-text {*
- The main point of the quotient package is to automatically ``lift'' theorems
- involving constants over the raw type to theorems involving constants over
- the quotient type. Before we can describe this lifting process, we need to impose
- two restrictions in form of proof obligations that arise during the
- lifting. The reason is that even if definitions for all raw constants
- can be given, \emph{not} all theorems can be lifted to the quotient type. Most
- notable is the bound variable function, that is the constant @{text bn}, defined
- for raw lambda-terms as follows
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- @{text "bn (x) \<equiv> \<emptyset>"}\hspace{4mm}
- @{text "bn (t\<^isub>1 t\<^isub>2) \<equiv> bn (t\<^isub>1) \<union> bn (t\<^isub>2)"}\hspace{4mm}
- @{text "bn (\<lambda>x. t) \<equiv> {x} \<union> bn (t)"}
- \end{isabelle}
-
- \noindent
- We can generate a definition for this constant using @{text ABS} and @{text REP}.
- But this constant does \emph{not} respect @{text "\<alpha>"}-equivalence and
- consequently no theorem involving this constant can be lifted to @{text
- "\<alpha>"}-equated lambda terms. Homeier formulates the restrictions in terms of
- the properties of \emph{respectfulness} and \emph{preservation}. We have
- to slightly extend Homeier's definitions in order to deal with quotient
- compositions.
-
-%%% FIXME: Reviewer 3 asks why are the definitions that follow enough to deal
-%%% with quotient composition.
-
- To formally define what respectfulness is, we have to first define
- the notion of \emph{aggregate equivalence relations} using the function @{text "REL(\<sigma>, \<tau>)"}
- The idea behind this function is to simultaneously descend into the raw types
- @{text \<sigma>} and quotient types @{text \<tau>}, and generate the appropriate
- quotient equivalence relations in places where the types differ and equalities
- elsewhere.
-
- \begin{center}
- \hfill
- \begin{tabular}{rcl}
- \multicolumn{3}{@ {\hspace{-4mm}}l}{equal types:}\\
- @{text "REL (\<sigma>, \<sigma>)"} & $\dn$ & @{text "= :: \<sigma> \<Rightarrow> \<sigma> \<Rightarrow> bool"}\smallskip\\
- \multicolumn{3}{@ {\hspace{-4mm}}l}{equal type constructors:}\\
- @{text "REL (\<sigma>s \<kappa>, \<tau>s \<kappa>)"} & $\dn$ & @{text "rel_\<kappa> (REL (\<sigma>s, \<tau>s))"}\smallskip\\
- \multicolumn{3}{@ {\hspace{-4mm}}l}{unequal type constructors with @{text "\<alpha>s
- \<kappa>\<^isub>q"} being the quotient of the raw type @{text "\<rho>s \<kappa>"}:}\smallskip\\
- @{text "REL (\<sigma>s \<kappa>, \<tau>s \<kappa>\<^isub>q)"} & $\dn$ & @{text "rel_\<kappa>\<^isub>q (REL (\<sigma>s', \<tau>s))"}\\
- \end{tabular}\hfill\numbered{REL}
- \end{center}
-
- \noindent
- The @{text "\<sigma>s'"} in the last clause are calculated as in \eqref{ABSREP}:
- we know that type @{text "\<alpha>s \<kappa>\<^isub>q"} is the quotient of the raw type
- @{text "\<rho>s \<kappa>"}. The @{text "\<sigma>s'"} are the substitutions for @{text "\<alpha>s"} obtained by matching
- @{text "\<rho>s \<kappa>"} and @{text "\<sigma>s \<kappa>"}.
-
- Let us return to the lifting procedure of theorems. Assume we have a theorem
- that contains the raw constant @{text "c\<^isub>r :: \<sigma>"} and which we want to
- lift to a theorem where @{text "c\<^isub>r"} is replaced by the corresponding
- constant @{text "c\<^isub>q :: \<tau>"} defined over a quotient type. In this situation
- we generate the following proof obligation
-
- @{text [display, indent=10] "REL (\<sigma>, \<tau>) c\<^isub>r c\<^isub>r"}
-
- \noindent
- Homeier calls these proof obligations \emph{respectfulness
- theorems}. However, unlike his quotient package, we might have several
- respectfulness theorems for one constant---he has at most one.
- The reason is that because of our quotient compositions, the types
- @{text \<sigma>} and @{text \<tau>} are not completely determined by @{text "c\<^bsub>r\<^esub>"}.
- And for every instantiation of the types, a corresponding
- respectfulness theorem is necessary.
-
- Before lifting a theorem, we require the user to discharge
- respectfulness proof obligations. In case of @{text bn}
- this obligation is as follows
-
- @{text [display, indent=10] "(\<approx>\<^isub>\<alpha> \<doublearr> =) bn bn"}
-
- \noindent
- and the point is that the user cannot discharge it: because it is not true. To see this,
- we can just unfold the definition of @{text "\<doublearr>"} \eqref{relfun}
- using extensionality to obtain the false statement
-
- @{text [display, indent=10] "\<forall>t\<^isub>1 t\<^isub>2. if t\<^isub>1 \<approx>\<^isub>\<alpha> t\<^isub>2 then bn(t\<^isub>1) = bn(t\<^isub>2)"}
-
- \noindent
- In contrast, if we lift a theorem about @{text "append"} to a theorem describing
- the union of finite sets, then we need to discharge the proof obligation
-
- @{text [display, indent=10] "(\<approx>\<^bsub>list\<^esub> \<doublearr> \<approx>\<^bsub>list\<^esub> \<doublearr> \<approx>\<^bsub>list\<^esub>) append append"}
-
- \noindent
- To do so, we have to establish
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
- if @{text "xs \<approx>\<^bsub>list\<^esub> ys"} and @{text "us \<approx>\<^bsub>list\<^esub> vs"}
- then @{text "xs @ us \<approx>\<^bsub>list\<^esub> ys @ vs"}
- \end{isabelle}
-
- \noindent
- which is straightforward given the definition shown in \eqref{listequiv}.
-
- The second restriction we have to impose arises from non-lifted polymorphic
- constants, which are instantiated to a type being quotient. For example,
- take the @{term "cons"}-constructor to add a pair of natural numbers to a
- list, whereby we assume the pair of natural numbers turns into an integer in
- the quotient construction. The point is that we still want to use @{text
- cons} for adding integers to lists---just with a different type. To be able
- to lift such theorems, we need a \emph{preservation property} for @{text
- cons}. Assuming we have a polymorphic raw constant @{text "c\<^isub>r :: \<sigma>"}
- and a corresponding quotient constant @{text "c\<^isub>q :: \<tau>"}, then a
- preservation property is as follows
-
-%%% FIXME: Reviewer 2 asks: You say what a preservation theorem is,
-%%% but not which preservation theorems you assume. Do you generate a
-%%% proof obligation for a preservation theorem for each raw constant
-%%% and its corresponding lifted constant?
-
-%%% Cezary: I think this would be a nice thing to do but we have not
-%%% done it, the theorems need to be 'guessed' from the remaining obligations
-
- @{text [display, indent=10] "Quotient R\<^bsub>\<alpha>s\<^esub> Abs\<^bsub>\<alpha>s\<^esub> Rep\<^bsub>\<alpha>s\<^esub> implies ABS (\<sigma>, \<tau>) c\<^isub>r = c\<^isub>r"}
-
- \noindent
- where the @{text "\<alpha>s"} stand for the type variables in the type of @{text "c\<^isub>r"}.
- In case of @{text cons} (which has type @{text "\<alpha> \<Rightarrow> \<alpha> list \<Rightarrow> \<alpha> list"}) we have
-
- @{text [display, indent=10] "(Rep ---> map_list Rep ---> map_list Abs) cons = cons"}
-
- \noindent
- under the assumption @{text "Quotient R Abs Rep"}. Interestingly, if we have
- an instance of @{text cons} where the type variable @{text \<alpha>} is instantiated
- with @{text "nat \<times> nat"} and we also quotient this type to yield integers,
- then we need to show the corresponding preservation property.
-
- %%%@ {thm [display, indent=10] insert_preserve2[no_vars]}
-
- %Given two quotients, one of which quotients a container, and the
- %other quotients the type in the container, we can write the
- %composition of those quotients. To compose two quotient theorems
- %we compose the relations with relation composition as defined above
- %and the abstraction and relation functions are the ones of the sub
- %quotients composed with the usual function composition.
- %The @ {term "Rep"} and @ {term "Abs"} functions that we obtain agree
- %with the definition of aggregate Abs/Rep functions and the
- %relation is the same as the one given by aggregate relations.
- %This becomes especially interesting
- %when we compose the quotient with itself, as there is no simple
- %intermediate step.
- %
- %Lets take again the example of @ {term flat}. To be able to lift
- %theorems that talk about it we provide the composition quotient
- %theorem which allows quotienting inside the container:
- %
- %If @ {term R} is an equivalence relation and @ {term "Quotient R Abs Rep"}
- %then
- %
- %@ {text [display, indent=10] "Quotient (list_rel R \<circ>\<circ>\<circ> \<approx>\<^bsub>list\<^esub>) (abs_fset \<circ> map_list Abs) (map_list Rep o rep_fset)"}
- %%%
- %%%\noindent
- %%%this theorem will then instantiate the quotients needed in the
- %%%injection and cleaning proofs allowing the lifting procedure to
- %%%proceed in an unchanged way.
-*}
-
-section {* Lifting of Theorems\label{sec:lift} *}
-
-text {*
-
-%%% FIXME Reviewer 3 asks: Section 5 shows the technicalities of
-%%% lifting theorems. But there is no clarification about the
-%%% correctness. A reader would also be interested in seeing some
-%%% discussions about the generality and limitation of the approach
-%%% proposed there
-
- The main benefit of a quotient package is to lift automatically theorems over raw
- types to theorems over quotient types. We will perform this lifting in
- three phases, called \emph{regularization},
- \emph{injection} and \emph{cleaning} according to procedures in Homeier's ML-code.
-
- The purpose of regularization is to change the quantifiers and abstractions
- in a ``raw'' theorem to quantifiers over variables that respect their respective relations
- (Definition \ref{def:respects} states what respects means). The purpose of injection is to add @{term Rep}
- and @{term Abs} of appropriate types in front of constants and variables
- of the raw type so that they can be replaced by the corresponding constants from the
- quotient type. The purpose of cleaning is to bring the theorem derived in the
- first two phases into the form the user has specified. Abstractly, our
- package establishes the following three proof steps:
-
-%%% FIXME: Reviewer 1 complains that the reader needs to guess the
-%%% meaning of reg_thm and inj_thm, as well as the arguments of REG
-%%% which are given above. I wouldn't change it.
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{4mm}}l}
- 1.) Regularization & @{text "raw_thm \<longrightarrow> reg_thm"}\\
- 2.) Injection & @{text "reg_thm \<longleftrightarrow> inj_thm"}\\
- 3.) Cleaning & @{text "inj_thm \<longleftrightarrow> quot_thm"}\\
- \end{tabular}
- \end{center}
-
- \noindent
- which means, stringed together, the raw theorem implies the quotient theorem.
- In contrast to other quotient packages, our package requires that the user specifies
- both, the @{text "raw_thm"} (as theorem) and the \emph{term} of the @{text "quot_thm"}.\footnote{Though we
- also provide a fully automated mode, where the @{text "quot_thm"} is guessed
- from the form of @{text "raw_thm"}.} As a result, the user has fine control
- over which parts of a raw theorem should be lifted.
-
- The second and third proof step performed in package will always succeed if the appropriate
- respectfulness and preservation theorems are given. In contrast, the first
- proof step can fail: a theorem given by the user does not always
- imply a regularized version and a stronger one needs to be proved. An example
- for this kind of failure is the simple statement for integers @{text "0 \<noteq> 1"}.
- One might hope that it can be proved by lifting @{text "(0, 0) \<noteq> (1, 0)"},
- but this raw theorem only shows that two particular elements in the
- equivalence classes are not equal. In order to obtain @{text "0 \<noteq> 1"}, a
- more general statement stipulating that the equivalence classes are not
- equal is necessary. This kind of failure is beyond the scope where the
- quotient package can help: the user has to provide a raw theorem that
- can be regularized automatically, or has to provide an explicit proof
- for the first proof step.
-
- In the following we will first define the statement of the
- regularized theorem based on @{text "raw_thm"} and
- @{text "quot_thm"}. Then we define the statement of the injected theorem, based
- on @{text "reg_thm"} and @{text "quot_thm"}. We then show the three proof steps,
- which can all be performed independently from each other.
-
- We first define the function @{text REG}, which takes the terms of the
- @{text "raw_thm"} and @{text "quot_thm"} as input and returns
- @{text "reg_thm"}. The idea
- behind this function is that it replaces quantifiers and
- abstractions involving raw types by bounded ones, and equalities
- involving raw types by appropriate aggregate
- equivalence relations. It is defined by simultaneously recursing on
- the structure of @{text "raw_thm"} and @{text "quot_thm"} as follows:
-
- \begin{center}
- \begin{tabular}{rcl}
- \multicolumn{3}{@ {}l}{abstractions:}\smallskip\\
- @{text "REG (\<lambda>x\<^sup>\<sigma>. t, \<lambda>x\<^sup>\<tau>. s)"} & $\dn$ &
- $\begin{cases}
- @{text "\<lambda>x\<^sup>\<sigma>. REG (t, s)"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
- @{text "\<lambda>x\<^sup>\<sigma> \<in> Respects (REL (\<sigma>, \<tau>)). REG (t, s)"}
- \end{cases}$\smallskip\\
- \\
- \multicolumn{3}{@ {}l}{universal quantifiers:}\\
- @{text "REG (\<forall>x\<^sup>\<sigma>. t, \<forall>x\<^sup>\<tau>. s)"} & $\dn$ &
- $\begin{cases}
- @{text "\<forall>x\<^sup>\<sigma>. REG (t, s)"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
- @{text "\<forall>x\<^sup>\<sigma> \<in> Respects (REL (\<sigma>, \<tau>)). REG (t, s)"}
- \end{cases}$\smallskip\\
- \multicolumn{3}{@ {}l}{equality:}\smallskip\\
- %% REL of two equal types is the equality so we do not need a separate case
- @{text "REG (=\<^bsup>\<sigma>\<Rightarrow>\<sigma>\<Rightarrow>bool\<^esup>, =\<^bsup>\<tau>\<Rightarrow>\<tau>\<Rightarrow>bool\<^esup>)"} & $\dn$ & @{text "REL (\<sigma>, \<tau>)"}\\\smallskip\\
- \multicolumn{3}{@ {}l}{applications, variables and constants:}\\
- @{text "REG (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2)"} & $\dn$ & @{text "REG (t\<^isub>1, s\<^isub>1) REG (t\<^isub>2, s\<^isub>2)"}\\
- @{text "REG (x\<^isub>1, x\<^isub>2)"} & $\dn$ & @{text "x\<^isub>1"}\\
- @{text "REG (c\<^isub>1, c\<^isub>2)"} & $\dn$ & @{text "c\<^isub>1"}\\
- \end{tabular}
- \end{center}
- %
- \noindent
- In the above definition we omitted the cases for existential quantifiers
- and unique existential quantifiers, as they are very similar to the cases
- for the universal quantifier.
-
- Next we define the function @{text INJ} which takes as argument
- @{text "reg_thm"} and @{text "quot_thm"} (both as
- terms) and returns @{text "inj_thm"}:
-
- \begin{center}
- \begin{tabular}{rcl}
- \multicolumn{3}{@ {\hspace{-4mm}}l}{abstractions:}\\
- @{text "INJ (\<lambda>x. t :: \<sigma>, \<lambda>x. s :: \<tau>) "} & $\dn$ &
- $\begin{cases}
- @{text "\<lambda>x. INJ (t, s)"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
- @{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) (\<lambda>x. INJ (t, s)))"}
- \end{cases}$\\
- @{text "INJ (\<lambda>x \<in> R. t :: \<sigma>, \<lambda>x. s :: \<tau>) "} & $\dn$
- & @{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) (\<lambda>x \<in> R. INJ (t, s)))"}\smallskip\\
- \multicolumn{3}{@ {\hspace{-4mm}}l}{universal quantifiers:}\\
- @{text "INJ (\<forall> t, \<forall> s) "} & $\dn$ & @{text "\<forall> INJ (t, s)"}\\
- @{text "INJ (\<forall> t \<in> R, \<forall> s) "} & $\dn$ & @{text "\<forall> INJ (t, s) \<in> R"}\smallskip\\
- \multicolumn{3}{@ {\hspace{-4mm}}l}{applications, variables and constants:}\smallskip\\
- @{text "INJ (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2) "} & $\dn$ & @{text " INJ (t\<^isub>1, s\<^isub>1) INJ (t\<^isub>2, s\<^isub>2)"}\\
- @{text "INJ (x\<^isub>1\<^sup>\<sigma>, x\<^isub>2\<^sup>\<tau>) "} & $\dn$ &
- $\begin{cases}
- @{text "x\<^isub>1"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
- @{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) x\<^isub>1)"}\\
- \end{cases}$\\
- @{text "INJ (c\<^isub>1\<^sup>\<sigma>, c\<^isub>2\<^sup>\<tau>) "} & $\dn$ &
- $\begin{cases}
- @{text "c\<^isub>1"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
- @{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) c\<^isub>1)"}\\
- \end{cases}$\\
- \end{tabular}
- \end{center}
-
- \noindent
- In this definition we again omitted the cases for existential and unique existential
- quantifiers.
-
-%%% FIXME: Reviewer2 citing following sentence: You mention earlier
-%%% that this implication may fail to be true. Does that meant that
-%%% the `first proof step' is a heuristic that proves the implication
-%%% raw_thm \implies reg_thm in some instances, but fails in others?
-%%% You should clarify under which circumstances the implication is
-%%% being proved here.
-%%% Cezary: It would be nice to cite Homeiers discussions in the
-%%% Quotient Package manual from HOL (the longer paper), do you agree?
-
- In the first proof step, establishing @{text "raw_thm \<longrightarrow> reg_thm"}, we always
- start with an implication. Isabelle provides \emph{mono} rules that can split up
- the implications into simpler implicational subgoals. This succeeds for every
- monotone connective, except in places where the function @{text REG} replaced,
- for instance, a quantifier by a bounded quantifier. In this case we have
- rules of the form
-
- @{text [display, indent=10] "(\<forall>x. R x \<longrightarrow> (P x \<longrightarrow> Q x)) \<longrightarrow> (\<forall>x. P x \<longrightarrow> \<forall>x \<in> R. Q x)"}
-
- \noindent
- They decompose a bounded quantifier on the right-hand side. We can decompose a
- bounded quantifier anywhere if R is an equivalence relation or
- if it is a relation over function types with the range being an equivalence
- relation. If @{text R} is an equivalence relation we can prove that
-
- @{text [display, indent=10] "\<forall>x \<in> Respects R. P x = \<forall>x. P x"}
-
- \noindent
- If @{term R\<^isub>2} is an equivalence relation, we can prove that for any predicate @{term P}
-
-%%% FIXME Reviewer 1 claims the theorem is obviously false so maybe we
-%%% should include a proof sketch?
-
- @{thm [display, indent=10] (concl) ball_reg_eqv_range[of R\<^isub>1 R\<^isub>2, no_vars]}
-
- \noindent
- The last theorem is new in comparison with Homeier's package. There the
- injection procedure would be used to prove such goals and
- the assumption about the equivalence relation would be used. We use the above theorem directly,
- because this allows us to completely separate the first and the second
- proof step into two independent ``units''.
-
- The second proof step, establishing @{text "reg_thm \<longleftrightarrow> inj_thm"}, starts with an equality
- between the terms of the regularized theorem and the injected theorem.
- The proof again follows the structure of the
- two underlying terms and is defined for a goal being a relation between these two terms.
-
- \begin{itemize}
- \item For two constants an appropriate respectfulness theorem is applied.
- \item For two variables, we use the assumptions proved in the regularization step.
- \item For two abstractions, we @{text "\<eta>"}-expand and @{text "\<beta>"}-reduce them.
- \item For two applications, we check that the right-hand side is an application of
- @{term Rep} to an @{term Abs} and @{term "Quotient R Rep Abs"} holds. If yes then we
- can apply the theorem:
-
- @{term [display, indent=10] "R x y \<longrightarrow> R x (Rep (Abs y))"}
-
- Otherwise we introduce an appropriate relation between the subterms
- and continue with two subgoals using the lemma:
-
- @{text [display, indent=10] "(R\<^isub>1 \<doublearr> R\<^isub>2) f g \<longrightarrow> R\<^isub>1 x y \<longrightarrow> R\<^isub>2 (f x) (g y)"}
- \end{itemize}
-
- We defined the theorem @{text "inj_thm"} in such a way that
- establishing the equivalence @{text "inj_thm \<longleftrightarrow> quot_thm"} can be
- achieved by rewriting @{text "inj_thm"} with the preservation theorems and quotient
- definitions. First the definitions of all lifted constants
- are used to fold the @{term Rep} with the raw constants. Next for
- all abstractions and quantifiers the lambda and
- quantifier preservation theorems are used to replace the
- variables that include raw types with respects by quantifiers
- over variables that include quotient types. We show here only
- the lambda preservation theorem. Given
- @{term "Quotient R\<^isub>1 Abs\<^isub>1 Rep\<^isub>1"} and @{term "Quotient R\<^isub>2 Abs\<^isub>2 Rep\<^isub>2"}, we have:
-
- @{thm [display, indent=10] (concl) lambda_prs[of _ "Abs\<^isub>1" "Rep\<^isub>1" _ "Abs\<^isub>2" "Rep\<^isub>2", no_vars]}
-
- \noindent
- Next, relations over lifted types can be rewritten to equalities
- over lifted type. Rewriting is performed with the following theorem,
- which has been shown by Homeier~\cite{Homeier05}:
-
- @{thm [display, indent=10] (concl) Quotient_rel_rep[no_vars]}
-
- \noindent
- Finally, we rewrite with the preservation theorems. This will result
- in two equal terms that can be solved by reflexivity.
- *}
-
-
-section {* Examples \label{sec:examples} *}
-
-text {*
-
-%%% FIXME Reviewer 1 would like an example of regularized and injected
-%%% statements. He asks for the examples twice, but I would still ignore
-%%% it due to lack of space...
-
- In this section we will show a sequence of declarations for defining the
- type of integers by quotienting pairs of natural numbers, and
- lifting one theorem.
-
- A user of our quotient package first needs to define a relation on
- the raw type with which the quotienting will be performed. We give
- the same integer relation as the one presented in \eqref{natpairequiv}:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
- \begin{tabular}{@ {}l}
- \isacommand{fun}~~@{text "int_rel :: (nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"}\\
- \isacommand{where}~~@{text "int_rel (m, n) (p, q) = (m + q = n + p)"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- Next the quotient type must be defined. This generates a proof obligation that the
- relation is an equivalence relation, which is solved automatically using the
- definition of equivalence and extensionality:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
- \begin{tabular}{@ {}l}
- \isacommand{quotient\_type}~~@{text "int"}~~\isacommand{=}~~@{text "(nat \<times> nat)"}~~\isacommand{/}~~@{text "int_rel"}\\
- \hspace{5mm}@{text "by (auto simp add: equivp_def expand_fun_eq)"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- The user can then specify the constants on the quotient type:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
- \begin{tabular}{@ {}l}
- \isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0 :: nat, 0 :: nat)"}\\[3mm]
- \isacommand{fun}~~@{text "add_pair"}~~\isacommand{where}~~%
- @{text "add_pair (m, n) (p, q) \<equiv> (m + p :: nat, n + q :: nat)"}\\
- \isacommand{quotient\_definition}~~@{text "+ :: int \<Rightarrow> int \<Rightarrow> int"}~~%
- \isacommand{is}~~@{text "add_pair"}\\
- \end{tabular}
- \end{isabelle}
-
- \noindent
- The following theorem about addition on the raw level can be proved.
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
- \isacommand{lemma}~~@{text "add_pair_zero: int_rel (add_pair (0, 0) x) x"}
- \end{isabelle}
-
- \noindent
- If the user lifts this theorem, the quotient package performs all the lifting
- automatically leaving the respectfulness proof for the constant @{text "add_pair"}
- as the only remaining proof obligation. This property needs to be proved by the user:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
- \begin{tabular}{@ {}l}
- \isacommand{lemma}~~@{text "[quot_respect]:"}\\
- @{text "(int_rel \<doublearr> int_rel \<doublearr> int_rel) add_pair add_pair"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- It can be discharged automatically by Isabelle when hinting to unfold the definition
- of @{text "\<doublearr>"}.
- After this, the user can prove the lifted lemma as follows:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
- \isacommand{lemma}~~@{text "0 + (x :: int) = x"}~~\isacommand{by}~~@{text "lifting add_pair_zero"}
- \end{isabelle}
-
- \noindent
- or by using the completely automated mode stating just:
-
- \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
- \isacommand{thm}~~@{text "add_pair_zero[quot_lifted]"}
- \end{isabelle}
-
- \noindent
- Both methods give the same result, namely
-
- @{text [display, indent=10] "0 + x = x"}
-
- \noindent
- where @{text x} is of type integer.
- Although seemingly simple, arriving at this result without the help of a quotient
- package requires a substantial reasoning effort (see \cite{Paulson06}).
-*}
-
-section {* Conclusion and Related Work\label{sec:conc}*}
-
-text {*
-
- The code of the quotient package and the examples described here are already
- included in the standard distribution of Isabelle.\footnote{Available from
- \href{http://isabelle.in.tum.de/}{http://isabelle.in.tum.de/}.} The package is
- heavily used in the new version of Nominal Isabelle, which provides a
- convenient reasoning infrastructure for programming language calculi
- involving general binders. To achieve this, it builds types representing
- @{text \<alpha>}-equivalent terms. Earlier versions of Nominal Isabelle have been
- used successfully in formalisations of an equivalence checking algorithm for
- LF \cite{UrbanCheneyBerghofer08}, Typed
- Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for concurrency
- \cite{BengtsonParow09} and a strong normalisation result for cut-elimination
- in classical logic \cite{UrbanZhu08}.
-
-
- There is a wide range of existing literature for dealing with quotients
- in theorem provers. Slotosch~\cite{Slotosch97} implemented a mechanism that
- automatically defines quotient types for Isabelle/HOL. But he did not
- include theorem lifting. Harrison's quotient package~\cite{harrison-thesis}
- is the first one that is able to automatically lift theorems, however only
- first-order theorems (that is theorems where abstractions, quantifiers and
- variables do not involve functions that include the quotient type). There is
- also some work on quotient types in non-HOL based systems and logical
- frameworks, including theory interpretations in
- PVS~\cite{PVS:Interpretations}, new types in MetaPRL~\cite{Nogin02}, and
- setoids in Coq \cite{ChicliPS02}. Paulson showed a construction of
- quotients that does not require the Hilbert Choice operator, but also only
- first-order theorems can be lifted~\cite{Paulson06}. The most related work
- to our package is the package for HOL4 by Homeier~\cite{Homeier05}. He
- introduced most of the abstract notions about quotients and also deals with
- lifting of higher-order theorems. However, he cannot deal with quotient
- compositions (needed for lifting theorems about @{text flat}). Also, a
- number of his definitions, like @{text ABS}, @{text REP} and @{text INJ} etc
- only exist in \cite{Homeier05} as ML-code, not included in the paper.
- Like Homeier's, our quotient package can deal with partial equivalence
- relations, but for lack of space we do not describe the mechanisms
- needed for this kind of quotient constructions.
-
-%%% FIXME Reviewer 3 would like to know more about the lifting in Coq and PVS,
-%%% and some comparison. I don't think we have the space for any additions...
-
- One feature of our quotient package is that when lifting theorems, the user
- can precisely specify what the lifted theorem should look like. This feature
- is necessary, for example, when lifting an induction principle for two
- lists. Assuming this principle has as the conclusion a predicate of the
- form @{text "P xs ys"}, then we can precisely specify whether we want to
- quotient @{text "xs"} or @{text "ys"}, or both. We found this feature very
- useful in the new version of Nominal Isabelle, where such a choice is
- required to generate a reasoning infrastructure for alpha-equated terms.
-%%
-%% give an example for this
-%%
- \medskip
-
- \noindent
- {\bf Acknowledgements:} We would like to thank Peter Homeier for the many
- discussions about his HOL4 quotient package and explaining to us
- some of its finer points in the implementation. Without his patient
- help, this work would have been impossible.
-
-*}
-
-
-
-(*<*)
-end
-(*>*)
--- a/Quotient-Paper/Paper.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1364 +0,0 @@
-(*<*)
-theory Paper
-imports "Quotient"
- "~~/src/HOL/Library/Quotient_Syntax"
- "~~/src/HOL/Library/LaTeXsugar"
- "~~/src/HOL/Quotient_Examples/FSet"
-begin
-
-(****
-
-** things to do for the next version
-*
-* - what are quot_thms?
-* - what do all preservation theorems look like,
- in particular preservation for quotient
- compositions
- - explain how Quotient R Abs Rep is proved (j-version)
- - give an example where precise specification helps (core Haskell in nominal?)
-
- - Mention Andreas Lochbiler in Acknowledgements and 'desceding'.
-
-*)
-
-notation (latex output)
- rel_conj ("_ \<circ>\<circ>\<circ> _" [53, 53] 52) and
- pred_comp ("_ \<circ>\<circ> _" [1, 1] 30) and
- implies (infix "\<longrightarrow>" 100) and
- "==>" (infix "\<Longrightarrow>" 100) and
- map_fun ("_ \<singlearr> _" 51) and
- fun_rel ("_ \<doublearr> _" 51) and
- list_eq (infix "\<approx>" 50) and (* Not sure if we want this notation...? *)
- empty_fset ("\<emptyset>") and
- union_fset ("_ \<union> _") and
- insert_fset ("{_} \<union> _") and
- Cons ("_::_") and
- concat ("flat") and
- concat_fset ("\<Union>") and
- Quotient ("Quot _ _ _")
-
-
-
-ML {*
-fun nth_conj n (_, r) = nth (HOLogic.dest_conj r) n;
-
-fun style_lhs_rhs proj = Scan.succeed (fn ctxt => fn t =>
- let
- val concl =
- Object_Logic.drop_judgment (ProofContext.theory_of ctxt) (Logic.strip_imp_concl t)
- in
- case concl of (_ $ l $ r) => proj (l, r)
- | _ => error ("Binary operator expected in term: " ^ Syntax.string_of_term ctxt concl)
- end);
-*}
-
-setup {*
- Term_Style.setup "rhs1" (style_lhs_rhs (nth_conj 0)) #>
- Term_Style.setup "rhs2" (style_lhs_rhs (nth_conj 1)) #>
- Term_Style.setup "rhs3" (style_lhs_rhs (nth_conj 2))
-*}
-
-lemma insert_preserve2:
- shows "((rep_fset ---> (map rep_fset \<circ> rep_fset) ---> (abs_fset \<circ> map abs_fset)) op #) =
- (id ---> rep_fset ---> abs_fset) op #"
- by (simp add: fun_eq_iff abs_o_rep[OF Quotient_fset] map_id Quotient_abs_rep[OF Quotient_fset])
-
-lemma list_all2_symp:
- assumes a: "equivp R"
- and b: "list_all2 R xs ys"
- shows "list_all2 R ys xs"
-using list_all2_lengthD[OF b] b
-apply(induct xs ys rule: list_induct2)
-apply(auto intro: equivp_symp[OF a])
-done
-
-lemma concat_rsp_unfolded:
- "\<lbrakk>list_all2 list_eq a ba; list_eq ba bb; list_all2 list_eq bb b\<rbrakk> \<Longrightarrow> list_eq (concat a) (concat b)"
-proof -
- fix a b ba bb
- assume a: "list_all2 list_eq a ba"
- assume b: "list_eq ba bb"
- assume c: "list_all2 list_eq bb b"
- have "\<forall>x. (\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" proof
- fix x
- show "(\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" proof
- assume d: "\<exists>xa\<in>set a. x \<in> set xa"
- show "\<exists>xa\<in>set b. x \<in> set xa" by (rule concat_rsp_pre[OF a b c d])
- next
- assume e: "\<exists>xa\<in>set b. x \<in> set xa"
- have a': "list_all2 list_eq ba a" by (rule list_all2_symp[OF list_eq_equivp, OF a])
- have b': "list_eq bb ba" by (rule equivp_symp[OF list_eq_equivp, OF b])
- have c': "list_all2 list_eq b bb" by (rule list_all2_symp[OF list_eq_equivp, OF c])
- show "\<exists>xa\<in>set a. x \<in> set xa" by (rule concat_rsp_pre[OF c' b' a' e])
- qed
- qed
- then show "list_eq (concat a) (concat b)" by auto
-qed
-
-(*>*)
-
-
-section {* Introduction *}
-
-text {*
- \noindent
- One might think quotients have been studied to death, but in the context of
- theorem provers many questions concerning them are far from settled. In
- this paper we address the question of how to establish a convenient reasoning
- infrastructure
- for quotient constructions in the Isabelle/HOL
- theorem prover. Higher-Order Logic (HOL) consists
- of a small number of axioms and inference rules over a simply-typed
- term-language. Safe reasoning in HOL is ensured by two very restricted
- mechanisms for extending the logic: one is the definition of new constants
- in terms of existing ones; the other is the introduction of new types by
- identifying non-empty subsets in existing types. Previous work has shown how
- to use both mechanisms for dealing with quotient constructions in HOL (see
- \cite{Homeier05,Paulson06}). For example the integers in Isabelle/HOL are
- constructed by a quotient construction over the type @{typ "nat \<times> nat"} and
- the equivalence relation
-
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "(n\<^isub>1, n\<^isub>2) \<approx> (m\<^isub>1, m\<^isub>2) \<equiv> n\<^isub>1 + m\<^isub>2 = m\<^isub>1 + n\<^isub>2"}\hfill\numbered{natpairequiv}
- \end{isabelle}
-
- \noindent
- This constructions yields the new type @{typ int}, and definitions for @{text
- "0"} and @{text "1"} of type @{typ int} can be given in terms of pairs of
- natural numbers (namely @{text "(0, 0)"} and @{text "(1, 0)"}). Operations
- such as @{text "add"} with type @{typ "int \<Rightarrow> int \<Rightarrow> int"} can be defined in
- terms of operations on pairs of natural numbers (namely @{text
- "add_pair (n\<^isub>1, m\<^isub>1) (n\<^isub>2,
- m\<^isub>2) \<equiv> (n\<^isub>1 + n\<^isub>2, m\<^isub>1 + m\<^isub>2)"}).
- Similarly one can construct %%the type of
- finite sets, written @{term "\<alpha> fset"},
- by quotienting the type @{text "\<alpha> list"} according to the equivalence relation
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "xs \<approx> ys \<equiv> (\<forall>x. memb x xs \<longleftrightarrow> memb x ys)"}\hfill\numbered{listequiv}
- \end{isabelle}
-
- \noindent
- which states that two lists are equivalent if every element in one list is
- also member in the other. The empty finite set, written @{term "{||}"}, can
- then be defined as the empty list and the union of two finite sets, written
- @{text "\<union>"}, as list append.
-
- Quotients are important in a variety of areas, but they are really ubiquitous in
- the area of reasoning about programming language calculi. A simple example
- is the lambda-calculus, whose raw terms are defined as
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "t ::= x | t t | \<lambda>x.t"}%\hfill\numbered{lambda}
- \end{isabelle}
-
- \noindent
- The problem with this definition arises, for instance, when one attempts to
- prove formally the substitution lemma \cite{Barendregt81} by induction
- over the structure of terms. This can be fiendishly complicated (see
- \cite[Pages 94--104]{CurryFeys58} for some ``rough'' sketches of a proof
- about raw lambda-terms). In contrast, if we reason about
- $\alpha$-equated lambda-terms, that means terms quotient according to
- $\alpha$-equivalence, then the reasoning infrastructure provided,
- for example, by Nominal Isabelle %%\cite{UrbanKaliszyk11}
- makes the formal
- proof of the substitution lemma almost trivial.
-
- The difficulty is that in order to be able to reason about integers, finite
- sets or $\alpha$-equated lambda-terms one needs to establish a reasoning
- infrastructure by transferring, or \emph{lifting}, definitions and theorems
- from the raw type @{typ "nat \<times> nat"} to the quotient type @{typ int}
- (similarly for finite sets and $\alpha$-equated lambda-terms). This lifting
- usually requires a \emph{lot} of tedious reasoning effort \cite{Paulson06}.
- In principle it is feasible to do this work manually, if one has only a few quotient
- constructions at hand. But if they have to be done over and over again, as in
- Nominal Isabelle, then manual reasoning is not an option.
-
- The purpose of a \emph{quotient package} is to ease the lifting of theorems
- and automate the reasoning as much as possible. In the
- context of HOL, there have been a few quotient packages already
- \cite{harrison-thesis,Slotosch97}. The most notable one is by Homeier
- \cite{Homeier05} implemented in HOL4. The fundamental construction these
- quotient packages perform can be illustrated by the following picture:
-
-%%% FIXME: Referee 1 says:
-%%% Diagram is unclear. Firstly, isn't an existing type a "set (not sets) of raw elements"?
-%%% Secondly, isn't the _set of_ equivalence classes mapped to and from the new type?
-%%% Thirdly, what do the words "non-empty subset" refer to ?
-
-%%% Cezary: I like the diagram, maybe 'new type' could be outside, but otherwise
-%%% I wouldn't change it.
-
- \begin{center}
- \mbox{}\hspace{20mm}\begin{tikzpicture}[scale=0.9]
- %%\draw[step=2mm] (-4,-1) grid (4,1);
-
- \draw[very thick] (0.7,0.3) circle (4.85mm);
- \draw[rounded corners=1mm, very thick] ( 0.0,-0.9) rectangle ( 1.8, 0.9);
- \draw[rounded corners=1mm, very thick] (-1.95,0.8) rectangle (-2.9,-0.195);
-
- \draw (-2.0, 0.8) -- (0.7,0.8);
- \draw (-2.0,-0.195) -- (0.7,-0.195);
-
- \draw ( 0.7, 0.23) node {\begin{tabular}{@ {}c@ {}}equiv-\\[-1mm]clas.\end{tabular}};
- \draw (-2.45, 0.35) node {\begin{tabular}{@ {}c@ {}}new\\[-1mm]type\end{tabular}};
- \draw (1.8, 0.35) node[right=-0.1mm]
- {\begin{tabular}{@ {}l@ {}}existing\\[-1mm] type\\ (sets of raw elements)\end{tabular}};
- \draw (0.9, -0.55) node {\begin{tabular}{@ {}l@ {}}non-empty\\[-1mm]subset\end{tabular}};
-
- \draw[->, very thick] (-1.8, 0.36) -- (-0.1,0.36);
- \draw[<-, very thick] (-1.8, 0.16) -- (-0.1,0.16);
- \draw (-0.95, 0.26) node[above=0.4mm] {@{text Rep}};
- \draw (-0.95, 0.26) node[below=0.4mm] {@{text Abs}};
- \end{tikzpicture}
- \end{center}
-
- \noindent
- The starting point is an existing type, to which we refer as the
- \emph{raw type} and over which an equivalence relation is given by the user.
- With this input the package introduces a new type, to which we
- refer as the \emph{quotient type}. This type comes with an
- \emph{abstraction} and a \emph{representation} function, written @{text Abs}
- and @{text Rep}.\footnote{Actually slightly more basic functions are given;
- the functions @{text Abs} and @{text Rep} need to be derived from them. We
- will show the details later. } They relate elements in the
- existing type to elements in the new type, % and vice versa,
- and can be uniquely
- identified by their quotient type. For example for the integer quotient construction
- the types of @{text Abs} and @{text Rep} are
-
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "Abs :: nat \<times> nat \<Rightarrow> int"}\hspace{10mm}@{text "Rep :: int \<Rightarrow> nat \<times> nat"}
- \end{isabelle}
-
- \noindent
- We therefore often write @{text Abs_int} and @{text Rep_int} if the
- typing information is important.
-
- Every abstraction and representation function stands for an isomorphism
- between the non-empty subset and elements in the new type. They are
- necessary for making definitions involving the new type. For example @{text
- "0"} and @{text "1"} of type @{typ int} can be defined as
-
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "0 \<equiv> Abs_int (0, 0)"}\hspace{10mm}@{text "1 \<equiv> Abs_int (1, 0)"}
- \end{isabelle}
-
- \noindent
- Slightly more complicated is the definition of @{text "add"} having type
- @{typ "int \<Rightarrow> int \<Rightarrow> int"}. Its definition is as follows
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "add n m \<equiv> Abs_int (add_pair (Rep_int n) (Rep_int m))"}
- \hfill\numbered{adddef}
- \end{isabelle}
-
- \noindent
- where we take the representation of the arguments @{text n} and @{text m},
- add them according to the function @{text "add_pair"} and then take the
- abstraction of the result. This is all straightforward and the existing
- quotient packages can deal with such definitions. But what is surprising is
- that none of them can deal with slightly more complicated definitions involving
- \emph{compositions} of quotients. Such compositions are needed for example
- in case of quotienting lists to yield finite sets and the operator that
- flattens lists of lists, defined as follows
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{thm concat.simps(1)[THEN eq_reflection]}\hspace{10mm}
- @{thm concat.simps(2)[THEN eq_reflection, no_vars]}
- \end{isabelle}
-
- \noindent
- where @{text "@"} is the usual
- list append. We expect that the corresponding
- operator on finite sets, written @{term "fconcat"},
- builds finite unions of finite sets:
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{thm concat_empty_fset[THEN eq_reflection, no_vars]}\hspace{10mm}
- @{thm concat_insert_fset[THEN eq_reflection, no_vars]}
- \end{isabelle}
-
- \noindent
- The quotient package should automatically provide us with a definition for @{text "\<Union>"} in
- terms of @{text flat}, @{text Rep_fset} and @{text Abs_fset}. The problem is
- that the method used in the existing quotient
- packages of just taking the representation of the arguments and then taking
- the abstraction of the result is \emph{not} enough. The reason is that in case
- of @{text "\<Union>"} we obtain the incorrect definition
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "\<Union> S \<equiv> Abs_fset (flat (Rep_fset S))"}
- \end{isabelle}
-
- \noindent
- where the right-hand side is not even typable! This problem can be remedied in the
- existing quotient packages by introducing an intermediate step and reasoning
- about flattening of lists of finite sets. However, this remedy is rather
- cumbersome and inelegant in light of our work, which can deal with such
- definitions directly. The solution is that we need to build aggregate
- representation and abstraction functions, which in case of @{text "\<Union>"}
- generate the %%%following
- definition
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "\<Union> S \<equiv> Abs_fset (flat ((map_list Rep_fset \<circ> Rep_fset) S))"}
- \end{isabelle}
-
- \noindent
- where @{term map_list} is the usual mapping function for lists. In this paper we
- will present a formal definition of our aggregate abstraction and
- representation functions (this definition was omitted in \cite{Homeier05}).
- They generate definitions, like the one above for @{text "\<Union>"},
- according to the type of the raw constant and the type
- of the quotient constant. This means we also have to extend the notions
- of \emph{aggregate equivalence relation}, \emph{respectfulness} and \emph{preservation}
- from Homeier \cite{Homeier05}.
-
- In addition we are able to clearly specify what is involved
- in the lifting process (this was only hinted at in \cite{Homeier05} and
- implemented as a ``rough recipe'' in ML-code). A pleasing side-result
- is that our procedure for lifting theorems is completely deterministic
- following the structure of the theorem being lifted and the theorem
- on the quotient level. Space constraints, unfortunately, allow us to only
- sketch this part of our work in Section 5 and we defer the reader to a longer
- version for the details. However, we will give in Section 3 and 4 all
- definitions that specify the input and output data of our three-step
- lifting procedure. Appendix A gives an example how our quotient
- package works in practise.
-*}
-
-section {* Preliminaries and General\\ Quotients\label{sec:prelims} *}
-
-text {*
- \noindent
- We will give in this section a crude overview of HOL and describe the main
- definitions given by Homeier for quotients \cite{Homeier05}.
-
- At its core, HOL is based on a simply-typed term language, where types are
- recorded in Church-style fashion (that means, we can always infer the type of
- a term and its subterms without any additional information). The grammars
- for types and terms are
-
- \begin{isabelle}\ \ \ \ \ %%%
- \begin{tabular}{@ {}c@ {\hspace{10mm}}c@ {}}
- @{text "\<sigma>, \<tau> ::= \<alpha> | (\<sigma>,\<dots>, \<sigma>) \<kappa>"} &
- @{text "t, s ::= x\<^isup>\<sigma> | c\<^isup>\<sigma> | t t | \<lambda>x\<^isup>\<sigma>. t"}\\
- \end{tabular}
- \end{isabelle}
-
- \noindent
- with types being either type variables or type constructors and terms
- being variables, constants, applications or abstractions.
- We often write just @{text \<kappa>} for @{text "() \<kappa>"}, and use @{text "\<alpha>s"} and
- @{text "\<sigma>s"} to stand for collections of type variables and types,
- respectively. The type of a term is often made explicit by writing @{text
- "t :: \<sigma>"}. HOL includes a type @{typ bool} for booleans and the function
- type, written @{text "\<sigma> \<Rightarrow> \<tau>"}. HOL also contains many primitive and defined
- constants; for example, a primitive constant is equality, with type @{text "= :: \<sigma> \<Rightarrow> \<sigma> \<Rightarrow>
- bool"}, and the identity function with type @{text "id :: \<sigma> \<Rightarrow> \<sigma>"} is
- defined as @{text "\<lambda>x\<^sup>\<sigma>. x\<^sup>\<sigma>"}.
-
- An important point to note is that theorems in HOL can be seen as a subset
- of terms that are constructed specially (namely through axioms and proof
- rules). As a result we are able to define automatic proof
- procedures showing that one theorem implies another by decomposing the term
- underlying the first theorem.
-
- Like Homeier's, our work relies on map-functions defined for every type
- constructor taking some arguments, for example @{text map_list} for lists. Homeier
- describes in \cite{Homeier05} map-functions for products, sums, options and
- also the following map for function types
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{thm map_fun_def[no_vars, THEN eq_reflection]}
- \end{isabelle}
-
- \noindent
- Using this map-function, we can give the following, equivalent, but more
- uniform definition for @{text add} shown in \eqref{adddef}:
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "add \<equiv> (Rep_int \<singlearr> Rep_int \<singlearr> Abs_int) add_pair"}
- \end{isabelle}
-
- \noindent
- Using extensionality and unfolding the definition of @{text "\<singlearr>"},
- we can get back to \eqref{adddef}.
- In what follows we shall use the convention to write @{text "map_\<kappa>"} for a map-function
- of the type-constructor @{text \<kappa>}.
- %% a general type for map all types is difficult to give (algebraic types are
- %% easy, but for example the function type is not algebraic
- %For a type @{text \<kappa>} with arguments @{text "\<alpha>\<^isub>1\<^isub>\<dots>\<^isub>n"} the
- %type of the function @{text "map_\<kappa>"} has to be @{text "\<alpha>\<^isub>1\<Rightarrow>\<dots>\<Rightarrow>\<alpha>\<^isub>n\<Rightarrow>\<alpha>\<^isub>1\<dots>\<alpha>\<^isub>n \<kappa>"}.
- %For example @{text "map_list"}
- %has to have the type @{text "\<alpha>\<Rightarrow>\<alpha> list"}.
- In our implementation we maintain
- a database of these map-functions that can be dynamically extended.
-
- It will also be necessary to have operators, referred to as @{text "rel_\<kappa>"},
- which define equivalence relations in terms of constituent equivalence
- relations. For example given two equivalence relations @{text "R\<^isub>1"}
- and @{text "R\<^isub>2"}, we can define an equivalence relations over
- products as %% follows
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "(R\<^isub>1 \<tripple> R\<^isub>2) (x\<^isub>1, x\<^isub>2) (y\<^isub>1, y\<^isub>2) \<equiv> R\<^isub>1 x\<^isub>1 y\<^isub>1 \<and> R\<^isub>2 x\<^isub>2 y\<^isub>2"}
- \end{isabelle}
-
- \noindent
- Homeier gives also the following operator for defining equivalence
- relations over function types
- %
- \begin{isabelle}\ \ \ \ \ %%%
- @{thm fun_rel_def[of "R\<^isub>1" "R\<^isub>2", no_vars, THEN eq_reflection]}
- \hfill\numbered{relfun}
- \end{isabelle}
-
- \noindent
- In the context of quotients, the following two notions from \cite{Homeier05}
- are needed later on.
-
- \begin{definition}[Respects]\label{def:respects}
- An element @{text "x"} respects a relation @{text "R"} provided @{text "R x x"}.
- \end{definition}
-
- \begin{definition}[Bounded $\forall$ and $\lambda$]\label{def:babs}
- @{text "\<forall>x \<in> S. P x"} holds if for all @{text x}, @{text "x \<in> S"} implies @{text "P x"};
- and @{text "(\<lambda>x \<in> S. f x) = f x"} provided @{text "x \<in> S"}.
- \end{definition}
-
- The central definition in Homeier's work \cite{Homeier05} relates equivalence
- relations, abstraction and representation functions:
-
- \begin{definition}[Quotient Types]
- Given a relation $R$, an abstraction function $Abs$
- and a representation function $Rep$, the predicate @{term "Quotient R Abs Rep"}
- holds if and only if
- \begin{isabelle}\ \ \ \ \ %%%%
- \begin{tabular}{rl}
- (i) & \begin{isa}@{thm (rhs1) Quotient_def[of "R", no_vars]}\end{isa}\\
- (ii) & \begin{isa}@{thm (rhs2) Quotient_def[of "R", no_vars]}\end{isa}\\
- (iii) & \begin{isa}@{thm (rhs3) Quotient_def[of "R", no_vars]}\end{isa}\\
- \end{tabular}
- \end{isabelle}
- \end{definition}
-
- \noindent
- The value of this definition lies in the fact that validity of @{term "Quotient R Abs Rep"} can
- often be proved in terms of the validity of @{term "Quot"} over the constituent
- types of @{text "R"}, @{text Abs} and @{text Rep}.
- For example Homeier proves the following property for higher-order quotient
- types:
-
- \begin{proposition}\label{funquot}
- \begin{isa}
- @{thm[mode=IfThen] fun_quotient[where ?R1.0="R\<^isub>1" and ?R2.0="R\<^isub>2"
- and ?abs1.0="Abs\<^isub>1" and ?abs2.0="Abs\<^isub>2" and ?rep1.0="Rep\<^isub>1" and ?rep2.0="Rep\<^isub>2"]}
- \end{isa}
- \end{proposition}
-
- \noindent
- As a result, Homeier is able to build an automatic prover that can nearly
- always discharge a proof obligation involving @{text "Quot"}. Our quotient
- package makes heavy
- use of this part of Homeier's work including an extension
- for dealing with \emph{conjugations} of equivalence relations\footnote{That are
- symmetric by definition.} defined as follows:
-
-%%% FIXME Referee 2 claims that composition-of-relations means OO, and this is also
-%%% what wikipedia says. Any idea for a different name? Conjugation of Relations?
-
- \begin{definition}%%[Composition of Relations]
- @{abbrev "rel_conj R\<^isub>1 R\<^isub>2"} where @{text "\<circ>\<circ>"} is the predicate
- composition defined by
- @{thm (concl) pred_compI[of "R\<^isub>1" "x" "y" "R\<^isub>2" "z"]}
- holds if and only if there exists a @{text y} such that @{thm (prem 1) pred_compI[of "R\<^isub>1" "x" "y" "R\<^isub>2" "z"]} and
- @{thm (prem 2) pred_compI[of "R\<^isub>1" "x" "y" "R\<^isub>2" "z"]}.
- \end{definition}
-
- \noindent
- Unfortunately a general quotient theorem for @{text "\<circ>\<circ>\<circ>"}, analogous to the one
- for @{text "\<singlearr>"} given in Proposition \ref{funquot}, would not be true
- in general. It cannot even be stated inside HOL, because of restrictions on types.
- However, we can prove specific instances of a
- quotient theorem for composing particular quotient relations.
- For example, to lift theorems involving @{term flat} the quotient theorem for
- composing @{text "\<approx>\<^bsub>list\<^esub>"} will be necessary: given @{term "Quotient R Abs Rep"}
- with @{text R} being an equivalence relation, then
-
- \begin{isabelle}\ \ \ \ \ %%%
- \begin{tabular}{r@ {\hspace{1mm}}l}
- @{text "Quot"} & @{text "(rel_list R \<circ>\<circ>\<circ> \<approx>\<^bsub>list\<^esub>)"}\\
- & @{text "(Abs_fset \<circ> map_list Abs)"} @{text "(map_list Rep \<circ> Rep_fset)"}\\
- \end{tabular}
- \end{isabelle}
-*}
-
-section {* Quotient Types and Quotient\\ Definitions\label{sec:type} *}
-
-text {*
- \noindent
- The first step in a quotient construction is to take a name for the new
- type, say @{text "\<kappa>\<^isub>q"}, and an equivalence relation, say @{text R},
- defined over a raw type, say @{text "\<sigma>"}. The type of the equivalence
- relation must be @{text "\<sigma> \<Rightarrow> \<sigma> \<Rightarrow> bool"}. The user-visible part of
- the quotient type declaration is therefore
-
- \begin{isabelle}\ \ \ \ \ %%%
- \isacommand{quotient\_type}~~@{text "\<alpha>s \<kappa>\<^isub>q = \<sigma> / R"}\hfill\numbered{typedecl}
- \end{isabelle}
-
- \noindent
- and a proof that @{text "R"} is indeed an equivalence relation. The @{text "\<alpha>s"}
- indicate the arity of the new type and the type-variables of @{text "\<sigma>"} can only
- be contained in @{text "\<alpha>s"}. Two concrete
- examples are
-
-
- \begin{isabelle}\ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- \isacommand{quotient\_type}~~@{text "int = nat \<times> nat / \<approx>\<^bsub>nat \<times> nat\<^esub>"}\\
- \isacommand{quotient\_type}~~@{text "\<alpha> fset = \<alpha> list / \<approx>\<^bsub>list\<^esub>"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- which introduce the type of integers and of finite sets using the
- equivalence relations @{text "\<approx>\<^bsub>nat \<times> nat\<^esub>"} and @{text
- "\<approx>\<^bsub>list\<^esub>"} defined in \eqref{natpairequiv} and
- \eqref{listequiv}, respectively (the proofs about being equivalence
- relations are omitted). Given this data, we define for declarations shown in
- \eqref{typedecl} the quotient types internally as
-
- \begin{isabelle}\ \ \ \ \ %%%
- \isacommand{typedef}~~@{text "\<alpha>s \<kappa>\<^isub>q = {c. \<exists>x. c = R x}"}
- \end{isabelle}
-
- \noindent
- where the right-hand side is the (non-empty) set of equivalence classes of
- @{text "R"}. The constraint in this declaration is that the type variables
- in the raw type @{text "\<sigma>"} must be included in the type variables @{text
- "\<alpha>s"} declared for @{text "\<kappa>\<^isub>q"}. HOL will then provide us with the following
- abstraction and representation functions
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "abs_\<kappa>\<^isub>q :: \<sigma> set \<Rightarrow> \<alpha>s \<kappa>\<^isub>q"}\hspace{10mm}@{text "rep_\<kappa>\<^isub>q :: \<alpha>s \<kappa>\<^isub>q \<Rightarrow> \<sigma> set"}
- \end{isabelle}
-
- \noindent
- As can be seen from the type, they relate the new quotient type and equivalence classes of the raw
- type. However, as Homeier \cite{Homeier05} noted, it is much more convenient
- to work with the following derived abstraction and representation functions
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "Abs_\<kappa>\<^isub>q x \<equiv> abs_\<kappa>\<^isub>q (R x)"}\hspace{10mm}@{text "Rep_\<kappa>\<^isub>q x \<equiv> \<epsilon> (rep_\<kappa>\<^isub>q x)"}
- \end{isabelle}
-
- \noindent
- on the expense of having to use Hilbert's choice operator @{text "\<epsilon>"} in the
- definition of @{text "Rep_\<kappa>\<^isub>q"}. These derived notions relate the
- quotient type and the raw type directly, as can be seen from their type,
- namely @{text "\<sigma> \<Rightarrow> \<alpha>s \<kappa>\<^isub>q"} and @{text "\<alpha>s \<kappa>\<^isub>q \<Rightarrow> \<sigma>"},
- respectively. Given that @{text "R"} is an equivalence relation, the
- following property holds for every quotient type
- (for the proof see \cite{Homeier05}).
-
- \begin{proposition}
- \begin{isa}@{term "Quotient R Abs_\<kappa>\<^isub>q Rep_\<kappa>\<^isub>q"}.\end{isa}
- \end{proposition}
-
- The next step in a quotient construction is to introduce definitions of new constants
- involving the quotient type. These definitions need to be given in terms of concepts
- of the raw type (remember this is the only way how to extend HOL
- with new definitions). For the user the visible part of such definitions is the declaration
-
- \begin{isabelle}\ \ \ \ \ %%%
- \isacommand{quotient\_definition}~~@{text "c :: \<tau>"}~~\isacommand{is}~~@{text "t :: \<sigma>"}
- \end{isabelle}
-
- \noindent
- where @{text t} is the definiens (its type @{text \<sigma>} can always be inferred)
- and @{text "c"} is the name of definiendum, whose type @{text "\<tau>"} needs to be
- given explicitly (the point is that @{text "\<tau>"} and @{text "\<sigma>"} can only differ
- in places where a quotient and raw type is involved). Two concrete examples are
-
- \begin{isabelle}\ \ \ \ \ %%%
- \begin{tabular}{@ {}l}
- \isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0::nat, 0::nat)"}\\
- \isacommand{quotient\_definition}~~@{text "\<Union> :: (\<alpha> fset) fset \<Rightarrow> \<alpha> fset"}~~%
- \isacommand{is}~~@{text "flat"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- The first one declares zero for integers and the second the operator for
- building unions of finite sets (@{text "flat"} having the type
- @{text "(\<alpha> list) list \<Rightarrow> \<alpha> list"}).
-
- From such declarations given by the user, the quotient package needs to derive proper
- definitions using @{text "Abs"} and @{text "Rep"}. The data we rely on is the given quotient type
- @{text "\<tau>"} and the raw type @{text "\<sigma>"}. They allow us to define \emph{aggregate
- abstraction} and \emph{representation functions} using the functions @{text "ABS (\<sigma>,
- \<tau>)"} and @{text "REP (\<sigma>, \<tau>)"} whose clauses we shall give below. The idea behind
- these two functions is to simultaneously descend into the raw types @{text \<sigma>} and
- quotient types @{text \<tau>}, and generate the appropriate
- @{text "Abs"} and @{text "Rep"} in places where the types differ. Therefore
- we generate just the identity whenever the types are equal. On the ``way'' down,
- however we might have to use map-functions to let @{text Abs} and @{text Rep} act
- over the appropriate types. In what follows we use the short-hand notation
- @{text "ABS (\<sigma>s, \<tau>s)"} to mean @{text "ABS (\<sigma>\<^isub>1, \<tau>\<^isub>1)\<dots>ABS (\<sigma>\<^isub>n, \<tau>\<^isub>n)"}; similarly
- for @{text REP}.
- %
- \begin{center}
- \hfill
- \begin{tabular}{@ {\hspace{2mm}}l@ {}}
- \multicolumn{1}{@ {}l}{equal types:}\\
- @{text "ABS (\<sigma>, \<sigma>)"} $\dn$ @{text "id :: \<sigma> \<Rightarrow> \<sigma>"}\hspace{5mm}%\\
- @{text "REP (\<sigma>, \<sigma>)"} $\dn$ @{text "id :: \<sigma> \<Rightarrow> \<sigma>"}\smallskip\\
- \multicolumn{1}{@ {}l}{function types:}\\
- @{text "ABS (\<sigma>\<^isub>1 \<Rightarrow> \<sigma>\<^isub>2, \<tau>\<^isub>1 \<Rightarrow> \<tau>\<^isub>2)"} $\dn$ @{text "REP (\<sigma>\<^isub>1, \<tau>\<^isub>1) \<singlearr> ABS (\<sigma>\<^isub>2, \<tau>\<^isub>2)"}\\
- @{text "REP (\<sigma>\<^isub>1 \<Rightarrow> \<sigma>\<^isub>2, \<tau>\<^isub>1 \<Rightarrow> \<tau>\<^isub>2)"} $\dn$ @{text "ABS (\<sigma>\<^isub>1, \<tau>\<^isub>1) \<singlearr> REP (\<sigma>\<^isub>2, \<tau>\<^isub>2)"}\smallskip\\
- \multicolumn{1}{@ {}l}{equal type constructors:}\\
- @{text "ABS (\<sigma>s \<kappa>, \<tau>s \<kappa>)"} $\dn$ @{text "map_\<kappa> (ABS (\<sigma>s, \<tau>s))"}\\
- @{text "REP (\<sigma>s \<kappa>, \<tau>s \<kappa>)"} $\dn$ @{text "map_\<kappa> (REP (\<sigma>s, \<tau>s))"}\smallskip\\
- \multicolumn{1}{@ {}l}{unequal type constructors:}\\
- @{text "ABS (\<sigma>s \<kappa>, \<tau>s \<kappa>\<^isub>q)"} $\dn$ @{text "Abs_\<kappa>\<^isub>q \<circ> (MAP(\<rho>s \<kappa>) (ABS (\<sigma>s', \<tau>s)))"}\\
- @{text "REP (\<sigma>s \<kappa>, \<tau>s \<kappa>\<^isub>q)"} $\dn$ @{text "(MAP(\<rho>s \<kappa>) (REP (\<sigma>s', \<tau>s))) \<circ> Rep_\<kappa>\<^isub>q"}
- \end{tabular}\hfill\numbered{ABSREP}
- \end{center}
- %
- \noindent
- In the last two clauses are subtle. We rely in them on the fact that the type @{text "\<alpha>s
- \<kappa>\<^isub>q"} is the quotient of the raw type @{text "\<rho>s \<kappa>"} (for example
- @{text "int"} and @{text "nat \<times> nat"}, or @{text "\<alpha> fset"} and @{text "\<alpha>
- list"}). This data is given by declarations shown in \eqref{typedecl}.
- The quotient construction ensures that the type variables in @{text
- "\<rho>s \<kappa>"} must be among the @{text "\<alpha>s"}. The @{text "\<sigma>s'"} are given by the
- substitutions for the @{text "\<alpha>s"} when matching @{text "\<sigma>s \<kappa>"} against
- @{text "\<rho>s \<kappa>"}. This calculation determines what are the types in place
- of the type variables @{text "\<alpha>s"} in the instance of
- quotient type @{text "\<alpha>s \<kappa>\<^isub>q"}---namely @{text "\<tau>s"}, and the corresponding
- types in place of the @{text "\<alpha>s"} in the raw type @{text "\<rho>s \<kappa>"}---namely @{text "\<sigma>s'"}. The
- function @{text "MAP"} calculates an \emph{aggregate map-function} for a raw
- type as follows:
- %
- \begin{center}
- \begin{tabular}{r@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
- @{text "MAP' (\<alpha>)"} & $\dn$ & @{text "a\<^sup>\<alpha>"}\\
- @{text "MAP' (\<kappa>)"} & $\dn$ & @{text "id :: \<kappa> \<Rightarrow> \<kappa>"}\\
- @{text "MAP' (\<sigma>s \<kappa>)"} & $\dn$ & @{text "map_\<kappa> (MAP'(\<sigma>s))"}\smallskip\\
- @{text "MAP (\<sigma>)"} & $\dn$ & @{text "\<lambda>as. MAP'(\<sigma>)"}
- \end{tabular}
- \end{center}
- %
- \noindent
- In this definition we rely on the fact that in the first clause we can interpret type-variables @{text \<alpha>} as
- term variables @{text a}. In the last clause we build an abstraction over all
- term-variables of the map-function generated by the auxiliary function
- @{text "MAP'"}.
- The need for aggregate map-functions can be seen in cases where we build quotients,
- say @{text "(\<alpha>, \<beta>) \<kappa>\<^isub>q"}, out of compound raw types, say @{text "(\<alpha> list) \<times> \<beta>"}.
- In this case @{text MAP} generates the
- aggregate map-function:
-
-%%% FIXME: Reviewer 2 asks: last two lines defining ABS and REP for
-%%% unequal type constructors: How are the $\varrho$s defined? The
-%%% following paragraph mentions them, but this paragraph is unclear,
-%%% since it then mentions $\alpha$s, which do not seem to be defined
-%%% either. As a result, I do not understand the first two sentences
-%%% in this paragraph. I can imagine roughly what the following
-%%% sentence `The $\sigma$s' are given by the matchers for the
-%%% $\alpha$s$ when matching $\varrho$s $\kappa$ against $\sigma$s
-%%% $\kappa$.' means, but also think that it is too vague.
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "\<lambda>a b. map_prod (map_list a) b"}
- \end{isabelle}
-
- \noindent
- which is essential in order to define the corresponding aggregate
- abstraction and representation functions.
-
- To see how these definitions pan out in practise, let us return to our
- example about @{term "concat"} and @{term "fconcat"}, where we have the raw type
- @{text "(\<alpha> list) list \<Rightarrow> \<alpha> list"} and the quotient type @{text "(\<alpha> fset) fset \<Rightarrow> \<alpha>
- fset"}. Feeding these types into @{text ABS} gives us (after some @{text "\<beta>"}-simplifications)
- the abstraction function
-
- \begin{isabelle}\ \ \ %%%
- \begin{tabular}{l}
- @{text "(map_list (map_list id \<circ> Rep_fset) \<circ> Rep_fset) \<singlearr>"}\\
- \mbox{}\hspace{4.5cm}@{text " Abs_fset \<circ> map_list id"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- In our implementation we further
- simplify this function by rewriting with the usual laws about @{text
- "map"}s and @{text "id"}, for example @{term "map_list id = id"} and @{text "f \<circ> id =
- id \<circ> f = f"}. This gives us the simpler abstraction function
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "(map_list Rep_fset \<circ> Rep_fset) \<singlearr> Abs_fset"}
- \end{isabelle}
-
- \noindent
- which we can use for defining @{term "fconcat"} as follows
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "\<Union> \<equiv> ((map_list Rep_fset \<circ> Rep_fset) \<singlearr> Abs_fset) flat"}
- \end{isabelle}
-
- \noindent
- Note that by using the operator @{text "\<singlearr>"} and special clauses
- for function types in \eqref{ABSREP}, we do not have to
- distinguish between arguments and results, but can deal with them uniformly.
- Consequently, all definitions in the quotient package
- are of the general form
-
- \begin{isabelle}\ \ \ \ \ %%%
- \mbox{@{text "c \<equiv> ABS (\<sigma>, \<tau>) t"}}
- \end{isabelle}
-
- \noindent
- where @{text \<sigma>} is the type of the definiens @{text "t"} and @{text "\<tau>"} the
- type of the defined quotient constant @{text "c"}. This data can be easily
- generated from the declaration given by the user.
- To increase the confidence in this way of making definitions, we can prove
- that the terms involved are all typable.
-
- \begin{lemma}
- If @{text "ABS (\<sigma>, \<tau>)"} returns some abstraction function @{text "Abs"}
- and @{text "REP (\<sigma>, \<tau>)"} some representation function @{text "Rep"},
- then @{text "Abs"} is of type @{text "\<sigma> \<Rightarrow> \<tau>"} and @{text "Rep"} of type
- @{text "\<tau> \<Rightarrow> \<sigma>"}.
- \end{lemma}
-
- \begin{proof}
- By mutual induction and analysing the definitions of @{text "ABS"} and @{text "REP"}.
- The cases of equal types and function types are
- straightforward (the latter follows from @{text "\<singlearr>"} having the
- type @{text "(\<alpha> \<Rightarrow> \<beta>) \<Rightarrow> (\<gamma> \<Rightarrow> \<delta>) \<Rightarrow> (\<beta> \<Rightarrow> \<gamma>) \<Rightarrow> (\<alpha> \<Rightarrow> \<delta>)"}). In case of equal type
- constructors we can observe that a map-function after applying the functions
- @{text "ABS (\<sigma>s, \<tau>s)"} produces a term of type @{text "\<sigma>s \<kappa> \<Rightarrow> \<tau>s \<kappa>"}. The
- interesting case is the one with unequal type constructors. Since we know
- the quotient is between @{text "\<alpha>s \<kappa>\<^isub>q"} and @{text "\<rho>s \<kappa>"}, we have
- that @{text "Abs_\<kappa>\<^isub>q"} is of type @{text "\<rho>s \<kappa> \<Rightarrow> \<alpha>s
- \<kappa>\<^isub>q"}. This type can be more specialised to @{text "\<rho>s[\<tau>s] \<kappa> \<Rightarrow> \<tau>s
- \<kappa>\<^isub>q"} where the type variables @{text "\<alpha>s"} are instantiated with the
- @{text "\<tau>s"}. The complete type can be calculated by observing that @{text
- "MAP (\<rho>s \<kappa>)"}, after applying the functions @{text "ABS (\<sigma>s', \<tau>s)"} to it,
- returns a term of type @{text "\<rho>s[\<sigma>s'] \<kappa> \<Rightarrow> \<rho>s[\<tau>s] \<kappa>"}. This type is
- equivalent to @{text "\<sigma>s \<kappa> \<Rightarrow> \<rho>s[\<tau>s] \<kappa>"}, which we just have to compose with
- @{text "\<rho>s[\<tau>s] \<kappa> \<Rightarrow> \<tau>s \<kappa>\<^isub>q"} according to the type of @{text "\<circ>"}.\qed
- \end{proof}
-*}
-
-section {* Respectfulness and\\ Preservation \label{sec:resp} *}
-
-text {*
- \noindent
- The main point of the quotient package is to automatically ``lift'' theorems
- involving constants over the raw type to theorems involving constants over
- the quotient type. Before we can describe this lifting process, we need to impose
- two restrictions in form of proof obligations that arise during the
- lifting. The reason is that even if definitions for all raw constants
- can be given, \emph{not} all theorems can be lifted to the quotient type. Most
- notable is the bound variable function, that is the constant @{text bn},
- defined
- for raw lambda-terms as follows
-
- \begin{isabelle}
- \begin{center}
- @{text "bn (x) \<equiv> \<emptyset>"}\hspace{4mm}
- @{text "bn (t\<^isub>1 t\<^isub>2) \<equiv> bn (t\<^isub>1) \<union> bn (t\<^isub>2)"}\smallskip\\
- @{text "bn (\<lambda>x. t) \<equiv> {x} \<union> bn (t)"}
- \end{center}
- \end{isabelle}
-
- \noindent
- We can generate a definition for this constant using @{text ABS} and @{text REP}.
- But this constant does \emph{not} respect @{text "\<alpha>"}-equivalence and
- consequently no theorem involving this constant can be lifted to @{text
- "\<alpha>"}-equated lambda terms. Homeier formulates the restrictions in terms of
- the properties of \emph{respectfulness} and \emph{preservation}. We have
- to slightly extend Homeier's definitions in order to deal with quotient
- compositions.
-
-%%% FIXME: Reviewer 3 asks why are the definitions that follow enough to deal
-%%% with quotient composition.
-
- To formally define what respectfulness is, we have to first define
- the notion of \emph{aggregate equivalence relations} using the function @{text "REL(\<sigma>, \<tau>)"}
- The idea behind this function is to simultaneously descend into the raw types
- @{text \<sigma>} and quotient types @{text \<tau>}, and generate the appropriate
- quotient equivalence relations in places where the types differ and equalities
- elsewhere.
-
- \begin{center}
- \hfill
- \begin{tabular}{l}
- \multicolumn{1}{@ {}l}{equal types:}
- @{text "REL (\<sigma>, \<sigma>)"} $\dn$ @{text "= :: \<sigma> \<Rightarrow> \<sigma> \<Rightarrow> bool"}\smallskip\\
- \multicolumn{1}{@ {}l}{equal type constructors:}\\
- @{text "REL (\<sigma>s \<kappa>, \<tau>s \<kappa>)"} $\dn$ @{text "rel_\<kappa> (REL (\<sigma>s, \<tau>s))"}\smallskip\\
- \multicolumn{1}{@ {}l}{unequal type constructors:}\\
- @{text "REL (\<sigma>s \<kappa>, \<tau>s \<kappa>\<^isub>q)"} $\dn$ @{text "rel_\<kappa>\<^isub>q (REL (\<sigma>s', \<tau>s))"}\\
- \end{tabular}\hfill\numbered{REL}
- \end{center}
-
- \noindent
- The @{text "\<sigma>s'"} in the last clause are calculated as in \eqref{ABSREP}:
- again we know that type @{text "\<alpha>s \<kappa>\<^isub>q"} is the quotient of the raw type
- @{text "\<rho>s \<kappa>"}. The @{text "\<sigma>s'"} are the substitutions for @{text "\<alpha>s"} obtained by matching
- @{text "\<rho>s \<kappa>"} and @{text "\<sigma>s \<kappa>"}.
-
- Let us return to the lifting procedure of theorems. Assume we have a theorem
- that contains the raw constant @{text "c\<^isub>r :: \<sigma>"} and which we want to
- lift to a theorem where @{text "c\<^isub>r"} is replaced by the corresponding
- constant @{text "c\<^isub>q :: \<tau>"} defined over a quotient type. In this situation
- we generate the following proof obligation
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "REL (\<sigma>, \<tau>) c\<^isub>r c\<^isub>r"}.
- \end{isabelle}
-
- \noindent
- Homeier calls these proof obligations \emph{respectfulness
- theorems}. However, unlike his quotient package, we might have several
- respectfulness theorems for one constant---he has at most one.
- The reason is that because of our quotient compositions, the types
- @{text \<sigma>} and @{text \<tau>} are not completely determined by @{text "c\<^bsub>r\<^esub>"}.
- And for every instantiation of the types, a corresponding
- respectfulness theorem is necessary.
-
- Before lifting a theorem, we require the user to discharge
- respectfulness proof obligations. In case of @{text bn}
- this obligation is %%as follows
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "(\<approx>\<^isub>\<alpha> \<doublearr> =) bn bn"}
- \end{isabelle}
-
- \noindent
- and the point is that the user cannot discharge it: because it is not true. To see this,
- we can just unfold the definition of @{text "\<doublearr>"} \eqref{relfun}
- using extensionality to obtain the false statement
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "\<forall>t\<^isub>1 t\<^isub>2. if t\<^isub>1 \<approx>\<^isub>\<alpha> t\<^isub>2 then bn(t\<^isub>1) = bn(t\<^isub>2)"}
- \end{isabelle}
-
- \noindent
- In contrast, lifting a theorem about @{text "append"} to a theorem describing
- the union of finite sets will mean to discharge the proof obligation
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "(\<approx>\<^bsub>list\<^esub> \<doublearr> \<approx>\<^bsub>list\<^esub> \<doublearr> \<approx>\<^bsub>list\<^esub>) append append"}
- \end{isabelle}
-
- \noindent
- To do so, we have to establish
-
- \begin{isabelle}\ \ \ \ \ %%%
- if @{text "xs \<approx>\<^bsub>list\<^esub> ys"} and @{text "us \<approx>\<^bsub>list\<^esub> vs"}
- then @{text "xs @ us \<approx>\<^bsub>list\<^esub> ys @ vs"}
- \end{isabelle}
-
- \noindent
- which is straightforward given the definition shown in \eqref{listequiv}.
-
- The second restriction we have to impose arises from non-lifted polymorphic
- constants, which are instantiated to a type being quotient. For example,
- take the @{term "cons"}-constructor to add a pair of natural numbers to a
- list, whereby we assume the pair of natural numbers turns into an integer in
- the quotient construction. The point is that we still want to use @{text
- cons} for adding integers to lists---just with a different type. To be able
- to lift such theorems, we need a \emph{preservation property} for @{text
- cons}. Assuming we have a polymorphic raw constant @{text "c\<^isub>r :: \<sigma>"}
- and a corresponding quotient constant @{text "c\<^isub>q :: \<tau>"}, then a
- preservation property is as follows
-
-%%% FIXME: Reviewer 2 asks: You say what a preservation theorem is,
-%%% but not which preservation theorems you assume. Do you generate a
-%%% proof obligation for a preservation theorem for each raw constant
-%%% and its corresponding lifted constant?
-
-%%% Cezary: I think this would be a nice thing to do but we have not
-%%% done it, the theorems need to be 'guessed' from the remaining obligations
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "Quot R\<^bsub>\<alpha>s\<^esub> Abs\<^bsub>\<alpha>s\<^esub> Rep\<^bsub>\<alpha>s\<^esub> implies ABS (\<sigma>, \<tau>) c\<^isub>r = c\<^isub>r"}
- \end{isabelle}
-
- \noindent
- where the @{text "\<alpha>s"} stand for the type variables in the type of @{text "c\<^isub>r"}.
- In case of @{text cons} (which has type @{text "\<alpha> \<Rightarrow> \<alpha> list \<Rightarrow> \<alpha> list"}) we have
-
- \begin{isabelle}\ \ \ \ \ %%%
- @{text "(Rep \<singlearr> map_list Rep \<singlearr> map_list Abs) cons = cons"}
- \end{isabelle}
-
- \noindent
- under the assumption @{term "Quotient R Abs Rep"}. The point is that if we have
- an instance of @{text cons} where the type variable @{text \<alpha>} is instantiated
- with @{text "nat \<times> nat"} and we also quotient this type to yield integers,
- then we need to show this preservation property.
-
- %%%@ {thm [display, indent=10] insert_preserve2[no_vars]}
-
- %Given two quotients, one of which quotients a container, and the
- %other quotients the type in the container, we can write the
- %composition of those quotients. To compose two quotient theorems
- %we compose the relations with relation composition as defined above
- %and the abstraction and relation functions are the ones of the sub
- %quotients composed with the usual function composition.
- %The @ {term "Rep"} and @ {term "Abs"} functions that we obtain agree
- %with the definition of aggregate Abs/Rep functions and the
- %relation is the same as the one given by aggregate relations.
- %This becomes especially interesting
- %when we compose the quotient with itself, as there is no simple
- %intermediate step.
- %
- %Lets take again the example of @ {term flat}. To be able to lift
- %theorems that talk about it we provide the composition quotient
- %theorem which allows quotienting inside the container:
- %
- %If @ {term R} is an equivalence relation and @ {term "Quotient R Abs Rep"}
- %then
- %
- %@ {text [display, indent=10] "Quotient (list_rel R \<circ>\<circ>\<circ> \<approx>\<^bsub>list\<^esub>) (abs_fset \<circ> map_list Abs) (map_list Rep o rep_fset)"}
- %%%
- %%%\noindent
- %%%this theorem will then instantiate the quotients needed in the
- %%%injection and cleaning proofs allowing the lifting procedure to
- %%%proceed in an unchanged way.
-*}
-
-section {* Lifting of Theorems\label{sec:lift} *}
-
-text {*
-
-%%% FIXME Reviewer 3 asks: Section 5 shows the technicalities of
-%%% lifting theorems. But there is no clarification about the
-%%% correctness. A reader would also be interested in seeing some
-%%% discussions about the generality and limitation of the approach
-%%% proposed there
-
- \noindent
- The main benefit of a quotient package is to lift automatically theorems over raw
- types to theorems over quotient types. We will perform this lifting in
- three phases, called \emph{regularization},
- \emph{injection} and \emph{cleaning} according to procedures in Homeier's ML-code.
- Space restrictions, unfortunately, prevent us from giving anything but a sketch of these three
- phases. However, we will precisely define the input and output data of these phases
- (this was omitted in \cite{Homeier05}).
-
- The purpose of regularization is to change the quantifiers and abstractions
- in a ``raw'' theorem to quantifiers over variables that respect their respective relations
- (Definition \ref{def:respects} states what respects means). The purpose of injection is to add @{term Rep}
- and @{term Abs} of appropriate types in front of constants and variables
- of the raw type so that they can be replaced by the corresponding constants from the
- quotient type. The purpose of cleaning is to bring the theorem derived in the
- first two phases into the form the user has specified. Abstractly, our
- package establishes the following three proof steps:
-
-%%% FIXME: Reviewer 1 complains that the reader needs to guess the
-%%% meaning of reg_thm and inj_thm, as well as the arguments of REG
-%%% which are given above. I wouldn't change it.
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{4mm}}l}
- 1.) Regularization & @{text "raw_thm \<longrightarrow> reg_thm"}\\
- 2.) Injection & @{text "reg_thm \<longleftrightarrow> inj_thm"}\\
- 3.) Cleaning & @{text "inj_thm \<longleftrightarrow> quot_thm"}\\
- \end{tabular}
- \end{center}
-
- \noindent
- which means, stringed together, the raw theorem implies the quotient theorem.
- In contrast to other quotient packages, our package requires that the user specifies
- both, the @{text "raw_thm"} (as theorem) and the \emph{term} of the @{text "quot_thm"}.\footnote{Though we
- also provide a fully automated mode, where the @{text "quot_thm"} is guessed
- from the form of @{text "raw_thm"}.} As a result, the user has fine control
- over which parts of a raw theorem should be lifted.
-
- The second and third proof step performed in package will always succeed if the appropriate
- respectfulness and preservation theorems are given. In contrast, the first
- proof step can fail: a theorem given by the user does not always
- imply a regularized version and a stronger one needs to be proved. An example
- for this kind of failure is the simple statement for integers @{text "0 \<noteq> 1"}.
- One might hope that it can be proved by lifting @{text "(0, 0) \<noteq> (1, 0)"},
- but this raw theorem only shows that two particular elements in the
- equivalence classes are not equal. In order to obtain @{text "0 \<noteq> 1"}, a
- more general statement stipulating that the equivalence classes are not
- equal is necessary. This kind of failure is beyond the scope where the
- quotient package can help: the user has to provide a raw theorem that
- can be regularized automatically, or has to provide an explicit proof
- for the first proof step. Homeier gives more details about this issue
- in the long version of \cite{Homeier05}.
-
- In the following we will first define the statement of the
- regularized theorem based on @{text "raw_thm"} and
- @{text "quot_thm"}. Then we define the statement of the injected theorem, based
- on @{text "reg_thm"} and @{text "quot_thm"}. We then show the three proof steps,
- which can all be performed independently from each other.
-
- We first define the function @{text REG}, which takes the terms of the
- @{text "raw_thm"} and @{text "quot_thm"} as input and returns
- @{text "reg_thm"}. The idea
- behind this function is that it replaces quantifiers and
- abstractions involving raw types by bounded ones, and equalities
- involving raw types by appropriate aggregate
- equivalence relations. It is defined by simultaneous recursion on
- the structure of the terms of @{text "raw_thm"} and @{text "quot_thm"} as follows:
- %
- \begin{center}
- \begin{tabular}{@ {}l@ {}}
- \multicolumn{1}{@ {}l@ {}}{abstractions:}\\%%\smallskip\\
- @{text "REG (\<lambda>x\<^sup>\<sigma>. t, \<lambda>x\<^sup>\<tau>. s)"} $\dn$
- $\begin{cases}
- @{text "\<lambda>x\<^sup>\<sigma>. REG (t, s)"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
- @{text "\<lambda>x\<^sup>\<sigma> \<in> Resp (REL (\<sigma>, \<tau>)). REG (t, s)"}
- \end{cases}$\\%%\smallskip\\
- \multicolumn{1}{@ {}l@ {}}{universal quantifiers:}\\
- @{text "REG (\<forall>x\<^sup>\<sigma>. t, \<forall>x\<^sup>\<tau>. s)"} $\dn$
- $\begin{cases}
- @{text "\<forall>x\<^sup>\<sigma>. REG (t, s)"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
- @{text "\<forall>x\<^sup>\<sigma> \<in> Resp (REL (\<sigma>, \<tau>)). REG (t, s)"}
- \end{cases}$\\%%\smallskip\\
- \multicolumn{1}{@ {}l@ {}}{equality: \hspace{3mm}%%}\smallskip\\
- %% REL of two equal types is the equality so we do not need a separate case
- @{text "REG (=\<^bsup>\<sigma>\<Rightarrow>\<sigma>\<Rightarrow>bool\<^esup>, =\<^bsup>\<tau>\<Rightarrow>\<tau>\<Rightarrow>bool\<^esup>)"} $\dn$ @{text "REL (\<sigma>, \<tau>)"}}\smallskip\\
- \multicolumn{1}{@ {}l@ {}}{applications, variables and constants:}\\
- @{text "REG (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2)"} $\dn$ @{text "REG (t\<^isub>1, s\<^isub>1) REG (t\<^isub>2, s\<^isub>2)"}\\
- @{text "REG (x\<^isub>1, x\<^isub>2)"} $\dn$ @{text "x\<^isub>1"}\\
- @{text "REG (c\<^isub>1, c\<^isub>2)"} $\dn$ @{text "c\<^isub>1"}\\
- \end{tabular}
- \end{center}
- %
- \noindent
- In the above definition we omitted the cases for existential quantifiers
- and unique existential quantifiers, as they are very similar to the cases
- for the universal quantifier.
-
- Next we define the function @{text INJ} which takes as argument
- @{text "reg_thm"} and @{text "quot_thm"} (both as
- terms) and returns @{text "inj_thm"}:
-
- \begin{center}
- \begin{tabular}{l}
- \multicolumn{1}{@ {}l}{abstractions:}\\
- @{text "INJ (\<lambda>x. t :: \<sigma>, \<lambda>x. s :: \<tau>) "} $\dn$\\
- \hspace{18mm}$\begin{cases}
- @{text "\<lambda>x. INJ (t, s)"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
- @{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) (\<lambda>x. INJ (t, s)))"}
- \end{cases}$\smallskip\\
- @{text "INJ (\<lambda>x \<in> R. t :: \<sigma>, \<lambda>x. s :: \<tau>) "} $\dn$\\
- \hspace{18mm}@{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) (\<lambda>x \<in> R. INJ (t, s)))"}\smallskip\\
- \multicolumn{1}{@ {}l}{universal quantifiers:}\\
- @{text "INJ (\<forall> t, \<forall> s) "} $\dn$ @{text "\<forall> INJ (t, s)"}\\
- @{text "INJ (\<forall> t \<in> R, \<forall> s) "} $\dn$ @{text "\<forall> INJ (t, s) \<in> R"}\smallskip\\
- \multicolumn{1}{@ {}l}{applications, variables and constants:}\smallskip\\
- @{text "INJ (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2) "} $\dn$ @{text " INJ (t\<^isub>1, s\<^isub>1) INJ (t\<^isub>2, s\<^isub>2)"}\\
- @{text "INJ (x\<^isub>1\<^sup>\<sigma>, x\<^isub>2\<^sup>\<tau>) "} $\dn$
- $\begin{cases}
- @{text "x\<^isub>1"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
- @{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) x\<^isub>1)"}\\
- \end{cases}$\\
- @{text "INJ (c\<^isub>1\<^sup>\<sigma>, c\<^isub>2\<^sup>\<tau>) "} $\dn$
- $\begin{cases}
- @{text "c\<^isub>1"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
- @{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) c\<^isub>1)"}\\
- \end{cases}$\\
- \end{tabular}
- \end{center}
-
- \noindent
- In this definition we again omitted the cases for existential and unique existential
- quantifiers.
-
-%%% FIXME: Reviewer2 citing following sentence: You mention earlier
-%%% that this implication may fail to be true. Does that meant that
-%%% the `first proof step' is a heuristic that proves the implication
-%%% raw_thm \implies reg_thm in some instances, but fails in others?
-%%% You should clarify under which circumstances the implication is
-%%% being proved here.
-%%% Cezary: It would be nice to cite Homeiers discussions in the
-%%% Quotient Package manual from HOL (the longer paper), do you agree?
-
- In the first phase, establishing @{text "raw_thm \<longrightarrow> reg_thm"}, we always
- start with an implication. Isabelle provides \emph{mono} rules that can split up
- the implications into simpler implicational subgoals. This succeeds for every
- monotone connective, except in places where the function @{text REG} replaced,
- for instance, a quantifier by a bounded quantifier. To decompose them, we have
- to prove that the relations involved are aggregate equivalence relations.
-
-
- %In this case we have
- %rules of the form
- %
- % \begin{isabelle}\ \ \ \ \ %%%
- %@{text "(\<forall>x. R x \<longrightarrow> (P x \<longrightarrow> Q x)) \<longrightarrow> (\<forall>x. P x \<longrightarrow> \<forall>x \<in> R. Q x)"}
- %\end{isabelle}
-
- %\noindent
- %They decompose a bounded quantifier on the right-hand side. We can decompose a
- %bounded quantifier anywhere if R is an equivalence relation or
- %if it is a relation over function types with the range being an equivalence
- %relation. If @{text R} is an equivalence relation we can prove that
-
- %\begin{isabelle}\ \ \ \ \ %%%
- %@{text "\<forall>x \<in> Resp R. P x = \<forall>x. P x"}
- %\end{isabelle}
-
- %\noindent
- %If @{term R\<^isub>2} is an equivalence relation, we can prove that for any predicate @{term P}
-
-%%% FIXME Reviewer 1 claims the theorem is obviously false so maybe we
-%%% should include a proof sketch?
-
- %\begin{isabelle}\ \ \ \ \ %%%
- %@{thm (concl) ball_reg_eqv_range[of R\<^isub>1 R\<^isub>2, no_vars]}
- %\end{isabelle}
-
- %\noindent
- %The last theorem is new in comparison with Homeier's package. There the
- %injection procedure would be used to prove such goals and
- %the assumption about the equivalence relation would be used. We use the above theorem directly,
- %because this allows us to completely separate the first and the second
- %proof step into two independent ``units''.
-
- The second phase, establishing @{text "reg_thm \<longleftrightarrow> inj_thm"}, starts with an equality
- between the terms of the regularized theorem and the injected theorem.
- The proof again follows the structure of the
- two underlying terms taking respectfulness theorems into account.
-
- %\begin{itemize}
- %\item For two constants an appropriate respectfulness theorem is applied.
- %\item For two variables, we use the assumptions proved in the regularization step.
- %\item For two abstractions, we @{text "\<eta>"}-expand and @{text "\<beta>"}-reduce them.
- %\item For two applications, we check that the right-hand side is an application of
- % @{term Rep} to an @{term Abs} and @{term "Quotient R Rep Abs"} holds. If yes then we
- % can apply the theorem:
-
- %\begin{isabelle}\ \ \ \ \ %%%
- % @{term "R x y \<longrightarrow> R x (Rep (Abs y))"}
- %\end{isabelle}
-
- % Otherwise we introduce an appropriate relation between the subterms
- % and continue with two subgoals using the lemma:
-
- %\begin{isabelle}\ \ \ \ \ %%%
- % @{text "(R\<^isub>1 \<doublearr> R\<^isub>2) f g \<longrightarrow> R\<^isub>1 x y \<longrightarrow> R\<^isub>2 (f x) (g y)"}
- %\end{isabelle}
- %\end{itemize}
-
- We defined the theorem @{text "inj_thm"} in such a way that
- establishing in the third phase the equivalence @{text "inj_thm \<longleftrightarrow> quot_thm"} can be
- achieved by rewriting @{text "inj_thm"} with the preservation theorems and quotient
- definitions. This step also requires that the definitions of all lifted constants
- are used to fold the @{term Rep} with the raw constants. We will give more details
- about our lifting procedure in a longer version of this paper.
-
- %Next for
- %all abstractions and quantifiers the lambda and
- %quantifier preservation theorems are used to replace the
- %variables that include raw types with respects by quantifiers
- %over variables that include quotient types. We show here only
- %the lambda preservation theorem. Given
- %@{term "Quotient R\<^isub>1 Abs\<^isub>1 Rep\<^isub>1"} and @{term "Quotient R\<^isub>2 Abs\<^isub>2 Rep\<^isub>2"}, we have:
-
- %\begin{isabelle}\ \ \ \ \ %%%
- %@{thm (concl) lambda_prs[of _ "Abs\<^isub>1" "Rep\<^isub>1" _ "Abs\<^isub>2" "Rep\<^isub>2", no_vars]}
- %\end{isabelle}
-
- %\noindent
- %Next, relations over lifted types can be rewritten to equalities
- %over lifted type. Rewriting is performed with the following theorem,
- %which has been shown by Homeier~\cite{Homeier05}:
-
- %\begin{isabelle}\ \ \ \ \ %%%
- %@{thm (concl) Quotient_rel_rep[no_vars]}
- %\end{isabelle}
-
-
- %Finally, we rewrite with the preservation theorems. This will result
- %in two equal terms that can be solved by reflexivity.
-*}
-
-section {* Conclusion and Related Work\label{sec:conc}*}
-
-text {*
-
- \noindent
- The code of the quotient package and the examples described here are already
- included in the standard distribution of Isabelle.
- \footnote{Available from \href{http://isabelle.in.tum.de/}{http://isabelle.in.tum.de/}.}
- The package is
- heavily used in the new version of Nominal Isabelle, which provides a
- convenient reasoning infrastructure for programming language calculi
- involving general binders. To achieve this, it builds types representing
- @{text \<alpha>}-equivalent terms. Earlier versions of Nominal Isabelle have been
- used successfully in formalisations of an equivalence checking algorithm for
- LF \cite{UrbanCheneyBerghofer08}, Typed
- Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for concurrency
- \cite{BengtsonParow09} and a strong normalisation result for cut-elimination
- in classical logic \cite{UrbanZhu08}.
-
-
- There is a wide range of existing literature for dealing with quotients
- in theorem provers. Slotosch~\cite{Slotosch97} implemented a mechanism that
- automatically defines quotient types for Isabelle/HOL. But he did not
- include theorem lifting. Harrison's quotient package~\cite{harrison-thesis}
- is the first one that is able to automatically lift theorems, however only
- first-order theorems (that is theorems where abstractions, quantifiers and
- variables do not involve functions that include the quotient type). There is
- also some work on quotient types in non-HOL based systems and logical
- frameworks, including theory interpretations in
- PVS~\cite{PVS:Interpretations}, new types in MetaPRL~\cite{Nogin02}, and
- setoids in Coq \cite{ChicliPS02}. Paulson showed a construction of
- quotients that does not require the Hilbert Choice operator, but also only
- first-order theorems can be lifted~\cite{Paulson06}. The most related work
- to our package is the package for HOL4 by Homeier~\cite{Homeier05}. He
- introduced most of the abstract notions about quotients and also deals with
- lifting of higher-order theorems. However, he cannot deal with quotient
- compositions (needed for lifting theorems about @{text flat}). Also, a
- number of his definitions, like @{text ABS}, @{text REP} and @{text INJ} etc
- only exist in \cite{Homeier05} as ML-code, not included in the paper.
- Like Homeier's, our quotient package can deal with partial equivalence
- relations, but for lack of space we do not describe the mechanisms
- needed for this kind of quotient constructions.
-
-%%% FIXME Reviewer 3 would like to know more about the lifting in Coq and PVS,
-%%% and some comparison. I don't think we have the space for any additions...
-
- One feature of our quotient package is that when lifting theorems, the user
- can precisely specify what the lifted theorem should look like. This feature
- is necessary, for example, when lifting an induction principle for two
- lists. Assuming this principle has as the conclusion a predicate of the
- form @{text "P xs ys"}, then we can precisely specify whether we want to
- quotient @{text "xs"} or @{text "ys"}, or both. We found this feature very
- useful in the new version of Nominal Isabelle, where such a choice is
- required to generate a reasoning infrastructure for alpha-equated terms.
-%%
-%% give an example for this
-%%
- \smallskip
-
- \noindent
- {\bf Acknowledgements:} We would like to thank Peter Homeier for the many
- discussions about his HOL4 quotient package and explaining to us
- some of its finer points in the implementation. Without his patient
- help, this work would have been impossible.
-*}
-
-text_raw {*
- %%\bibliographystyle{abbrv}
- \bibliography{root}
-
- \appendix
-
-*}
-
-section {* Examples \label{sec:examples} *}
-
-text {*
-
-%%% FIXME Reviewer 1 would like an example of regularized and injected
-%%% statements. He asks for the examples twice, but I would still ignore
-%%% it due to lack of space...
-
- \noindent
- In this appendix we will show a sequence of declarations for defining the
- type of integers by quotienting pairs of natural numbers, and
- lifting one theorem.
-
- A user of our quotient package first needs to define a relation on
- the raw type with which the quotienting will be performed. We give
- the same integer relation as the one presented in \eqref{natpairequiv}:
-
- \begin{isabelle}\ \ \ \ \ %
- \begin{tabular}{@ {}l}
- \isacommand{fun}~~@{text "int_rel :: (nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"}\\
- \isacommand{where}~~@{text "int_rel (m, n) (p, q) = (m + q = n + p)"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- Next the quotient type must be defined. This generates a proof obligation that the
- relation is an equivalence relation, which is solved automatically using the
- definition of equivalence and extensionality:
-
- \begin{isabelle}\ \ \ \ \ %
- \begin{tabular}{@ {}l}
- \isacommand{quotient\_type}~~@{text "int"}~~\isacommand{=}~~@{text "(nat \<times> nat)"}~~\isacommand{/}~~@{text "int_rel"}\\
- \hspace{5mm}@{text "by (auto simp add: equivp_def expand_fun_eq)"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- The user can then specify the constants on the quotient type:
-
- \begin{isabelle}\ \ \ \ \ %
- \begin{tabular}{@ {}l}
- \isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0 :: nat, 0 :: nat)"}\\[3mm]
- \isacommand{fun}~~@{text "add_pair"}\\
- \isacommand{where}~~%
- @{text "add_pair (m, n) (p, q) \<equiv> (m + p :: nat, n + q :: nat)"}\\
- \isacommand{quotient\_definition}~~@{text "+ :: int \<Rightarrow> int \<Rightarrow> int"}~~%
- \isacommand{is}~~@{text "add_pair"}\\
- \end{tabular}
- \end{isabelle}
-
- \noindent
- The following theorem about addition on the raw level can be proved.
-
- \begin{isabelle}\ \ \ \ \ %
- \isacommand{lemma}~~@{text "add_pair_zero: int_rel (add_pair (0, 0) x) x"}
- \end{isabelle}
-
- \noindent
- If the user lifts this theorem, the quotient package performs all the lifting
- automatically leaving the respectfulness proof for the constant @{text "add_pair"}
- as the only remaining proof obligation. This property needs to be proved by the user:
-
- \begin{isabelle}\ \ \ \ \ %
- \begin{tabular}{@ {}l}
- \isacommand{lemma}~~@{text "[quot_respect]:"}\\
- @{text "(int_rel \<doublearr> int_rel \<doublearr> int_rel) add_pair add_pair"}
- \end{tabular}
- \end{isabelle}
-
- \noindent
- It can be discharged automatically by Isabelle when hinting to unfold the definition
- of @{text "\<doublearr>"}.
- After this, the user can prove the lifted lemma as follows:
-
- \begin{isabelle}\ \ \ \ \ %
- \isacommand{lemma}~~@{text "0 + (x :: int) = x"}~~\isacommand{by}~~@{text "lifting add_pair_zero"}
- \end{isabelle}
-
- \noindent
- or by using the completely automated mode stating just:
-
- \begin{isabelle}\ \ \ \ \ %
- \isacommand{thm}~~@{text "add_pair_zero[quot_lifted]"}
- \end{isabelle}
-
- \noindent
- Both methods give the same result, namely @{text "0 + x = x"}
- where @{text x} is of type integer.
- Although seemingly simple, arriving at this result without the help of a quotient
- package requires a substantial reasoning effort (see \cite{Paulson06}).
-*}
-
-
-
-(*<*)
-end
-(*>*)
--- a/Quotient-Paper/ROOT.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,5 +0,0 @@
-quick_and_dirty := true;
-no_document use_thys ["Quotient",
- "~~/src/HOL/Library/LaTeXsugar",
- "~~/src/HOL/Quotient_Examples/FSet" ];
-use_thys ["Paper"];
--- a/Quotient-Paper/document/llncs.cls Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1189 +0,0 @@
-% LLNCS DOCUMENT CLASS -- version 2.13 (28-Jan-2002)
-% Springer Verlag LaTeX2e support for Lecture Notes in Computer Science
-%
-%%
-%% \CharacterTable
-%% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
-%% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
-%% Digits \0\1\2\3\4\5\6\7\8\9
-%% Exclamation \! Double quote \" Hash (number) \#
-%% Dollar \$ Percent \% Ampersand \&
-%% Acute accent \' Left paren \( Right paren \)
-%% Asterisk \* Plus \+ Comma \,
-%% Minus \- Point \. Solidus \/
-%% Colon \: Semicolon \; Less than \<
-%% Equals \= Greater than \> Question mark \?
-%% Commercial at \@ Left bracket \[ Backslash \\
-%% Right bracket \] Circumflex \^ Underscore \_
-%% Grave accent \` Left brace \{ Vertical bar \|
-%% Right brace \} Tilde \~}
-%%
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-\ProvidesClass{llncs}[2002/01/28 v2.13
-^^J LaTeX document class for Lecture Notes in Computer Science]
-% Options
-\let\if@envcntreset\iffalse
-\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue}
-\DeclareOption{citeauthoryear}{\let\citeauthoryear=Y}
-\DeclareOption{oribibl}{\let\oribibl=Y}
-\let\if@custvec\iftrue
-\DeclareOption{orivec}{\let\if@custvec\iffalse}
-\let\if@envcntsame\iffalse
-\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue}
-\let\if@envcntsect\iffalse
-\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue}
-\let\if@runhead\iffalse
-\DeclareOption{runningheads}{\let\if@runhead\iftrue}
-
-\let\if@openbib\iffalse
-\DeclareOption{openbib}{\let\if@openbib\iftrue}
-
-% languages
-\let\switcht@@therlang\relax
-\def\ds@deutsch{\def\switcht@@therlang{\switcht@deutsch}}
-\def\ds@francais{\def\switcht@@therlang{\switcht@francais}}
-
-\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}}
-
-\ProcessOptions
-
-\LoadClass[twoside]{article}
-\RequirePackage{multicol} % needed for the list of participants, index
-
-\setlength{\textwidth}{12.2cm}
-\setlength{\textheight}{19.3cm}
-\renewcommand\@pnumwidth{2em}
-\renewcommand\@tocrmarg{3.5em}
-%
-\def\@dottedtocline#1#2#3#4#5{%
- \ifnum #1>\c@tocdepth \else
- \vskip \z@ \@plus.2\p@
- {\leftskip #2\relax \rightskip \@tocrmarg \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \parindent #2\relax\@afterindenttrue
- \interlinepenalty\@M
- \leavevmode
- \@tempdima #3\relax
- \advance\leftskip \@tempdima \null\nobreak\hskip -\leftskip
- {#4}\nobreak
- \leaders\hbox{$\m@th
- \mkern \@dotsep mu\hbox{.}\mkern \@dotsep
- mu$}\hfill
- \nobreak
- \hb@xt@\@pnumwidth{\hfil\normalfont \normalcolor #5}%
- \par}%
- \fi}
-%
-\def\switcht@albion{%
-\def\abstractname{Abstract.}
-\def\ackname{Acknowledgement.}
-\def\andname{and}
-\def\lastandname{\unskip, and}
-\def\appendixname{Appendix}
-\def\chaptername{Chapter}
-\def\claimname{Claim}
-\def\conjecturename{Conjecture}
-\def\contentsname{Table of Contents}
-\def\corollaryname{Corollary}
-\def\definitionname{Definition}
-\def\examplename{Example}
-\def\exercisename{Exercise}
-\def\figurename{Fig.}
-\def\keywordname{{\bf Key words:}}
-\def\indexname{Index}
-\def\lemmaname{Lemma}
-\def\contriblistname{List of Contributors}
-\def\listfigurename{List of Figures}
-\def\listtablename{List of Tables}
-\def\mailname{{\it Correspondence to\/}:}
-\def\noteaddname{Note added in proof}
-\def\notename{Note}
-\def\partname{Part}
-\def\problemname{Problem}
-\def\proofname{Proof}
-\def\propertyname{Property}
-\def\propositionname{Proposition}
-\def\questionname{Question}
-\def\remarkname{Remark}
-\def\seename{see}
-\def\solutionname{Solution}
-\def\subclassname{{\it Subject Classifications\/}:}
-\def\tablename{Table}
-\def\theoremname{Theorem}}
-\switcht@albion
-% Names of theorem like environments are already defined
-% but must be translated if another language is chosen
-%
-% French section
-\def\switcht@francais{%\typeout{On parle francais.}%
- \def\abstractname{R\'esum\'e.}%
- \def\ackname{Remerciements.}%
- \def\andname{et}%
- \def\lastandname{ et}%
- \def\appendixname{Appendice}
- \def\chaptername{Chapitre}%
- \def\claimname{Pr\'etention}%
- \def\conjecturename{Hypoth\`ese}%
- \def\contentsname{Table des mati\`eres}%
- \def\corollaryname{Corollaire}%
- \def\definitionname{D\'efinition}%
- \def\examplename{Exemple}%
- \def\exercisename{Exercice}%
- \def\figurename{Fig.}%
- \def\keywordname{{\bf Mots-cl\'e:}}
- \def\indexname{Index}
- \def\lemmaname{Lemme}%
- \def\contriblistname{Liste des contributeurs}
- \def\listfigurename{Liste des figures}%
- \def\listtablename{Liste des tables}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}%
- \def\notename{Remarque}%
- \def\partname{Partie}%
- \def\problemname{Probl\`eme}%
- \def\proofname{Preuve}%
- \def\propertyname{Caract\'eristique}%
-%\def\propositionname{Proposition}%
- \def\questionname{Question}%
- \def\remarkname{Remarque}%
- \def\seename{voir}
- \def\solutionname{Solution}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tableau}%
- \def\theoremname{Th\'eor\`eme}%
-}
-%
-% German section
-\def\switcht@deutsch{%\typeout{Man spricht deutsch.}%
- \def\abstractname{Zusammenfassung.}%
- \def\ackname{Danksagung.}%
- \def\andname{und}%
- \def\lastandname{ und}%
- \def\appendixname{Anhang}%
- \def\chaptername{Kapitel}%
- \def\claimname{Behauptung}%
- \def\conjecturename{Hypothese}%
- \def\contentsname{Inhaltsverzeichnis}%
- \def\corollaryname{Korollar}%
-%\def\definitionname{Definition}%
- \def\examplename{Beispiel}%
- \def\exercisename{\"Ubung}%
- \def\figurename{Abb.}%
- \def\keywordname{{\bf Schl\"usselw\"orter:}}
- \def\indexname{Index}
-%\def\lemmaname{Lemma}%
- \def\contriblistname{Mitarbeiter}
- \def\listfigurename{Abbildungsverzeichnis}%
- \def\listtablename{Tabellenverzeichnis}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Nachtrag}%
- \def\notename{Anmerkung}%
- \def\partname{Teil}%
-%\def\problemname{Problem}%
- \def\proofname{Beweis}%
- \def\propertyname{Eigenschaft}%
-%\def\propositionname{Proposition}%
- \def\questionname{Frage}%
- \def\remarkname{Anmerkung}%
- \def\seename{siehe}
- \def\solutionname{L\"osung}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tabelle}%
-%\def\theoremname{Theorem}%
-}
-
-% Ragged bottom for the actual page
-\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil
-\global\let\@textbottom\relax}}
-
-\renewcommand\small{%
- \@setfontsize\small\@ixpt{11}%
- \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
- \abovedisplayshortskip \z@ \@plus2\p@
- \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
- \def\@listi{\leftmargin\leftmargini
- \parsep 0\p@ \@plus1\p@ \@minus\p@
- \topsep 8\p@ \@plus2\p@ \@minus4\p@
- \itemsep0\p@}%
- \belowdisplayskip \abovedisplayskip
-}
-
-\frenchspacing
-\widowpenalty=10000
-\clubpenalty=10000
-
-\setlength\oddsidemargin {63\p@}
-\setlength\evensidemargin {63\p@}
-\setlength\marginparwidth {90\p@}
-
-\setlength\headsep {16\p@}
-
-\setlength\footnotesep{7.7\p@}
-\setlength\textfloatsep{8mm\@plus 2\p@ \@minus 4\p@}
-\setlength\intextsep {8mm\@plus 2\p@ \@minus 2\p@}
-
-\setcounter{secnumdepth}{2}
-
-\newcounter {chapter}
-\renewcommand\thechapter {\@arabic\c@chapter}
-
-\newif\if@mainmatter \@mainmattertrue
-\newcommand\frontmatter{\cleardoublepage
- \@mainmatterfalse\pagenumbering{Roman}}
-\newcommand\mainmatter{\cleardoublepage
- \@mainmattertrue\pagenumbering{arabic}}
-\newcommand\backmatter{\if@openright\cleardoublepage\else\clearpage\fi
- \@mainmatterfalse}
-
-\renewcommand\part{\cleardoublepage
- \thispagestyle{empty}%
- \if@twocolumn
- \onecolumn
- \@tempswatrue
- \else
- \@tempswafalse
- \fi
- \null\vfil
- \secdef\@part\@spart}
-
-\def\@part[#1]#2{%
- \ifnum \c@secnumdepth >-2\relax
- \refstepcounter{part}%
- \addcontentsline{toc}{part}{\thepart\hspace{1em}#1}%
- \else
- \addcontentsline{toc}{part}{#1}%
- \fi
- \markboth{}{}%
- {\centering
- \interlinepenalty \@M
- \normalfont
- \ifnum \c@secnumdepth >-2\relax
- \huge\bfseries \partname~\thepart
- \par
- \vskip 20\p@
- \fi
- \Huge \bfseries #2\par}%
- \@endpart}
-\def\@spart#1{%
- {\centering
- \interlinepenalty \@M
- \normalfont
- \Huge \bfseries #1\par}%
- \@endpart}
-\def\@endpart{\vfil\newpage
- \if@twoside
- \null
- \thispagestyle{empty}%
- \newpage
- \fi
- \if@tempswa
- \twocolumn
- \fi}
-
-\newcommand\chapter{\clearpage
- \thispagestyle{empty}%
- \global\@topnum\z@
- \@afterindentfalse
- \secdef\@chapter\@schapter}
-\def\@chapter[#1]#2{\ifnum \c@secnumdepth >\m@ne
- \if@mainmatter
- \refstepcounter{chapter}%
- \typeout{\@chapapp\space\thechapter.}%
- \addcontentsline{toc}{chapter}%
- {\protect\numberline{\thechapter}#1}%
- \else
- \addcontentsline{toc}{chapter}{#1}%
- \fi
- \else
- \addcontentsline{toc}{chapter}{#1}%
- \fi
- \chaptermark{#1}%
- \addtocontents{lof}{\protect\addvspace{10\p@}}%
- \addtocontents{lot}{\protect\addvspace{10\p@}}%
- \if@twocolumn
- \@topnewpage[\@makechapterhead{#2}]%
- \else
- \@makechapterhead{#2}%
- \@afterheading
- \fi}
-\def\@makechapterhead#1{%
-% \vspace*{50\p@}%
- {\centering
- \ifnum \c@secnumdepth >\m@ne
- \if@mainmatter
- \large\bfseries \@chapapp{} \thechapter
- \par\nobreak
- \vskip 20\p@
- \fi
- \fi
- \interlinepenalty\@M
- \Large \bfseries #1\par\nobreak
- \vskip 40\p@
- }}
-\def\@schapter#1{\if@twocolumn
- \@topnewpage[\@makeschapterhead{#1}]%
- \else
- \@makeschapterhead{#1}%
- \@afterheading
- \fi}
-\def\@makeschapterhead#1{%
-% \vspace*{50\p@}%
- {\centering
- \normalfont
- \interlinepenalty\@M
- \Large \bfseries #1\par\nobreak
- \vskip 40\p@
- }}
-
-\renewcommand\section{\@startsection{section}{1}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {12\p@ \@plus 4\p@ \@minus 4\p@}%
- {\normalfont\large\bfseries\boldmath
- \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsection{\@startsection{subsection}{2}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {8\p@ \@plus 4\p@ \@minus 4\p@}%
- {\normalfont\normalsize\bfseries\boldmath
- \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {-0.5em \@plus -0.22em \@minus -0.1em}%
- {\normalfont\normalsize\bfseries\boldmath}}
-\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}%
- {-12\p@ \@plus -4\p@ \@minus -4\p@}%
- {-0.5em \@plus -0.22em \@minus -0.1em}%
- {\normalfont\normalsize\itshape}}
-\renewcommand\subparagraph[1]{\typeout{LLNCS warning: You should not use
- \string\subparagraph\space with this class}\vskip0.5cm
-You should not use \verb|\subparagraph| with this class.\vskip0.5cm}
-
-\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00}
-\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01}
-\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02}
-\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03}
-\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04}
-\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05}
-\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06}
-\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07}
-\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08}
-\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09}
-\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A}
-
-\let\footnotesize\small
-
-\if@custvec
-\def\vec#1{\mathchoice{\mbox{\boldmath$\displaystyle#1$}}
-{\mbox{\boldmath$\textstyle#1$}}
-{\mbox{\boldmath$\scriptstyle#1$}}
-{\mbox{\boldmath$\scriptscriptstyle#1$}}}
-\fi
-
-\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}}
-\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil
-\penalty50\hskip1em\null\nobreak\hfil\squareforqed
-\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi}
-
-\def\getsto{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr\gets\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-\gets\cr\to\cr}}}}}
-\def\lid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr<\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr<\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr<\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-<\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\gid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr>\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr>\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\grole{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr
->\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.8pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.3pt}<\cr}}}}}
-\def\bbbr{{\rm I\!R}} %reelle Zahlen
-\def\bbbm{{\rm I\!M}}
-\def\bbbn{{\rm I\!N}} %natuerliche Zahlen
-\def\bbbf{{\rm I\!F}}
-\def\bbbh{{\rm I\!H}}
-\def\bbbk{{\rm I\!K}}
-\def\bbbp{{\rm I\!P}}
-\def\bbbone{{\mathchoice {\rm 1\mskip-4mu l} {\rm 1\mskip-4mu l}
-{\rm 1\mskip-4.5mu l} {\rm 1\mskip-5mu l}}}
-\def\bbbc{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbq{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}}}
-\def\bbbt{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-T$}\hbox{\hbox to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbs{{\mathchoice
-{\setbox0=\hbox{$\displaystyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.5\wd0\vrule height0.45\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.4\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.55\wd0\vrule height0.45\ht0\hss}\box0}}}}
-\def\bbbz{{\mathchoice {\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\scriptstyle Z\kern-0.3em Z$}}
-{\hbox{$\mathsf\scriptscriptstyle Z\kern-0.2em Z$}}}}
-
-\let\ts\,
-
-\setlength\leftmargini {17\p@}
-\setlength\leftmargin {\leftmargini}
-\setlength\leftmarginii {\leftmargini}
-\setlength\leftmarginiii {\leftmargini}
-\setlength\leftmarginiv {\leftmargini}
-\setlength \labelsep {.5em}
-\setlength \labelwidth{\leftmargini}
-\addtolength\labelwidth{-\labelsep}
-
-\def\@listI{\leftmargin\leftmargini
- \parsep 0\p@ \@plus1\p@ \@minus\p@
- \topsep 8\p@ \@plus2\p@ \@minus4\p@
- \itemsep0\p@}
-\let\@listi\@listI
-\@listi
-\def\@listii {\leftmargin\leftmarginii
- \labelwidth\leftmarginii
- \advance\labelwidth-\labelsep
- \topsep 0\p@ \@plus2\p@ \@minus\p@}
-\def\@listiii{\leftmargin\leftmarginiii
- \labelwidth\leftmarginiii
- \advance\labelwidth-\labelsep
- \topsep 0\p@ \@plus\p@\@minus\p@
- \parsep \z@
- \partopsep \p@ \@plus\z@ \@minus\p@}
-
-\renewcommand\labelitemi{\normalfont\bfseries --}
-\renewcommand\labelitemii{$\m@th\bullet$}
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\def\tableofcontents{\chapter*{\contentsname\@mkboth{{\contentsname}}%
- {{\contentsname}}}
- \def\authcount##1{\setcounter{auco}{##1}\setcounter{@auth}{1}}
- \def\lastand{\ifnum\value{auco}=2\relax
- \unskip{} \andname\
- \else
- \unskip \lastandname\
- \fi}%
- \def\and{\stepcounter{@auth}\relax
- \ifnum\value{@auth}=\value{auco}%
- \lastand
- \else
- \unskip,
- \fi}%
- \@starttoc{toc}\if@restonecol\twocolumn\fi}
-
-\def\l@part#1#2{\addpenalty{\@secpenalty}%
- \addvspace{2em plus\p@}% % space above part line
- \begingroup
- \parindent \z@
- \rightskip \z@ plus 5em
- \hrule\vskip5pt
- \large % same size as for a contribution heading
- \bfseries\boldmath % set line in boldface
- \leavevmode % TeX command to enter horizontal mode.
- #1\par
- \vskip5pt
- \hrule
- \vskip1pt
- \nobreak % Never break after part entry
- \endgroup}
-
-\def\@dotsep{2}
-
-\def\hyperhrefextend{\ifx\hyper@anchor\@undefined\else
-{chapter.\thechapter}\fi}
-
-\def\addnumcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{\protect\numberline
- {\thechapter}#3}{\thepage}\hyperhrefextend}}
-\def\addcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{\thepage}\hyperhrefextend}}
-\def\addcontentsmarkwop#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{0}\hyperhrefextend}}
-
-\def\@adcmk[#1]{\ifcase #1 \or
-\def\@gtempa{\addnumcontentsmark}%
- \or \def\@gtempa{\addcontentsmark}%
- \or \def\@gtempa{\addcontentsmarkwop}%
- \fi\@gtempa{toc}{chapter}}
-\def\addtocmark{\@ifnextchar[{\@adcmk}{\@adcmk[3]}}
-
-\def\l@chapter#1#2{\addpenalty{-\@highpenalty}
- \vskip 1.0em plus 1pt \@tempdima 1.5em \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- {\large\bfseries\boldmath#1}\ifx0#2\hfil\null
- \else
- \nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}%
- \fi\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@title#1#2{\addpenalty{-\@highpenalty}
- \addvspace{8pt plus 1pt}
- \@tempdima \z@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- #1\nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@author#1#2{\addpenalty{\@highpenalty}
- \@tempdima=\z@ %15\p@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima %\hskip -\leftskip
- \textit{#1}\par
- \penalty\@highpenalty \endgroup}
-
-%\setcounter{tocdepth}{0}
-\newdimen\tocchpnum
-\newdimen\tocsecnum
-\newdimen\tocsectotal
-\newdimen\tocsubsecnum
-\newdimen\tocsubsectotal
-\newdimen\tocsubsubsecnum
-\newdimen\tocsubsubsectotal
-\newdimen\tocparanum
-\newdimen\tocparatotal
-\newdimen\tocsubparanum
-\tocchpnum=\z@ % no chapter numbers
-\tocsecnum=15\p@ % section 88. plus 2.222pt
-\tocsubsecnum=23\p@ % subsection 88.8 plus 2.222pt
-\tocsubsubsecnum=27\p@ % subsubsection 88.8.8 plus 1.444pt
-\tocparanum=35\p@ % paragraph 88.8.8.8 plus 1.666pt
-\tocsubparanum=43\p@ % subparagraph 88.8.8.8.8 plus 1.888pt
-\def\calctocindent{%
-\tocsectotal=\tocchpnum
-\advance\tocsectotal by\tocsecnum
-\tocsubsectotal=\tocsectotal
-\advance\tocsubsectotal by\tocsubsecnum
-\tocsubsubsectotal=\tocsubsectotal
-\advance\tocsubsubsectotal by\tocsubsubsecnum
-\tocparatotal=\tocsubsubsectotal
-\advance\tocparatotal by\tocparanum}
-\calctocindent
-
-\def\l@section{\@dottedtocline{1}{\tocchpnum}{\tocsecnum}}
-\def\l@subsection{\@dottedtocline{2}{\tocsectotal}{\tocsubsecnum}}
-\def\l@subsubsection{\@dottedtocline{3}{\tocsubsectotal}{\tocsubsubsecnum}}
-\def\l@paragraph{\@dottedtocline{4}{\tocsubsubsectotal}{\tocparanum}}
-\def\l@subparagraph{\@dottedtocline{5}{\tocparatotal}{\tocsubparanum}}
-
-\def\listoffigures{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}}
- \@starttoc{lof}\if@restonecol\twocolumn\fi}
-\def\l@figure{\@dottedtocline{1}{0em}{1.5em}}
-
-\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}}
- \@starttoc{lot}\if@restonecol\twocolumn\fi}
-\let\l@table\l@figure
-
-\renewcommand\listoffigures{%
- \section*{\listfigurename
- \@mkboth{\listfigurename}{\listfigurename}}%
- \@starttoc{lof}%
- }
-
-\renewcommand\listoftables{%
- \section*{\listtablename
- \@mkboth{\listtablename}{\listtablename}}%
- \@starttoc{lot}%
- }
-
-\ifx\oribibl\undefined
-\ifx\citeauthoryear\undefined
-\renewenvironment{thebibliography}[1]
- {\section*{\refname}
- \def\@biblabel##1{##1.}
- \small
- \list{\@biblabel{\@arabic\c@enumiv}}%
- {\settowidth\labelwidth{\@biblabel{#1}}%
- \leftmargin\labelwidth
- \advance\leftmargin\labelsep
- \if@openbib
- \advance\leftmargin\bibindent
- \itemindent -\bibindent
- \listparindent \itemindent
- \parsep \z@
- \fi
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{\@arabic\c@enumiv}}%
- \if@openbib
- \renewcommand\newblock{\par}%
- \else
- \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
- \fi
- \sloppy\clubpenalty4000\widowpenalty4000%
- \sfcode`\.=\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
-\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw
- {\let\protect\noexpand\immediate
- \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
-\newcount\@tempcntc
-\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi
- \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do
- {\@ifundefined
- {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries
- ?}\@warning
- {Citation `\@citeb' on page \thepage \space undefined}}%
- {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}%
- \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne
- \@citea\def\@citea{,}\hbox{\csname b@\@citeb\endcsname}%
- \else
- \advance\@tempcntb\@ne
- \ifnum\@tempcntb=\@tempcntc
- \else\advance\@tempcntb\m@ne\@citeo
- \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}}
-\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else
- \@citea\def\@citea{,\,\hskip\z@skip}%
- \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else
- {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else
- \def\@citea{--}\fi
- \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi}
-\else
-\renewenvironment{thebibliography}[1]
- {\section*{\refname}
- \small
- \list{}%
- {\settowidth\labelwidth{}%
- \leftmargin\parindent
- \itemindent=-\parindent
- \labelsep=\z@
- \if@openbib
- \advance\leftmargin\bibindent
- \itemindent -\bibindent
- \listparindent \itemindent
- \parsep \z@
- \fi
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{}}%
- \if@openbib
- \renewcommand\newblock{\par}%
- \else
- \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
- \fi
- \sloppy\clubpenalty4000\widowpenalty4000%
- \sfcode`\.=\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
- \def\@cite#1{#1}%
- \def\@lbibitem[#1]#2{\item[]\if@filesw
- {\def\protect##1{\string ##1\space}\immediate
- \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
- \fi
-\else
-\@cons\@openbib@code{\noexpand\small}
-\fi
-
-\def\idxquad{\hskip 10\p@}% space that divides entry from number
-
-\def\@idxitem{\par\hangindent 10\p@}
-
-\def\subitem{\par\setbox0=\hbox{--\enspace}% second order
- \noindent\hangindent\wd0\box0}% index entry
-
-\def\subsubitem{\par\setbox0=\hbox{--\,--\enspace}% third
- \noindent\hangindent\wd0\box0}% order index entry
-
-\def\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax}
-
-\renewenvironment{theindex}
- {\@mkboth{\indexname}{\indexname}%
- \thispagestyle{empty}\parindent\z@
- \parskip\z@ \@plus .3\p@\relax
- \let\item\par
- \def\,{\relax\ifmmode\mskip\thinmuskip
- \else\hskip0.2em\ignorespaces\fi}%
- \normalfont\small
- \begin{multicols}{2}[\@makeschapterhead{\indexname}]%
- }
- {\end{multicols}}
-
-\renewcommand\footnoterule{%
- \kern-3\p@
- \hrule\@width 2truecm
- \kern2.6\p@}
- \newdimen\fnindent
- \fnindent1em
-\long\def\@makefntext#1{%
- \parindent \fnindent%
- \leftskip \fnindent%
- \noindent
- \llap{\hb@xt@1em{\hss\@makefnmark\ }}\ignorespaces#1}
-
-\long\def\@makecaption#1#2{%
- \vskip\abovecaptionskip
- \sbox\@tempboxa{{\bfseries #1.} #2}%
- \ifdim \wd\@tempboxa >\hsize
- {\bfseries #1.} #2\par
- \else
- \global \@minipagefalse
- \hb@xt@\hsize{\hfil\box\@tempboxa\hfil}%
- \fi
- \vskip\belowcaptionskip}
-
-\def\fps@figure{htbp}
-\def\fnum@figure{\figurename\thinspace\thefigure}
-\def \@floatboxreset {%
- \reset@font
- \small
- \@setnobreak
- \@setminipage
-}
-\def\fps@table{htbp}
-\def\fnum@table{\tablename~\thetable}
-\renewenvironment{table}
- {\setlength\abovecaptionskip{0\p@}%
- \setlength\belowcaptionskip{10\p@}%
- \@float{table}}
- {\end@float}
-\renewenvironment{table*}
- {\setlength\abovecaptionskip{0\p@}%
- \setlength\belowcaptionskip{10\p@}%
- \@dblfloat{table}}
- {\end@dblfloat}
-
-\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname
- ext@#1\endcsname}{#1}{\protect\numberline{\csname
- the#1\endcsname}{\ignorespaces #2}}\begingroup
- \@parboxrestore
- \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
- \endgroup}
-
-% LaTeX does not provide a command to enter the authors institute
-% addresses. The \institute command is defined here.
-
-\newcounter{@inst}
-\newcounter{@auth}
-\newcounter{auco}
-\newdimen\instindent
-\newbox\authrun
-\newtoks\authorrunning
-\newtoks\tocauthor
-\newbox\titrun
-\newtoks\titlerunning
-\newtoks\toctitle
-
-\def\clearheadinfo{\gdef\@author{No Author Given}%
- \gdef\@title{No Title Given}%
- \gdef\@subtitle{}%
- \gdef\@institute{No Institute Given}%
- \gdef\@thanks{}%
- \global\titlerunning={}\global\authorrunning={}%
- \global\toctitle={}\global\tocauthor={}}
-
-\def\institute#1{\gdef\@institute{#1}}
-
-\def\institutename{\par
- \begingroup
- \parskip=\z@
- \parindent=\z@
- \setcounter{@inst}{1}%
- \def\and{\par\stepcounter{@inst}%
- \noindent$^{\the@inst}$\enspace\ignorespaces}%
- \setbox0=\vbox{\def\thanks##1{}\@institute}%
- \ifnum\c@@inst=1\relax
- \gdef\fnnstart{0}%
- \else
- \xdef\fnnstart{\c@@inst}%
- \setcounter{@inst}{1}%
- \noindent$^{\the@inst}$\enspace
- \fi
- \ignorespaces
- \@institute\par
- \endgroup}
-
-\def\@fnsymbol#1{\ensuremath{\ifcase#1\or\star\or{\star\star}\or
- {\star\star\star}\or \dagger\or \ddagger\or
- \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger
- \or \ddagger\ddagger \else\@ctrerr\fi}}
-
-\def\inst#1{\unskip$^{#1}$}
-\def\fnmsep{\unskip$^,$}
-\def\email#1{{\tt#1}}
-\AtBeginDocument{\@ifundefined{url}{\def\url#1{#1}}{}%
-\@ifpackageloaded{babel}{%
-\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}%
-\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}%
-\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}%
-}{\switcht@@therlang}%
-}
-\def\homedir{\~{ }}
-
-\def\subtitle#1{\gdef\@subtitle{#1}}
-\clearheadinfo
-
-\renewcommand\maketitle{\newpage
- \refstepcounter{chapter}%
- \stepcounter{section}%
- \setcounter{section}{0}%
- \setcounter{subsection}{0}%
- \setcounter{figure}{0}
- \setcounter{table}{0}
- \setcounter{equation}{0}
- \setcounter{footnote}{0}%
- \begingroup
- \parindent=\z@
- \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
- \if@twocolumn
- \ifnum \col@number=\@ne
- \@maketitle
- \else
- \twocolumn[\@maketitle]%
- \fi
- \else
- \newpage
- \global\@topnum\z@ % Prevents figures from going at top of page.
- \@maketitle
- \fi
- \thispagestyle{empty}\@thanks
-%
- \def\\{\unskip\ \ignorespaces}\def\inst##1{\unskip{}}%
- \def\thanks##1{\unskip{}}\def\fnmsep{\unskip}%
- \instindent=\hsize
- \advance\instindent by-\headlineindent
-% \if!\the\toctitle!\addcontentsline{toc}{title}{\@title}\else
-% \addcontentsline{toc}{title}{\the\toctitle}\fi
- \if@runhead
- \if!\the\titlerunning!\else
- \edef\@title{\the\titlerunning}%
- \fi
- \global\setbox\titrun=\hbox{\small\rm\unboldmath\ignorespaces\@title}%
- \ifdim\wd\titrun>\instindent
- \typeout{Title too long for running head. Please supply}%
- \typeout{a shorter form with \string\titlerunning\space prior to
- \string\maketitle}%
- \global\setbox\titrun=\hbox{\small\rm
- Title Suppressed Due to Excessive Length}%
- \fi
- \xdef\@title{\copy\titrun}%
- \fi
-%
- \if!\the\tocauthor!\relax
- {\def\and{\noexpand\protect\noexpand\and}%
- \protected@xdef\toc@uthor{\@author}}%
- \else
- \def\\{\noexpand\protect\noexpand\newline}%
- \protected@xdef\scratch{\the\tocauthor}%
- \protected@xdef\toc@uthor{\scratch}%
- \fi
-% \addcontentsline{toc}{author}{\toc@uthor}%
- \if@runhead
- \if!\the\authorrunning!
- \value{@inst}=\value{@auth}%
- \setcounter{@auth}{1}%
- \else
- \edef\@author{\the\authorrunning}%
- \fi
- \global\setbox\authrun=\hbox{\small\unboldmath\@author\unskip}%
- \ifdim\wd\authrun>\instindent
- \typeout{Names of authors too long for running head. Please supply}%
- \typeout{a shorter form with \string\authorrunning\space prior to
- \string\maketitle}%
- \global\setbox\authrun=\hbox{\small\rm
- Authors Suppressed Due to Excessive Length}%
- \fi
- \xdef\@author{\copy\authrun}%
- \markboth{\@author}{\@title}%
- \fi
- \endgroup
- \setcounter{footnote}{\fnnstart}%
- \clearheadinfo}
-%
-\def\@maketitle{\newpage
- \markboth{}{}%
- \def\lastand{\ifnum\value{@inst}=2\relax
- \unskip{} \andname\
- \else
- \unskip \lastandname\
- \fi}%
- \def\and{\stepcounter{@auth}\relax
- \ifnum\value{@auth}=\value{@inst}%
- \lastand
- \else
- \unskip,
- \fi}%
- \begin{center}%
- \let\newline\\
- {\Large \bfseries\boldmath
- \pretolerance=10000
- \@title \par}\vskip .8cm
-\if!\@subtitle!\else {\large \bfseries\boldmath
- \vskip -.65cm
- \pretolerance=10000
- \@subtitle \par}\vskip .8cm\fi
- \setbox0=\vbox{\setcounter{@auth}{1}\def\and{\stepcounter{@auth}}%
- \def\thanks##1{}\@author}%
- \global\value{@inst}=\value{@auth}%
- \global\value{auco}=\value{@auth}%
- \setcounter{@auth}{1}%
-{\lineskip .5em
-\noindent\ignorespaces
-\@author\vskip.35cm}
- {\small\institutename}
- \end{center}%
- }
-
-% definition of the "\spnewtheorem" command.
-%
-% Usage:
-%
-% \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font}
-% or \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font}
-% or \spnewtheorem*{env_nam}{caption}{cap_font}{body_font}
-%
-% New is "cap_font" and "body_font". It stands for
-% fontdefinition of the caption and the text itself.
-%
-% "\spnewtheorem*" gives a theorem without number.
-%
-% A defined spnewthoerem environment is used as described
-% by Lamport.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\def\@thmcountersep{}
-\def\@thmcounterend{.}
-
-\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}}
-
-% definition of \spnewtheorem with number
-
-\def\@spnthm#1#2{%
- \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}}
-\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}}
-
-\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}\@addtoreset{#1}{#3}%
- \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand
- \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}%
- \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spothm#1[#2]#3#4#5{%
- \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}%
- {\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{the#1}{\@nameuse{the#2}}%
- \expandafter\xdef\csname #1name\endcsname{#3}%
- \global\@namedef{#1}{\@spthm{#2}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}}
-
-\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\refstepcounter{#1}%
-\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}}
-
-\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}%
- \ignorespaces}
-
-\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname
- the#1\endcsname}{#5}{#3}{#4}\ignorespaces}
-
-\def\@spbegintheorem#1#2#3#4{\trivlist
- \item[\hskip\labelsep{#3#1\ #2\@thmcounterend}]#4}
-
-\def\@spopargbegintheorem#1#2#3#4#5{\trivlist
- \item[\hskip\labelsep{#4#1\ #2}]{#4(#3)\@thmcounterend\ }#5}
-
-% definition of \spnewtheorem* without number
-
-\def\@sthm#1#2{\@Ynthm{#1}{#2}}
-
-\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}}
-
-\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces}
-
-\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1}
- {#4}{#2}{#3}\ignorespaces}
-
-\def\@Begintheorem#1#2#3{#3\trivlist
- \item[\hskip\labelsep{#2#1\@thmcounterend}]}
-
-\def\@Opargbegintheorem#1#2#3#4{#4\trivlist
- \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }}
-
-\if@envcntsect
- \def\@thmcountersep{.}
- \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape}
-\else
- \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape}
- \if@envcntreset
- \@addtoreset{theorem}{section}
- \else
- \@addtoreset{theorem}{chapter}
- \fi
-\fi
-
-%definition of divers theorem environments
-\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily}
-\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily}
-\if@envcntsame % alle Umgebungen wie Theorem.
- \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}}
-\else % alle Umgebungen mit eigenem Zaehler
- \if@envcntsect % mit section numeriert
- \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}}
- \else % nicht mit section numeriert
- \if@envcntreset
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{section}}
- \else
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{chapter}}%
- \fi
- \fi
-\fi
-\spn@wtheorem{case}{Case}{\itshape}{\rmfamily}
-\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily}
-\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape}
-\spn@wtheorem{definition}{Definition}{\bfseries}{\itshape}
-\spn@wtheorem{example}{Example}{\itshape}{\rmfamily}
-\spn@wtheorem{exercise}{Exercise}{\itshape}{\rmfamily}
-\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape}
-\spn@wtheorem{note}{Note}{\itshape}{\rmfamily}
-\spn@wtheorem{problem}{Problem}{\itshape}{\rmfamily}
-\spn@wtheorem{property}{Property}{\itshape}{\rmfamily}
-\spn@wtheorem{proposition}{Proposition}{\bfseries}{\rmfamily}
-\spn@wtheorem{question}{Question}{\itshape}{\rmfamily}
-\spn@wtheorem{solution}{Solution}{\itshape}{\rmfamily}
-\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily}
-
-\def\@takefromreset#1#2{%
- \def\@tempa{#1}%
- \let\@tempd\@elt
- \def\@elt##1{%
- \def\@tempb{##1}%
- \ifx\@tempa\@tempb\else
- \@addtoreset{##1}{#2}%
- \fi}%
- \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
- \expandafter\def\csname cl@#2\endcsname{}%
- \@tempc
- \let\@elt\@tempd}
-
-\def\theopargself{\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5}
- \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
- \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }}
- }
-
-\renewenvironment{abstract}{%
- \list{}{\advance\topsep by0.35cm\relax\small
- \leftmargin=1cm
- \labelwidth=\z@
- \listparindent=\z@
- \itemindent\listparindent
- \rightmargin\leftmargin}\item[\hskip\labelsep
- \bfseries\abstractname]}
- {\endlist}
-
-\newdimen\headlineindent % dimension for space between
-\headlineindent=1.166cm % number and text of headings.
-
-\def\ps@headings{\let\@mkboth\@gobbletwo
- \let\@oddfoot\@empty\let\@evenfoot\@empty
- \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
- \leftmark\hfil}
- \def\@oddhead{\normalfont\small\hfil\rightmark\hspace{\headlineindent}%
- \llap{\thepage}}
- \def\chaptermark##1{}%
- \def\sectionmark##1{}%
- \def\subsectionmark##1{}}
-
-\def\ps@titlepage{\let\@mkboth\@gobbletwo
- \let\@oddfoot\@empty\let\@evenfoot\@empty
- \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
- \hfil}
- \def\@oddhead{\normalfont\small\hfil\hspace{\headlineindent}%
- \llap{\thepage}}
- \def\chaptermark##1{}%
- \def\sectionmark##1{}%
- \def\subsectionmark##1{}}
-
-\if@runhead\ps@headings\else
-\ps@empty\fi
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\endinput
-%end of file llncs.cls
--- a/Quotient-Paper/document/mathpartir.sty Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,388 +0,0 @@
-% Mathpartir --- Math Paragraph for Typesetting Inference Rules
-%
-% Copyright (C) 2001, 2002, 2003 Didier Rémy
-%
-% Author : Didier Remy
-% Version : 1.1.1
-% Bug Reports : to author
-% Web Site : http://pauillac.inria.fr/~remy/latex/
-%
-% WhizzyTeX 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}
- [2003/07/10 version 1.1.1 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
-}
-
-%% Part III -- Type inference rules
-
-\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}
-
-\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@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 P 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{$\displaystyle {##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
- }{}
-
-
-\def \mpr@@fraction #1#2{\hbox {\advance \hsize by -0.5em
- $\displaystyle {#1\@@over #2}$}}
-\let \mpr@fraction \mpr@@fraction
-\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
- \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.
-%
-
-\define@key {mprset}{flushleft}[]{\mpr@centerfalse}
-\define@key {mprset}{center}[]{\mpr@centertrue}
-\def \mprset #1{\setkeys{mprset}{#1}}
-
-\newbox \mpr@right
-\define@key {mpr}{flushleft}[]{\mpr@centerfalse}
-\define@key {mpr}{center}[]{\mpr@centertrue}
-\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 \tir@name #1{\hbox {\small \sc #1}}
-\let \TirName \tir@name
-\let \RefTirName \tir@name
-
-%%% Other Exports
-
-% \let \listcons \mpr@cons
-% \let \listsnoc \mpr@snoc
-% \let \listhead \mpr@head
-% \let \listmake \mpr@makelist
-
-
-
-
-\endinput
--- a/Quotient-Paper/document/root-lncs.tex Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,69 +0,0 @@
-%\documentclass{svjour3}
-\documentclass{llncs}
-\usepackage{times}
-\usepackage{isabelle}
-\usepackage{isabellesym}
-\usepackage{amsmath}
-\usepackage{amssymb}
-\usepackage{pdfsetup}
-\usepackage{tikz}
-\usepackage{pgf}
-\usepackage{verbdef}
-\usepackage{longtable}
-\usepackage{mathpartir}
-
-\urlstyle{rm}
-\isabellestyle{it}
-\renewcommand{\isastyle}{\isastyleminor}
-
-\def\dn{\,\stackrel{\mbox{\scriptsize def}}{=}\,}
-\verbdef\singlearr|--->|
-\verbdef\doublearr|===>|
-\verbdef\tripple|###|
-
-\renewcommand{\isasymequiv}{$\dn$}
-\renewcommand{\isasymemptyset}{$\varnothing$}
-\renewcommand{\isacharunderscore}{\mbox{$\_\!\_$}}
-\renewcommand{\isasymUnion}{$\bigcup$}
-
-\newcommand{\isasymsinglearr}{\singlearr}
-\newcommand{\isasymdoublearr}{\doublearr}
-\newcommand{\isasymtripple}{\tripple}
-
-\newcommand{\numbered}[1]{\refstepcounter{equation}{\rm(\arabic{equation})}\label{#1}}
-
-\begin{document}
-
-\title{Quotients Revisited for Isabelle/HOL}
-\author{Cezary Kaliszyk$^*$ and Christian Urban$^*$}
-\institute{$^*$ Technical University of Munich, Germany}
-\maketitle
-
-\begin{abstract}
-Higher-Order Logic (HOL) is based on a small logic kernel, whose only
-mechanism for extension is the introduction of safe definitions and of
-non-empty types. Both extensions are often performed in quotient
-constructions. To ease the work involved with such quotient constructions, we
-re-implemented in Isabelle/HOL the quotient package by Homeier. In doing so we
-extended his work in order to deal with compositions of quotients. Like his
-package, we designed our quotient package so that every step in a quotient construction
-can be performed separately and as a result we are able to specify completely
-the procedure of lifting theorems from the raw level to the quotient level.
-The importance for programming language research is that many properties of
-programming language calculi are easier to verify over $\alpha$-equated, or
-$\alpha$-quotient, terms, than over raw terms.
-\end{abstract}
-
-% generated text of all theories
-\input{session}
-
-% optional bibliography
-\bibliographystyle{abbrv}
-\bibliography{root}
-
-\end{document}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
--- a/Quotient-Paper/document/root-sac.tex Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,92 +0,0 @@
-\documentclass{sig-alternate}
- \pdfpagewidth=8.5truein
- \pdfpageheight=11truein
-\usepackage{times}
-\usepackage{isabelle}
-\usepackage{isabellesym}
-\usepackage{amsmath}
-\usepackage{amssymb}
-\usepackage{pdfsetup}
-\usepackage{tikz}
-\usepackage{pgf}
-\usepackage{verbdef}
-\usepackage{longtable}
-\usepackage{mathpartir}
-\newtheorem{definition}{Definition}
-\newtheorem{proposition}{Proposition}
-\newtheorem{lemma}{Lemma}
-
-\urlstyle{rm}
-\isabellestyle{it}
-\renewcommand{\isastyleminor}{\it}%
-\renewcommand{\isastyle}{\normalsize\rm}%
-
-\def\dn{\,\stackrel{\mbox{\scriptsize def}}{=}\,}
-\verbdef\singlearr|--->|
-\verbdef\doublearr|===>|
-\verbdef\tripple|###|
-
-\renewcommand{\isasymequiv}{$\dn$}
-\renewcommand{\isasymemptyset}{$\varnothing$}
-\renewcommand{\isacharunderscore}{\mbox{$\_\!\_$}}
-\renewcommand{\isasymUnion}{$\bigcup$}
-
-\newcommand{\isasymsinglearr}{\singlearr}
-\newcommand{\isasymdoublearr}{\doublearr}
-\newcommand{\isasymtripple}{\tripple}
-
-\newcommand{\numbered}[1]{\refstepcounter{equation}{\rm(\arabic{equation})}\label{#1}}
-
-\begin{document}
-
-\conferenceinfo{SAC'11}{March 21-25, 2011, TaiChung, Taiwan.}
-\CopyrightYear{2011}
-\crdata{978-1-4503-0113-8/11/03}
-
-\title{Quotients Revisited for Isabelle/HOL}
-\numberofauthors{2}
-\author{
-\alignauthor
-Cezary Kaliszyk\\
- \affaddr{University of Tsukuba, Japan}\\
- \email{kaliszyk@score.cs.tsukuba.ac.jp}
-\alignauthor
-Christian Urban\\
- \affaddr{Technical University of Munich, Germany}\\
- \email{urbanc@in.tum.de}
-}
-
-\maketitle
-
-\begin{abstract}
-Higher-Order Logic (HOL) is based on a small logic kernel, whose only
-mechanism for extension is the introduction of safe definitions and of
-non-empty types. Both extensions are often performed in quotient
-constructions. To ease the work involved with such quotient constructions, we
-re-implemented in Isabelle/HOL the quotient package by Homeier. In doing so we
-extended his work in order to deal with compositions of quotients. Like his
-package, we designed our quotient package so that every step in a quotient construction
-can be performed separately and as a result we are able to specify completely
-the procedure of lifting theorems from the raw level to the quotient level.
-The importance for programming language research is that many properties of
-programming language calculi are easier to verify over $\alpha$-equated, or
-$\alpha$-quotient, terms, than over raw terms.
-\end{abstract}
-
-\category{D.??}{TODO}{TODO}
-
-\keywords{quotients, isabelle, higher order logic}
-
-% generated text of all theories
-\input{session}
-
-% optional bibliography
-\bibliographystyle{abbrv}
-\bibliography{root}
-
-\end{document}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
--- a/Quotient-Paper/document/root.bib Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,203 +0,0 @@
-@inproceedings{Nogin02,
- author = {Aleksey Nogin},
- title = {Quotient Types: A Modular Approach},
- booktitle = {Proc.~of the 15th TPHOLs conference},
- year = {2002},
- pages = {263-280},
- series = {LNCS},
- volume = {2646}
-}
-
-@proceedings{DBLP:conf/tphol/2002,
- editor = {Victor Carre{\~n}o and
- C{\'e}sar Mu{\~n}oz and
- Sofi{\`e}ne Tahar},
- title = {Theorem Proving in Higher Order Logics, 15th International
- Conference, TPHOLs 2002, Hampton, VA, USA, August 20-23,
- 2002, Proceedings},
- booktitle = {TPHOLs},
- publisher = {Springer},
- series = {Lecture Notes in Computer Science},
- volume = {2410},
- year = {2002},
- isbn = {3-540-44039-9},
- bibsource = {DBLP, http://dblp.uni-trier.de}
-}
-
-@techreport{PVS:Interpretations,
- Author= {S. Owre and N. Shankar},
- Title= {{T}heory {I}nterpretations in {PVS}},
- Number= {SRI-CSL-01-01},
- Institution= {Computer Science Laboratory, SRI International},
- Address= {Menlo Park, CA},
- Month= {April},
- Year= {2001}}
-
-@inproceedings{ChicliPS02,
- author = {Laurent Chicli and
- Loic Pottier and
- Carlos Simpson},
- title = {{M}athematical {Q}uotients and {Q}uotient {T}ypes in {C}oq},
- booktitle = {Proc of the TYPES workshop},
- year = {2002},
- pages = {95-107},
- series = {LNCS},
- volume = {2646}
-}
-
-@proceedings{DBLP:conf/types/2002,
- editor = {Herman Geuvers and
- Freek Wiedijk},
- title = {Types for Proofs and Programs, Second International Workshop,
- TYPES 2002, Berg en Dal, The Netherlands, April 24-28, 2002,
- Selected Papers},
- booktitle = {TYPES},
- publisher = {Springer},
- series = {Lecture Notes in Computer Science},
- volume = {2646},
- year = {2003},
- isbn = {3-540-14031-X},
- bibsource = {DBLP, http://dblp.uni-trier.de}
-}
-
-@article{Paulson06,
- author = {Lawrence C. Paulson},
- title = {Defining functions on equivalence classes},
- journal = {ACM Trans. Comput. Log.},
- volume = {7},
- number = {4},
- year = {2006},
- pages = {658-675},
- ee = {http://doi.acm.org/10.1145/1183278.1183280},
- bibsource = {DBLP, http://dblp.uni-trier.de}
-}
-
-@inproceedings{Slotosch97,
- author = {Oscar Slotosch},
- title = {Higher Order Quotients and their Implementation in Isabelle
- HOL},
- booktitle = {Proc.~of the 10th TPHOLs conference},
- year = {1997},
- pages = {291-306},
- series = {LNCS},
- volume = {1275}
-}
-
-@proceedings{DBLP:conf/tphol/1997,
- editor = {Elsa L. Gunter and
- Amy P. Felty},
- title = {Theorem Proving in Higher Order Logics, 10th International
- Conference, TPHOLs'97, Murray Hill, NJ, USA, August 19-22,
- 1997, Proceedings},
- booktitle = {TPHOLs},
- publisher = {Springer},
- series = {Lecture Notes in Computer Science},
- volume = {1275},
- year = {1997},
- isbn = {3-540-63379-0},
- bibsource = {DBLP, http://dblp.uni-trier.de}
-}
-
-@inproceedings{Homeier05,
- author = {Peter V. Homeier},
- title = {A Design Structure for Higher Order Quotients},
- booktitle = {Proc of the 18th TPHOLs conference},
- year = {2005},
- pages = {130-146},
- series = {LNCS},
- volume = {3603}
-}
-
-@proceedings{DBLP:conf/tphol/2005,
- editor = {Joe Hurd and
- Thomas F. Melham},
- title = {Theorem Proving in Higher Order Logics, 18th International
- Conference, TPHOLs 2005, Oxford, UK, August 22-25, 2005,
- Proceedings},
- booktitle = {TPHOLs},
- publisher = {Springer},
- series = {Lecture Notes in Computer Science},
- volume = {3603},
- year = {2005},
- isbn = {3-540-28372-2},
- bibsource = {DBLP, http://dblp.uni-trier.de}
-}
-
-@BOOK{harrison-thesis,
- author = "John Harrison",
- title = "Theorem Proving with the Real Numbers",
- publisher = "Springer-Verlag",
- year = 1998}
-
-@BOOK{Barendregt81,
- AUTHOR = "H.~Barendregt",
- TITLE = "{T}he {L}ambda {C}alculus: {I}ts {S}yntax and {S}emantics",
- PUBLISHER = "North-Holland",
- YEAR = 1981,
- VOLUME = 103,
- SERIES = "Studies in Logic and the Foundations of Mathematics"
-}
-
-@BOOK{CurryFeys58,
- AUTHOR = "H.~B.~Curry and R.~Feys",
- TITLE = "{C}ombinatory {L}ogic",
- PUBLISHER = "North-Holland",
- YEAR = "1958",
- VOLUME = 1,
- SERIES = "Studies in Logic and the Foundations of Mathematics"
-}
-
-@Unpublished{UrbanKaliszyk11,
- author = {C.~Urban and C.~Kaliszyk},
- title = {{G}eneral {B}indings and {A}lpha-{E}quivalence in {N}ominal {I}sabelle},
- note = {submitted for publication},
- month = {July},
- year = {2010},
-}
-
-@InProceedings{BengtsonParrow07,
- author = {J.~Bengtson and J.~Parrow},
- title = {Formalising the pi-{C}alculus using {N}ominal {L}ogic},
- booktitle = {Proc.~of the 10th FOSSACS Conference},
- year = 2007,
- pages = {63--77},
- series = {LNCS},
- volume = {4423}
-}
-
-@inproceedings{BengtsonParow09,
- author = {J.~Bengtson and J.~Parrow},
- title = {{P}si-{C}alculi in {I}sabelle},
- booktitle = {Proc of the 22nd TPHOLs Conference},
- year = 2009,
- pages = {99--114},
- series = {LNCS},
- volume = {5674}
-}
-
-@inproceedings{TobinHochstadtFelleisen08,
- author = {S.~Tobin-Hochstadt and M.~Felleisen},
- booktitle = {Proc.~of the 35rd POPL Symposium},
- title = {{T}he {D}esign and {I}mplementation of {T}yped {S}cheme},
- publisher = {ACM},
- year = {2008},
- pages = {395--406}
-}
-
-@InProceedings{UrbanCheneyBerghofer08,
- author = "C.~Urban and J.~Cheney and S.~Berghofer",
- title = "{M}echanizing the {M}etatheory of {LF}",
- pages = "45--56",
- year = 2008,
- booktitle = "Proc.~of the 23rd LICS Symposium"
-}
-
-@InProceedings{UrbanZhu08,
- title = "{R}evisiting {C}ut-{E}limination: {O}ne {D}ifficult {P}roof is {R}eally a {P}roof",
- author = "C.~Urban and B.~Zhu",
- booktitle = "Proc.~of the 9th RTA Conference",
- year = "2008",
- pages = "409--424",
- series = "LNCS",
- volume = 5117
-}
\ No newline at end of file
--- a/Quotient-Paper/document/root.tex Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,92 +0,0 @@
-\documentclass{sig-alternate}
- \pdfpagewidth=8.5truein
- \pdfpageheight=11truein
-\usepackage{times}
-\usepackage{isabelle}
-\usepackage{isabellesym}
-\usepackage{amsmath}
-\usepackage{amssymb}
-\usepackage{pdfsetup}
-\usepackage{tikz}
-\usepackage{pgf}
-\usepackage{stmaryrd}
-\usepackage{verbdef}
-\usepackage{longtable}
-\usepackage{mathpartir}
-\newtheorem{definition}{Definition}
-\newtheorem{proposition}{Proposition}
-\newtheorem{lemma}{Lemma}
-
-\urlstyle{rm}
-\isabellestyle{rm}
-\renewcommand{\isastyleminor}{\rm}%
-\renewcommand{\isastyle}{\normalsize\rm}%
-\renewcommand{\isastylescript}{\it}
-\def\dn{\,\triangleq\,}
-\verbdef\singlearr|---->|
-\verbdef\doublearr|===>|
-\verbdef\tripple|###|
-
-\renewcommand{\isasymequiv}{$\triangleq$}
-\renewcommand{\isasymemptyset}{$\varnothing$}
-%%\renewcommand{\isacharunderscore}{\mbox{$\_\!\_$}}
-\renewcommand{\isasymUnion}{$\bigcup$}
-
-\newcommand{\isasymsinglearr}{$\mapsto$}
-\newcommand{\isasymdoublearr}{$\Mapsto$}
-\newcommand{\isasymtripple}{\tripple}
-
-\newcommand{\numbered}[1]{\refstepcounter{equation}{\rm(\arabic{equation})}\label{#1}}
-
-\begin{document}
-
-\conferenceinfo{SAC'11}{March 21-25, 2011, TaiChung, Taiwan.}
-\CopyrightYear{2011}
-\crdata{978-1-4503-0113-8/11/03}
-
-\title{Quotients Revisited for Isabelle/HOL}
-\numberofauthors{2}
-\author{
-\alignauthor
-Cezary Kaliszyk\\
- \affaddr{University of Tsukuba, Japan}\\
- \email{kaliszyk@cs.tsukuba.ac.jp}
-\alignauthor
-Christian Urban\\
- \affaddr{Technical University of Munich, Germany}\\
- \email{urbanc@in.tum.de}
-}
-
-\maketitle
-
-\begin{abstract}
-Higher-Order Logic (HOL) is based on a small logic kernel, whose only
-mechanism for extension is the introduction of safe definitions and of
-non-empty types. Both extensions are often performed in quotient
-constructions. To ease the work involved with such quotient constructions, we
-re-implemented in the %popular
-Isabelle/HOL theorem prover the quotient
-package by Homeier. In doing so we extended his work in order to deal with
-compositions of quotients and also specified completely the procedure
-of lifting theorems from the raw level to the quotient level.
-The importance for theorem proving is that many formal
-verifications, in order to be feasible, require a convenient reasoning infrastructure
-for quotient constructions.
-\end{abstract}
-
-%\category{D.??}{TODO}{TODO}
-
-\keywords{Quotients, Isabelle theorem prover, Higher-Order Logic}
-
-% generated text of all theories
-\bibliographystyle{abbrv}
-\input{session}
-
-
-
-\end{document}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
--- a/Quotient-Paper/document/sig-alternate.cls Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1649 +0,0 @@
-% SIG-ALTERNATE.CLS - VERSION 2.4
-% "COMPATIBLE" WITH THE "ACM_PROC_ARTICLE-SP.CLS" V3.2SP
-% Gerald Murray - April 22nd. 2009
-%
-% ---- Start of 'updates' ----
-%
-% Changed $5 fee to $10 - Gerry
-% April 22nd. 2009 - Fixed 'Natbib' incompatibility problem - Gerry
-% April 22nd. 2009 - Fixed 'Babel' incompatibility problem - Gerry
-% April 22nd. 2009 - Inserted various bug-fixes and improvements - Gerry
-%
-% To produce Type 1 fonts in the document plus allow for 'normal LaTeX accenting' in the critical areas;
-% title, author block, section-heads, confname, etc. etc.
-% i.e. the whole purpose of this version update is to NOT resort to 'inelegant accent patches'.
-% After much research, three extra .sty packages were added to the the tail (ae, aecompl, aeguill) to solve,
-% in particular, the accenting problem(s). We _could_ ask authors (via instructions/sample file) to 'include' these in
-% the source .tex file - in the preamble - but if everything is already provided ('behind the scenes' - embedded IN the .cls)
-% then this is less work for authors and also makes everything appear 'vanilla'.
-% NOTE: all 'patchwork accenting" has been commented out (here) and is no longer 'used' in the sample .tex file (either).
-% Gerry June 2007
-%
-% Patch for accenting in conference name/location. Gerry May 3rd. 2007
-% Rule widths changed to .5, author count (>6) fixed, roll-back for Type 3 problem. Gerry March 20th. 2007
-% Changes made to 'modernize' the fontnames but esp. for MikTeX users V2.4/2.5 - Nov. 30th. 2006
-% Updated the \email definition to allow for its use inside of 'shared affiliations' - Nov. 30th. 2006
-% Fixed the 'section number depth value' - Nov. 30th. 2006
-%
-% Footnotes inside table cells using \minipage (Oct. 2002)
-% Georgia fixed bug in sub-sub-section numbering in paragraphs (July 29th. 2002)
-% JS/GM fix to vertical spacing before Proofs (July 30th. 2002)
-%
-% Made the Permission Statement / Conference Info / Copyright Info
-% 'user definable' in the source .tex file OR automatic if
-% not specified.
-%
-% Allowance made to switch default fonts between those systems using
-% normal/modern font names and those using 'Type 1' or 'Truetype' fonts.
-% See LINE NUMBER 255 for details.
-% Also provided for enumerated/annotated Corollaries 'surrounded' by
-% enumerated Theorems (line 848).
-% Gerry November 11th. 1999
-%
-% ---- End of 'updates' ----
-%
-\def\fileversion{v2.4} % for ACM's tracking purposes
-\def\filedate{April 22, 2009} % Gerry Murray's tracking data
-\def\docdate {Wednesday 22nd. April 2009} % Gerry Murray (with deltas to doc}
-\usepackage{epsfig}
-\usepackage{amssymb}
-\usepackage{amsmath}
-\usepackage{amsfonts}
-% Need this for accents in Arial/Helvetica
-%\usepackage[T1]{fontenc} % Gerry March 12, 2007 - causes Type 3 problems (body text)
-%\usepackage{textcomp}
-%
-% SIG-ALTERNATE DOCUMENT STYLE
-% G.K.M. Tobin August-October 1999
-% adapted from ARTICLE document style by Ken Traub, Olin Shivers
-% also using elements of esub2acm.cls
-% HEAVILY MODIFIED, SUBSEQUENTLY, BY GERRY MURRAY 2000
-% ARTICLE DOCUMENT STYLE -- Released 16 March 1988
-% for LaTeX version 2.09
-% Copyright (C) 1988 by Leslie Lamport
-%
-%
-%%% sig-alternate.cls is an 'ALTERNATE' document style for producing
-%%% two-column camera-ready pages for ACM conferences.
-%%% THIS FILE DOES NOT STRICTLY ADHERE TO THE SIGS (BOARD-ENDORSED)
-%%% PROCEEDINGS STYLE. It has been designed to produce a 'tighter'
-%%% paper in response to concerns over page budgets.
-%%% The main features of this style are:
-%%%
-%%% 1) Two columns.
-%%% 2) Side and top margins of 4.5pc, bottom margin of 6pc, column gutter of
-%%% 2pc, hence columns are 20pc wide and 55.5pc tall. (6pc =3D 1in, approx)
-%%% 3) First page has title information, and an extra 6pc of space at the
-%%% bottom of the first column for the ACM copyright notice.
-%%% 4) Text is 9pt on 10pt baselines; titles (except main) are 9pt bold.
-%%%
-%%%
-%%% There are a few restrictions you must observe:
-%%%
-%%% 1) You cannot change the font size; ACM wants you to use 9pt.
-%%% 3) You must start your paper with the \maketitle command. Prior to the
-%%% \maketitle you must have \title and \author commands. If you have a
-%%% \date command it will be ignored; no date appears on the paper, since
-%%% the proceedings will have a date on the front cover.
-%%% 4) Marginal paragraphs, tables of contents, lists of figures and tables,
-%%% and page headings are all forbidden.
-%%% 5) The `figure' environment will produce a figure one column wide; if you
-%%% want one that is two columns wide, use `figure*'.
-%%%
-%
-%%% Copyright Space:
-%%% This style automatically reserves 1" blank space at the bottom of page 1/
-%%% column 1. This space can optionally be filled with some text using the
-%%% \toappear{...} command. If used, this command must be BEFORE the \maketitle
-%%% command. If this command is defined AND [preprint] is on, then the
-%%% space is filled with the {...} text (at the bottom); otherwise, it is
-%%% blank. If you use \toappearbox{...} instead of \toappear{...} then a
-%%% box will be drawn around the text (if [preprint] is on).
-%%%
-%%% A typical usage looks like this:
-%%% \toappear{To appear in the Ninth AES Conference on Medievil Lithuanian
-%%% Embalming Technique, June 1991, Alfaretta, Georgia.}
-%%% This will be included in the preprint, and left out of the conference
-%%% version.
-%%%
-%%% WARNING:
-%%% Some dvi-ps converters heuristically allow chars to drift from their
-%%% true positions a few pixels. This may be noticeable with the 9pt sans-serif
-%%% bold font used for section headers.
-%%% You may turn this hackery off via the -e option:
-%%% dvips -e 0 foo.dvi >foo.ps
-%%%
-\typeout{Document Class 'sig-alternate' <22nd. April '09>. Modified by G.K.M. Tobin/Gerry Murray}
-\typeout{Based in part upon document Style `acmconf' <22 May 89>. Hacked 4/91 by}
-\typeout{shivers@cs.cmu.edu, 4/93 by theobald@cs.mcgill.ca}
-\typeout{Excerpts were taken from (Journal Style) 'esub2acm.cls'.}
-\typeout{****** Bugs/comments/suggestions/technicalities to Gerry Murray -- murray@hq.acm.org ******}
-\typeout{Questions on the style, SIGS policies, etc. to Adrienne Griscti griscti@acm.org}
-\oddsidemargin 4.5pc
-\evensidemargin 4.5pc
-\advance\oddsidemargin by -1in % Correct for LaTeX gratuitousness
-\advance\evensidemargin by -1in % Correct for LaTeX gratuitousness
-\marginparwidth 0pt % Margin pars are not allowed.
-\marginparsep 11pt % Horizontal space between outer margin and
- % marginal note
-
- % Top of page:
-\topmargin 4.5pc % Nominal distance from top of page to top of
- % box containing running head.
-\advance\topmargin by -1in % Correct for LaTeX gratuitousness
-\headheight 0pt % Height of box containing running head.
-\headsep 0pt % Space between running head and text.
- % Bottom of page:
-\footskip 30pt % Distance from baseline of box containing foot
- % to baseline of last line of text.
-\@ifundefined{footheight}{\newdimen\footheight}{}% this is for LaTeX2e
-\footheight 12pt % Height of box containing running foot.
-
-%% Must redefine the top margin so there's room for headers and
-%% page numbers if you are using the preprint option. Footers
-%% are OK as is. Olin.
-\advance\topmargin by -37pt % Leave 37pt above text for headers
-\headheight 12pt % Height of box containing running head.
-\headsep 25pt % Space between running head and text.
-
-\textheight 666pt % 9 1/4 column height
-\textwidth 42pc % Width of text line.
- % For two-column mode:
-\columnsep 2pc % Space between columns
-\columnseprule 0pt % Width of rule between columns.
-\hfuzz 1pt % Allow some variation in column width, otherwise it's
- % too hard to typeset in narrow columns.
-
-\footnotesep 5.6pt % Height of strut placed at the beginning of every
- % footnote =3D height of normal \footnotesize strut,
- % so no extra space between footnotes.
-
-\skip\footins 8.1pt plus 4pt minus 2pt % Space between last line of text and
- % top of first footnote.
-\floatsep 11pt plus 2pt minus 2pt % Space between adjacent floats moved
- % to top or bottom of text page.
-\textfloatsep 18pt plus 2pt minus 4pt % Space between main text and floats
- % at top or bottom of page.
-\intextsep 11pt plus 2pt minus 2pt % Space between in-text figures and
- % text.
-\@ifundefined{@maxsep}{\newdimen\@maxsep}{}% this is for LaTeX2e
-\@maxsep 18pt % The maximum of \floatsep,
- % \textfloatsep and \intextsep (minus
- % the stretch and shrink).
-\dblfloatsep 11pt plus 2pt minus 2pt % Same as \floatsep for double-column
- % figures in two-column mode.
-\dbltextfloatsep 18pt plus 2pt minus 4pt% \textfloatsep for double-column
- % floats.
-\@ifundefined{@dblmaxsep}{\newdimen\@dblmaxsep}{}% this is for LaTeX2e
-\@dblmaxsep 18pt % The maximum of \dblfloatsep and
- % \dbltexfloatsep.
-\@fptop 0pt plus 1fil % Stretch at top of float page/column. (Must be
- % 0pt plus ...)
-\@fpsep 8pt plus 2fil % Space between floats on float page/column.
-\@fpbot 0pt plus 1fil % Stretch at bottom of float page/column. (Must be
- % 0pt plus ... )
-\@dblfptop 0pt plus 1fil % Stretch at top of float page. (Must be 0pt plus ...)
-\@dblfpsep 8pt plus 2fil % Space between floats on float page.
-\@dblfpbot 0pt plus 1fil % Stretch at bottom of float page. (Must be
- % 0pt plus ... )
-\marginparpush 5pt % Minimum vertical separation between two marginal
- % notes.
-
-\parskip 0pt plus 1pt % Extra vertical space between paragraphs.
-\parindent 9pt % GM July 2000 / was 0pt - width of paragraph indentation.
-\partopsep 2pt plus 1pt minus 1pt% Extra vertical space, in addition to
- % \parskip and \topsep, added when user
- % leaves blank line before environment.
-
-\@lowpenalty 51 % Produced by \nopagebreak[1] or \nolinebreak[1]
-\@medpenalty 151 % Produced by \nopagebreak[2] or \nolinebreak[2]
-\@highpenalty 301 % Produced by \nopagebreak[3] or \nolinebreak[3]
-
-\@beginparpenalty -\@lowpenalty % Before a list or paragraph environment.
-\@endparpenalty -\@lowpenalty % After a list or paragraph environment.
-\@itempenalty -\@lowpenalty % Between list items.
-
-%\@namedef{ds@10pt}{\@latexerr{The `10pt' option is not allowed in the `acmconf'
-\@namedef{ds@10pt}{\ClassError{The `10pt' option is not allowed in the `acmconf' % January 2008
- document style.}\@eha}
-%\@namedef{ds@11pt}{\@latexerr{The `11pt' option is not allowed in the `acmconf'
-\@namedef{ds@11pt}{\ClassError{The `11pt' option is not allowed in the `acmconf' % January 2008
- document style.}\@eha}
-%\@namedef{ds@12pt}{\@latexerr{The `12pt' option is not allowed in the `acmconf'
-\@namedef{ds@12pt}{\ClassError{The `12pt' option is not allowed in the `acmconf' % January 2008
- document style.}\@eha}
-
-\@options
-
-\lineskip 2pt % \lineskip is 1pt for all font sizes.
-\normallineskip 2pt
-\def\baselinestretch{1}
-
-\abovedisplayskip 9pt plus2pt minus4.5pt%
-\belowdisplayskip \abovedisplayskip
-\abovedisplayshortskip \z@ plus3pt%
-\belowdisplayshortskip 5.4pt plus3pt minus3pt%
-\let\@listi\@listI % Setting of \@listi added 9 Jun 87
-
-\def\small{\@setsize\small{9pt}\viiipt\@viiipt
-\abovedisplayskip 7.6pt plus 3pt minus 4pt%
-\belowdisplayskip \abovedisplayskip
-\abovedisplayshortskip \z@ plus2pt%
-\belowdisplayshortskip 3.6pt plus2pt minus 2pt
-\def\@listi{\leftmargin\leftmargini %% Added 22 Dec 87
-\topsep 4pt plus 2pt minus 2pt\parsep 2pt plus 1pt minus 1pt
-\itemsep \parsep}}
-
-\def\footnotesize{\@setsize\footnotesize{9pt}\ixpt\@ixpt
-\abovedisplayskip 6.4pt plus 2pt minus 4pt%
-\belowdisplayskip \abovedisplayskip
-\abovedisplayshortskip \z@ plus 1pt%
-\belowdisplayshortskip 2.7pt plus 1pt minus 2pt
-\def\@listi{\leftmargin\leftmargini %% Added 22 Dec 87
-\topsep 3pt plus 1pt minus 1pt\parsep 2pt plus 1pt minus 1pt
-\itemsep \parsep}}
-
-\newcount\aucount
-\newcount\originalaucount
-\newdimen\auwidth
-\auwidth=\textwidth
-\newdimen\auskip
-\newcount\auskipcount
-\newdimen\auskip
-\global\auskip=1pc
-\newdimen\allauboxes
-\allauboxes=\auwidth
-\newtoks\addauthors
-\newcount\addauflag
-\global\addauflag=0 %Haven't shown additional authors yet
-
-\newtoks\subtitletext
-\gdef\subtitle#1{\subtitletext={#1}}
-
-\gdef\additionalauthors#1{\addauthors={#1}}
-
-\gdef\numberofauthors#1{\global\aucount=#1
-\ifnum\aucount>3\global\originalaucount=\aucount \global\aucount=3\fi %g} % 3 OK - Gerry March 2007
-\global\auskipcount=\aucount\global\advance\auskipcount by 1
-\global\multiply\auskipcount by 2
-\global\multiply\auskip by \auskipcount
-\global\advance\auwidth by -\auskip
-\global\divide\auwidth by \aucount}
-
-% \and was modified to count the number of authors. GKMT 12 Aug 1999
-\def\alignauthor{% % \begin{tabular}
-\end{tabular}%
- \begin{tabular}[t]{p{\auwidth}}\centering}%
-
-% *** NOTE *** NOTE *** NOTE *** NOTE ***
-% If you have 'font problems' then you may need
-% to change these, e.g. 'arialb' instead of "arialbd".
-% Gerry Murray 11/11/1999
-% *** OR ** comment out block A and activate block B or vice versa.
-% **********************************************
-%
-% -- Start of block A -- (Type 1 or Truetype fonts)
-%\newfont{\secfnt}{timesbd at 12pt} % was timenrb originally - now is timesbd
-%\newfont{\secit}{timesbi at 12pt} %13 Jan 00 gkmt
-%\newfont{\subsecfnt}{timesi at 11pt} % was timenrri originally - now is timesi
-%\newfont{\subsecit}{timesbi at 11pt} % 13 Jan 00 gkmt -- was times changed to timesbi gm 2/4/2000
-% % because "normal" is italic, "italic" is Roman
-%\newfont{\ttlfnt}{arialbd at 18pt} % was arialb originally - now is arialbd
-%\newfont{\ttlit}{arialbi at 18pt} % 13 Jan 00 gkmt
-%\newfont{\subttlfnt}{arial at 14pt} % was arialr originally - now is arial
-%\newfont{\subttlit}{ariali at 14pt} % 13 Jan 00 gkmt
-%\newfont{\subttlbf}{arialbd at 14pt} % 13 Jan 00 gkmt
-%\newfont{\aufnt}{arial at 12pt} % was arialr originally - now is arial
-%\newfont{\auit}{ariali at 12pt} % 13 Jan 00 gkmt
-%\newfont{\affaddr}{arial at 10pt} % was arialr originally - now is arial
-%\newfont{\affaddrit}{ariali at 10pt} %13 Jan 00 gkmt
-%\newfont{\eaddfnt}{arial at 12pt} % was arialr originally - now is arial
-%\newfont{\ixpt}{times at 9pt} % was timenrr originally - now is times
-%\newfont{\confname}{timesi at 8pt} % was timenrri - now is timesi
-%\newfont{\crnotice}{times at 8pt} % was timenrr originally - now is times
-%\newfont{\ninept}{times at 9pt} % was timenrr originally - now is times
-
-% *********************************************
-% -- End of block A --
-%
-%
-% -- Start of block B -- UPDATED FONT NAMES
-% *********************************************
-% Gerry Murray 11/30/2006
-% *********************************************
-\newfont{\secfnt}{ptmb8t at 12pt}
-\newfont{\secit}{ptmbi8t at 12pt} %13 Jan 00 gkmt
-\newfont{\subsecfnt}{ptmri8t at 11pt}
-\newfont{\subsecit}{ptmbi8t at 11pt} %
-\newfont{\ttlfnt}{phvb8t at 18pt}
-\newfont{\ttlit}{phvbo8t at 18pt} % GM 2/4/2000
-\newfont{\subttlfnt}{phvr8t at 14pt}
-\newfont{\subttlit}{phvro8t at 14pt} % GM 2/4/2000
-\newfont{\subttlbf}{phvb8t at 14pt} % 13 Jan 00 gkmt
-\newfont{\aufnt}{phvr8t at 12pt}
-\newfont{\auit}{phvro8t at 12pt} % GM 2/4/2000
-\newfont{\affaddr}{phvr8t at 10pt}
-\newfont{\affaddrit}{phvro8t at 10pt} % GM 2/4/2000
-\newfont{\eaddfnt}{phvr8t at 12pt}
-\newfont{\ixpt}{ptmr8t at 9pt}
-\newfont{\confname}{ptmri8t at 8pt}
-\newfont{\crnotice}{ptmr8t at 8pt}
-\newfont{\ninept}{ptmr8t at 9pt}
-% +++++++++++++++++++++++++++++++++++++++++++++
-% -- End of block B --
-
-%\def\email#1{{{\eaddfnt{\vskip 4pt#1}}}}
-% If we have an email, inside a "shared affiliation" then we need the following instead
-\def\email#1{{{\eaddfnt{\par #1}}}} % revised - GM - 11/30/2006
-
-\def\addauthorsection{\ifnum\originalaucount>6 % was 3 - Gerry March 2007
- \section{Additional Authors}\the\addauthors
- \fi}
-
-\newcount\savesection
-\newcount\sectioncntr
-\global\sectioncntr=1
-
-\setcounter{secnumdepth}{3}
-
-\def\appendix{\par
-\section*{APPENDIX}
-\setcounter{section}{0}
- \setcounter{subsection}{0}
- \def\thesection{\Alph{section}} }
-
-\leftmargini 22.5pt
-\leftmarginii 19.8pt % > \labelsep + width of '(m)'
-\leftmarginiii 16.8pt % > \labelsep + width of 'vii.'
-\leftmarginiv 15.3pt % > \labelsep + width of 'M.'
-\leftmarginv 9pt
-\leftmarginvi 9pt
-
-\leftmargin\leftmargini
-\labelsep 4.5pt
-\labelwidth\leftmargini\advance\labelwidth-\labelsep
-
-\def\@listI{\leftmargin\leftmargini \parsep 3.6pt plus 2pt minus 1pt%
-\topsep 7.2pt plus 2pt minus 4pt%
-\itemsep 3.6pt plus 2pt minus 1pt}
-
-\let\@listi\@listI
-\@listi
-
-\def\@listii{\leftmargin\leftmarginii
- \labelwidth\leftmarginii\advance\labelwidth-\labelsep
- \topsep 3.6pt plus 2pt minus 1pt
- \parsep 1.8pt plus 0.9pt minus 0.9pt
- \itemsep \parsep}
-
-\def\@listiii{\leftmargin\leftmarginiii
- \labelwidth\leftmarginiii\advance\labelwidth-\labelsep
- \topsep 1.8pt plus 0.9pt minus 0.9pt
- \parsep \z@ \partopsep 1pt plus 0pt minus 1pt
- \itemsep \topsep}
-
-\def\@listiv{\leftmargin\leftmarginiv
- \labelwidth\leftmarginiv\advance\labelwidth-\labelsep}
-
-\def\@listv{\leftmargin\leftmarginv
- \labelwidth\leftmarginv\advance\labelwidth-\labelsep}
-
-\def\@listvi{\leftmargin\leftmarginvi
- \labelwidth\leftmarginvi\advance\labelwidth-\labelsep}
-
-\def\labelenumi{\theenumi.}
-\def\theenumi{\arabic{enumi}}
-
-\def\labelenumii{(\theenumii)}
-\def\theenumii{\alph{enumii}}
-\def\p@enumii{\theenumi}
-
-\def\labelenumiii{\theenumiii.}
-\def\theenumiii{\roman{enumiii}}
-\def\p@enumiii{\theenumi(\theenumii)}
-
-\def\labelenumiv{\theenumiv.}
-\def\theenumiv{\Alph{enumiv}}
-\def\p@enumiv{\p@enumiii\theenumiii}
-
-\def\labelitemi{$\bullet$}
-\def\labelitemii{\bf --}
-\def\labelitemiii{$\ast$}
-\def\labelitemiv{$\cdot$}
-
-\def\verse{\let\\=\@centercr
- \list{}{\itemsep\z@ \itemindent -1.5em\listparindent \itemindent
- \rightmargin\leftmargin\advance\leftmargin 1.5em}\item[]}
-\let\endverse\endlist
-
-\def\quotation{\list{}{\listparindent 1.5em
- \itemindent\listparindent
- \rightmargin\leftmargin \parsep 0pt plus 1pt}\item[]}
-\let\endquotation=\endlist
-
-\def\quote{\list{}{\rightmargin\leftmargin}\item[]}
-\let\endquote=\endlist
-
-\def\descriptionlabel#1{\hspace\labelsep \bf #1}
-\def\description{\list{}{\labelwidth\z@ \itemindent-\leftmargin
- \let\makelabel\descriptionlabel}}
-
-\let\enddescription\endlist
-
-\def\theequation{\arabic{equation}}
-
-\arraycolsep 4.5pt % Half the space between columns in an array environment.
-\tabcolsep 5.4pt % Half the space between columns in a tabular environment.
-\arrayrulewidth .5pt % Width of rules in array and tabular environment. % (was .4) updated Gerry March 20 2007
-\doublerulesep 1.8pt % Space between adjacent rules in array or tabular env.
-
-\tabbingsep \labelsep % Space used by the \' command. (See LaTeX manual.)
-
-\skip\@mpfootins =\skip\footins
-
-\fboxsep =2.7pt % Space left between box and text by \fbox and \framebox.
-\fboxrule =.5pt % Width of rules in box made by \fbox and \framebox. % (was .4) updated Gerry March 20 2007
-
-\def\thepart{\Roman{part}} % Roman numeral part numbers.
-\def\thesection {\arabic{section}}
-\def\thesubsection {\thesection.\arabic{subsection}}
-%\def\thesubsubsection {\thesubsection.\arabic{subsubsection}} % GM 7/30/2002
-%\def\theparagraph {\thesubsubsection.\arabic{paragraph}} % GM 7/30/2002
-\def\thesubparagraph {\theparagraph.\arabic{subparagraph}}
-
-\def\@pnumwidth{1.55em}
-\def\@tocrmarg {2.55em}
-\def\@dotsep{4.5}
-\setcounter{tocdepth}{3}
-
-%\def\tableofcontents{\@latexerr{\tableofcontents: Tables of contents are not
-% allowed in the `acmconf' document style.}\@eha}
-
-\def\tableofcontents{\ClassError{%
- \string\tableofcontents\space is not allowed in the `acmconf' document % January 2008
- style}\@eha}
-
-\def\l@part#1#2{\addpenalty{\@secpenalty}
- \addvspace{2.25em plus 1pt} % space above part line
- \begingroup
- \@tempdima 3em % width of box holding part number, used by
- \parindent \z@ \rightskip \@pnumwidth %% \numberline
- \parfillskip -\@pnumwidth
- {\large \bf % set line in \large boldface
- \leavevmode % TeX command to enter horizontal mode.
- #1\hfil \hbox to\@pnumwidth{\hss #2}}\par
- \nobreak % Never break after part entry
- \endgroup}
-
-\def\l@section#1#2{\addpenalty{\@secpenalty} % good place for page break
- \addvspace{1.0em plus 1pt} % space above toc entry
- \@tempdima 1.5em % width of box holding section number
- \begingroup
- \parindent \z@ \rightskip \@pnumwidth
- \parfillskip -\@pnumwidth
- \bf % Boldface.
- \leavevmode % TeX command to enter horizontal mode.
- \advance\leftskip\@tempdima %% added 5 Feb 88 to conform to
- \hskip -\leftskip %% 25 Jan 88 change to \numberline
- #1\nobreak\hfil \nobreak\hbox to\@pnumwidth{\hss #2}\par
- \endgroup}
-
-
-\def\l@subsection{\@dottedtocline{2}{1.5em}{2.3em}}
-\def\l@subsubsection{\@dottedtocline{3}{3.8em}{3.2em}}
-\def\l@paragraph{\@dottedtocline{4}{7.0em}{4.1em}}
-\def\l@subparagraph{\@dottedtocline{5}{10em}{5em}}
-
-%\def\listoffigures{\@latexerr{\listoffigures: Lists of figures are not
-% allowed in the `acmconf' document style.}\@eha}
-
-\def\listoffigures{\ClassError{%
- \string\listoffigures\space is not allowed in the `acmconf' document % January 2008
- style}\@eha}
-
-\def\l@figure{\@dottedtocline{1}{1.5em}{2.3em}}
-
-%\def\listoftables{\@latexerr{\listoftables: Lists of tables are not
-% allowed in the `acmconf' document style.}\@eha}
-%\let\l@table\l@figure
-
-\def\listoftables{\ClassError{%
- \string\listoftables\space is not allowed in the `acmconf' document % January 2008
- style}\@eha}
- \let\l@table\l@figure
-
-\def\footnoterule{\kern-3\p@
- \hrule width .5\columnwidth % (was .4) updated Gerry March 20 2007
- \kern 2.6\p@} % The \hrule has default height of .4pt % (was .4) updated Gerry March 20 2007
-% ------
-\long\def\@makefntext#1{\noindent
-%\hbox to .5em{\hss$^{\@thefnmark}$}#1} % original
-\hbox to .5em{\hss\textsuperscript{\@thefnmark}}#1} % C. Clifton / GM Oct. 2nd. 2002
-% -------
-
-\long\def\@maketntext#1{\noindent
-#1}
-
-\long\def\@maketitlenotetext#1#2{\noindent
- \hbox to 1.8em{\hss$^{#1}$}#2}
-
-\setcounter{topnumber}{2}
-\def\topfraction{.7}
-\setcounter{bottomnumber}{1}
-\def\bottomfraction{.3}
-\setcounter{totalnumber}{3}
-\def\textfraction{.2}
-\def\floatpagefraction{.5}
-\setcounter{dbltopnumber}{2}
-\def\dbltopfraction{.7}
-\def\dblfloatpagefraction{.5}
-
-%
-\long\def\@makecaption#1#2{
- \vskip \baselineskip
- \setbox\@tempboxa\hbox{\textbf{#1: #2}}
- \ifdim \wd\@tempboxa >\hsize % IF longer than one line:
- \textbf{#1: #2}\par % THEN set as ordinary paragraph.
- \else % ELSE center.
- \hbox to\hsize{\hfil\box\@tempboxa\hfil}\par
- \fi}
-
-%
-
-\long\def\@makecaption#1#2{
- \vskip 10pt
- \setbox\@tempboxa\hbox{\textbf{#1: #2}}
- \ifdim \wd\@tempboxa >\hsize % IF longer than one line:
- \textbf{#1: #2}\par % THEN set as ordinary paragraph.
- \else % ELSE center.
- \hbox to\hsize{\hfil\box\@tempboxa\hfil}
- \fi}
-
-\@ifundefined{figure}{\newcounter {figure}} % this is for LaTeX2e
-
-\def\fps@figure{tbp}
-\def\ftype@figure{1}
-\def\ext@figure{lof}
-\def\fnum@figure{Figure \thefigure}
-\def\figure{\@float{figure}}
-%\let\endfigure\end@float
-\def\endfigure{\end@float} % Gerry January 2008
-\@namedef{figure*}{\@dblfloat{figure}}
-\@namedef{endfigure*}{\end@dblfloat}
-
-\@ifundefined{table}{\newcounter {table}} % this is for LaTeX2e
-
-\def\fps@table{tbp}
-\def\ftype@table{2}
-\def\ext@table{lot}
-\def\fnum@table{Table \thetable}
-\def\table{\@float{table}}
-%\let\endtable\end@float
-\def\endtable{\end@float} % Gerry January 2008
-\@namedef{table*}{\@dblfloat{table}}
-\@namedef{endtable*}{\end@dblfloat}
-
-\newtoks\titleboxnotes
-\newcount\titleboxnoteflag
-
-\def\maketitle{\par
- \begingroup
- \def\thefootnote{\fnsymbol{footnote}}
- \def\@makefnmark{\hbox
- to 0pt{$^{\@thefnmark}$\hss}}
- \twocolumn[\@maketitle]
-\@thanks
- \endgroup
- \setcounter{footnote}{0}
- \let\maketitle\relax
- \let\@maketitle\relax
- \gdef\@thanks{}\gdef\@author{}\gdef\@title{}\gdef\@subtitle{}\let\thanks\relax
- \@copyrightspace}
-
-%% CHANGES ON NEXT LINES
-\newif\if@ll % to record which version of LaTeX is in use
-
-\expandafter\ifx\csname LaTeXe\endcsname\relax % LaTeX2.09 is used
-\else% LaTeX2e is used, so set ll to true
-\global\@lltrue
-\fi
-
-\if@ll
- \NeedsTeXFormat{LaTeX2e}
- \ProvidesClass{sig-alternate} [2009/04/22 - V2.4 - based on acmproc.cls V1.3 <Nov. 30 '99>]
- \RequirePackage{latexsym}% QUERY: are these two really needed?
- \let\dooptions\ProcessOptions
-\else
- \let\dooptions\@options
-\fi
-%% END CHANGES
-
-\def\@height{height}
-\def\@width{width}
-\def\@minus{minus}
-\def\@plus{plus}
-\def\hb@xt@{\hbox to}
-\newif\if@faircopy
-\@faircopyfalse
-\def\ds@faircopy{\@faircopytrue}
-
-\def\ds@preprint{\@faircopyfalse}
-
-\@twosidetrue
-\@mparswitchtrue
-\def\ds@draft{\overfullrule 5\p@}
-%% CHANGE ON NEXT LINE
-\dooptions
-
-\lineskip \p@
-\normallineskip \p@
-\def\baselinestretch{1}
-\def\@ptsize{0} %needed for amssymbols.sty
-
-%% CHANGES ON NEXT LINES
-\if@ll% allow use of old-style font change commands in LaTeX2e
-\@maxdepth\maxdepth
-%
-\DeclareOldFontCommand{\rm}{\ninept\rmfamily}{\mathrm}
-\DeclareOldFontCommand{\sf}{\normalfont\sffamily}{\mathsf}
-\DeclareOldFontCommand{\tt}{\normalfont\ttfamily}{\mathtt}
-\DeclareOldFontCommand{\bf}{\normalfont\bfseries}{\mathbf}
-\DeclareOldFontCommand{\it}{\normalfont\itshape}{\mathit}
-\DeclareOldFontCommand{\sl}{\normalfont\slshape}{\@nomath\sl}
-\DeclareOldFontCommand{\sc}{\normalfont\scshape}{\@nomath\sc}
-\DeclareRobustCommand*{\cal}{\@fontswitch{\relax}{\mathcal}}
-\DeclareRobustCommand*{\mit}{\@fontswitch{\relax}{\mathnormal}}
-\fi
-%
-\if@ll
- \renewcommand{\rmdefault}{cmr} % was 'ttm'
-% Note! I have also found 'mvr' to work ESPECIALLY well.
-% Gerry - October 1999
-% You may need to change your LV1times.fd file so that sc is
-% mapped to cmcsc - -for smallcaps -- that is if you decide
-% to change {cmr} to {times} above. (Not recommended)
- \renewcommand{\@ptsize}{}
- \renewcommand{\normalsize}{%
- \@setfontsize\normalsize\@ixpt{10.5\p@}%\ninept%
- \abovedisplayskip 6\p@ \@plus2\p@ \@minus\p@
- \belowdisplayskip \abovedisplayskip
- \abovedisplayshortskip 6\p@ \@minus 3\p@
- \belowdisplayshortskip 6\p@ \@minus 3\p@
- \let\@listi\@listI
- }
-\else
- \def\@normalsize{%changed next to 9 from 10
- \@setsize\normalsize{9\p@}\ixpt\@ixpt
- \abovedisplayskip 6\p@ \@plus2\p@ \@minus\p@
- \belowdisplayskip \abovedisplayskip
- \abovedisplayshortskip 6\p@ \@minus 3\p@
- \belowdisplayshortskip 6\p@ \@minus 3\p@
- \let\@listi\@listI
- }%
-\fi
-\if@ll
- \newcommand\scriptsize{\@setfontsize\scriptsize\@viipt{8\p@}}
- \newcommand\tiny{\@setfontsize\tiny\@vpt{6\p@}}
- \newcommand\large{\@setfontsize\large\@xiipt{14\p@}}
- \newcommand\Large{\@setfontsize\Large\@xivpt{18\p@}}
- \newcommand\LARGE{\@setfontsize\LARGE\@xviipt{20\p@}}
- \newcommand\huge{\@setfontsize\huge\@xxpt{25\p@}}
- \newcommand\Huge{\@setfontsize\Huge\@xxvpt{30\p@}}
-\else
- \def\scriptsize{\@setsize\scriptsize{8\p@}\viipt\@viipt}
- \def\tiny{\@setsize\tiny{6\p@}\vpt\@vpt}
- \def\large{\@setsize\large{14\p@}\xiipt\@xiipt}
- \def\Large{\@setsize\Large{18\p@}\xivpt\@xivpt}
- \def\LARGE{\@setsize\LARGE{20\p@}\xviipt\@xviipt}
- \def\huge{\@setsize\huge{25\p@}\xxpt\@xxpt}
- \def\Huge{\@setsize\Huge{30\p@}\xxvpt\@xxvpt}
-\fi
-\normalsize
-
-% make aubox hsize/number of authors up to 3, less gutter
-% then showbox gutter showbox gutter showbox -- GKMT Aug 99
-\newbox\@acmtitlebox
-\def\@maketitle{\newpage
- \null
- \setbox\@acmtitlebox\vbox{%
-\baselineskip 20pt
-\vskip 2em % Vertical space above title.
- \begin{center}
- {\ttlfnt \@title\par} % Title set in 18pt Helvetica (Arial) bold size.
- \vskip 1.5em % Vertical space after title.
-%This should be the subtitle.
-{\subttlfnt \the\subtitletext\par}\vskip 1.25em%\fi
- {\baselineskip 16pt\aufnt % each author set in \12 pt Arial, in a
- \lineskip .5em % tabular environment
- \begin{tabular}[t]{c}\@author
- \end{tabular}\par}
- \vskip 1.5em % Vertical space after author.
- \end{center}}
- \dimen0=\ht\@acmtitlebox
- \advance\dimen0 by -12.75pc\relax % Increased space for title box -- KBT
- \unvbox\@acmtitlebox
- \ifdim\dimen0<0.0pt\relax\vskip-\dimen0\fi}
-
-
-\newcount\titlenotecount
-\global\titlenotecount=0
-\newtoks\tntoks
-\newtoks\tntokstwo
-\newtoks\tntoksthree
-\newtoks\tntoksfour
-\newtoks\tntoksfive
-
-\def\abstract{
-\ifnum\titlenotecount>0 % was =1
- \insert\footins{%
- \reset@font\footnotesize
- \interlinepenalty\interfootnotelinepenalty
- \splittopskip\footnotesep
- \splitmaxdepth \dp\strutbox \floatingpenalty \@MM
- \hsize\columnwidth \@parboxrestore
- \protected@edef\@currentlabel{%
- }%
- \color@begingroup
-\ifnum\titlenotecount=1
- \@maketntext{%
- \raisebox{4pt}{$\ast$}\rule\z@\footnotesep\ignorespaces\the\tntoks\@finalstrut\strutbox}%
-\fi
-\ifnum\titlenotecount=2
- \@maketntext{%
- \raisebox{4pt}{$\ast$}\rule\z@\footnotesep\ignorespaces\the\tntoks\par\@finalstrut\strutbox}%
-\@maketntext{%
- \raisebox{4pt}{$\dagger$}\rule\z@\footnotesep\ignorespaces\the\tntokstwo\@finalstrut\strutbox}%
-\fi
-\ifnum\titlenotecount=3
- \@maketntext{%
- \raisebox{4pt}{$\ast$}\rule\z@\footnotesep\ignorespaces\the\tntoks\par\@finalstrut\strutbox}%
-\@maketntext{%
- \raisebox{4pt}{$\dagger$}\rule\z@\footnotesep\ignorespaces\the\tntokstwo\par\@finalstrut\strutbox}%
-\@maketntext{%
- \raisebox{4pt}{$\ddagger$}\rule\z@\footnotesep\ignorespaces\the\tntoksthree\@finalstrut\strutbox}%
-\fi
-\ifnum\titlenotecount=4
- \@maketntext{%
- \raisebox{4pt}{$\ast$}\rule\z@\footnotesep\ignorespaces\the\tntoks\par\@finalstrut\strutbox}%
-\@maketntext{%
- \raisebox{4pt}{$\dagger$}\rule\z@\footnotesep\ignorespaces\the\tntokstwo\par\@finalstrut\strutbox}%
-\@maketntext{%
- \raisebox{4pt}{$\ddagger$}\rule\z@\footnotesep\ignorespaces\the\tntoksthree\par\@finalstrut\strutbox}%
-\@maketntext{%
- \raisebox{4pt}{$\S$}\rule\z@\footnotesep\ignorespaces\the\tntoksfour\@finalstrut\strutbox}%
-\fi
-\ifnum\titlenotecount=5
- \@maketntext{%
- \raisebox{4pt}{$\ast$}\rule\z@\footnotesep\ignorespaces\the\tntoks\par\@finalstrut\strutbox}%
-\@maketntext{%
- \raisebox{4pt}{$\dagger$}\rule\z@\footnotesep\ignorespaces\the\tntokstwo\par\@finalstrut\strutbox}%
-\@maketntext{%
- \raisebox{4pt}{$\ddagger$}\rule\z@\footnotesep\ignorespaces\the\tntoksthree\par\@finalstrut\strutbox}%
-\@maketntext{%
- \raisebox{4pt}{$\S$}\rule\z@\footnotesep\ignorespaces\the\tntoksfour\par\@finalstrut\strutbox}%
-\@maketntext{%
- \raisebox{4pt}{$\P$}\rule\z@\footnotesep\ignorespaces\the\tntoksfive\@finalstrut\strutbox}%
-\fi
- \color@endgroup} %g}
-\fi
-\setcounter{footnote}{0}
-\section*{ABSTRACT}\normalsize%\ninept
-}
-
-\def\endabstract{\if@twocolumn\else\endquotation\fi}
-
-\def\keywords{\if@twocolumn
-\section*{Keywords}
-\else \small
-\quotation
-\fi}
-
-\def\terms{\if@twocolumn
-\section*{General Terms}
-\else \small
-\quotation
-\fi}
-
-% -- Classification needs to be a bit smart due to optionals - Gerry/Georgia November 2nd. 1999
-\newcount\catcount
-\global\catcount=1
-
-\def\category#1#2#3{%
-\ifnum\catcount=1
-\section*{Categories and Subject Descriptors}
-\advance\catcount by 1\else{\unskip; }\fi
- \@ifnextchar [{\@category{#1}{#2}{#3}}{\@category{#1}{#2}{#3}[]}%
-}
-
-\def\@category#1#2#3[#4]{%
- \begingroup
- \let\and\relax
- #1 [\textbf{#2}]%
- \if!#4!%
- \if!#3!\else : #3\fi
- \else
- :\space
- \if!#3!\else #3\kern\z@---\hskip\z@\fi
- \textit{#4}%
- \fi
- \endgroup
-}
-%
-
-%%% This section (written by KBT) handles the 1" box in the lower left
-%%% corner of the left column of the first page by creating a picture,
-%%% and inserting the predefined string at the bottom (with a negative
-%%% displacement to offset the space allocated for a non-existent
-%%% caption).
-%%%
-\newtoks\copyrightnotice
-\def\ftype@copyrightbox{8}
-\def\@copyrightspace{
-\@float{copyrightbox}[b]
-\begin{center}
-\setlength{\unitlength}{1pc}
-\begin{picture}(20,6) %Space for copyright notice
-\put(0,-0.95){\crnotice{\@toappear}}
-\end{picture}
-\end{center}
-\end@float}
-
-\def\@toappear{} % Default setting blank - commands below change this.
-\long\def\toappear#1{\def\@toappear{\parbox[b]{20pc}{\baselineskip 9pt#1}}}
-\def\toappearbox#1{\def\@toappear{\raisebox{5pt}{\framebox[20pc]{\parbox[b]{19pc}{#1}}}}}
-
-\newtoks\conf
-\newtoks\confinfo
-\def\conferenceinfo#1#2{\global\conf={#1}\global\confinfo{#2}}
-
-
-%\def\marginpar{\@latexerr{The \marginpar command is not allowed in the
-% `acmconf' document style.}\@eha}
-
-\def\marginpar{\ClassError{%
- \string\marginpar\space is not allowed in the `acmconf' document % January 2008
- style}\@eha}
-
-\mark{{}{}} % Initializes TeX's marks
-
-\def\today{\ifcase\month\or
- January\or February\or March\or April\or May\or June\or
- July\or August\or September\or October\or November\or December\fi
- \space\number\day, \number\year}
-
-\def\@begintheorem#1#2{%
- \parskip 0pt % GM July 2000 (for tighter spacing)
- \trivlist
- \item[%
- \hskip 10\p@
- \hskip \labelsep
- {{\sc #1}\hskip 5\p@\relax#2.}%
- ]
- \it
-}
-\def\@opargbegintheorem#1#2#3{%
- \parskip 0pt % GM July 2000 (for tighter spacing)
- \trivlist
- \item[%
- \hskip 10\p@
- \hskip \labelsep
- {\sc #1\ #2\ % This mod by Gerry to enumerate corollaries
- \setbox\@tempboxa\hbox{(#3)} % and bracket the 'corollary title'
- \ifdim \wd\@tempboxa>\z@ % and retain the correct numbering of e.g. theorems
- \hskip 5\p@\relax % if they occur 'around' said corollaries.
- \box\@tempboxa % Gerry - Nov. 1999.
- \fi.}%
- ]
- \it
-}
-\newif\if@qeded
-\global\@qededfalse
-
-% -- original
-%\def\proof{%
-% \vspace{-\parskip} % GM July 2000 (for tighter spacing)
-% \global\@qededfalse
-% \@ifnextchar[{\@xproof}{\@proof}%
-%}
-% -- end of original
-
-% (JSS) Fix for vertical spacing bug - Gerry Murray July 30th. 2002
-\def\proof{%
-\vspace{-\lastskip}\vspace{-\parsep}\penalty-51%
-\global\@qededfalse
-\@ifnextchar[{\@xproof}{\@proof}%
-}
-
-\def\endproof{%
- \if@qeded\else\qed\fi
- \endtrivlist
-}
-\def\@proof{%
- \trivlist
- \item[%
- \hskip 10\p@
- \hskip \labelsep
- {\sc Proof.}%
- ]
- \ignorespaces
-}
-\def\@xproof[#1]{%
- \trivlist
- \item[\hskip 10\p@\hskip \labelsep{\sc Proof #1.}]%
- \ignorespaces
-}
-\def\qed{%
- \unskip
- \kern 10\p@
- \begingroup
- \unitlength\p@
- \linethickness{.4\p@}%
- \framebox(6,6){}%
- \endgroup
- \global\@qededtrue
-}
-
-\def\newdef#1#2{%
- \expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}%
- \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
- \global\@namedef{#1}{\@defthm{#1}{#2}}%
- \global\@namedef{end#1}{\@endtheorem}%
- }%
-}
-\def\@defthm#1#2{%
- \refstepcounter{#1}%
- \@ifnextchar[{\@ydefthm{#1}{#2}}{\@xdefthm{#1}{#2}}%
-}
-\def\@xdefthm#1#2{%
- \@begindef{#2}{\csname the#1\endcsname}%
- \ignorespaces
-}
-\def\@ydefthm#1#2[#3]{%
- \trivlist
- \item[%
- \hskip 10\p@
- \hskip \labelsep
- {\it #2%
-% \savebox\@tempboxa{#3}%
- \saveb@x\@tempboxa{#3}% % January 2008
- \ifdim \wd\@tempboxa>\z@
- \ \box\@tempboxa
- \fi.%
- }]%
- \ignorespaces
-}
-\def\@begindef#1#2{%
- \trivlist
- \item[%
- \hskip 10\p@
- \hskip \labelsep
- {\it #1\ \rm #2.}%
- ]%
-}
-\def\theequation{\arabic{equation}}
-
-\newcounter{part}
-\newcounter{section}
-\newcounter{subsection}[section]
-\newcounter{subsubsection}[subsection]
-\newcounter{paragraph}[subsubsection]
-\def\thepart{\Roman{part}}
-\def\thesection{\arabic{section}}
-\def\thesubsection{\thesection.\arabic{subsection}}
-\def\thesubsubsection{\thesubsection.\arabic{subsubsection}} %removed \subsecfnt 29 July 2002 gkmt
-\def\theparagraph{\thesubsubsection.\arabic{paragraph}} %removed \subsecfnt 29 July 2002 gkmt
-\newif\if@uchead
-\@ucheadfalse
-
-%% CHANGES: NEW NOTE
-%% NOTE: OK to use old-style font commands below, since they were
-%% suitably redefined for LaTeX2e
-%% END CHANGES
-\setcounter{secnumdepth}{3}
-\def\part{%
- \@startsection{part}{9}{\z@}{-10\p@ \@plus -4\p@ \@minus -2\p@}
- {4\p@}{\normalsize\@ucheadtrue}%
-}
-\def\section{%
- \@startsection{section}{1}{\z@}{-10\p@ \@plus -4\p@ \@minus -2\p@}% GM
- {4\p@}{\baselineskip 14pt\secfnt\@ucheadtrue}%
-}
-
-\def\subsection{%
- \@startsection{subsection}{2}{\z@}{-8\p@ \@plus -2\p@ \@minus -\p@}
- {4\p@}{\secfnt}%
-}
-\def\subsubsection{%
- \@startsection{subsubsection}{3}{\z@}{-8\p@ \@plus -2\p@ \@minus -\p@}%
- {4\p@}{\subsecfnt}%
-}
-%\def\paragraph{%
-% \vskip 12pt\@startsection{paragraph}{3}{\z@}{6\p@ \@plus \p@}% original
-% {-5\p@}{\subsecfnt}%
-%}
-% If one wants sections, subsections and subsubsections numbered,
-% but not paragraphs, one usually sets secnumepth to 3.
-% For that, the "depth" of paragraphs must be given correctly
-% in the definition (``4'' instead of ``3'' as second argument
-% of @startsection):
-\def\paragraph{%
- \vskip 12pt\@startsection{paragraph}{4}{\z@}{6\p@ \@plus \p@}% % GM and Wolfgang May - 11/30/06
- {-5\p@}{\subsecfnt}%
-}
-\let\@period=.
-\def\@startsection#1#2#3#4#5#6{%
- \if@noskipsec %gkmt, 11 aug 99
- \global\let\@period\@empty
- \leavevmode
- \global\let\@period.%
- \fi
- \par %
- \@tempskipa #4\relax
- \@afterindenttrue
- \ifdim \@tempskipa <\z@
- \@tempskipa -\@tempskipa
- \@afterindentfalse
- \fi
- \if@nobreak
- \everypar{}%
- \else
- \addpenalty\@secpenalty
- \addvspace\@tempskipa
- \fi
-\parskip=0pt % GM July 2000 (non numbered) section heads
- \@ifstar
- {\@ssect{#3}{#4}{#5}{#6}}
- {\@dblarg{\@sect{#1}{#2}{#3}{#4}{#5}{#6}}}%
-}
-\def\@sect#1#2#3#4#5#6[#7]#8{%
- \ifnum #2>\c@secnumdepth
- \let\@svsec\@empty
- \else
- \refstepcounter{#1}%
- \edef\@svsec{%
- \begingroup
- %\ifnum#2>2 \noexpand\rm \fi % changed to next 29 July 2002 gkmt
- \ifnum#2>2 \noexpand#6 \fi
- \csname the#1\endcsname
- \endgroup
- \ifnum #2=1\relax .\fi
- \hskip 1em
- }%
- \fi
- \@tempskipa #5\relax
- \ifdim \@tempskipa>\z@
- \begingroup
- #6\relax
- \@hangfrom{\hskip #3\relax\@svsec}%
- \begingroup
- \interlinepenalty \@M
- \if@uchead
- \uppercase{#8}%
- \else
- #8%
- \fi
- \par
- \endgroup
- \endgroup
- \csname #1mark\endcsname{#7}%
- \vskip -12pt %gkmt, 11 aug 99 and GM July 2000 (was -14) - numbered section head spacing
-\addcontentsline{toc}{#1}{%
- \ifnum #2>\c@secnumdepth \else
- \protect\numberline{\csname the#1\endcsname}%
- \fi
- #7%
- }%
- \else
- \def\@svsechd{%
- #6%
- \hskip #3\relax
- \@svsec
- \if@uchead
- \uppercase{#8}%
- \else
- #8%
- \fi
- \csname #1mark\endcsname{#7}%
- \addcontentsline{toc}{#1}{%
- \ifnum #2>\c@secnumdepth \else
- \protect\numberline{\csname the#1\endcsname}%
- \fi
- #7%
- }%
- }%
- \fi
- \@xsect{#5}\hskip 1pt
- \par
-}
-\def\@xsect#1{%
- \@tempskipa #1\relax
- \ifdim \@tempskipa>\z@
- \par
- \nobreak
- \vskip \@tempskipa
- \@afterheading
- \else
- \global\@nobreakfalse
- \global\@noskipsectrue
- \everypar{%
- \if@noskipsec
- \global\@noskipsecfalse
- \clubpenalty\@M
- \hskip -\parindent
- \begingroup
- \@svsechd
- \@period
- \endgroup
- \unskip
- \@tempskipa #1\relax
- \hskip -\@tempskipa
- \else
- \clubpenalty \@clubpenalty
- \everypar{}%
- \fi
- }%
- \fi
- \ignorespaces
-}
-\def\@trivlist{%
- \@topsepadd\topsep
- \if@noskipsec
- \global\let\@period\@empty
- \leavevmode
- \global\let\@period.%
- \fi
- \ifvmode
- \advance\@topsepadd\partopsep
- \else
- \unskip
- \par
- \fi
- \if@inlabel
- \@noparitemtrue
- \@noparlisttrue
- \else
- \@noparlistfalse
- \@topsep\@topsepadd
- \fi
- \advance\@topsep \parskip
- \leftskip\z@skip
- \rightskip\@rightskip
- \parfillskip\@flushglue
- \@setpar{\if@newlist\else{\@@par}\fi}
- \global\@newlisttrue
- \@outerparskip\parskip
-}
-
-%%% Actually, 'abbrev' works just fine as the default
-%%% Bibliography style.
-
-\typeout{Using 'Abbrev' bibliography style}
-\newcommand\bibyear[2]{%
- \unskip\quad\ignorespaces#1\unskip
- \if#2..\quad \else \quad#2 \fi
-}
-\newcommand{\bibemph}[1]{{\em#1}}
-\newcommand{\bibemphic}[1]{{\em#1\/}}
-\newcommand{\bibsc}[1]{{\sc#1}}
-\def\@normalcite{%
- \def\@cite##1##2{[##1\if@tempswa , ##2\fi]}%
-}
-\def\@citeNB{%
- \def\@cite##1##2{##1\if@tempswa , ##2\fi}%
-}
-\def\@citeRB{%
- \def\@cite##1##2{##1\if@tempswa , ##2\fi]}%
-}
-\def\start@cite#1#2{%
- \edef\citeauthoryear##1##2##3{%
- ###1%
- \ifnum#2=\z@ \else\ ###2\fi
- }%
- \ifnum#1=\thr@@
- \let\@@cite\@citeyear
- \else
- \let\@@cite\@citenormal
- \fi
- \@ifstar{\@citeNB\@@cite}{\@normalcite\@@cite}%
-}
-%\def\cite{\start@cite23}
-\DeclareRobustCommand\cite{\start@cite23} % January 2008
-\def\citeNP{\cite*} % No Parentheses e.g. 5
-%\def\citeA{\start@cite10}
-\DeclareRobustCommand\citeA{\start@cite10} % January 2008
-\def\citeANP{\citeA*}
-%\def\shortcite{\start@cite23}
-\DeclareRobustCommand\shortcite{\start@cite23} % January 2008
-\def\shortciteNP{\shortcite*}
-%\def\shortciteA{\start@cite20}
-\DeclareRobustCommand\shortciteA{\start@cite20} % January 2008
-\def\shortciteANP{\shortciteA*}
-%\def\citeyear{\start@cite30}
-\DeclareRobustCommand\citeyear{\start@cite30} % January 2008
-\def\citeyearNP{\citeyear*}
-%\def\citeN{%
-\DeclareRobustCommand\citeN{% % January 2008
- \@citeRB
- \def\citeauthoryear##1##2##3{##1\ [##3%
- \def\reserved@a{##1}%
- \def\citeauthoryear####1####2####3{%
- \def\reserved@b{####1}%
- \ifx\reserved@a\reserved@b
- ####3%
- \else
- \errmessage{Package acmart Error: author mismatch
- in \string\citeN^^J^^J%
- See the acmart package documentation for explanation}%
- \fi
- }%
- }%
- \@ifstar\@citeyear\@citeyear
-}
-%\def\shortciteN{%
-\DeclareRobustCommand\shortciteN{% % January 2008
- \@citeRB
- \def\citeauthoryear##1##2##3{##2\ [##3%
- \def\reserved@a{##2}%
- \def\citeauthoryear####1####2####3{%
- \def\reserved@b{####2}%
- \ifx\reserved@a\reserved@b
- ####3%
- \else
- \errmessage{Package acmart Error: author mismatch
- in \string\shortciteN^^J^^J%
- See the acmart package documentation for explanation}%
- \fi
- }%
- }%
- \@ifstar\@citeyear\@citeyear % GM July 2000
-}
-
-\def\@citenormal{%
- \@ifnextchar [{\@tempswatrue\@citex;}%
-% original {\@tempswafalse\@citex,[]}% was ; Gerry 2/24/00
-{\@tempswafalse\@citex[]}% % GERRY FIX FOR BABEL 3/20/2009
-}
-
-\def\@citeyear{%
- \@ifnextchar [{\@tempswatrue\@citex,}%
-% original {\@tempswafalse\@citex,[]}%
-{\@tempswafalse\@citex[]}% % GERRY FIX FOR BABEL 3/20/2009
-}
-
-\def\@citex#1[#2]#3{%
- \let\@citea\@empty
- \@cite{%
- \@for\@citeb:=#3\do{%
- \@citea
-% original \def\@citea{#1 }%
- \def\@citea{#1, }% % GERRY FIX FOR BABEL 3/20/2009 -- SO THAT YOU GET [1, 2] IN THE BODY TEXT
- \edef\@citeb{\expandafter\@iden\@citeb}%
- \if@filesw
- \immediate\write\@auxout{\string\citation{\@citeb}}%
- \fi
- \@ifundefined{b@\@citeb}{%
- {\bf ?}%
- \@warning{%
- Citation `\@citeb' on page \thepage\space undefined%
- }%
- }%
- {\csname b@\@citeb\endcsname}%
- }%
- }{#2}%
-}
-%\let\@biblabel\@gobble % Dec. 2008 - Gerry
-% ----
-\def\@biblabelnum#1{[#1]} % Gerry's solution #1 - for Natbib -- April 2009
-\let\@biblabel=\@biblabelnum % Gerry's solution #1 - for Natbib -- April 2009
-\def\newblock{\relax} % Gerry Dec. 2008
-% ---
-\newdimen\bibindent
-\setcounter{enumi}{1}
-\bibindent=0em
-\def\thebibliography#1{%
-\ifnum\addauflag=0\addauthorsection\global\addauflag=1\fi
- \section[References]{% <=== OPTIONAL ARGUMENT ADDED HERE
- {References} % was uppercased but this affects pdf bookmarks (SP/GM October 2004)
- {\vskip -9pt plus 1pt} % GM Nov. 2006 / GM July 2000 (for somewhat tighter spacing)
- \@mkboth{{\refname}}{{\refname}}%
- }%
- \list{[\arabic{enumi}]}{%
- \settowidth\labelwidth{[#1]}%
- \leftmargin\labelwidth
- \advance\leftmargin\labelsep
- \advance\leftmargin\bibindent
- \parsep=0pt\itemsep=1pt % GM July 2000
- \itemindent -\bibindent
- \listparindent \itemindent
- \usecounter{enumi}
- }%
- \let\newblock\@empty
- \raggedright % GM July 2000
- \sloppy
- \sfcode`\.=1000\relax
-}
-
-
-\gdef\balancecolumns
-{\vfill\eject
-\global\@colht=\textheight
-\global\ht\@cclv=\textheight
-}
-
-\newcount\colcntr
-\global\colcntr=0
-%\newbox\savebox
-\newbox\saveb@x % January 2008
-
-\gdef \@makecol {%
-\global\advance\colcntr by 1
-\ifnum\colcntr>2 \global\colcntr=1\fi
- \ifvoid\footins
- \setbox\@outputbox \box\@cclv
- \else
- \setbox\@outputbox \vbox{%
-\boxmaxdepth \@maxdepth
- \@tempdima\dp\@cclv
- \unvbox \@cclv
- \vskip-\@tempdima
- \vskip \skip\footins
- \color@begingroup
- \normalcolor
- \footnoterule
- \unvbox \footins
- \color@endgroup
- }%
- \fi
- \xdef\@freelist{\@freelist\@midlist}%
- \global \let \@midlist \@empty
- \@combinefloats
- \ifvbox\@kludgeins
- \@makespecialcolbox
- \else
- \setbox\@outputbox \vbox to\@colht {%
-\@texttop
- \dimen@ \dp\@outputbox
- \unvbox \@outputbox
- \vskip -\dimen@
- \@textbottom
- }%
- \fi
- \global \maxdepth \@maxdepth
-}
-\def\titlenote{\@ifnextchar[\@xtitlenote{\stepcounter\@mpfn
-\global\advance\titlenotecount by 1
-\ifnum\titlenotecount=1
- \raisebox{9pt}{$\ast$}
-\fi
-\ifnum\titlenotecount=2
- \raisebox{9pt}{$\dagger$}
-\fi
-\ifnum\titlenotecount=3
- \raisebox{9pt}{$\ddagger$}
-\fi
-\ifnum\titlenotecount=4
-\raisebox{9pt}{$\S$}
-\fi
-\ifnum\titlenotecount=5
-\raisebox{9pt}{$\P$}
-\fi
- \@titlenotetext
-}}
-
-\long\def\@titlenotetext#1{\insert\footins{%
-\ifnum\titlenotecount=1\global\tntoks={#1}\fi
-\ifnum\titlenotecount=2\global\tntokstwo={#1}\fi
-\ifnum\titlenotecount=3\global\tntoksthree={#1}\fi
-\ifnum\titlenotecount=4\global\tntoksfour={#1}\fi
-\ifnum\titlenotecount=5\global\tntoksfive={#1}\fi
- \reset@font\footnotesize
- \interlinepenalty\interfootnotelinepenalty
- \splittopskip\footnotesep
- \splitmaxdepth \dp\strutbox \floatingpenalty \@MM
- \hsize\columnwidth \@parboxrestore
- \protected@edef\@currentlabel{%
- }%
- \color@begingroup
- \color@endgroup}}
-
-%%%%%%%%%%%%%%%%%%%%%%%%%
-\ps@plain
-\baselineskip=11pt
-\let\thepage\relax % For NO page numbers - GM Nov. 30th. 1999 and July 2000
-\def\setpagenumber#1{\global\setcounter{page}{#1}}
-%\pagenumbering{arabic} % Arabic page numbers GM July 2000
-\twocolumn % Double column.
-\flushbottom % Even bottom -- alas, does not balance columns at end of document
-\pagestyle{plain}
-
-% Need Copyright Year and Copyright Data to be user definable (in .tex file).
-% Gerry Nov. 30th. 1999
-\newtoks\copyrtyr
-\newtoks\acmcopyr
-\newtoks\boilerplate
-\global\acmcopyr={X-XXXXX-XX-X/XX/XX} % Default - 5/11/2001 *** Gerry
-\global\copyrtyr={200X} % Default - 3/3/2003 *** Gerry
-\def\CopyrightYear#1{\global\copyrtyr{#1}}
-\def\crdata#1{\global\acmcopyr{#1}}
-\def\permission#1{\global\boilerplate{#1}}
-%
-\global\boilerplate={Permission to make digital or hard copies of all or part of this work for personal or classroom use is granted without fee provided that copies are not made or distributed for profit or commercial advantage and that copies bear this notice and the full citation on the first page. To copy otherwise, to republish, to post on servers or to redistribute to lists, requires prior specific permission and/or a fee.}
-\newtoks\copyrightetc
-\global\copyrightetc{Copyright \the\copyrtyr\ ACM \the\acmcopyr\ ...\$10.00}
-\toappear{\the\boilerplate\par
-{\confname{\the\conf}} \the\confinfo\par \the\copyrightetc.}
-%\DeclareFixedFont{\altcrnotice}{OT1}{tmr}{m}{n}{8} % << patch needed for accenting e.g. Montreal - Gerry, May 2007
-%\DeclareFixedFont{\altconfname}{OT1}{tmr}{m}{it}{8} % << patch needed for accenting in italicized confname - Gerry, May 2007
-%
-%{\altconfname{{\the\conf}}} {\altcrnotice\the\confinfo\par} \the\copyrightetc.} % << Gerry, May 2007
-%
-% The following section (i.e. 3 .sty inclusions) was added in May 2007 so as to fix the problems that many
-% authors were having with accents. Sometimes accents would occur, but the letter-character would be of a different
-% font. Conversely the letter-character font would be correct but, e.g. a 'bar' would appear superimposed on the
-% character instead of, say, an unlaut/diaresis. Sometimes the letter-character would NOT appear at all.
-% Using [T1]{fontenc} outright was not an option as this caused 99% of the authors to 'produce' a Type-3 (bitmapped)
-% PDF file - useless for production.
-%
-% For proper (font) accenting we NEED these packages to be part of the .cls file i.e. 'ae', 'aecompl' and 'aeguil'
-% ~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
-%% This is file `ae.sty'
-\def\fileversion{1.3}
-\def\filedate{2001/02/12}
-\NeedsTeXFormat{LaTeX2e}
-%\ProvidesPackage{ae}[\filedate\space\fileversion\space % GM
-% Almost European Computer Modern] % GM - keeping the log file clean(er)
-\newif\if@ae@slides \@ae@slidesfalse
-\DeclareOption{slides}{\@ae@slidestrue}
-\ProcessOptions
-\fontfamily{aer}
-\RequirePackage[T1]{fontenc}
-\if@ae@slides
- \renewcommand{\sfdefault}{laess}
- \renewcommand{\rmdefault}{laess} % no roman
- \renewcommand{\ttdefault}{laett}
-\else
- \renewcommand{\sfdefault}{aess}
- \renewcommand{\rmdefault}{aer}
- \renewcommand{\ttdefault}{aett}
-\fi
-\endinput
-%%
-%% End of file `ae.sty'.
-%
-%
-\def\fileversion{0.9}
-\def\filedate{1998/07/23}
-\NeedsTeXFormat{LaTeX2e}
-%\ProvidesPackage{aecompl}[\filedate\space\fileversion\space % GM
-%T1 Complements for AE fonts (D. Roegel)] % GM -- keeping the log file clean(er)
-
-\def\@ae@compl#1{{\fontencoding{T1}\fontfamily{cmr}\selectfont\symbol{#1}}}
-\def\guillemotleft{\@ae@compl{19}}
-\def\guillemotright{\@ae@compl{20}}
-\def\guilsinglleft{\@ae@compl{14}}
-\def\guilsinglright{\@ae@compl{15}}
-\def\TH{\@ae@compl{222}}
-\def\NG{\@ae@compl{141}}
-\def\ng{\@ae@compl{173}}
-\def\th{\@ae@compl{254}}
-\def\DJ{\@ae@compl{208}}
-\def\dj{\@ae@compl{158}}
-\def\DH{\@ae@compl{208}}
-\def\dh{\@ae@compl{240}}
-\def\@perthousandzero{\@ae@compl{24}}
-\def\textperthousand{\%\@perthousandzero}
-\def\textpertenthousand{\%\@perthousandzero\@perthousandzero}
-\endinput
-%
-%
-%% This is file `aeguill.sty'
-% This file gives french guillemets (and not guillemots!)
-% built with the Polish CMR fonts (default), WNCYR fonts, the LASY fonts
-% or with the EC fonts.
-% This is useful in conjunction with the ae package
-% (this package loads the ae package in case it has not been loaded)
-% and with or without the french(le) package.
-%
-% In order to get the guillemets, it is necessary to either type
-% \guillemotleft and \guillemotright, or to use an 8 bit encoding
-% (such as ISO-Latin1) which selects these two commands,
-% or, if you use the french package (but not the frenchle package),
-% to type << or >>.
-%
-% By default, you get the Polish CMR guillemets; if this package is loaded
-% with the `cm' option, you get the LASY guillemets; with `ec,' you
-% get the EC guillemets, and with `cyr,' you get the cyrillic guillemets.
-%
-% In verbatim mode, you always get the EC/TT guillemets.
-%
-% The default option is interesting in conjunction with PDF,
-% because there is a Type 1 version of the Polish CMR fonts
-% and these guillemets are very close in shape to the EC guillemets.
-% There are no free Type 1 versions of the EC fonts.
-%
-% Support for Polish CMR guillemets was kindly provided by
-% Rolf Niepraschk <niepraschk@ptb.de> in version 0.99 (2000/05/22).
-% Bernd Raichle provided extensive simplifications to the code
-% for version 1.00.
-%
-% This package is released under the LPPL.
-%
-% Changes:
-% Date version
-% 2001/04/12 1.01 the frenchle and french package are now distinguished.
-%
-\def\fileversion{1.01}
-\def\filedate{2001/04/12}
-\NeedsTeXFormat{LaTeX2e}
-%\ProvidesPackage{aeguill}[2001/04/12 1.01 % % GM
-%AE fonts with french guillemets (D. Roegel)] % GM - keeping the log file clean(er)
-%\RequirePackage{ae} % GM May 2007 - already embedded here
-
-\newcommand{\@ae@switch}[4]{#4}
-\DeclareOption{ec}{\renewcommand\@ae@switch[4]{#1}}
-\DeclareOption{cm}{\renewcommand\@ae@switch[4]{#2}}
-\DeclareOption{cyr}{\renewcommand\@ae@switch[4]{#3}}
-\DeclareOption{pl}{\renewcommand\@ae@switch[4]{#4}}
-\ExecuteOptions{pl}
-\ProcessOptions
-
-%
-% Load necessary packages
-%
-\@ae@switch{% ec
- % do nothing
-}{% cm
- \RequirePackage{latexsym}% GM - May 2007 - already 'mentioned as required' up above
-}{% cyr
- \RequirePackage[OT2,T1]{fontenc}%
-}{% pl
- \RequirePackage[OT4,T1]{fontenc}%
-}
-
-% The following command will be compared to \frenchname,
-% as defined in french.sty and frenchle.sty.
-\def\aeguillfrenchdefault{french}%
-
-\let\guill@verbatim@font\verbatim@font
-\def\verbatim@font{\guill@verbatim@font\ecguills{cmtt}%
- \let\guillemotleft\@oguills\let\guillemotright\@fguills}
-
-\begingroup \catcode`\<=13 \catcode`\>=13
-\def\x{\endgroup
- \def\ae@lfguill{<<}%
- \def\ae@rfguill{>>}%
-}\x
-
-\newcommand{\ecguills}[1]{%
- \def\selectguillfont{\fontencoding{T1}\fontfamily{#1}\selectfont}%
- \def\@oguills{{\selectguillfont\symbol{19}}}%
- \def\@fguills{{\selectguillfont\symbol{20}}}%
- }
-
-\newcommand{\aeguills}{%
- \ae@guills
- % We redefine \guillemotleft and \guillemotright
- % in order to catch them when they are used
- % with \DeclareInputText (in latin1.def for instance)
- % We use \auxWARNINGi as a safe indicator that french.sty is used.
- \gdef\guillemotleft{\ifx\auxWARNINGi\undefined
- \@oguills % neither french.sty nor frenchle.sty
- \else
- \ifx\aeguillfrenchdefault\frenchname
- \ae@lfguill % french.sty
- \else
- \@oguills % frenchle.sty
- \fi
- \fi}%
- \gdef\guillemotright{\ifx\auxWARNINGi\undefined
- \@fguills % neither french.sty nor frenchle.sty
- \else
- \ifx\aeguillfrenchdefault\frenchname
- \ae@rfguill % french.sty
- \else
- \@fguills % frenchle.sty
- \fi
- \fi}%
- }
-
-%
-% Depending on the class option
-% define the internal command \ae@guills
-\@ae@switch{% ec
- \newcommand{\ae@guills}{%
- \ecguills{cmr}}%
-}{% cm
- \newcommand{\ae@guills}{%
- \def\selectguillfont{\fontencoding{U}\fontfamily{lasy}%
- \fontseries{m}\fontshape{n}\selectfont}%
- \def\@oguills{\leavevmode\nobreak
- \hbox{\selectguillfont (\kern-.20em(\kern.20em}\nobreak}%
- \def\@fguills{\leavevmode\nobreak
- \hbox{\selectguillfont \kern.20em)\kern-.2em)}%
- \ifdim\fontdimen\@ne\font>\z@\/\fi}}%
-}{% cyr
- \newcommand{\ae@guills}{%
- \def\selectguillfont{\fontencoding{OT2}\fontfamily{wncyr}\selectfont}%
- \def\@oguills{{\selectguillfont\symbol{60}}}%
- \def\@fguills{{\selectguillfont\symbol{62}}}}
-}{% pl
- \newcommand{\ae@guills}{%
- \def\selectguillfont{\fontencoding{OT4}\fontfamily{cmr}\selectfont}%
- \def\@oguills{{\selectguillfont\symbol{174}}}%
- \def\@fguills{{\selectguillfont\symbol{175}}}}
-}
-
-
-\AtBeginDocument{%
- \ifx\GOfrench\undefined
- \aeguills
- \else
- \let\aeguill@GOfrench\GOfrench
- \gdef\GOfrench{\aeguill@GOfrench \aeguills}%
- \fi
- }
-
-\endinput
-%
-
--- a/Quotient-Paper/document/svglov3.clo Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,101 +0,0 @@
-% SVJour3 DOCUMENT CLASS OPTION SVGLOV3 -- for standardised journals
-%
-% This is an enhancement for the LaTeX
-% SVJour3 document class for Springer journals
-%
-%%
-%%
-%% \CharacterTable
-%% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
-%% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
-%% Digits \0\1\2\3\4\5\6\7\8\9
-%% Exclamation \! Double quote \" Hash (number) \#
-%% Dollar \$ Percent \% Ampersand \&
-%% Acute accent \' Left paren \( Right paren \)
-%% Asterisk \* Plus \+ Comma \,
-%% Minus \- Point \. Solidus \/
-%% Colon \: Semicolon \; Less than \<
-%% Equals \= Greater than \> Question mark \?
-%% Commercial at \@ Left bracket \[ Backslash \\
-%% Right bracket \] Circumflex \^ Underscore \_
-%% Grave accent \` Left brace \{ Vertical bar \|
-%% Right brace \} Tilde \~}
-\ProvidesFile{svglov3.clo}
- [2006/02/03 v3.1
- style option for standardised journals]
-\typeout{SVJour Class option: svglov3.clo for standardised journals}
-\def\validfor{svjour3}
-\ExecuteOptions{final,10pt,runningheads}
-% No size changing allowed, hence a "copy" of size10.clo is included
-\renewcommand\normalsize{%
-\if@twocolumn
- \@setfontsize\normalsize\@xpt{12.5pt}%
-\else
- \if@smallext
- \@setfontsize\normalsize\@xpt\@xiipt
- \else
- \@setfontsize\normalsize{9.5pt}{11.5pt}%
- \fi
-\fi
- \abovedisplayskip=3 mm plus6pt minus 4pt
- \belowdisplayskip=3 mm plus6pt minus 4pt
- \abovedisplayshortskip=0.0 mm plus6pt
- \belowdisplayshortskip=2 mm plus4pt minus 4pt
- \let\@listi\@listI}
-\normalsize
-\newcommand\small{%
-\if@twocolumn
- \@setfontsize\small{8.5pt}\@xpt
-\else
- \if@smallext
- \@setfontsize\small\@viiipt{9.5pt}%
- \else
- \@setfontsize\small\@viiipt{9.25pt}%
- \fi
-\fi
- \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
- \abovedisplayshortskip \z@ \@plus2\p@
- \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
- \def\@listi{\leftmargin\leftmargini
- \parsep 0\p@ \@plus1\p@ \@minus\p@
- \topsep 4\p@ \@plus2\p@ \@minus4\p@
- \itemsep0\p@}%
- \belowdisplayskip \abovedisplayskip
-}
-\let\footnotesize\small
-\newcommand\scriptsize{\@setfontsize\scriptsize\@viipt\@viiipt}
-\newcommand\tiny{\@setfontsize\tiny\@vpt\@vipt}
-\if@twocolumn
- \newcommand\large{\@setfontsize\large\@xiipt\@xivpt}
- \newcommand\LARGE{\@setfontsize\LARGE{16pt}{18pt}}
-\else
- \newcommand\large{\@setfontsize\large\@xipt\@xiipt}
- \newcommand\LARGE{\@setfontsize\LARGE{13pt}{15pt}}
-\fi
-\newcommand\Large{\@setfontsize\Large\@xivpt{16dd}}
-\newcommand\huge{\@setfontsize\huge\@xxpt{25}}
-\newcommand\Huge{\@setfontsize\Huge\@xxvpt{30}}
-%
-\def\runheadhook{\rlap{\smash{\lower6.5pt\hbox to\textwidth{\hrulefill}}}}
-\if@twocolumn
-\setlength{\textwidth}{17.4cm}
-\setlength{\textheight}{234mm}
-\AtEndOfClass{\setlength\columnsep{6mm}}
-\else
- \if@smallext
- \setlength{\textwidth}{11.9cm}
- \setlength{\textheight}{19.4cm}
- \else
- \setlength{\textwidth}{12.2cm}
- \setlength{\textheight}{19.8cm}
- \fi
-\fi
-%
-\AtBeginDocument{%
-\@ifundefined{@journalname}
- {\typeout{Unknown journal: specify \string\journalname\string{%
-<name of your journal>\string} in preambel^^J}}{}}
-%
-\endinput
-%%
-%% End of file `svglov3.clo'.
--- a/Quotient-Paper/document/svjour3.cls Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1431 +0,0 @@
-% SVJour3 DOCUMENT CLASS -- version 3.2 for LaTeX2e
-%
-% LaTeX document class for Springer journals
-%
-%%
-%%
-%% \CharacterTable
-%% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
-%% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
-%% Digits \0\1\2\3\4\5\6\7\8\9
-%% Exclamation \! Double quote \" Hash (number) \#
-%% Dollar \$ Percent \% Ampersand \&
-%% Acute accent \' Left paren \( Right paren \)
-%% Asterisk \* Plus \+ Comma \,
-%% Minus \- Point \. Solidus \/
-%% Colon \: Semicolon \; Less than \<
-%% Equals \= Greater than \> Question mark \?
-%% Commercial at \@ Left bracket \[ Backslash \\
-%% Right bracket \] Circumflex \^ Underscore \_
-%% Grave accent \` Left brace \{ Vertical bar \|
-%% Right brace \} Tilde \~}
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-\ProvidesClass{svjour3}[2007/05/08 v3.2
-^^JLaTeX document class for Springer journals]
-\newcommand\@ptsize{}
-\newif\if@restonecol
-\newif\if@titlepage
-\@titlepagefalse
-\DeclareOption{a4paper}
- {\setlength\paperheight {297mm}%
- \setlength\paperwidth {210mm}}
-\DeclareOption{10pt}{\renewcommand\@ptsize{0}}
-\DeclareOption{twoside}{\@twosidetrue \@mparswitchtrue}
-\DeclareOption{draft}{\setlength\overfullrule{5pt}}
-\DeclareOption{final}{\setlength\overfullrule{0pt}}
-\DeclareOption{fleqn}{\input{fleqn.clo}\AtBeginDocument{\mathindent\z@}%
-\AtBeginDocument{\@ifpackageloaded{amsmath}{\@mathmargin\z@}{}}%
-\PassOptionsToPackage{fleqn}{amsmath}}
-%%%
-\DeclareOption{onecolumn}{}
-\DeclareOption{smallcondensed}{}
-\DeclareOption{twocolumn}{\@twocolumntrue\ExecuteOptions{fleqn}}
-\newif\if@smallext\@smallextfalse
-\DeclareOption{smallextended}{\@smallexttrue}
-\let\if@mathematic\iftrue
-\let\if@numbook\iffalse
-\DeclareOption{numbook}{\let\if@envcntsect\iftrue
- \AtEndOfPackage{%
- \renewcommand\thefigure{\thesection.\@arabic\c@figure}%
- \renewcommand\thetable{\thesection.\@arabic\c@table}%
- \renewcommand\theequation{\thesection.\@arabic\c@equation}%
- \@addtoreset{figure}{section}%
- \@addtoreset{table}{section}%
- \@addtoreset{equation}{section}%
- }%
-}
-\DeclareOption{openbib}{%
- \AtEndOfPackage{%
- \renewcommand\@openbib@code{%
- \advance\leftmargin\bibindent
- \itemindent -\bibindent
- \listparindent \itemindent
- \parsep \z@
- }%
- \renewcommand\newblock{\par}}%
-}
-\DeclareOption{natbib}{%
-\AtEndOfClass{\RequirePackage{natbib}%
-% Changing some parameters of NATBIB
-\setlength{\bibhang}{\parindent}%
-%\setlength{\bibsep}{0mm}%
-\let\bibfont=\small
-\def\@biblabel#1{#1.}%
-\newcommand{\etal}{et al.}%
-\bibpunct{(}{)}{;}{a}{}{,}}}
-%
-\let\if@runhead\iffalse
-\DeclareOption{runningheads}{\let\if@runhead\iftrue}
-\let\if@smartrunh\iffalse
-\DeclareOption{smartrunhead}{\let\if@smartrunh\iftrue}
-\DeclareOption{nosmartrunhead}{\let\if@smartrunh\iffalse}
-\let\if@envcntreset\iffalse
-\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue}
-\let\if@envcntsame\iffalse
-\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue}
-\let\if@envcntsect\iffalse
-\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue}
-\let\if@referee\iffalse
-\DeclareOption{referee}{\let\if@referee\iftrue}
-\def\makereferee{\def\baselinestretch{2}}
-\let\if@instindent\iffalse
-\DeclareOption{instindent}{\let\if@instindent\iftrue}
-\let\if@smartand\iffalse
-\DeclareOption{smartand}{\let\if@smartand\iftrue}
-\let\if@spthms\iftrue
-\DeclareOption{nospthms}{\let\if@spthms\iffalse}
-%
-% language and babel dependencies
-\DeclareOption{deutsch}{\def\switcht@@therlang{\switcht@deutsch}%
-\gdef\svlanginfo{\typeout{Man spricht deutsch.}\global\let\svlanginfo\relax}}
-\DeclareOption{francais}{\def\switcht@@therlang{\switcht@francais}%
-\gdef\svlanginfo{\typeout{On parle francais.}\global\let\svlanginfo\relax}}
-\let\switcht@@therlang\relax
-\let\svlanginfo\relax
-%
-\AtBeginDocument{\@ifpackageloaded{babel}{%
-\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}%
-\@ifundefined{extrasUKenglish}{}{\addto\extrasUKenglish{\switcht@albion}}%
-\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}%
-\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}%
-\@ifundefined{extrasngerman}{}{\addto\extrasngerman{\switcht@deutsch}}%
-}{\switcht@@therlang}%
-}
-%
-\def\ClassInfoNoLine#1#2{%
- \ClassInfo{#1}{#2\@gobble}%
-}
-\let\journalopt\@empty
-\DeclareOption*{%
-\InputIfFileExists{sv\CurrentOption.clo}{%
-\global\let\journalopt\CurrentOption}{%
-\ClassWarning{Springer-SVJour3}{Specified option or subpackage
-"\CurrentOption" not found -}\OptionNotUsed}}
-\ExecuteOptions{a4paper,twoside,10pt,instindent}
-\ProcessOptions
-%
-\ifx\journalopt\@empty\relax
-\ClassInfoNoLine{Springer-SVJour3}{extra/valid Springer sub-package (-> *.clo)
-\MessageBreak not found in option list of \string\documentclass
-\MessageBreak - autoactivating "global" style}{}
-\input{svglov3.clo}
-\else
-\@ifundefined{validfor}{%
-\ClassError{Springer-SVJour3}{Possible option clash for sub-package
-\MessageBreak "sv\journalopt.clo" - option file not valid
-\MessageBreak for this class}{Perhaps you used an option of the old
-Springer class SVJour!}
-}{}
-\fi
-%
-\if@smartrunh\AtEndDocument{\islastpageeven\getlastpagenumber}\fi
-%
-\newcommand{\twocoltest}[2]{\if@twocolumn\def\@gtempa{#2}\else\def\@gtempa{#1}\fi
-\@gtempa\makeatother}
-\newcommand{\columncase}{\makeatletter\twocoltest}
-%
-\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00}
-\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01}
-\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02}
-\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03}
-\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04}
-\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05}
-\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06}
-\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07}
-\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08}
-\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09}
-\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A}
-%
-\setlength\parindent{15\p@}
-\setlength\smallskipamount{3\p@ \@plus 1\p@ \@minus 1\p@}
-\setlength\medskipamount{6\p@ \@plus 2\p@ \@minus 2\p@}
-\setlength\bigskipamount{12\p@ \@plus 4\p@ \@minus 4\p@}
-\setlength\headheight{12\p@}
-\setlength\headsep {14.50dd}
-\setlength\topskip {10\p@}
-\setlength\footskip{30\p@}
-\setlength\maxdepth{.5\topskip}
-%
-\@settopoint\textwidth
-\setlength\marginparsep {10\p@}
-\setlength\marginparpush{5\p@}
-\setlength\topmargin{-10pt}
-\if@twocolumn
- \setlength\oddsidemargin {-30\p@}
- \setlength\evensidemargin{-30\p@}
-\else
- \setlength\oddsidemargin {\z@}
- \setlength\evensidemargin{\z@}
-\fi
-\setlength\marginparwidth {48\p@}
-\setlength\footnotesep{8\p@}
-\setlength{\skip\footins}{9\p@ \@plus 4\p@ \@minus 2\p@}
-\setlength\floatsep {12\p@ \@plus 2\p@ \@minus 2\p@}
-\setlength\textfloatsep{20\p@ \@plus 2\p@ \@minus 4\p@}
-\setlength\intextsep {20\p@ \@plus 2\p@ \@minus 2\p@}
-\setlength\dblfloatsep {12\p@ \@plus 2\p@ \@minus 2\p@}
-\setlength\dbltextfloatsep{20\p@ \@plus 2\p@ \@minus 4\p@}
-\setlength\@fptop{0\p@}
-\setlength\@fpsep{12\p@ \@plus 2\p@ \@minus 2\p@}
-\setlength\@fpbot{0\p@ \@plus 1fil}
-\setlength\@dblfptop{0\p@}
-\setlength\@dblfpsep{12\p@ \@plus 2\p@ \@minus 2\p@}
-\setlength\@dblfpbot{0\p@ \@plus 1fil}
-\setlength\partopsep{2\p@ \@plus 1\p@ \@minus 1\p@}
-\def\@listi{\leftmargin\leftmargini
- \parsep \z@
- \topsep 6\p@ \@plus2\p@ \@minus4\p@
- \itemsep\parsep}
-\let\@listI\@listi
-\@listi
-\def\@listii {\leftmargin\leftmarginii
- \labelwidth\leftmarginii
- \advance\labelwidth-\labelsep
- \topsep \z@
- \parsep \topsep
- \itemsep \parsep}
-\def\@listiii{\leftmargin\leftmarginiii
- \labelwidth\leftmarginiii
- \advance\labelwidth-\labelsep
- \topsep \z@
- \parsep \topsep
- \itemsep \parsep}
-\def\@listiv {\leftmargin\leftmarginiv
- \labelwidth\leftmarginiv
- \advance\labelwidth-\labelsep}
-\def\@listv {\leftmargin\leftmarginv
- \labelwidth\leftmarginv
- \advance\labelwidth-\labelsep}
-\def\@listvi {\leftmargin\leftmarginvi
- \labelwidth\leftmarginvi
- \advance\labelwidth-\labelsep}
-%
-\setlength\lineskip{1\p@}
-\setlength\normallineskip{1\p@}
-\renewcommand\baselinestretch{}
-\setlength\parskip{0\p@ \@plus \p@}
-\@lowpenalty 51
-\@medpenalty 151
-\@highpenalty 301
-\setcounter{topnumber}{4}
-\renewcommand\topfraction{.9}
-\setcounter{bottomnumber}{2}
-\renewcommand\bottomfraction{.7}
-\setcounter{totalnumber}{6}
-\renewcommand\textfraction{.1}
-\renewcommand\floatpagefraction{.85}
-\setcounter{dbltopnumber}{3}
-\renewcommand\dbltopfraction{.85}
-\renewcommand\dblfloatpagefraction{.85}
-\def\ps@headings{%
- \let\@oddfoot\@empty\let\@evenfoot\@empty
- \def\@evenhead{\small\csname runheadhook\endcsname
- \rlap{\thepage}\hfil\leftmark\unskip}%
- \def\@oddhead{\small\csname runheadhook\endcsname
- \ignorespaces\rightmark\hfil\llap{\thepage}}%
- \let\@mkboth\@gobbletwo
- \let\sectionmark\@gobble
- \let\subsectionmark\@gobble
- }
-% make indentations changeable
-\def\setitemindent#1{\settowidth{\labelwidth}{#1}%
- \leftmargini\labelwidth
- \advance\leftmargini\labelsep
- \def\@listi{\leftmargin\leftmargini
- \labelwidth\leftmargini\advance\labelwidth by -\labelsep
- \parsep=\parskip
- \topsep=\medskipamount
- \itemsep=\parskip \advance\itemsep by -\parsep}}
-\def\setitemitemindent#1{\settowidth{\labelwidth}{#1}%
- \leftmarginii\labelwidth
- \advance\leftmarginii\labelsep
-\def\@listii{\leftmargin\leftmarginii
- \labelwidth\leftmarginii\advance\labelwidth by -\labelsep
- \parsep=\parskip
- \topsep=\z@
- \itemsep=\parskip \advance\itemsep by -\parsep}}
-% labels of description
-\def\descriptionlabel#1{\hspace\labelsep #1\hfil}
-% adjusted environment "description"
-% if an optional parameter (at the first two levels of lists)
-% is present, its width is considered to be the widest mark
-% throughout the current list.
-\def\description{\@ifnextchar[{\@describe}{\list{}{\labelwidth\z@
- \itemindent-\leftmargin \let\makelabel\descriptionlabel}}}
-\let\enddescription\endlist
-%
-\def\describelabel#1{#1\hfil}
-\def\@describe[#1]{\relax\ifnum\@listdepth=0
-\setitemindent{#1}\else\ifnum\@listdepth=1
-\setitemitemindent{#1}\fi\fi
-\list{--}{\let\makelabel\describelabel}}
-%
-\newdimen\logodepth
-\logodepth=1.2cm
-\newdimen\headerboxheight
-\headerboxheight=163pt % 18 10.5dd-lines - 2\baselineskip
-\if@twocolumn\else\advance\headerboxheight by-14.5mm\fi
-\newdimen\betweenumberspace % dimension for space between
-\betweenumberspace=3.33pt % number and text of titles.
-\newdimen\aftertext % dimension for space after
-\aftertext=5pt % text of title.
-\newdimen\headlineindent % dimension for space between
-\headlineindent=1.166cm % number and text of headings.
-\if@mathematic
- \def\runinend{} % \enspace}
- \def\floatcounterend{\enspace}
- \def\sectcounterend{}
-\else
- \def\runinend{.}
- \def\floatcounterend{.\ }
- \def\sectcounterend{.}
-\fi
-\def\email#1{\emailname: #1}
-\def\keywords#1{\par\addvspace\medskipamount{\rightskip=0pt plus1cm
-\def\and{\ifhmode\unskip\nobreak\fi\ $\cdot$
-}\noindent\keywordname\enspace\ignorespaces#1\par}}
-%
-\def\subclassname{{\bfseries Mathematics Subject Classification
-(2000)}\enspace}
-\def\subclass#1{\par\addvspace\medskipamount{\rightskip=0pt plus1cm
-\def\and{\ifhmode\unskip\nobreak\fi\ $\cdot$
-}\noindent\subclassname\ignorespaces#1\par}}
-%
-\def\PACSname{\textbf{PACS}\enspace}
-\def\PACS#1{\par\addvspace\medskipamount{\rightskip=0pt plus1cm
-\def\and{\ifhmode\unskip\nobreak\fi\ $\cdot$
-}\noindent\PACSname\ignorespaces#1\par}}
-%
-\def\CRclassname{{\bfseries CR Subject Classification}\enspace}
-\def\CRclass#1{\par\addvspace\medskipamount{\rightskip=0pt plus1cm
-\def\and{\ifhmode\unskip\nobreak\fi\ $\cdot$
-}\noindent\CRclassname\ignorespaces#1\par}}
-%
-\def\ESMname{\textbf{Electronic Supplementary Material}\enspace}
-\def\ESM#1{\par\addvspace\medskipamount
-\noindent\ESMname\ignorespaces#1\par}
-%
-\newcounter{inst}
-\newcounter{auth}
-\def\authdepth{2}
-\newdimen\instindent
-\newbox\authrun
-\newtoks\authorrunning
-\newbox\titrun
-\newtoks\titlerunning
-\def\authorfont{\bfseries}
-
-\def\combirunning#1{\gdef\@combi{#1}}
-\def\@combi{}
-\newbox\combirun
-%
-\def\ps@last{\def\@evenhead{\small\rlap{\thepage}\hfil
-\lastevenhead}}
-\newcounter{lastpage}
-\def\islastpageeven{\@ifundefined{lastpagenumber}
-{\setcounter{lastpage}{0}}{\setcounter{lastpage}{\lastpagenumber}}
-\ifnum\value{lastpage}>0
- \ifodd\value{lastpage}%
- \else
- \if@smartrunh
- \thispagestyle{last}%
- \fi
- \fi
-\fi}
-\def\getlastpagenumber{\clearpage
-\addtocounter{page}{-1}%
- \immediate\write\@auxout{\string\gdef\string\lastpagenumber{\thepage}}%
- \immediate\write\@auxout{\string\newlabel{LastPage}{{}{\thepage}}}%
- \addtocounter{page}{1}}
-
-\def\journalname#1{\gdef\@journalname{#1}}
-
-\def\dedication#1{\gdef\@dedic{#1}}
-\def\@dedic{}
-
-\let\@date\undefined
-\def\notused{~}
-
-\def\institute#1{\gdef\@institute{#1}}
-
-\def\offprints#1{\begingroup
-\def\protect{\noexpand\protect\noexpand}\xdef\@thanks{\@thanks
-\protect\footnotetext[0]{\unskip\hskip-15pt{\itshape Send offprint requests
-to\/}: \ignorespaces#1}}\endgroup\ignorespaces}
-
-%\def\mail#1{\gdef\@mail{#1}}
-%\def\@mail{}
-
-\def\@thanks{}
-
-\def\@fnsymbol#1{\ifcase#1\or\star\or{\star\star}\or{\star\star\star}%
- \or \dagger\or \ddagger\or
- \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger
- \or \ddagger\ddagger \else\@ctrerr\fi\relax}
-%
-%\def\invthanks#1{\footnotetext[0]{\kern-\bibindent#1}}
-%
-\def\nothanksmarks{\def\thanks##1{\protected@xdef\@thanks{\@thanks
- \protect\footnotetext[0]{\kern-\bibindent##1}}}}
-%
-\def\subtitle#1{\gdef\@subtitle{#1}}
-\def\@subtitle{}
-
-\def\headnote#1{\gdef\@headnote{#1}}
-\def\@headnote{}
-
-\def\papertype#1{\gdef\paper@type{\MakeUppercase{#1}}}
-\def\paper@type{}
-
-\def\ch@ckobl#1#2{\@ifundefined{@#1}
- {\typeout{SVJour3 warning: Missing
-\expandafter\string\csname#1\endcsname}%
- \csname #1\endcsname{#2}}
- {}}
-%
-\def\ProcessRunnHead{%
- \def\\{\unskip\ \ignorespaces}%
- \def\thanks##1{\unskip{}}%
- \instindent=\textwidth
- \advance\instindent by-\headlineindent
- \if!\the\titlerunning!\else
- \edef\@title{\the\titlerunning}%
- \fi
- \global\setbox\titrun=\hbox{\small\rmfamily\unboldmath\ignorespaces\@title
- \unskip}%
- \ifdim\wd\titrun>\instindent
- \typeout{^^JSVJour3 Warning: Title too long for running head.}%
- \typeout{Please supply a shorter form with \string\titlerunning
- \space prior to \string\maketitle}%
- \global\setbox\titrun=\hbox{\small\rmfamily
- Title Suppressed Due to Excessive Length}%
- \fi
- \xdef\@title{\copy\titrun}%
-%
- \if!\the\authorrunning!
- \else
- \setcounter{auth}{1}%
- \edef\@author{\the\authorrunning}%
- \fi
- \ifnum\value{inst}>\authdepth
- \def\stripauthor##1\and##2\endauthor{%
- \protected@xdef\@author{##1\unskip\unskip\if!##2!\else\ et al.\fi}}%
- \expandafter\stripauthor\@author\and\endauthor
- \else
- \gdef\and{\unskip, \ignorespaces}%
- {\def\and{\noexpand\protect\noexpand\and}%
- \protected@xdef\@author{\@author}}
- \fi
- \global\setbox\authrun=\hbox{\small\rmfamily\unboldmath\ignorespaces
- \@author\unskip}%
- \ifdim\wd\authrun>\instindent
- \typeout{^^JSVJour3 Warning: Author name(s) too long for running head.
- ^^JPlease supply a shorter form with \string\authorrunning
- \space prior to \string\maketitle}%
- \global\setbox\authrun=\hbox{\small\rmfamily Please give a shorter version
- with: {\tt\string\authorrunning\space and
- \string\titlerunning\space prior to \string\maketitle}}%
- \fi
- \xdef\@author{\copy\authrun}%
- \markboth{\@author}{\@title}%
-}
-%
-\let\orithanks=\thanks
-\def\thanks#1{\ClassWarning{SVJour3}{\string\thanks\space may only be
-used inside of \string\title, \string\author,\MessageBreak
-and \string\date\space prior to \string\maketitle}}
-%
-\def\maketitle{\par\let\thanks=\orithanks
-\ch@ckobl{}{Noname}
-\ch@ckobl{date}{the date of receipt and acceptance should be inserted
-later}
-\ch@ckobl{title}{A title should be given}
-\ch@ckobl{author}{Name(s) and initial(s) of author(s) should be given}
-\ch@ckobl{institute}{Address(es) of author(s) should be given}
-\begingroup
-%
- \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
- \def\@makefnmark{$^{\@thefnmark}$}%
- \renewcommand\@makefntext[1]{%
- \noindent
- \hb@xt@\bibindent{\hss\@makefnmark\enspace}##1\vrule height0pt
- width0pt depth8pt}
-%
- \def\lastand{\ifnum\value{inst}=2\relax
- \unskip{} \andname\
- \else
- \unskip, \andname\
- \fi}%
- \def\and{\stepcounter{auth}\relax
- \if@smartand
- \ifnum\value{auth}=\value{inst}%
- \lastand
- \else
- \unskip,
- \fi
- \else
- \unskip,
- \fi}%
- \thispagestyle{empty}
- \ifnum \col@number=\@ne
- \@maketitle
- \else
- \twocolumn[\@maketitle]%
- \fi
-%
- \global\@topnum\z@
- \if!\@thanks!\else
- \@thanks
-\insert\footins{\vskip-3pt\hrule\@width\if@twocolumn\columnwidth
-\else 38mm\fi\vskip3pt}%
- \fi
- {\def\thanks##1{\unskip{}}%
- \def\iand{\\[5pt]\let\and=\nand}%
- \def\nand{\ifhmode\unskip\nobreak\fi\ $\cdot$ }%
- \let\and=\nand
- \def\at{\\\let\and=\iand}%
- \footnotetext[0]{\kern-\bibindent
- \ignorespaces\@institute}\vspace{5dd}}%
-%\if!\@mail!\else
-% \footnotetext[0]{\kern-\bibindent\mailname\
-% \ignorespaces\@mail}%
-%\fi
-%
- \if@runhead
- \ProcessRunnHead
- \fi
-%
- \endgroup
- \setcounter{footnote}{0}
- \global\let\thanks\relax
- \global\let\maketitle\relax
- \global\let\@maketitle\relax
- \global\let\@thanks\@empty
- \global\let\@author\@empty
- \global\let\@date\@empty
- \global\let\@title\@empty
- \global\let\@subtitle\@empty
- \global\let\title\relax
- \global\let\author\relax
- \global\let\date\relax
- \global\let\and\relax}
-
-\def\makeheadbox{{%
-\hbox to0pt{\vbox{\baselineskip=10dd\hrule\hbox
-to\hsize{\vrule\kern3pt\vbox{\kern3pt
-\hbox{J Automated Reasoning}
-\hbox{}
-\kern3pt}\hfil\kern3pt\vrule}\hrule}%
-\hss}}}
-%
-\def\rubric{\setbox0=\hbox{\small\strut}\@tempdima=\ht0\advance
-\@tempdima\dp0\advance\@tempdima2\fboxsep\vrule\@height\@tempdima
-\@width\z@}
-\newdimen\rubricwidth
-%
-\def\@maketitle{\newpage
-\normalfont
-\vbox to0pt{\if@twocolumn\vskip-39pt\else\vskip-49pt\fi
-\nointerlineskip
-\makeheadbox\vss}\nointerlineskip
-\vbox to 0pt{\offinterlineskip\rubricwidth=\columnwidth
-%%%%\vskip-12.5pt % -12.5pt
-\if@twocolumn\else % one column journal
- \divide\rubricwidth by144\multiply\rubricwidth by89 % perform golden section
- \vskip-\topskip
-\fi
-\hrule\@height0.35mm\noindent
-\advance\fboxsep by.25mm
-\global\advance\rubricwidth by0pt
-\rubric
-\vss}\vskip19.5pt % war 9pt
-%
-\if@twocolumn\else
- \gdef\footnoterule{%
- \kern-3\p@
- \hrule\@width38mm % \columnwidth \rubricwidth
- \kern2.6\p@}
-\fi
-%
- \setbox\authrun=\vbox\bgroup
- \if@twocolumn
- \hrule\@height10.5mm\@width0\p@
- \else
- \hrule\@height 2mm\@width0\p@
- \fi
- \pretolerance=10000
- \rightskip=0pt plus 4cm
- \nothanksmarks
-% \if!\@headnote!\else
-% \noindent
-% {\LARGE\normalfont\itshape\ignorespaces\@headnote\par}\vskip 3.5mm
-% \fi
- {\LARGE\bfseries
- \noindent\ignorespaces
- \@title \par}\vskip 17pt\relax
- \if!\@subtitle!\else
- {\large\bfseries
- \pretolerance=10000
- \rightskip=0pt plus 3cm
- \vskip-12pt
-% \noindent\ignorespaces\@subtitle \par}\vskip 11.24pt\relax
- \noindent\ignorespaces\@subtitle \par}\vskip 17pt\relax
- \fi
- {\authorfont
- \setbox0=\vbox{\setcounter{auth}{1}\def\and{\stepcounter{auth} }%
- \hfuzz=2\textwidth\def\thanks##1{}\@author}%
- \setcounter{footnote}{0}%
- \global\value{inst}=\value{auth}%
- \setcounter{auth}{1}%
- \if@twocolumn
- \rightskip43mm plus 4cm minus 3mm
- \else % one column journal
- \rightskip=\linewidth
- \advance\rightskip by-\rubricwidth
- \advance\rightskip by0pt plus 4cm minus 3mm
- \fi
-%
-\def\and{\unskip\nobreak\enskip{\boldmath$\cdot$}\enskip\ignorespaces}%
- \noindent\ignorespaces\@author\vskip7.23pt}
-%
- \small
- \if!\@dedic!\else
- \par
- \normalsize\it
- \addvspace\baselineskip
- \noindent\@dedic
- \fi
- \egroup % end of header box
- \@tempdima=\headerboxheight
- \advance\@tempdima by-\ht\authrun
- \unvbox\authrun
- \ifdim\@tempdima>0pt
- \vrule width0pt height\@tempdima\par
- \fi
- \noindent{\small\@date\if@twocolumn\vskip 7.2mm\else\vskip 5.2mm\fi}
- \global\@minipagetrue
- \global\everypar{\global\@minipagefalse\global\everypar{}}%
-%\vskip22.47pt
-}
-%
-\if@mathematic
- \def\vec#1{\ensuremath{\mathchoice
- {\mbox{\boldmath$\displaystyle\mathbf{#1}$}}
- {\mbox{\boldmath$\textstyle\mathbf{#1}$}}
- {\mbox{\boldmath$\scriptstyle\mathbf{#1}$}}
- {\mbox{\boldmath$\scriptscriptstyle\mathbf{#1}$}}}}
-\else
- \def\vec#1{\ensuremath{\mathchoice
- {\mbox{\boldmath$\displaystyle#1$}}
- {\mbox{\boldmath$\textstyle#1$}}
- {\mbox{\boldmath$\scriptstyle#1$}}
- {\mbox{\boldmath$\scriptscriptstyle#1$}}}}
-\fi
-%
-\def\tens#1{\ensuremath{\mathsf{#1}}}
-%
-\setcounter{secnumdepth}{3}
-\newcounter {section}
-\newcounter {subsection}[section]
-\newcounter {subsubsection}[subsection]
-\newcounter {paragraph}[subsubsection]
-\newcounter {subparagraph}[paragraph]
-\renewcommand\thesection {\@arabic\c@section}
-\renewcommand\thesubsection {\thesection.\@arabic\c@subsection}
-\renewcommand\thesubsubsection{\thesubsection.\@arabic\c@subsubsection}
-\renewcommand\theparagraph {\thesubsubsection.\@arabic\c@paragraph}
-\renewcommand\thesubparagraph {\theparagraph.\@arabic\c@subparagraph}
-%
-\def\@hangfrom#1{\setbox\@tempboxa\hbox{#1}%
- \hangindent \z@\noindent\box\@tempboxa}
-%
-\def\@seccntformat#1{\csname the#1\endcsname\sectcounterend
-\hskip\betweenumberspace}
-%
-% \newif\if@sectrule
-% \if@twocolumn\else\let\@sectruletrue=\relax\fi
-% \if@avier\let\@sectruletrue=\relax\fi
-% \def\makesectrule{\if@sectrule\global\@sectrulefalse\null\vglue-\topskip
-% \hrule\nobreak\parskip=5pt\relax\fi}
-% %
-% \let\makesectruleori=\makesectrule
-% \def\restoresectrule{\global\let\makesectrule=\makesectruleori\global\@sectrulefalse}
-% \def\nosectrule{\let\makesectrule=\restoresectrule}
-%
-\def\@startsection#1#2#3#4#5#6{%
- \if@noskipsec \leavevmode \fi
- \par
- \@tempskipa #4\relax
- \@afterindenttrue
- \ifdim \@tempskipa <\z@
- \@tempskipa -\@tempskipa \@afterindentfalse
- \fi
- \if@nobreak
- \everypar{}%
- \else
- \addpenalty\@secpenalty\addvspace\@tempskipa
- \fi
-% \ifnum#2=1\relax\@sectruletrue\fi
- \@ifstar
- {\@ssect{#3}{#4}{#5}{#6}}%
- {\@dblarg{\@sect{#1}{#2}{#3}{#4}{#5}{#6}}}}
-%
-\def\@sect#1#2#3#4#5#6[#7]#8{%
- \ifnum #2>\c@secnumdepth
- \let\@svsec\@empty
- \else
- \refstepcounter{#1}%
- \protected@edef\@svsec{\@seccntformat{#1}\relax}%
- \fi
- \@tempskipa #5\relax
- \ifdim \@tempskipa>\z@
- \begingroup
- #6{% \makesectrule
- \@hangfrom{\hskip #3\relax\@svsec}%
- \raggedright
- \hyphenpenalty \@M%
- \interlinepenalty \@M #8\@@par}%
- \endgroup
- \csname #1mark\endcsname{#7}%
- \addcontentsline{toc}{#1}{%
- \ifnum #2>\c@secnumdepth \else
- \protect\numberline{\csname the#1\endcsname\sectcounterend}%
- \fi
- #7}%
- \else
- \def\@svsechd{%
- #6{\hskip #3\relax
- \@svsec #8\/\hskip\aftertext}%
- \csname #1mark\endcsname{#7}%
- \addcontentsline{toc}{#1}{%
- \ifnum #2>\c@secnumdepth \else
- \protect\numberline{\csname the#1\endcsname}%
- \fi
- #7}}%
- \fi
- \@xsect{#5}}
-%
-\def\@ssect#1#2#3#4#5{%
- \@tempskipa #3\relax
- \ifdim \@tempskipa>\z@
- \begingroup
- #4{% \makesectrule
- \@hangfrom{\hskip #1}%
- \interlinepenalty \@M #5\@@par}%
- \endgroup
- \else
- \def\@svsechd{#4{\hskip #1\relax #5}}%
- \fi
- \@xsect{#3}}
-
-%
-% measures and setting of sections
-%
-\def\section{\@startsection{section}{1}{\z@}%
- {-21dd plus-8pt minus-4pt}{10.5dd}
- {\normalsize\bfseries\boldmath}}
-\def\subsection{\@startsection{subsection}{2}{\z@}%
- {-21dd plus-8pt minus-4pt}{10.5dd}
- {\normalsize\upshape}}
-\def\subsubsection{\@startsection{subsubsection}{3}{\z@}%
- {-13dd plus-8pt minus-4pt}{10.5dd}
- {\normalsize\itshape}}
-\def\paragraph{\@startsection{paragraph}{4}{\z@}%
- {-13pt plus-8pt minus-4pt}{\z@}{\normalsize\itshape}}
-
-\setlength\leftmargini {\parindent}
-\leftmargin \leftmargini
-\setlength\leftmarginii {\parindent}
-\setlength\leftmarginiii {1.87em}
-\setlength\leftmarginiv {1.7em}
-\setlength\leftmarginv {.5em}
-\setlength\leftmarginvi {.5em}
-\setlength \labelsep {.5em}
-\setlength \labelwidth{\leftmargini}
-\addtolength\labelwidth{-\labelsep}
-\@beginparpenalty -\@lowpenalty
-\@endparpenalty -\@lowpenalty
-\@itempenalty -\@lowpenalty
-\renewcommand\theenumi{\@arabic\c@enumi}
-\renewcommand\theenumii{\@alph\c@enumii}
-\renewcommand\theenumiii{\@roman\c@enumiii}
-\renewcommand\theenumiv{\@Alph\c@enumiv}
-\newcommand\labelenumi{\theenumi.}
-\newcommand\labelenumii{(\theenumii)}
-\newcommand\labelenumiii{\theenumiii.}
-\newcommand\labelenumiv{\theenumiv.}
-\renewcommand\p@enumii{\theenumi}
-\renewcommand\p@enumiii{\theenumi(\theenumii)}
-\renewcommand\p@enumiv{\p@enumiii\theenumiii}
-\newcommand\labelitemi{\normalfont\bfseries --}
-\newcommand\labelitemii{\normalfont\bfseries --}
-\newcommand\labelitemiii{$\m@th\bullet$}
-\newcommand\labelitemiv{$\m@th\cdot$}
-
-\if@spthms
-% definition of the "\spnewtheorem" command.
-%
-% Usage:
-%
-% \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font}
-% or \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font}
-% or \spnewtheorem*{env_nam}{caption}{cap_font}{body_font}
-%
-% New is "cap_font" and "body_font". It stands for
-% fontdefinition of the caption and the text itself.
-%
-% "\spnewtheorem*" gives a theorem without number.
-%
-% A defined spnewthoerem environment is used as described
-% by Lamport.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\def\@thmcountersep{}
-\def\@thmcounterend{}
-\newcommand\nocaption{\noexpand\@gobble}
-\newdimen\spthmsep \spthmsep=5pt
-
-\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}}
-
-% definition of \spnewtheorem with number
-
-\def\@spnthm#1#2{%
- \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}}
-\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}}
-
-\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}\@addtoreset{#1}{#3}%
- \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand
- \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}%
- \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spothm#1[#2]#3#4#5{%
- \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}%
- {\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{the#1}{\@nameuse{the#2}}%
- \expandafter\xdef\csname #1name\endcsname{#3}%
- \global\@namedef{#1}{\@spthm{#2}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}}
-
-\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\labelsep=\spthmsep\refstepcounter{#1}%
-\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}}
-
-\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}%
- \ignorespaces}
-
-\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname
- the#1\endcsname}{#5}{#3}{#4}\ignorespaces}
-
-\def\normalthmheadings{\def\@spbegintheorem##1##2##3##4{\trivlist\normalfont
- \item[\hskip\labelsep{##3##1\ ##2\@thmcounterend}]##4}
-\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##1\ ##2}]{##4(##3)\@thmcounterend\ }##5}}
-\normalthmheadings
-
-\def\reversethmheadings{\def\@spbegintheorem##1##2##3##4{\trivlist\normalfont
- \item[\hskip\labelsep{##3##2\ ##1\@thmcounterend}]##4}
-\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##2\ ##1}]{##4(##3)\@thmcounterend\ }##5}}
-
-% definition of \spnewtheorem* without number
-
-\def\@sthm#1#2{\@Ynthm{#1}{#2}}
-
-\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}}
-
-\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces}
-
-\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1}
- {#4}{#2}{#3}\ignorespaces}
-
-\def\@Begintheorem#1#2#3{#3\trivlist
- \item[\hskip\labelsep{#2#1\@thmcounterend}]}
-
-\def\@Opargbegintheorem#1#2#3#4{#4\trivlist
- \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }}
-
-% initialize theorem environment
-
-\if@envcntsect
- \def\@thmcountersep{.}
- \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape}
-\else
- \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape}
- \if@envcntreset
- \@addtoreset{theorem}{section}
- \else
- \@addtoreset{theorem}{chapter}
- \fi
-\fi
-
-%definition of divers theorem environments
-\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily}
-\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily}
-\if@envcntsame % all environments like "Theorem" - using its counter
- \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}}
-\else % all environments with their own counter
- \if@envcntsect % show section counter
- \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}}
- \else % not numbered with section
- \if@envcntreset
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{section}}
- \else
- \let\spn@wtheorem=\@spynthm
- \fi
- \fi
-\fi
-%
-\let\spdefaulttheorem=\spn@wtheorem
-%
-\spn@wtheorem{case}{Case}{\itshape}{\rmfamily}
-\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily}
-\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape}
-\spn@wtheorem{definition}{Definition}{\bfseries}{\rmfamily}
-\spn@wtheorem{example}{Example}{\itshape}{\rmfamily}
-\spn@wtheorem{exercise}{Exercise}{\bfseries}{\rmfamily}
-\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape}
-\spn@wtheorem{note}{Note}{\itshape}{\rmfamily}
-\spn@wtheorem{problem}{Problem}{\bfseries}{\rmfamily}
-\spn@wtheorem{property}{Property}{\itshape}{\rmfamily}
-\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape}
-\spn@wtheorem{question}{Question}{\itshape}{\rmfamily}
-\spn@wtheorem{solution}{Solution}{\bfseries}{\rmfamily}
-\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily}
-%
-\newenvironment{theopargself}
- {\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5}
- \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
- \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }}}{}
-\newenvironment{theopargself*}
- {\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##1\ ##2}]{\hspace*{-\labelsep}##4##3\@thmcounterend}##5}
- \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
- \item[\hskip\labelsep{##3##1}]{\hspace*{-\labelsep}##3##2\@thmcounterend}}}{}
-%
-\fi
-
-\def\@takefromreset#1#2{%
- \def\@tempa{#1}%
- \let\@tempd\@elt
- \def\@elt##1{%
- \def\@tempb{##1}%
- \ifx\@tempa\@tempb\else
- \@addtoreset{##1}{#2}%
- \fi}%
- \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
- \expandafter\def\csname cl@#2\endcsname{}%
- \@tempc
- \let\@elt\@tempd}
-
-\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}}
-\def\qed{\ifmmode\else\unskip\quad\fi\squareforqed}
-\def\smartqed{\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil
-\penalty50\hskip1em\null\nobreak\hfil\squareforqed
-\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi}}
-
-% Define `abstract' environment
-\def\abstract{\topsep=0pt\partopsep=0pt\parsep=0pt\itemsep=0pt\relax
-\trivlist\item[\hskip\labelsep
-{\bfseries\abstractname}]\if!\abstractname!\hskip-\labelsep\fi}
-\if@twocolumn
-% \if@avier
-% \def\endabstract{\endtrivlist\addvspace{5mm}\strich}
-% \def\strich{\hrule\vskip1ptplus12pt}
-% \else
- \def\endabstract{\endtrivlist\addvspace{3mm}}
-% \fi
-\else
-\fi
-%
-\newenvironment{verse}
- {\let\\\@centercr
- \list{}{\itemsep \z@
- \itemindent -1.5em%
- \listparindent\itemindent
- \rightmargin \leftmargin
- \advance\leftmargin 1.5em}%
- \item\relax}
- {\endlist}
-\newenvironment{quotation}
- {\list{}{\listparindent 1.5em%
- \itemindent \listparindent
- \rightmargin \leftmargin
- \parsep \z@ \@plus\p@}%
- \item\relax}
- {\endlist}
-\newenvironment{quote}
- {\list{}{\rightmargin\leftmargin}%
- \item\relax}
- {\endlist}
-\newcommand\appendix{\par\small
- \setcounter{section}{0}%
- \setcounter{subsection}{0}%
- \renewcommand\thesection{\@Alph\c@section}}
-\setlength\arraycolsep{1.5\p@}
-\setlength\tabcolsep{6\p@}
-\setlength\arrayrulewidth{.4\p@}
-\setlength\doublerulesep{2\p@}
-\setlength\tabbingsep{\labelsep}
-\skip\@mpfootins = \skip\footins
-\setlength\fboxsep{3\p@}
-\setlength\fboxrule{.4\p@}
-\renewcommand\theequation{\@arabic\c@equation}
-\newcounter{figure}
-\renewcommand\thefigure{\@arabic\c@figure}
-\def\fps@figure{tbp}
-\def\ftype@figure{1}
-\def\ext@figure{lof}
-\def\fnum@figure{\figurename~\thefigure}
-\newenvironment{figure}
- {\@float{figure}}
- {\end@float}
-\newenvironment{figure*}
- {\@dblfloat{figure}}
- {\end@dblfloat}
-\newcounter{table}
-\renewcommand\thetable{\@arabic\c@table}
-\def\fps@table{tbp}
-\def\ftype@table{2}
-\def\ext@table{lot}
-\def\fnum@table{\tablename~\thetable}
-\newenvironment{table}
- {\@float{table}}
- {\end@float}
-\newenvironment{table*}
- {\@dblfloat{table}}
- {\end@dblfloat}
-%
-\def \@floatboxreset {%
- \reset@font
- \small
- \@setnobreak
- \@setminipage
-}
-%
-\newcommand{\tableheadseprule}{\noalign{\hrule height.375mm}}
-%
-\newlength\abovecaptionskip
-\newlength\belowcaptionskip
-\setlength\abovecaptionskip{10\p@}
-\setlength\belowcaptionskip{0\p@}
-\newcommand\leftlegendglue{}
-
-\def\fig@type{figure}
-
-\newdimen\figcapgap\figcapgap=3pt
-\newdimen\tabcapgap\tabcapgap=5.5pt
-
-\@ifundefined{floatlegendstyle}{\def\floatlegendstyle{\bfseries}}{}
-
-\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname
- ext@#1\endcsname}{#1}{\protect\numberline{\csname
- the#1\endcsname}{\ignorespaces #2}}\begingroup
- \@parboxrestore
- \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
- \endgroup}
-
-\def\capstrut{\vrule\@width\z@\@height\topskip}
-
-\@ifundefined{captionstyle}{\def\captionstyle{\normalfont\small}}{}
-
-\long\def\@makecaption#1#2{%
- \captionstyle
- \ifx\@captype\fig@type
- \vskip\figcapgap
- \fi
- \setbox\@tempboxa\hbox{{\floatlegendstyle #1\floatcounterend}%
- \capstrut #2}%
- \ifdim \wd\@tempboxa >\hsize
- {\floatlegendstyle #1\floatcounterend}\capstrut #2\par
- \else
- \hbox to\hsize{\leftlegendglue\unhbox\@tempboxa\hfil}%
- \fi
- \ifx\@captype\fig@type\else
- \vskip\tabcapgap
- \fi}
-
-\newdimen\figgap\figgap=1cc
-\long\def\@makesidecaption#1#2{%
- \parbox[b]{\@tempdimb}{\captionstyle{\floatlegendstyle
- #1\floatcounterend}#2}}
-\def\sidecaption#1\caption{%
-\setbox\@tempboxa=\hbox{#1\unskip}%
-\if@twocolumn
- \ifdim\hsize<\textwidth\else
- \ifdim\wd\@tempboxa<\columnwidth
- \typeout{Double column float fits into single column -
- ^^Jyou'd better switch the environment. }%
- \fi
- \fi
-\fi
-\@tempdimb=\hsize
-\advance\@tempdimb by-\figgap
-\advance\@tempdimb by-\wd\@tempboxa
-\ifdim\@tempdimb<3cm
- \typeout{\string\sidecaption: No sufficient room for the legend;
- using normal \string\caption. }%
- \unhbox\@tempboxa
- \let\@capcommand=\@caption
-\else
- \let\@capcommand=\@sidecaption
- \leavevmode
- \unhbox\@tempboxa
- \hfill
-\fi
-\refstepcounter\@captype
-\@dblarg{\@capcommand\@captype}}
-
-\long\def\@sidecaption#1[#2]#3{\addcontentsline{\csname
- ext@#1\endcsname}{#1}{\protect\numberline{\csname
- the#1\endcsname}{\ignorespaces #2}}\begingroup
- \@parboxrestore
- \@makesidecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
- \endgroup}
-
-% Define `acknowledgement' environment
-\def\acknowledgement{\par\addvspace{17pt}\small\rmfamily
-\trivlist\if!\ackname!\item[]\else
-\item[\hskip\labelsep
-{\bfseries\ackname}]\fi}
-\def\endacknowledgement{\endtrivlist\addvspace{6pt}}
-\newenvironment{acknowledgements}{\begin{acknowledgement}}
-{\end{acknowledgement}}
-% Define `noteadd' environment
-\def\noteadd{\par\addvspace{17pt}\small\rmfamily
-\trivlist\item[\hskip\labelsep
-{\itshape\noteaddname}]}
-\def\endnoteadd{\endtrivlist\addvspace{6pt}}
-
-\DeclareOldFontCommand{\rm}{\normalfont\rmfamily}{\mathrm}
-\DeclareOldFontCommand{\sf}{\normalfont\sffamily}{\mathsf}
-\DeclareOldFontCommand{\tt}{\normalfont\ttfamily}{\mathtt}
-\DeclareOldFontCommand{\bf}{\normalfont\bfseries}{\mathbf}
-\DeclareOldFontCommand{\it}{\normalfont\itshape}{\mathit}
-\DeclareOldFontCommand{\sl}{\normalfont\slshape}{\@nomath\sl}
-\DeclareOldFontCommand{\sc}{\normalfont\scshape}{\@nomath\sc}
-\DeclareRobustCommand*\cal{\@fontswitch\relax\mathcal}
-\DeclareRobustCommand*\mit{\@fontswitch\relax\mathnormal}
-\newcommand\@pnumwidth{1.55em}
-\newcommand\@tocrmarg{2.55em}
-\newcommand\@dotsep{4.5}
-\setcounter{tocdepth}{1}
-\newcommand\tableofcontents{%
- \section*{\contentsname}%
- \@starttoc{toc}%
- \addtocontents{toc}{\begingroup\protect\small}%
- \AtEndDocument{\addtocontents{toc}{\endgroup}}%
- }
-\newcommand*\l@part[2]{%
- \ifnum \c@tocdepth >-2\relax
- \addpenalty\@secpenalty
- \addvspace{2.25em \@plus\p@}%
- \begingroup
- \setlength\@tempdima{3em}%
- \parindent \z@ \rightskip \@pnumwidth
- \parfillskip -\@pnumwidth
- {\leavevmode
- \large \bfseries #1\hfil \hb@xt@\@pnumwidth{\hss #2}}\par
- \nobreak
- \if@compatibility
- \global\@nobreaktrue
- \everypar{\global\@nobreakfalse\everypar{}}%
- \fi
- \endgroup
- \fi}
-\newcommand*\l@section{\@dottedtocline{1}{0pt}{1.5em}}
-\newcommand*\l@subsection{\@dottedtocline{2}{1.5em}{2.3em}}
-\newcommand*\l@subsubsection{\@dottedtocline{3}{3.8em}{3.2em}}
-\newcommand*\l@paragraph{\@dottedtocline{4}{7.0em}{4.1em}}
-\newcommand*\l@subparagraph{\@dottedtocline{5}{10em}{5em}}
-\newcommand\listoffigures{%
- \section*{\listfigurename
- \@mkboth{\listfigurename}%
- {\listfigurename}}%
- \@starttoc{lof}%
- }
-\newcommand*\l@figure{\@dottedtocline{1}{1.5em}{2.3em}}
-\newcommand\listoftables{%
- \section*{\listtablename
- \@mkboth{\listtablename}{\listtablename}}%
- \@starttoc{lot}%
- }
-\let\l@table\l@figure
-\newdimen\bibindent
-\setlength\bibindent{\parindent}
-\def\@biblabel#1{#1.}
-\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw
- {\let\protect\noexpand
- \immediate
- \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
-\newenvironment{thebibliography}[1]
- {\section*{\refname
- \@mkboth{\refname}{\refname}}\small
- \list{\@biblabel{\@arabic\c@enumiv}}%
- {\settowidth\labelwidth{\@biblabel{#1}}%
- \leftmargin\labelwidth
- \advance\leftmargin\labelsep
- \@openbib@code
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{\@arabic\c@enumiv}}%
- \sloppy\clubpenalty4000\widowpenalty4000%
- \sfcode`\.\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
-%
-\newcount\@tempcntc
-\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi
- \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do
- {\@ifundefined
- {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries
- ?}\@warning
- {Citation `\@citeb' on page \thepage \space undefined}}%
- {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}%
- \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne
- \@citea\def\@citea{,\hskip0.1em\ignorespaces}\hbox{\csname b@\@citeb\endcsname}%
- \else
- \advance\@tempcntb\@ne
- \ifnum\@tempcntb=\@tempcntc
- \else\advance\@tempcntb\m@ne\@citeo
- \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}}
-\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else
- \@citea\def\@citea{,\hskip0.1em\ignorespaces}%
- \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else
- {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else \def\@citea{--}\fi
- \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi}
-%
-\newcommand\newblock{\hskip .11em\@plus.33em\@minus.07em}
-\let\@openbib@code\@empty
-\newenvironment{theindex}
- {\if@twocolumn
- \@restonecolfalse
- \else
- \@restonecoltrue
- \fi
- \columnseprule \z@
- \columnsep 35\p@
- \twocolumn[\section*{\indexname}]%
- \@mkboth{\indexname}{\indexname}%
- \thispagestyle{plain}\parindent\z@
- \parskip\z@ \@plus .3\p@\relax
- \let\item\@idxitem}
- {\if@restonecol\onecolumn\else\clearpage\fi}
-\newcommand\@idxitem{\par\hangindent 40\p@}
-\newcommand\subitem{\@idxitem \hspace*{20\p@}}
-\newcommand\subsubitem{\@idxitem \hspace*{30\p@}}
-\newcommand\indexspace{\par \vskip 10\p@ \@plus5\p@ \@minus3\p@\relax}
-
-\if@twocolumn
- \renewcommand\footnoterule{%
- \kern-3\p@
- \hrule\@width\columnwidth
- \kern2.6\p@}
-\else
- \renewcommand\footnoterule{%
- \kern-3\p@
- \hrule\@width.382\columnwidth
- \kern2.6\p@}
-\fi
-\newcommand\@makefntext[1]{%
- \noindent
- \hb@xt@\bibindent{\hss\@makefnmark\enspace}#1}
-%
-\def\trans@english{\switcht@albion}
-\def\trans@french{\switcht@francais}
-\def\trans@german{\switcht@deutsch}
-\newenvironment{translation}[1]{\if!#1!\else
-\@ifundefined{selectlanguage}{\csname trans@#1\endcsname}{\selectlanguage{#1}}%
-\fi}{}
-% languages
-% English section
-\def\switcht@albion{%\typeout{English spoken.}%
- \def\abstractname{Abstract}%
- \def\ackname{Acknowledgements}%
- \def\andname{and}%
- \def\lastandname{, and}%
- \def\appendixname{Appendix}%
- \def\chaptername{Chapter}%
- \def\claimname{Claim}%
- \def\conjecturename{Conjecture}%
- \def\contentsname{Contents}%
- \def\corollaryname{Corollary}%
- \def\definitionname{Definition}%
- \def\emailname{E-mail}%
- \def\examplename{Example}%
- \def\exercisename{Exercise}%
- \def\figurename{Fig.}%
- \def\keywordname{{\bfseries Keywords}}%
- \def\indexname{Index}%
- \def\lemmaname{Lemma}%
- \def\contriblistname{List of Contributors}%
- \def\listfigurename{List of Figures}%
- \def\listtablename{List of Tables}%
- \def\mailname{{\itshape Correspondence to\/}:}%
- \def\noteaddname{Note added in proof}%
- \def\notename{Note}%
- \def\partname{Part}%
- \def\problemname{Problem}%
- \def\proofname{Proof}%
- \def\propertyname{Property}%
- \def\questionname{Question}%
- \def\refname{References}%
- \def\remarkname{Remark}%
- \def\seename{see}%
- \def\solutionname{Solution}%
- \def\tablename{Table}%
- \def\theoremname{Theorem}%
-}\switcht@albion % make English default
-%
-% French section
-\def\switcht@francais{\svlanginfo
-%\typeout{On parle francais.}%
- \def\abstractname{R\'esum\'e\runinend}%
- \def\ackname{Remerciements\runinend}%
- \def\andname{et}%
- \def\lastandname{ et}%
- \def\appendixname{Appendice}%
- \def\chaptername{Chapitre}%
- \def\claimname{Pr\'etention}%
- \def\conjecturename{Hypoth\`ese}%
- \def\contentsname{Table des mati\`eres}%
- \def\corollaryname{Corollaire}%
- \def\definitionname{D\'efinition}%
- \def\emailname{E-mail}%
- \def\examplename{Exemple}%
- \def\exercisename{Exercice}%
- \def\figurename{Fig.}%
- \def\keywordname{{\bfseries Mots-cl\'e\runinend}}%
- \def\indexname{Index}%
- \def\lemmaname{Lemme}%
- \def\contriblistname{Liste des contributeurs}%
- \def\listfigurename{Liste des figures}%
- \def\listtablename{Liste des tables}%
- \def\mailname{{\itshape Correspondence to\/}:}%
- \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}%
- \def\notename{Remarque}%
- \def\partname{Partie}%
- \def\problemname{Probl\`eme}%
- \def\proofname{Preuve}%
- \def\propertyname{Caract\'eristique}%
-%\def\propositionname{Proposition}%
- \def\questionname{Question}%
- \def\refname{Bibliographie}%
- \def\remarkname{Remarque}%
- \def\seename{voyez}%
- \def\solutionname{Solution}%
-%\def\subclassname{{\it Subject Classifications\/}:}%
- \def\tablename{Tableau}%
- \def\theoremname{Th\'eor\`eme}%
-}
-%
-% German section
-\def\switcht@deutsch{\svlanginfo
-%\typeout{Man spricht deutsch.}%
- \def\abstractname{Zusammenfassung\runinend}%
- \def\ackname{Danksagung\runinend}%
- \def\andname{und}%
- \def\lastandname{ und}%
- \def\appendixname{Anhang}%
- \def\chaptername{Kapitel}%
- \def\claimname{Behauptung}%
- \def\conjecturename{Hypothese}%
- \def\contentsname{Inhaltsverzeichnis}%
- \def\corollaryname{Korollar}%
-%\def\definitionname{Definition}%
- \def\emailname{E-Mail}%
- \def\examplename{Beispiel}%
- \def\exercisename{\"Ubung}%
- \def\figurename{Abb.}%
- \def\keywordname{{\bfseries Schl\"usselw\"orter\runinend}}%
- \def\indexname{Index}%
-%\def\lemmaname{Lemma}%
- \def\contriblistname{Mitarbeiter}%
- \def\listfigurename{Abbildungsverzeichnis}%
- \def\listtablename{Tabellenverzeichnis}%
- \def\mailname{{\itshape Correspondence to\/}:}%
- \def\noteaddname{Nachtrag}%
- \def\notename{Anmerkung}%
- \def\partname{Teil}%
-%\def\problemname{Problem}%
- \def\proofname{Beweis}%
- \def\propertyname{Eigenschaft}%
-%\def\propositionname{Proposition}%
- \def\questionname{Frage}%
- \def\refname{Literatur}%
- \def\remarkname{Anmerkung}%
- \def\seename{siehe}%
- \def\solutionname{L\"osung}%
-%\def\subclassname{{\it Subject Classifications\/}:}%
- \def\tablename{Tabelle}%
-%\def\theoremname{Theorem}%
-}
-\newcommand\today{}
-\edef\today{\ifcase\month\or
- January\or February\or March\or April\or May\or June\or
- July\or August\or September\or October\or November\or December\fi
- \space\number\day, \number\year}
-\setlength\columnsep{1.5cc}
-\setlength\columnseprule{0\p@}
-%
-\frenchspacing
-\clubpenalty=10000
-\widowpenalty=10000
-\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil
-\global\let\@textbottom\relax}}
-\pagestyle{headings}
-\pagenumbering{arabic}
-\if@twocolumn
- \twocolumn
-\fi
-%\if@avier
-% \onecolumn
-% \setlength{\textwidth}{156mm}
-% \setlength{\textheight}{226mm}
-%\fi
-\if@referee
- \makereferee
-\fi
-\flushbottom
-\endinput
-%%
-%% End of file `svjour3.cls'.
--- a/README Sat Dec 17 16:58:11 2011 +0000
+++ b/README Sat Dec 17 17:03:01 2011 +0000
@@ -5,27 +5,5 @@
Subdirectories:
===============
-Nominal ... main files for new Nominal Isabelle
-
-Nominal/Ex ... examples for new implementation
-
-
-
-
-Outher Subdirectories:
-======================
-
-Attic ... old version of the quotient package (is now
- part of the Isabelle distribution)
-
-Literature ... some relevant papers about binders and
- Core-Haskell
-
-Paper ... submitted to ESOP
-
-Pearl ... accepted at ITP
-Pearl-jv ... journal version
-
-Quotient-Paper .. accepted to SAC
-
-Slides ... various talks Christian gave recently
\ No newline at end of file
+Nominal ... main files
+Nominal/Ex ... examples
--- a/Slides/ROOT1.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,6 +0,0 @@
-(* show_question_marks := false; *)
-quick_and_dirty := true;
-
-no_document use_thy "~~/src/HOL/Library/LaTeXsugar";
-
-use_thy "Slides1"
\ No newline at end of file
--- a/Slides/ROOT2.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(* show_question_marks := false; *)
-
-quick_and_dirty := true;
-
-
-no_document use_thy "~~/src/HOL/Library/LaTeXsugar";
-
-use_thy "Slides2"
\ No newline at end of file
--- a/Slides/ROOT3.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,6 +0,0 @@
-(*show_question_marks := false;*)
-quick_and_dirty := true;
-
-no_document use_thy "~~/src/HOL/Library/LaTeXsugar";
-
-use_thy "Slides3"
\ No newline at end of file
--- a/Slides/ROOT4.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,6 +0,0 @@
-(*show_question_marks := false;*)
-quick_and_dirty := true;
-
-no_document use_thy "~~/src/HOL/Library/LaTeXsugar";
-
-use_thy "Slides4"
\ No newline at end of file
--- a/Slides/ROOT5.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,6 +0,0 @@
-(* show_question_marks := false; *)
-quick_and_dirty := true;
-
-no_document use_thy "~~/src/HOL/Library/LaTeXsugar";
-
-use_thy "Slides5"
\ No newline at end of file
--- a/Slides/ROOT6.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,6 +0,0 @@
-(*show_question_marks := false;*)
-quick_and_dirty := true;
-
-no_document use_thy "~~/src/HOL/Library/LaTeXsugar";
-
-use_thy "Slides6"
\ No newline at end of file
--- a/Slides/ROOT7.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,6 +0,0 @@
-(*show_question_marks := false;*)
-quick_and_dirty := true;
-
-no_document use_thy "~~/src/HOL/Library/LaTeXsugar";
-
-use_thy "Slides7"
\ No newline at end of file
--- a/Slides/ROOT8.ML Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,6 +0,0 @@
-(*show_question_marks := false;*)
-quick_and_dirty := true;
-
-no_document use_thy "~~/src/HOL/Library/LaTeXsugar";
-
-use_thy "Slides8"
\ No newline at end of file
--- a/Slides/Slides1.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1122 +0,0 @@
-(*<*)
-theory Slides1
-imports "~~/src/HOL/Library/LaTeXsugar" "Nominal"
-begin
-
-notation (latex output)
- set ("_") and
- Cons ("_::/_" [66,65] 65)
-
-(*>*)
-
-
-text_raw {*
- %%\renewcommand{\slidecaption}{Cambridge, 8.~June 2010}
- \renewcommand{\slidecaption}{Uppsala, 3.~March 2011}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[t]
- \frametitle{%
- \begin{tabular}{@ {\hspace{-3mm}}c@ {}}
- \\
- \huge Nominal Isabelle 2\\[-2mm]
- \large Or, How to Reason Conveniently\\[-5mm]
- \large with General Bindings\\[5mm]
- \end{tabular}}
- \begin{center}
- Christian Urban
- \end{center}
- \begin{center}
- joint work with {\bf Cezary Kaliszyk}\\[0mm]
- \end{center}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>
- \frametitle{\begin{tabular}{c}Binding in Old Nominal\end{tabular}}
- \mbox{}\\[-6mm]
-
- \begin{itemize}
- \item the old Nominal Isabelle provided a reasoning infrastructure for single binders\medskip
-
- \begin{center}
- Lam [a].(Var a)
- \end{center}\bigskip
-
- \item<2-> but representing
-
- \begin{center}
- $\forall\{a_1,\ldots,a_n\}.\; T$
- \end{center}\medskip
-
- with single binders and reasoning about it is a \alert{\bf major} pain;
- take my word for it!
- \end{itemize}
-
- \only<1>{
- \begin{textblock}{6}(1.5,11)
- \small
- for example\\
- \begin{tabular}{l@ {\hspace{2mm}}l}
- & a $\fresh$ Lam [a]. t\\
- & Lam [a]. (Var a) \alert{$=$} Lam [b]. (Var b)\\
- & Barendregt-style reasoning about bound variables\\
- \end{tabular}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-4>
- \frametitle{\begin{tabular}{c}Binding Sets of Names\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item binding sets of names has some interesting properties:\medskip
-
- \begin{center}
- \begin{tabular}{l}
- $\forall\{x, y\}.\, x \rightarrow y \;\;\approx_\alpha\;\; \forall\{y, x\}.\, y \rightarrow x$
- \bigskip\smallskip\\
-
- \onslide<2->{%
- $\forall\{x, y\}.\, x \rightarrow y \;\;\not\approx_\alpha\;\; \forall\{z\}.\, z \rightarrow z$
- }\bigskip\smallskip\\
-
- \onslide<3->{%
- $\forall\{x\}.\, x \rightarrow y \;\;\approx_\alpha\;\; \forall\{x, \alert{z}\}.\, x \rightarrow y$
- }\medskip\\
- \onslide<3->{\hspace{4cm}\small provided $z$ is fresh for the type}
- \end{tabular}
- \end{center}
- \end{itemize}
-
- \begin{textblock}{8}(2,14.5)
- \footnotesize $^*$ $x$, $y$, $z$ are assumed to be distinct
- \end{textblock}
-
- \only<4>{
- \begin{textblock}{6}(2.5,4)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=3mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\normalsize\color{darkgray}
- \begin{minipage}{8cm}\raggedright
- For type-schemes the order of bound names does not matter, and
- alpha-equivalence is preserved under \alert{vacuous} binders.
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>
- \frametitle{\begin{tabular}{c}Other Binding Modes\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item alpha-equivalence being preserved under vacuous binders is \underline{not} always
- wanted:\bigskip\bigskip\normalsize
-
- \begin{tabular}{@ {\hspace{-8mm}}l}
- $\text{let}\;x = 3\;\text{and}\;y = 2\;\text{in}\;x - y\;\text{end}$\medskip\\
- \onslide<2->{$\;\;\;\only<2>{\approx_\alpha}\only<3>{\alert{\not\approx_\alpha}}
- \text{let}\;y = 2\;\text{and}\;x = 3\only<3->{\alert{\;\text{and}
- \;z = \text{loop}}}\;\text{in}\;x - y\;\text{end}$}
- \end{tabular}
-
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>
- \frametitle{\begin{tabular}{c}\LARGE{}Even Another Binding Mode\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item sometimes one wants to abstract more than one name, but the order \underline{does} matter\bigskip
-
- \begin{center}
- \begin{tabular}{@ {\hspace{-8mm}}l}
- $\text{let}\;(x, y) = (3, 2)\;\text{in}\;x - y\;\text{end}$\medskip\\
- $\;\;\;\not\approx_\alpha
- \text{let}\;(y, x) = (3, 2)\;\text{in}\;x - y\;\text{end}$
- \end{tabular}
- \end{center}
-
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>
- \frametitle{\begin{tabular}{c}\LARGE{}Three Binding Modes\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item the order does not matter and alpha-equivelence is preserved under
- vacuous binders \textcolor{gray}{(restriction)}\medskip
-
- \item the order does not matter, but the cardinality of the binders
- must be the same \textcolor{gray}{(abstraction)}\medskip
-
- \item the order does matter \textcolor{gray}{(iterated single binders)}
- \end{itemize}
-
- \onslide<2->{
- \begin{center}
- \isacommand{bind (set+)}\hspace{6mm}
- \isacommand{bind (set)}\hspace{6mm}
- \isacommand{bind}
- \end{center}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>
- \frametitle{\begin{tabular}{c}Specification of Binding\end{tabular}}
- \mbox{}\\[-6mm]
-
- \mbox{}\hspace{10mm}
- \begin{tabular}{ll}
- \multicolumn{2}{l}{\isacommand{nominal\_datatype} trm $=$}\\
- \hspace{5mm}\phantom{$|$} Var name\\
- \hspace{5mm}$|$ App trm trm\\
- \hspace{5mm}$|$ Lam \only<2->{x::}name \only<2->{t::}trm
- & \onslide<2->{\isacommand{bind} x \isacommand{in} t}\\
- \hspace{5mm}$|$ Let \only<2->{as::}assn \only<2->{t::}trm
- & \onslide<2->{\isacommand{bind} bn(as) \isacommand{in} t}\\
- \multicolumn{2}{l}{\isacommand{and} assn $=$}\\
- \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} ANil}\\
- \multicolumn{2}{l}{\hspace{5mm}$|$ ACons name trm assn}\\
- \multicolumn{2}{l}{\onslide<3->{\isacommand{binder} bn \isacommand{where}}}\\
- \multicolumn{2}{l}{\onslide<3->{\hspace{5mm}\phantom{$|$} bn(ANil) $=$ []}}\\
- \multicolumn{2}{l}{\onslide<3->{\hspace{5mm}$|$ bn(ACons a t as) $=$ [a] @ bn(as)}}\\
- \end{tabular}
-
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-5>
- \frametitle{\begin{tabular}{c}Inspiration from Ott\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item this way of specifying binding is inspired by
- {\bf Ott}\onslide<2->{, \alert{\bf but} we made some adjustments:}\medskip
-
-
- \only<2>{
- \begin{itemize}
- \item Ott allows specifications like\smallskip
- \begin{center}
- $t ::= t\;t\; |\;\lambda x.t$
- \end{center}
- \end{itemize}}
-
- \only<3-4>{
- \begin{itemize}
- \item whether something is bound can depend in Ott on other bound things\smallskip
- \begin{center}
- \begin{tikzpicture}
- \node (A) at (-0.5,1) {Foo $(\lambda y. \lambda x. t)$};
- \node (B) at ( 1.1,1) {$s$};
- \onslide<4>{\node (C) at (0.5,0) {$\{y, x\}$};}
- \onslide<4>{\draw[->,red,line width=1mm] (A) -- (C);}
- \onslide<4>{\draw[->,red,line width=1mm] (C) -- (B);}
- \end{tikzpicture}
- \end{center}
- \onslide<4>{this might make sense for ``raw'' terms, but not at all
- for $\alpha$-equated terms}
- \end{itemize}}
-
- \only<5>{
- \begin{itemize}
- \item we allow multiple ``binders'' and ``bodies''\smallskip
- \begin{center}
- \begin{tabular}{l}
- \isacommand{bind} a b c \ldots \isacommand{in} x y z \ldots\\
- \isacommand{bind (set)} a b c \ldots \isacommand{in} x y z \ldots\\
- \isacommand{bind (set+)} a b c \ldots \isacommand{in} x y z \ldots
- \end{tabular}
- \end{center}\bigskip\medskip
- the reason is that with our definition of $\alpha$-equivalence\medskip
- \begin{center}
- \begin{tabular}{l}
- \isacommand{bind (set+)} as \isacommand{in} x y $\not\Leftrightarrow$\\
- \hspace{8mm}\isacommand{bind (set+)} as \isacommand{in} x, \isacommand{bind (set+)} as \isacommand{in} y
- \end{tabular}
- \end{center}\medskip
-
- same with \isacommand{bind (set)}
- \end{itemize}}
- \end{itemize}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>
- \frametitle{\begin{tabular}{c}Alpha-Equivalence\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item in the old Nominal Isabelle, we represented single binders as partial functions:\bigskip
-
- \begin{center}
- \begin{tabular}{l}
- Lam [$a$].\,$t$ $\;{^\text{``}}\!\dn{}\!^{\text{''}}$\\[2mm]
- \;\;\;\;$\lambda b.$\;$\text{if}\;a = b\;\text{then}\;t\;\text{else}$\\
- \phantom{\;\;\;\;$\lambda b.$\;\;\;\;\;\;}$\text{if}\;b \fresh t\;
- \text{then}\;(a\;b)\act t\;\text{else}\;\text{error}$
- \end{tabular}
- \end{center}
- \end{itemize}
-
- \begin{textblock}{10}(2,14)
- \footnotesize $^*$ alpha-equality coincides with equality on functions
- \end{textblock}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{\begin{tabular}{c}New Design\end{tabular}}
- \mbox{}\\[4mm]
-
- \begin{center}
- \begin{tikzpicture}
- {\draw (0,0) node[inner sep=3mm, ultra thick, draw=fg, rounded corners=2mm]
- (A) {\begin{minipage}{1.1cm}bind.\\spec.\end{minipage}};}
-
- {\draw (3,0) node[inner sep=3mm, ultra thick, draw=fg, rounded corners=2mm]
- (B) {\begin{minipage}{1.1cm}raw\\terms\end{minipage}};}
-
- \alt<2>
- {\draw (6,0) node[inner sep=3mm, ultra thick, draw=red, rounded corners=2mm]
- (C) {\textcolor{red}{\begin{minipage}{1.1cm}$\alpha$-\\equiv.\end{minipage}}};}
- {\draw (6,0) node[inner sep=3mm, ultra thick, draw=fg, rounded corners=2mm]
- (C) {\begin{minipage}{1.1cm}$\alpha$-\\equiv.\end{minipage}};}
-
- {\draw (0,-3) node[inner sep=3mm, ultra thick, draw=fg, rounded corners=2mm]
- (D) {\begin{minipage}{1.1cm}quot.\\type\end{minipage}};}
-
- {\draw (3,-3) node[inner sep=3mm, ultra thick, draw=fg, rounded corners=2mm]
- (E) {\begin{minipage}{1.1cm}lift\\thms\end{minipage}};}
-
- {\draw (6,-3) node[inner sep=3mm, ultra thick, draw=fg, rounded corners=2mm]
- (F) {\begin{minipage}{1.1cm}add.\\thms\end{minipage}};}
-
- \draw[->,fg!50,line width=1mm] (A) -- (B);
- \draw[->,fg!50,line width=1mm] (B) -- (C);
- \draw[->,fg!50,line width=1mm, line join=round, rounded corners=2mm]
- (C) -- (8,0) -- (8,-1.5) -- (-2,-1.5) -- (-2,-3) -- (D);
- \draw[->,fg!50,line width=1mm] (D) -- (E);
- \draw[->,fg!50,line width=1mm] (E) -- (F);
- \end{tikzpicture}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-8>
- \frametitle{\begin{tabular}{c}Alpha-Equivalence\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item lets first look at pairs\bigskip\medskip
-
- \begin{tabular}{@ {\hspace{1cm}}l}
- $(as, x) \onslide<2->{\approx\!}\makebox[0mm][l]{\only<2-6>{${}_{\text{set}}$}%
- \only<7>{${}_{\text{\alert{list}}}$}%
- \only<8>{${}_{\text{\alert{set+}}}$}}%
- \onslide<3->{^{R,\text{fv}}}\,\onslide<2->{(bs,y)}$
- \end{tabular}\bigskip
- \end{itemize}
-
- \only<1>{
- \begin{textblock}{8}(3,8.5)
- \begin{tabular}{l@ {\hspace{2mm}}p{8cm}}
- & $as$ is a set of names\ldots the binders\\
- & $x$ is the body (might be a tuple)\\
- & $\approx_{\text{set}}$ is where the cardinality
- of the binders has to be the same\\
- \end{tabular}
- \end{textblock}}
-
- \only<4->{
- \begin{textblock}{12}(5,8)
- \begin{tabular}{ll@ {\hspace{1mm}}l}
- $\dn$ & \onslide<5->{$\exists \pi.\,$} & $\text{fv}(x) - as = \text{fv}(y) - bs$\\[1mm]
- & \onslide<5->{$\;\;\;\wedge$} & \onslide<5->{$\text{fv}(x) - as \fresh^* \pi$}\\[1mm]
- & \onslide<5->{$\;\;\;\wedge$} & \onslide<5->{$(\pi \act x)\;R\;y$}\\[1mm]
- & \onslide<6-7>{$\;\;\;\wedge$} & \onslide<6-7>{$\pi \act as = bs$}\\
- \end{tabular}
- \end{textblock}}
-
- \only<7>{
- \begin{textblock}{7}(3,13.8)
- \footnotesize $^*$ $as$ and $bs$ are \alert{lists} of names
- \end{textblock}}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>
- \frametitle{\begin{tabular}{c}Examples\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item lets look at ``type-schemes'':\medskip\medskip
-
- \begin{center}
- $(as, x) \approx\!\makebox[0mm][l]{${}_{\text{set}}$}\only<1>{{}^{R,\text{fv}}}\only<2->{{}^{\alert{=},\alert{\text{fv}}}} (bs, y)$
- \end{center}\medskip
-
- \onslide<2->{
- \begin{center}
- \begin{tabular}{l}
- $\text{fv}(x) = \{x\}$\\[1mm]
- $\text{fv}(T_1 \rightarrow T_2) = \text{fv}(T_1) \cup \text{fv}(T_2)$\\
- \end{tabular}
- \end{center}}
- \end{itemize}
-
-
- \only<3->{
- \begin{textblock}{4}(0.3,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{set+:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- \\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
- \only<3->{
- \begin{textblock}{4}(5.2,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{set:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- $\wedge$ & $\pi \cdot as = bs$\\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
- \only<3->{
- \begin{textblock}{4}(10.2,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{list:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- $\wedge$ & $\pi \cdot as = bs$\\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>
- \frametitle{\begin{tabular}{c}Examples\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{center}
- \only<1>{$(\{x, y\}, x \rightarrow y) \approx_? (\{x, y\}, y \rightarrow x)$}
- \only<2>{$([x, y], x \rightarrow y) \approx_? ([x, y], y \rightarrow x)$}
- \end{center}
-
- \begin{itemize}
- \item $\approx_{\text{set+}}$, $\approx_{\text{set}}$%
- \only<2>{, \alert{$\not\approx_{\text{list}}$}}
- \end{itemize}
-
-
- \only<1->{
- \begin{textblock}{4}(0.3,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{set+:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- \\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
- \only<1->{
- \begin{textblock}{4}(5.2,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{set:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- $\wedge$ & $\pi \cdot as = bs$\\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
- \only<1->{
- \begin{textblock}{4}(10.2,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{list:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- $\wedge$ & $\pi \cdot as = bs$\\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>
- \frametitle{\begin{tabular}{c}Examples\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{center}
- \only<1>{$(\{x\}, x) \approx_? (\{x, y\}, x)$}
- \end{center}
-
- \begin{itemize}
- \item $\approx_{\text{set+}}$, $\not\approx_{\text{set}}$,
- $\not\approx_{\text{list}}$
- \end{itemize}
-
-
- \only<1->{
- \begin{textblock}{4}(0.3,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{set+:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- \\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
- \only<1->{
- \begin{textblock}{4}(5.2,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{set:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- $\wedge$ & $\pi \cdot as = bs$\\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
- \only<1->{
- \begin{textblock}{4}(10.2,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{list:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- $\wedge$ & $\pi \cdot as = bs$\\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \only<2>{
- \begin{textblock}{6}(2.5,4)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=5mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\normalsize
- \begin{minipage}{8cm}\raggedright
- \begin{itemize}
- \item \color{darkgray}$\alpha$-equivalences coincide when a single name is
- abstracted
- \item \color{darkgray}in that case they are equivalent to ``old-fashioned'' definitions of $\alpha$
- \end{itemize}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>
- \frametitle{\begin{tabular}{c}General Abstractions\end{tabular}}
- \mbox{}\\[-7mm]
-
- \begin{itemize}
- \item we take $(as, x) \approx\!\makebox[0mm][l]{${}_{{}*{}}$}^{=,\text{supp}} (bs, y)$\medskip
- \item they are equivalence relations\medskip
- \item we can therefore use the quotient package to introduce the
- types $\beta\;\text{abs}_*$\bigskip
- \begin{center}
- \only<1>{$[as].\,x$}
- \only<2>{$\text{supp}([as].x) = \text{supp}(x) - as$}
- \only<3>{%
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {\hspace{-7mm}}l}{$[as]. x \alert{=} [bs].y\;\;\;\text{if\!f}$}\\[2mm]
- $\exists \pi.$ & $\text{supp}(x) - as = \text{supp}(y) - bs$\\
- $\wedge$ & $\text{supp}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \act x = y $\\
- $(\wedge$ & $\pi \act as = bs)\;^*$\\
- \end{tabular}}
- \end{center}
- \end{itemize}
-
- \only<1->{
- \begin{textblock}{8}(12,3.8)
- \footnotesize $^*$ set, set+, list
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>
- \frametitle{\begin{tabular}{c}A Problem\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{center}
- $\text{let}\;x_1=t_1 \ldots x_n=t_n\;\text{in}\;s$
- \end{center}
-
- \begin{itemize}
- \item we cannot represent this as\medskip
- \begin{center}
- $\text{let}\;[x_1,\ldots,x_n]\alert{.}s\;\;[t_1,\ldots,t_n]$
- \end{center}\bigskip
-
- because\medskip
- \begin{center}
- $\text{let}\;[x].s\;\;[t_1,t_2]$
- \end{center}
- \end{itemize}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{\begin{tabular}{c}Our Specifications\end{tabular}}
- \mbox{}\\[-6mm]
-
- \mbox{}\hspace{10mm}
- \begin{tabular}{ll}
- \multicolumn{2}{l}{\isacommand{nominal\_datatype} trm $=$}\\
- \hspace{5mm}\phantom{$|$} Var name\\
- \hspace{5mm}$|$ App trm trm\\
- \hspace{5mm}$|$ Lam x::name t::trm
- & \isacommand{bind} x \isacommand{in} t\\
- \hspace{5mm}$|$ Let as::assn t::trm
- & \isacommand{bind} bn(as) \isacommand{in} t\\
- \multicolumn{2}{l}{\isacommand{and} assn $=$}\\
- \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} ANil}\\
- \multicolumn{2}{l}{\hspace{5mm}$|$ ACons name trm assn}\\
- \multicolumn{2}{l}{\isacommand{binder} bn \isacommand{where}}\\
- \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} bn(ANil) $=$ $[]$}\\
- \multicolumn{2}{l}{\hspace{5mm}$|$ bn(ACons a t as) $=$ $[$a$]$ @ bn(as)}\\
- \end{tabular}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>
- \frametitle{\begin{tabular}{c}``Raw'' Definitions\end{tabular}}
- \mbox{}\\[-6mm]
-
- \mbox{}\hspace{10mm}
- \begin{tabular}{ll}
- \multicolumn{2}{l}{\isacommand{datatype} trm $=$}\\
- \hspace{5mm}\phantom{$|$} Var name\\
- \hspace{5mm}$|$ App trm trm\\
- \hspace{5mm}$|$ Lam name trm\\
- \hspace{5mm}$|$ Let assn trm\\
- \multicolumn{2}{l}{\isacommand{and} assn $=$}\\
- \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} ANil}\\
- \multicolumn{2}{l}{\hspace{5mm}$|$ ACons name trm assn}\\[5mm]
- \multicolumn{2}{l}{\isacommand{function} bn \isacommand{where}}\\
- \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} bn(ANil) $=$ $[]$}\\
- \multicolumn{2}{l}{\hspace{5mm}$|$ bn(ACons a t as) $=$ $[$a$]$ @ bn(as)}\\
- \end{tabular}
-
- \only<2>{
- \begin{textblock}{5}(10,5)
- $+$ \begin{tabular}{l}automatically\\
- generate fv's\end{tabular}
- \end{textblock}}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>
- \frametitle{\begin{tabular}{c}\LARGE``Raw'' Alpha-Equivalence\end{tabular}}
- \mbox{}\\[6mm]
-
- \begin{center}
- Lam x::name t::trm \hspace{10mm}\isacommand{bind} x \isacommand{in} t\\
- \end{center}
-
-
- \[
- \infer[\text{Lam-}\!\approx_\alpha]
- {\text{Lam}\;x\;t \approx_\alpha \text{Lam}\;x'\;t'}
- {([x], t) \approx\!\makebox[0mm][l]{${}_{\text{list}}$}
- ^{\approx_\alpha,\text{fv}} ([x'], t')}
- \]
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>
- \frametitle{\begin{tabular}{c}\LARGE``Raw'' Alpha-Equivalence\end{tabular}}
- \mbox{}\\[6mm]
-
- \begin{center}
- Lam x::name y::name t::trm s::trm \hspace{5mm}\isacommand{bind} x y \isacommand{in} t s\\
- \end{center}
-
-
- \[
- \infer[\text{Lam-}\!\approx_\alpha]
- {\text{Lam}\;x\;y\;t\;s \approx_\alpha \text{Lam}\;x'\;y'\;t'\;s'}
- {([x, y], (t, s)) \approx\!\makebox[0mm][l]{${}_{\text{list}}$}
- ^{R, fv} ([x', y'], (t', s'))}
- \]
-
- \footnotesize
- where $R =\;\approx_\alpha\times\approx_\alpha$ and $fv = \text{fv}\cup\text{fv}$
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>
- \frametitle{\begin{tabular}{c}\LARGE``Raw'' Alpha-Equivalence\end{tabular}}
- \mbox{}\\[6mm]
-
- \begin{center}
- Let as::assn t::trm \hspace{10mm}\isacommand{bind} bn(as) \isacommand{in} t\\
- \end{center}
-
-
- \[
- \infer[\text{Let-}\!\approx_\alpha]
- {\text{Let}\;as\;t \approx_\alpha \text{Let}\;as'\;t'}
- {(\text{bn}(as), t) \approx\!\makebox[0mm][l]{${}_{\text{list}}$}
- ^{\approx_\alpha,\text{fv}} (\text{bn}(as'), t') &
- \onslide<2->{as \approx_\alpha^{\text{bn}} as'}}
- \]\bigskip
-
-
- \onslide<1->{\small{}bn-function $\Rightarrow$ \alert{deep binders}}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{\begin{tabular}{c}\LARGE{}$\alpha$ for Binding Functions\end{tabular}}
- \mbox{}\\[-6mm]
-
- \mbox{}\hspace{10mm}
- \begin{tabular}{l}
- \ldots\\
- \isacommand{binder} bn \isacommand{where}\\
- \phantom{$|$} bn(ANil) $=$ $[]$\\
- $|$ bn(ACons a t as) $=$ $[$a$]$ @ bn(as)\\
- \end{tabular}\bigskip
-
- \begin{center}
- \mbox{\infer{\text{ANil} \approx_\alpha^{\text{bn}} \text{ANil}}{}}\bigskip
-
- \mbox{\infer{\text{ACons}\;a\;t\;as \approx_\alpha^{\text{bn}} \text{ACons}\;a'\;t'\;as'}
- {t \approx_\alpha t' & as \approx_\alpha^{bn} as'}}
- \end{center}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>
- \frametitle{\begin{tabular}{c}\LARGE``Raw'' Alpha-Equivalence\end{tabular}}
- \mbox{}\\[6mm]
-
- \begin{center}
- LetRec as::assn t::trm \hspace{10mm}\isacommand{bind} bn(as) \isacommand{in} t \alert{as}\\
- \end{center}
-
-
- \[\mbox{}\hspace{-4mm}
- \infer[\text{LetRec-}\!\approx_\alpha]
- {\text{LetRec}\;as\;t \approx_\alpha \text{LetRec}\;as'\;t'}
- {(\text{bn}(as), (t, as)) \approx\!\makebox[0mm][l]{${}_{\text{list}}$}
- ^{R,\text{fv}} (\text{bn}(as'), (t', as'))}
- \]\bigskip
-
- \onslide<1->{\alert{deep recursive binders}}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{\begin{tabular}{c}Restrictions\end{tabular}}
- \mbox{}\\[-6mm]
-
- Our restrictions on binding specifications:
-
- \begin{itemize}
- \item a body can only occur once in a list of binding clauses\medskip
- \item you can only have one binding function for a deep binder\medskip
- \item binding functions can return: the empty set, singletons, unions (similarly for lists)
- \end{itemize}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{\begin{tabular}{c}Automatic Proofs\end{tabular}}
- \mbox{}\\[-6mm]
-
- \begin{itemize}
- \item we can show that $\alpha$'s are equivalence relations\medskip
- \item as a result we can use our quotient package to introduce the type(s)
- of $\alpha$-equated terms
-
- \[
- \infer
- {\text{Lam}\;x\;t \alert{=} \text{Lam}\;x'\;t'}
- {\only<1>{([x], t) \approx\!\makebox[0mm][l]{${}_{\text{list}}$}
- ^{=,\text{supp}} ([x'], t')}%
- \only<2>{[x].t = [x'].t'}}
- \]
-
-
- \item the properties for support are implied by the properties of $[\_].\_$
- \item we can derive strong induction principles
- \end{itemize}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[t]
- \frametitle{\begin{tabular}{c}Runtime is Acceptable\end{tabular}}
- \mbox{}\\[-7mm]\mbox{}
-
- \footnotesize
- \begin{center}
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=2mm, ultra thick, draw=fg, rounded corners=2mm]
- (A) {\begin{minipage}{0.8cm}bind.\\spec.\end{minipage}};
-
- \draw (2,0) node[inner sep=2mm, ultra thick, draw=fg, rounded corners=2mm]
- (B) {\begin{minipage}{0.8cm}raw\\terms\end{minipage}};
-
- \draw (4,0) node[inner sep=2mm, ultra thick, draw=fg, rounded corners=2mm]
- (C) {\begin{minipage}{0.8cm}$\alpha$-\\equiv.\end{minipage}};
-
- \draw (0,-2) node[inner sep=2mm, ultra thick, draw=fg, rounded corners=2mm]
- (D) {\begin{minipage}{0.8cm}quot.\\type\end{minipage}};
-
- \draw (2,-2) node[inner sep=2mm, ultra thick, draw=fg, rounded corners=2mm]
- (E) {\begin{minipage}{0.8cm}lift\\thms\end{minipage}};
-
- \draw (4,-2) node[inner sep=2mm, ultra thick, draw=fg, rounded corners=2mm]
- (F) {\begin{minipage}{0.8cm}add.\\thms\end{minipage}};
-
- \draw[->,fg!50,line width=1mm] (A) -- (B);
- \draw[->,fg!50,line width=1mm] (B) -- (C);
- \draw[->,fg!50,line width=1mm, line join=round, rounded corners=2mm]
- (C) -- (5,0) -- (5,-1) -- (-1,-1) -- (-1,-2) -- (D);
- \draw[->,fg!50,line width=1mm] (D) -- (E);
- \draw[->,fg!50,line width=1mm] (E) -- (F);
- \end{tikzpicture}
- \end{center}
-
- \begin{itemize}
- \item Core Haskell: 11 types, 49 term-constructors, 7 binding functions
- \begin{center}
- $\sim$ 2 mins
- \end{center}
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{\begin{tabular}{c}Interesting Phenomenon\end{tabular}}
- \mbox{}\\[-6mm]
-
- \small
- \mbox{}\hspace{20mm}
- \begin{tabular}{ll}
- \multicolumn{2}{l}{\isacommand{nominal\_datatype} trm $=$}\\
- \hspace{5mm}\phantom{$|$} Var name\\
- \hspace{5mm}$|$ App trm trm\\
- \hspace{5mm}$|$ Lam x::name t::trm
- & \isacommand{bind} x \isacommand{in} t\\
- \hspace{5mm}$|$ Let as::assn t::trm
- & \isacommand{bind} bn(as) \isacommand{in} t\\
- \multicolumn{2}{l}{\isacommand{and} assn $=$}\\
- \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} ANil}\\
- \multicolumn{2}{l}{\hspace{5mm}$|$ ACons name trm assn}\\
- \multicolumn{2}{l}{\isacommand{binder} bn \isacommand{where}}\\
- \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} bn(ANil) $=$ $[]$}\\
- \multicolumn{2}{l}{\hspace{5mm}$|$ bn(ACons a t as) $=$ $[$a$]$ @ bn(as)}\\
- \end{tabular}\bigskip\medskip
-
- we cannot quotient assn: ACons a \ldots $\not\approx_\alpha$ ACons b \ldots
-
- \only<1->{
- \begin{textblock}{8}(0.2,7.3)
- \alert{\begin{tabular}{p{2.6cm}}
- \raggedright\footnotesize{}Should a ``naked'' assn be quotient?
- \end{tabular}\hspace{-3mm}
- $\begin{cases}
- \mbox{} \\ \mbox{}
- \end{cases}$}
- \end{textblock}}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{\begin{tabular}{c}Conclusion\end{tabular}}
- \mbox{}\\[-6mm]
-
- \begin{itemize}
- \item the user does not see anything of the raw level\medskip
- \only<1>{\begin{center}
- Lam a (Var a) \alert{$=$} Lam b (Var b)
- \end{center}\bigskip}
-
- \item<2-> we have not yet done function definitions (will come soon and
- we hope to make improvements over the old way there too)\medskip
- \item<3-> it took quite some time to get here, but it seems worthwhile
- (Barendregt's variable convention is unsound in general,
- found bugs in two paper proofs, quotient package, POPL 2011 tutorial)\medskip
- \end{itemize}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{\begin{tabular}{c}Future Work\end{tabular}}
- \mbox{}\\[-6mm]
-
- \begin{itemize}
- \item Function definitions
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{\begin{tabular}{c}Questions?\end{tabular}}
- \mbox{}\\[-6mm]
-
- \begin{center}
- \alert{\huge{Thanks!}}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>[c]
- \frametitle{\begin{tabular}{c}Examples\end{tabular}}
- \mbox{}\\[-6mm]
-
- \begin{center}
- $(\{a,b\}, a \rightarrow b) \approx_\alpha (\{a, b\}, a \rightarrow b)$
- $(\{a,b\}, a \rightarrow b) \approx_\alpha (\{a, b\}, b \rightarrow a)$
- \end{center}
-
- \begin{center}
- $(\{a,b\}, (a \rightarrow b, a \rightarrow b))$\\
- \hspace{17mm}$\not\approx_\alpha (\{a, b\}, (a \rightarrow b, b \rightarrow a))$
- \end{center}
-
- \onslide<2->
- {1.) \hspace{3mm}\isacommand{bind (set)} as \isacommand{in} $\tau_1$,
- \isacommand{bind (set)} as \isacommand{in} $\tau_2$\medskip
-
- 2.) \hspace{3mm}\isacommand{bind (set)} as \isacommand{in} $\tau_1$ $\tau_2$
- }
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-(*<*)
-end
-(*>*)
\ No newline at end of file
--- a/Slides/Slides2.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,449 +0,0 @@
-(*<*)
-theory Slides2
-imports "~~/src/HOL/Library/LaTeXsugar" "Nominal"
-begin
-
-notation (latex output)
- set ("_") and
- Cons ("_::/_" [66,65] 65)
-
-(*>*)
-
-
-text_raw {*
- \renewcommand{\slidecaption}{Edinburgh, 11.~July 2010}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[t]
- \frametitle{%
- \begin{tabular}{@ {\hspace{-3mm}}c@ {}}
- \\
- \LARGE Proof Pearl:\\[-0mm]
- \LARGE A New Foundation for\\[-2mm]
- \LARGE Nominal Isabelle\\[12mm]
- \end{tabular}}
- \begin{center}
- Brian Huf\!fman and {\bf Christian Urban}\\[0mm]
- \end{center}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>[c]
- \frametitle{Nominal Isabelle}
-
- \begin{itemize}
- \item \ldots is a definitional extension of Isabelle/HOL
- (let-polymorphism and type classes)\medskip
-
- \item \ldots provides a convenient reasoning infrastructure for
- terms involving binders (e.g.~lambda calculus, variable convention)\medskip
-
- \item<2-> \ldots mainly used to find errors in my own
- (published) paper proofs and in those of others \alert{\bf ;o)}
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>[c]
- \frametitle{Nominal Theory}
-
- \ldots by Pitts; at its core are:\bigskip
-
- \begin{itemize}
- \item sorted atoms and
- \item sort-respecting permutations
- \end{itemize}
-
- \onslide<2->{
- \begin{textblock}{8}(4,11)
- \onslide<3->{$\text{inv\_of\_}\pi \,\act\, ($}\onslide<2->{$\pi \,\act\, x$}%
- \onslide<3->{$) \,=\, x$}
- \end{textblock}}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-4>[t]
- \frametitle{The ``Old Way''}
-
- \begin{itemize}
- \item sorted atoms\\ \alert{$\quad\mapsto$ separate types}\; (``copies'' of nat)\bigskip
- \item sort-respecting permutations\\ \alert{$\quad\mapsto$ lists of pairs of atoms (list swappings)}
- \onslide<2->{
- \begin{center}
- \begin{tabular}{c@ {\hspace{7mm}}c}
- $\text{[]} \,\act\, c = c$ &
- $\swap{a}{b}\!::\!\pi \,\act\, c =
- \begin{cases}
- b & \text{if}\, \pi\act c = a\\
- a & \text{if}\, \pi\act c = b\\
- \pi\act c & \text{otherwise}
- \end{cases}$
- \end{tabular}
- \end{center}}
- \end{itemize}
-
- \only<3>{
- \begin{textblock}{14}(1,12.5)
- \alert{The big benefit:} the type system takes care of the sort-respecting requirement.
- \end{textblock}}
- \only<4>{
- \begin{textblock}{14}(1,12.5)
- \alert{A small benefit:} permutation composition is \alert{list append}
- and permutation inversion is \alert{list reversal}.
- \end{textblock}
- }
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-4>
- \frametitle{Problems}
-
- \begin{itemize}
- \item @{text "_ \<bullet> _ :: \<alpha> perm \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}\bigskip
-
- \item @{text "supp _ :: \<beta> \<Rightarrow> \<alpha> set"}
-
- \begin{center}
- $\text{finite} (\text{supp}\;x)_{\,\alpha_1\,\text{set}}$ \ldots
- $\text{finite} (\text{supp}\;x)_{\,\alpha_n\,\text{set}}$
- \end{center}\bigskip
-
- \item $\forall \pi_{\alpha_1} \ldots \pi_{\alpha_n}\;.\; P$\bigskip
-
- \item type-classes
- \onslide<3>{\small\hspace{5mm}can only have \alert{\bf one} type parameter}
- \begin{itemize}
- \item<2-> $\text{[]}\,\act\, x = x$
- \item<2-> $(\pi_1 @ \pi_2) \,\act \, x = \pi_1 \,\act\,(\pi_2 \,\act\, x)$
- \item<2-> if $\pi_1 \sim \pi_2$ then $\pi_1 \,\act\, x = \pi_2 \,\act\, x$
- \item<2-> if $\pi_1$, $\pi_2$ have dif$\!$f.~type, then
- $\pi_1 \act (\pi_2 \act x) = \pi_2 \act (\pi_1 \act x)$
- \end{itemize}
- \end{itemize}
-
- \only<4->{
- \begin{textblock}{9}(3,7)\begin{tikzpicture}
- \draw (0,0) node[inner sep=3mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\normalsize\color{darkgray}
- \begin{quote}
- \begin{itemize}
- \item \alert{lots} of ML-code
- \item \alert{not} pretty
- \item \alert{not a {\bf proof pearl}} :o(
- \end{itemize}
- \end{quote}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-4>
- \frametitle{A Better Way}
-*}
-datatype atom = Atom string nat
-
-text_raw {*
- \mbox{}\bigskip
- \begin{itemize}
- \item<2-> permutations are (restricted) bijective functions from @{text "atom \<Rightarrow> atom"}
-
- \begin{itemize}
- \item sort-respecting \hspace{5mm}($\,\forall a.\;\text{sort}(\pi a) = \text{sort}(a)$)
- \item finite domain \hspace{5mm}($\text{finite} \{a.\;\pi a \not= a\}$)
- \end{itemize}\medskip
-
- \item<3-> \alert{What about {\bf swappings}?}
- \small
- \[
- \begin{array}{l@ {\hspace{1mm}}l}
- (a\;b) \dn & \text{if}\;\text{sort}(a) = \text{sort}(b)\\
- & \text{then}\;\lambda c. \text{if}\;a = c\;\text{then}\;b\;\text{else}\;
- \text{if}\;b = c\;\text{then}\;a\;\text{else}\;c\\
- & \text{else}\;\only<3>{\raisebox{-5mm}{\textcolor{red}{\huge\bf $?$}}}
- \only<4->{\text{\alert{\bf id}}}
- \end{array}
- \]
- \end{itemize}
-
- \only<2>{
- \begin{textblock}{7}(4,11)
- @{text "_ \<bullet> _ :: perm \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-7>
- \frametitle{A Smoother Nominal Theory}
-
- From there it is essentially plain sailing:\bigskip
-
- \begin{itemize}
- \item<2-> $(a\;b) = (b\;a) \onslide<4->{= (a\;c) + (b\;c) + (a\;c)}$\bigskip
-
- \item<3-> permutations are an instance of Isabelle's\\
- group\_add ($0$, $\pi_1 + \pi_2$, $- \pi$)\bigskip
-
- \item<6-> $\_\;\act\;\_ :: \text{perm} \Rightarrow \alpha \Rightarrow \alpha$\medskip
-
- \begin{itemize}
- \item $0\;\act\;x = x$\\
- \item $(\pi_1 + \pi_2)\;\act\;x = \pi_1\;\act\;(\pi_2\;\act\;x)$
- \end{itemize}\medskip
-
- \onslide<7->{\alert{$\;\mapsto\;$}only one type class needed, $\text{finite}(\text{supp}\;x)$,\\
- \hspace{9mm}$\forall \pi. P$}
- \end{itemize}
-
- \only<5>{
- \begin{textblock}{6}(2.5,11)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=3mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\normalsize\color{darkgray}
- \begin{minipage}{8cm}\raggedright
- This is slightly odd, since in general:
- \begin{center}$\pi_1 + \pi_2 \alert{\not=} \pi_2 + \pi_1$\end{center}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{One Snatch}
-*}
-
-datatype \<iota>atom = \<iota>Atom string nat
-
-
-text_raw {*
- \isanewline\isanewline
- \begin{itemize}
- \item You like to get the advantages of the old way back: you
- \alert{cannot mix} atoms of dif$\!$ferent sort:\bigskip
-
- \small
- e.g.~LF-objects:
- \begin{center}
- $\quad M ::= c \mid x \mid \lambda x\!:\!A. M \mid M_1\;M_2$
- \end{center}
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>
- \frametitle{Our Solution}
-
-
- \begin{itemize}
- \item \underline{concrete} atoms:
- \end{itemize}
-*}
-(*<*)
-consts sort :: "atom \<Rightarrow> string"
-(*>*)
-
-typedef name = "{a :: atom. sort a = ''name''}" (*<*)sorry(*>*)
-typedef ident = "{a :: atom. sort a = ''ident''}" (*<*)sorry(*>*)
-
-text_raw {*
- \mbox{}\bigskip\bigskip
- \begin{itemize}
- \item they are a ``subtype'' of the generic atom type
- \item there is an overloaded function \alert{\bf atom}, which injects concrete
- atoms into generic ones\medskip
- \begin{center}
- \begin{tabular}{l}
- $\text{atom}(a) \fresh x$\\
- $(a \leftrightarrow b) \dn (\text{atom}(a)\;\;\text{atom}(b))$
- \end{tabular}
- \end{center}
- \pause
- One would like to have $a \fresh x$, $\swap{a}{b}$, \ldots
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-4>
- \frametitle{Sorts Reloaded}
-*}
-datatype atom\<iota> = Atom\<iota> string nat
-
-text_raw {*
- \isanewline\isanewline
- \pause
- \alert{Problem}: HOL-binders or Church-style lambda-terms
- \begin{center}
- $\lambda x_\alpha.\, x_\alpha\;x_\beta$
- \end{center}
- \pause
- \isanewline\isanewline
-
-\isacommand{datatype} ty = TVar string $\mid$ ty $\rightarrow$ ty\\
-\isacommand{datatype} var = Var name ty\\
-\pause
-$(x \leftrightarrow y) \,\act\, (x_\alpha, x_\beta) = (y_\alpha, y_\beta)$
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>
- \frametitle{Non-Working Solution}
-
-Instead of\isanewline\isanewline
-*}
- datatype atom\<iota>\<iota> = Atom\<iota>\<iota> string nat
-
-text_raw {*
-\isanewline\isanewline
-have
-\isanewline\isanewline
-*}
-
- datatype 'a atom\<iota>\<iota>\<iota> = Atom\<iota>\<iota>\<iota> 'a nat
-text_raw {*
- \pause
- \isanewline\isanewline
- But then
- \begin{center}
- @{text "_ \<bullet> _ :: \<alpha> perm \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
- \end{center}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-(*<*)
-hide_const sort
-(*>*)
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-4>
- \frametitle{A Working Solution}
-
-*}
- datatype sort = Sort string "sort list"
- datatype atom\<iota>\<iota>\<iota>\<iota> = Atom\<iota>\<iota>\<iota>\<iota> sort nat
-(*<*)
-consts sort_ty :: "nat \<Rightarrow> sort"
-(*>*)
-text_raw {*
- \pause
- \isanewline
-
- {\small\begin{tabular}{rcl}
- @{text "sort_ty (TVar x)"} & @{text "\<equiv>"} & @{text "Sort ''TVar'' [Sort x []]"} \\
- @{text "sort_ty (\<tau>\<^isub>1 \<rightarrow> \<tau>\<^isub>2)"} & @{text "\<equiv>"} & @{text "Sort ''Fun'' [sort_ty \<tau>\<^isub>1, sort_ty \<tau>\<^isub>2]"}\\
- \end{tabular}}
- \pause
- \isanewline\isanewline
- \isacommand{typedef} @{text "var = {a :: atom. sort a \<in> range sort_ty}"}
- \pause
- \isanewline\isanewline
- \small
- \begin{minipage}{12cm}
- @{text "Var x \<tau> \<equiv> \<lceil> Atom (sort_ty \<tau>) x \<rceil>"}\isanewline
-
- @{text "(Var x \<tau> \<leftrightarrow> Var y \<tau>) \<bullet> Var x \<tau> = Var y \<tau>"}\\
- @{text "(Var x \<tau> \<leftrightarrow> Var y \<tau>) \<bullet> Var x \<tau>' = Var x \<tau>'"}\\
- \end{minipage}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-4>[c]
- \frametitle{Conclusion}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item the formalised version of the nominal theory is now much nicer to
- work with (sorts are occasionally explicit, \alert{$\forall \pi. P$})\medskip
-
- \item permutations: ``be as abstract as you can'' (group\_add is a slight oddity)\medskip
-
- \item the crucial insight: allow sort-disrespecting swappings%
- \onslide<2->{ \ldots just define them as the identity}\\%
- \onslide<3->{ (a referee called this a \alert{``hack''})}\medskip
-
- \item<4-> there will be a hands-on tutorial about Nominal Isabelle at \alert{POPL'11}
- in Austin Texas
- \end{itemize}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{
- \begin{tabular}{c}
- \mbox{}\\[23mm]
- \alert{\LARGE Thank you very much}\\
- \alert{\Large Questions?}
- \end{tabular}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-(*<*)
-end
-(*>*)
\ No newline at end of file
--- a/Slides/Slides3.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1243 +0,0 @@
-(*<*)
-theory Slides3
-imports "~~/src/HOL/Library/LaTeXsugar" "Nominal"
-begin
-
-declare [[show_question_marks = false]]
-
-notation (latex output)
- set ("_") and
- Cons ("_::/_" [66,65] 65)
-
-(*>*)
-
-text_raw {*
- \renewcommand{\slidecaption}{UNIF, Edinburgh, 14.~July 2010}
-
- \newcommand{\abst}[2]{#1.#2}% atom-abstraction
- \newcommand{\pair}[2]{\langle #1,#2\rangle} % pairing
- \newcommand{\susp}{{\boldsymbol{\cdot}}}% for suspensions
- \newcommand{\unit}{\langle\rangle}% unit
- \newcommand{\app}[2]{#1\,#2}% application
- \newcommand{\eqprob}{\mathrel{{\approx}?}}
- \newcommand{\freshprob}{\mathrel{\#?}}
- \newcommand{\redu}[1]{\stackrel{#1}{\Longrightarrow}}% reduction
- \newcommand{\id}{\varepsilon}% identity substitution
-
- \pgfdeclareradialshading{smallbluesphere}{\pgfpoint{0.5mm}{0.5mm}}%
- {rgb(0mm)=(0,0,0.9);
- rgb(0.9mm)=(0,0,0.7);
- rgb(1.3mm)=(0,0,0.5);
- rgb(1.4mm)=(1,1,1)}
-
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{Quiz}
-
- Assuming that \smath{a} and \smath{b} are distinct variables,\\
- is it possible to find $\lambda$-terms \smath{M_1} to \smath{M_7}
- that make the following pairs \alert{$\alpha$-equivalent}?
-
- \begin{tabular}{@ {\hspace{14mm}}p{12cm}}
- \begin{itemize}
- \item \smath{\lambda a.\lambda b. (M_1\,b)\;} and
- \smath{\lambda b.\lambda a. (a\,M_1)\;}
-
- \item \smath{\lambda a.\lambda b. (M_2\,b)\;} and
- \smath{\lambda b.\lambda a. (a\,M_3)\;}
-
- \item \smath{\lambda a.\lambda b. (b\,M_4)\;} and
- \smath{\lambda b.\lambda a. (a\,M_5)\;}
-
- \item \smath{\lambda a.\lambda b. (b\,M_6)\;} and
- \smath{\lambda a.\lambda a. (a\,M_7)\;}
- \end{itemize}
- \end{tabular}
-
- If there is one solution for a pair, can you describe all its solutions?
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[t]
- \frametitle{%
- \begin{tabular}{@ {\hspace{-3mm}}c@ {}}
- \\
- \huge Nominal Unification\\[-2mm]
- \Large Hitting a Sweet Spot\\[5mm]
- \end{tabular}}
- \begin{center}
- Christian Urban
- \end{center}
- \begin{center}
- \small initial spark from Roy Dyckhoff in November 2001\\[0mm]
- \small joint work with Andy Pitts and Jamie Gabbay\\[0mm]
- \end{center}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-4>[c]
- \frametitle{One Motivation}
-
- \onslide<2->{Typing implemented in Prolog \textcolor{darkgray}{(from a textbook)}}\bigskip\\
-
- \onslide<3->{\color{darkgray}
- \begin{tabular}{l}
- type (Gamma, var(X), T) :- member (X,T) Gamma.\smallskip\medskip\\
-
- type (Gamma, app(M, N), T') :-\\
- \hspace{3cm}type (Gamma, M, arrow(T, T')),\\
- \hspace{3cm}type (Gamma, N, T).\smallskip\medskip\\
-
- type (Gamma, lam(X, M), arrow(T, T')) :-\\
- \hspace{3cm}type ((X, T)::Gamma, M, T').\smallskip\medskip\\
-
- member X X::Tail.\\
- member X Y::Tail :- member X Tail.\\
- \end{tabular}}
-
- \only<4>{
- \begin{textblock}{6}(2.5,2)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=3mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\color{darkgray}
- \begin{minipage}{8cm}\raggedright
- The problem is that \smath{\lambda x.\lambda x. (x\;x)}
- will have the types
- \begin{center}
- \begin{tabular}{l}
- \smath{T\rightarrow (T\rightarrow S) \rightarrow S} and\\
- \smath{(T\rightarrow S)\rightarrow T \rightarrow S}\\
- \end{tabular}
- \end{center}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{Higher-Order Unification}
-
- State of the art at the time:
-
- \begin{itemize}
- \item Lambda Prolog with full Higher-Order Unification\\
- \textcolor{darkgray}{(no mgus, undecidable, modulo $\alpha\beta$)}\bigskip
- \item Higher-Order Pattern Unification\\
- \textcolor{darkgray}{(has mgus, decidable, some restrictions, modulo $\alpha\beta_0$)}
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-10>[t]
- \frametitle{Underlying Ideas}
-
- \begin{itemize}
- \item<1-> Unification (\alert{only}) up to $\alpha$
-
- \item<2-> Swappings / Permutations
-
- \only<2-5>{
- \begin{center}
- \begin{tabular}{r@ {\hspace{1mm}}l@ {\hspace{12mm}}r@ {\hspace{1mm}}l}
- \\
- \only<2>{\smath{\textcolor{white}{[b\!:=\!a]}}}%
- \only<3>{\smath{[b\!:=\!a]}}%
- \only<4-5>{\smath{\alert{\swap{a}{b}\,\act}}} &
- \onslide<2-5>{\smath{\lambda a.b}} &
-
- \only<2>{\smath{\textcolor{white}{[b\!:=\!a]}}}%
- \only<3>{\smath{[b\!:=\!a]}}%
- \only<4-5>{\smath{\alert{\swap{a}{b}\,\act}}} &
- \onslide<2-5>{\smath{\lambda c.b}}\\
-
- \onslide<3-5>{\smath{=}} & \only<3>{\smath{\lambda a.a}}\only<4-5>{\smath{\lambda b.a}} &
- \onslide<3-5>{\smath{=}} & \only<3>{\smath{\lambda c.a}}\only<4-5>{\smath{\lambda c.a}}\\
- \end{tabular}
- \end{center}\bigskip
-
- \onslide<4-5>{
- \begin{center}
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=0mm,fill=cream, ultra thick, draw=cream]
- {\begin{minipage}{8cm}
- \begin{tabular}{r@ {\hspace{3mm}}l}
- \smath{\swap{a}{b}\act t} $\;\dn$ & \alert{swap} {\bf all} occurrences of\\
- & \smath{b} and \smath{a} in \smath{t}
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{center}}\bigskip
-
- \onslide<5>{
- Unlike for \smath{[b\!:=\!a]\act(-)}, for \smath{\swap{a}{b}\act (-)} we do
- have if \smath{t =_\alpha t'} then \smath{\pi \act t =_\alpha \pi \act t'.}}}
-
- \item<6-> Variables (or holes)\bigskip
-
- \begin{center}
- \onslide<7->{\mbox{}\hspace{-25mm}\smath{\lambda x\hspace{-0.5mm}s .}}
- \onslide<8-9>{\raisebox{-1.7mm}{\huge\smath{(}}}\raisebox{-4mm}{\begin{tikzpicture}
- \fill[blue] (0, 0) circle (5mm);
- \end{tikzpicture}}
- \onslide<8-9>{\smath{y\hspace{-0.5mm}s}{\raisebox{-1.7mm}{\huge\smath{)}}}}\bigskip
- \end{center}
-
- \only<8-9>{\smath{y\hspace{-0.5mm}s} are the parameters the hole can depend on\onslide<9->{, but
- then you need $\beta_0$-reduction\medskip
- \begin{center}
- \smath{(\lambda x. t) y \longrightarrow_{\beta_0} t[x:=y]}
- \end{center}}}
-
- \only<10>{we will record the information about which parameters a hole
- \alert{\bf cannot} depend on}
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-4>[c]
- \frametitle{Terms}
-
- \begin{tabular}{lll @ {\hspace{10mm}}lll}
-
- \onslide<1->{\pgfuseshading{smallbluesphere}} &
- \onslide<1->{\colorbox{cream}{\smath{\unit}}} &
- \onslide<1->{Units} &
-
- \onslide<2->{\pgfuseshading{smallbluesphere}} &
- \onslide<2->{\colorbox{cream}{\smath{a}}} &
- \onslide<2->{Atoms} \\[5mm]
-
- \onslide<1->{\pgfuseshading{smallbluesphere}} &
- \onslide<1->{\colorbox{cream}{\smath{\pair{t}{t'}}}} &
- \onslide<1->{Pairs} &
-
- \onslide<3->{\pgfuseshading{smallbluesphere}} &
- \onslide<3->{\colorbox{cream}{\smath{\abst{a}{t}}}} &
- \onslide<3->{Abstractions}\\[5mm]
-
- \onslide<1->{\pgfuseshading{smallbluesphere}} &
- \onslide<1->{\colorbox{cream}{\smath{\app{F}{t}}}} &
- \onslide<1->{Funct.} &
-
- \onslide<4->{\pgfuseshading{smallbluesphere}} &
- \onslide<4->{\colorbox{cream}{\smath{\pi\susp X}}} &
- \onslide<4->{Suspensions}
- \end{tabular}
-
- \only<2>{
- \begin{textblock}{13}(1.5,12)
- \small Atoms are constants \textcolor{darkgray}{(infinitely many of them)}
- \end{textblock}}
-
- \only<3>{
- \begin{textblock}{13}(1.5,12)
- \small \smath{\ulcorner \lambda\abst{a}{a}\urcorner \mapsto \text{fn\ }\abst{a}{a}}\\
- \small constructions like \smath{\text{fn\ }\abst{X}{X}} are not allowed
- \end{textblock}}
-
- \only<4>{
- \begin{textblock}{13}(1.5,12)
- \small \smath{X} is a variable standing for a term\\
- \small \smath{\pi} is an explicit permutation \smath{\swap{a_1}{b_1}\ldots\swap{a_n}{b_n}},
- waiting to be applied to the term that is substituted for \smath{X}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>[c]
- \frametitle{Permutations}
-
- a permutation applied to a term
-
- \begin{center}
- \begin{tabular}{lrcl}
- \pgfuseshading{smallbluesphere} &
- \smath{[]\act c} & \smath{\dn} & \smath{c} \\
-
- \pgfuseshading{smallbluesphere} &
- \smath{\swap{a}{b}\!::\!\pi\act c} & \smath{\dn} &
- \smath{\begin{cases}
- a & \text{if}\;\pi\act c = b\\
- b & \text{if}\;\pi\act c = a\\
- \pi\act c & \text{otherwise}
- \end{cases}}\\
-
- \onslide<2->{\pgfuseshading{smallbluesphere}} &
- \onslide<2->{\smath{\pi\act\abst{a}{t}}} & \onslide<2->{\smath{\dn}} &
- \onslide<2->{\smath{\abst{\pi\act a}{\pi\act t}}}\\
-
- \onslide<3->{\pgfuseshading{smallbluesphere}} &
- \onslide<3->{\smath{\pi\act\pi'\act X}} & \onslide<3->{\smath{\dn}} &
- \onslide<3->{\smath{(\pi @ \pi')\act X}}\\
- \end{tabular}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>[c]
- \frametitle{Freshness Constraints}
-
- Recall \smath{\lambda a. \raisebox{-0.7mm}{\tikz \fill[blue] (0, 0) circle (2.5mm);}}
- \bigskip\pause
-
- We therefore will identify
-
- \begin{center}
- \smath{\text{fn\ } a. X \;\approx\; \text{fn\ } b. \alert<3->{\swap{a}{b}}\act X}
- \end{center}
-
- provided that `\smath{b} is fresh for \smath{X} --- (\smath{b\fresh X})',
- i.e., does not occur freely in any ground term that might be substituted for
- \smath{X}.\bigskip\pause
-
- If we know more about \smath{X}, e.g., if we knew that \smath{a\fresh X} and
- \smath{b\fresh X}, then we can replace\\ \smath{\swap{a}{b}\act X} by
- \smath{X}.
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-4>[c]
- \frametitle{Equivalence Judgements}
-
- \alt<1>{Our equality is {\bf not} just}{but judgements}
-
- \begin{center}
- \begin{tabular}{rl}
- \colorbox{cream}{\smath{\onslide<2->{\nabla \vdash} t \approx t'}} & \alert{$\alpha$-equivalence}\\[1mm]
- \onslide<4->{\colorbox{cream}{\smath{\onslide<2->{\nabla \vdash} a \fresh t}}} &
- \onslide<4->{\alert{freshness}}
- \end{tabular}
- \end{center}
-
- \onslide<2->{
- where
- \begin{center}
- \smath{\nabla = \{a_1\fresh X_1,\ldots, a_n\fresh X_n\}}
- \end{center}
- is a finite set of \alert{freshness assumptions}.}
-
- \onslide<3->{
- \begin{center}
- \smath{\{a\fresh X,b\fresh X\} \vdash \text{fn\ } a. X \approx \text{fn\ } b. X}
- \end{center}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{Rules for Equivalence}
-
- \begin{center}
- \begin{tabular}{c}
- Excerpt\\
- (i.e.~only the interesting rules)
- \end{tabular}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{Rules for Equivalence}
-
- \begin{center}
- \begin{tabular}{c}
- \colorbox{cream}{\smath{\infer{\nabla \vdash a \approx a}{}}}\\[8mm]
-
- \colorbox{cream}{%
- \smath{\infer{\nabla \vdash \abst{a}{t} \approx \abst{a}{t'}}
- {\nabla \vdash t \approx t'}}}\\[8mm]
-
- \colorbox{cream}{%
- \smath{\infer{\nabla \vdash \abst{a}{t} \approx \abst{b}{t'}}
- {a\not=b\;\; & \nabla \vdash t \approx \swap{a}{b}\act t'\;\;& \nabla \vdash a\fresh t'}}}
- \end{tabular}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>[c]
- \frametitle{Rules for Equivalence}
-
- \begin{center}
- \colorbox{cream}{%
- \smath{%
- \infer{\nabla \vdash \pi\act X \approx \pi'\act X}
- {\begin{array}{c}
- (a\fresh X)\in\nabla\\
- \text{for all}\; a \;\text{with}\;\pi\act a \not= \pi'\act a
- \end{array}
- }}}
- \end{center}
-
- \onslide<2->{
- for example\\[4mm]
-
- \alt<2>{%
- \begin{center}
- \smath{\{a\fresh\!X, b\fresh\!X\} \vdash X \approx \swap{a}{b}\act X}
- \end{center}}
- {%
- \begin{center}
- \smath{\{a\fresh\!X, c\fresh\!X\} \vdash \swap{a}{c}\swap{a}{b}\act X \approx \swap{b}{c}\act X}
- \end{center}}
-
- \onslide<3->{
- \begin{tabular}{@ {}lllll@ {}}
- because &
- \smath{\swap{a}{c}\swap{a}{b}}: &
- \smath{a\mapsto b} &
- \smath{\swap{b}{c}}: &
- \smath{a\mapsto a}\\
- & & \smath{b\mapsto c} & & \smath{b\mapsto c}\\
- & & \smath{c\mapsto a} & & \smath{c\mapsto b}\\
- \end{tabular}
- disagree at \smath{a} and \smath{c}.}
- }
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{Rules for Freshness}
-
- \begin{center}
- \begin{tabular}{c}
- Excerpt\\
- (i.e.~only the interesting rules)
- \end{tabular}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{Rules for Freshness}
-
- \begin{center}
- \begin{tabular}{c}
- \colorbox{cream}{%
- \smath{\infer{\nabla \vdash a\fresh b}{a\not= b}}}\\[5mm]
-
- \colorbox{cream}{%
- \smath{\infer{\nabla \vdash a\fresh\abst{a}{t}}{}}}\hspace{7mm}
- \colorbox{cream}{%
- \smath{\infer{\nabla \vdash a\fresh\abst{b}{t}}
- {a\not= b\;\; & \nabla \vdash a\fresh t}}}\\[5mm]
-
- \colorbox{cream}{%
- \smath{\infer{\nabla \vdash a\fresh \pi\act X}
- {(\pi^{-1}\act a\fresh X)\in\nabla}}}
- \end{tabular}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-4>[t]
- \frametitle{$\approx$ is an Equivalence}
- \mbox{}\\[5mm]
-
- \begin{center}
- \colorbox{cream}{\alert{Theorem:}
- $\approx$ is an equivalence relation.}
- \end{center}\bigskip
-
- \only<1>{%
- \begin{tabular}{ll}
- (Reflexivity) & $\smath{\nabla\vdash t\approx t}$\\[2mm]
- (Symmetry) & if $\smath{\nabla\vdash t_1\approx t_2}\;$
- then $\;\smath{\nabla\vdash t_2\approx t_1}$\\[2mm]
- (Transitivity) & if $\smath{\nabla\vdash t_1\approx t_2}\;$ and
- $\;\smath{\nabla\vdash t_2\approx t_3}$\\
- & then $\smath{\nabla\vdash t_1\approx t_3}$\\
- \end{tabular}}
-
- \only<2->{%
- \begin{itemize}
- \item<2-> \smath{\nabla \vdash t\approx t'} then \smath{\nabla \vdash \pi\act t\approx \pi\act t'}
-
- \item<2-> \smath{\nabla \vdash a\fresh t} then
- \smath{\nabla \vdash \pi\act a\fresh \pi\act t}
-
- \item<3-> \smath{\nabla \vdash t\approx \pi\act t'} then
- \smath{\nabla \vdash (\pi^{-1})\act t\approx t'}
-
- \item<3-> \smath{\nabla \vdash a\fresh \pi\act t} then
- \smath{\nabla \vdash (\pi^{-1})\act a\fresh t}
-
- \item<4-> \smath{\nabla \vdash a\fresh t} and \smath{\nabla \vdash t\approx t'} then
- \smath{\nabla \vdash a\fresh t'}
- \end{itemize}
- }
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-4>
- \frametitle{Comparison $=_\alpha$}
-
- Traditionally \smath{=_\alpha} is defined as
-
- \begin{center}
- \colorbox{cream}{%
- \begin{minipage}{9cm}
- \raggedright least congruence which identifies \smath{\abst{a}{t}}
- with \smath{\abst{b}{[a:=b]t}} provided \smath{b} is not free
- in \smath{t}
- \end{minipage}}
- \end{center}
-
- where \smath{[a:=b]t} replaces all free occurrences of\\
- \smath{a} by \smath{b} in \smath{t}.
- \bigskip
-
- \only<2>{%
- \begin{textblock}{13}(1.2,10)
- For \alert{ground} terms:
-
- \begin{center}
- \colorbox{cream}{%
- \begin{minipage}{9.0cm}
- \begin{tabular}{@ {}rl}
- \underline{Theorem:}
- & \smath{t=_\alpha t'\;\;} if\hspace{-0.5mm}f~\smath{\;\;\emptyset \vdash t\approx t'}\\[2mm]
- & \smath{a\not\in F\hspace{-0.9mm}A(t)\;\;} if\hspace{-0.5mm}f~\smath{\;\;\emptyset\vdash a\fresh t}
- \end{tabular}
- \end{minipage}}
- \end{center}
- \end{textblock}}
-
- \only<3>{%
- \begin{textblock}{13}(1.2,10)
- In general \smath{=_\alpha} and \smath{\approx} are distinct!
- \begin{center}
- \colorbox{cream}{%
- \begin{minipage}{6.0cm}
- \smath{\abst{a}{X}=_\alpha \abst{b}{X}\;} but not\\[2mm]
- \smath{\emptyset \vdash \abst{a}{X} \approx \abst{b}{X}\;} (\smath{a\not=b})
- \end{minipage}}
- \end{center}
- \end{textblock}}
-
- \only<4>{
- \begin{textblock}{6}(1,2)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=3mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\color{darkgray}
- \begin{minipage}{10cm}\raggedright
- That is a crucial point: if we had\\[-2mm]
- \[\smath{\emptyset \vdash \abst{a}{X}\approx \abst{b}{X}}\mbox{,}\]
- then applying $\smath{[X:=a]}$, $\smath{[X:=b]}$, $\ldots$\\
- give two terms that are {\bf not} $\alpha$-equivalent.\\[3mm]
- The freshness constraints $\smath{a\fresh X}$ and $\smath{b\fresh X}$
- rule out the problematic substitutions. Therefore
-
- \[\smath{\{a\fresh X,b\fresh X\} \vdash \abst{a}{X}\approx \abst{b}{X}}\]
-
- does hold.
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-9>
- \frametitle{Substitution}
-
- \begin{tabular}{l@ {\hspace{8mm}}r@ {\hspace{1.5mm}}c@ {\hspace{1.5mm}}l@ {}}
- \pgfuseshading{smallbluesphere} &
- \smath{\sigma(\abst{a}{t})} & \smath{\dn} & \smath{\abst{a}{\sigma(t)}}\\[2mm]
-
- \pgfuseshading{smallbluesphere} &
- \smath{\sigma(\pi\act X)} & \smath{\dn} &
- \smath{\begin{cases}%
- \pi\;\act\;\sigma(X) & \!\!\text{if\ } \sigma(X)\not=X\\
- \pi\act X & \!\!\text{otherwise}%
- \end{cases}}\\[6mm]
- \end{tabular}\bigskip\bigskip
-
- \pause
- \only<2-5>{
- \only<2->{for example}
- \def\arraystretch{1.3}
- \begin{tabular}{@ {\hspace{14mm}}l@ {\hspace{3mm}}l}
- \onslide<2->{\textcolor{white}{$\Rightarrow$}} &
- \onslide<2->{\alt<3>{\smath{\underline{\abst{a}{\swap{a}{b}\act X}\;\,[X:=\pair{b}{Y}]}}}
- {\smath{\abst{a}{\swap{a}{b}\act X}\;\,[X:=\pair{b}{Y}]}}}\\
- \onslide<3->{\smath{\Rightarrow}} &
- \onslide<3->{\alt<3,4>{\smath{\abst{a}{\underline{\swap{a}{b}\act X[X:=\pair{b}{Y}]}}}}
- {\smath{\abst{a}{\swap{a}{b}\act X}[X:=\pair{b}{Y}]}}}\\
- \onslide<4->{\smath{\Rightarrow}} &
- \onslide<4->{\alt<4>{\smath{\abst{a}{\swap{a}{b}\act \underline{\pair{b}{Y}}}}}
- {\smath{\abst{a}{\underline{\swap{a}{b}}\act \pair{b}{Y}}}}}\\
- \onslide<5->{\smath{\Rightarrow}} &
- \onslide<5->{\smath{\abst{a}{\pair{a}{\swap{a}{b}\act Y}}}}
- \end{tabular}}
-
- \only<6->
- {\begin{tabular}{l@ {\hspace{8mm}}l@ {}}
- \pgfuseshading{smallbluesphere} &
- if \smath{\nabla\vdash t\approx t'} and\hspace{-2mm}\mbox{}
- \raisebox{-2.7mm}{
- \alt<7>{\begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, very thick, draw=red, rounded corners=3mm]
- {\smath{\;\nabla'\vdash\sigma(\nabla)\;}};
- \end{tikzpicture}}
- {\begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=white, very thick, draw=white, rounded corners=3mm]
- {\smath{\;\nabla'\vdash\sigma(\nabla)\;}};
- \end{tikzpicture}}}\\
- & then \smath{\nabla'\vdash\sigma(t)\approx\sigma(t')}
- \end{tabular}}
-
- \only<9>
- {\begin{tabular}{l@ {\hspace{8mm}}l@ {}}
- \\[-4mm]
- \pgfuseshading{smallbluesphere} &
- \smath{\sigma(\pi\act t)=\pi\act\sigma(t)}
- \end{tabular}}
-
-
- \only<7>{
- \begin{textblock}{6}(10,10.5)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, very thick, draw=red, rounded corners=2mm]
- {\color{darkgray}
- \begin{minipage}{3.8cm}\raggedright
- this means\\[1mm]
- \smath{\nabla'\vdash a\fresh\sigma(X)}\\[1mm]
- holds for all\\[1mm]
- \smath{(a\fresh X)\in\nabla}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{Equational Problems}
-
- An equational problem
- \[
- \colorbox{cream}{\smath{t \eqprob t'}}
- \]
- is \alert{solved} by
-
- \begin{center}
- \begin{tabular}{ll}
- \pgfuseshading{smallbluesphere} & a substitution \smath{\sigma} (terms for variables)\\[3mm]
- \pgfuseshading{smallbluesphere} & {\bf and} a set of freshness assumptions \smath{\nabla}
- \end{tabular}
- \end{center}
-
- so that \smath{\nabla\vdash \sigma(t)\approx \sigma(t')}.
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
-
- Unifying equations may entail solving
- \alert{freshness problems}.
-
- \bigskip
-
- E.g.~assuming that \smath{a\not=a'}, then
- \[
- \smath{\abst{a}{t}\eqprob \abst{a'}{t'}}
- \]
- can only be solved if
- \[
- \smath{t\eqprob \swap{a}{a'}\act t'} \quad\text{\emph{and}}\quad
- \smath{a\freshprob t'}
- \]
- can be solved.
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{Freshness Problems}
-
- A freshness problem
- \[
- \colorbox{cream}{\smath{a \freshprob t}}
- \]
- is \alert{solved} by
-
- \begin{center}
- \begin{tabular}{ll}
- \pgfuseshading{smallbluesphere} & a substitution \smath{\sigma}\\[3mm]
- \pgfuseshading{smallbluesphere} & and a set of freshness assumptions \smath{\nabla}
- \end{tabular}
- \end{center}
-
- so that \smath{\nabla\vdash a \fresh \sigma(t)}.
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>
- \frametitle{Existence of MGUs}
-
- \underline{Theorem}: There is an algorithm which, given a nominal
- unification problem \smath{P}, decides whether\\
- or not it has a solution \smath{(\sigma,\nabla)}, and returns a \\
- \alert{most general} one if it does.\bigskip\bigskip
-
- \only<3>{
- Proof: one can reduce all the equations to `solved form'
- first (creating a substitution), and then solve the freshness
- problems (easy).}
-
- \only<2>{
- \begin{textblock}{6}(2.5,9.5)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=3mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\color{darkgray}
- \begin{minipage}{8cm}\raggedright
- \alert{most general:}\\
- straightforward definition\\
- ``if\hspace{-0.5mm}f there exists a \smath{\tau} such that \ldots''
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>
- \frametitle{Remember the Quiz?}
-
- \textcolor{gray}{Assuming that $a$ and $b$ are distinct variables,\\
- is it possible to find $\lambda$-terms $M_1$ to $M_7$
- that make the following pairs $\alpha$-equivalent?}
-
- \begin{tabular}{@ {\hspace{14mm}}p{12cm}}
- \begin{itemize}
- \item \smath{\lambda a.\lambda b. (M_1\,b)\;} and
- \smath{\lambda b.\lambda a. (a\,M_1)\;}
-
- \item \textcolor{gray}{$\lambda a.\lambda b. (M_2\,b)\;$ and
- $\lambda b.\lambda a. (a\,M_3)\;$}
-
- \item \textcolor{gray}{$\lambda a.\lambda b. (b\,M_4)\;$ and
- $\lambda b.\lambda a. (a\,M_5)\;$}
-
- \item \smath{\lambda a.\lambda b. (b\,M_6)\;} and
- \smath{\lambda a.\lambda a. (a\,M_7)\;}
- \end{itemize}
- \end{tabular}
-
- \textcolor{gray}{If there is one solution for a pair, can you
- describe all its solutions?}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{Answers to the Quiz}
- \small
- \def\arraystretch{1.6}
- \begin{tabular}{c@ {\hspace{2mm}}l}
- & \only<1>{\smath{\lambda a.\lambda b. (M_1\,b)\;} and \smath{\;\lambda b.\lambda a. (a\,M_1)}}%
- \only<2->{\smath{\abst{a}{\abst{b}{\pair{M_1}{b}}} \;\eqprob\; \abst{b}{\abst{a}{\pair{a}{M_1}}}}}\\
-
- \onslide<3->{\smath{\redu{\id}}} &
- \only<3>{\smath{\abst{b}{\pair{M_1}{b}} \eqprob
- \alert{\swap{a}{b}} \act \abst{a}{\pair{a}{M_1}}\;,\;a\freshprob \abst{a}{\pair{a}{M_1}}}}%
- \only<4->{\smath{\abst{b}{\pair{M_1}{b}} \eqprob \abst{b}{\pair{b}{\swap{a}{b}\act M_1}}\;,\
- a\freshprob \abst{a}{\pair{a}{M_1}}}}\\
-
- \onslide<5->{\smath{\redu{\id}}} &
- \only<5->{\smath{\pair{M_1}{b} \eqprob \pair{b}{\swap{a}{b}\act M_1}\;,\;%
- a\freshprob \abst{a}{\pair{a}{M_1}}}}\\
-
- \onslide<6->{\smath{\redu{\id}}} &
- \only<6->{\smath{M_1 \eqprob b \;,\; b \eqprob \swap{a}{b}\act M_1\;,\;%
- a\freshprob \abst{a}{\pair{a}{M_1}}}}\\
-
- \onslide<7->{\smath{\redu{[M_1:=b]}}} &
- \only<7>{\smath{b \eqprob \swap{a}{b}\act \alert{b}\;,\;%
- a\freshprob \abst{a}{\pair{a}{\alert{b}}}}}%
- \only<8->{\smath{b \eqprob a\;,\; a\freshprob \abst{a}{\pair{a}{b}}}}\\
-
- \onslide<9->{\smath{\redu{}}} &
- \only<9->{\smath{F\hspace{-0.5mm}AIL}}
- \end{tabular}
-
- \only<10>{
- \begin{textblock}{6}(2,11)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=3mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\color{darkgray}
- \begin{minipage}{9cm}\raggedright
- \smath{\lambda a.\lambda b. (M_1\,b)} \smath{=_\alpha}
- \smath{\lambda b.\lambda a. (a\,M_1)} has no solution
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{Answers to the Quiz}
- \small
- \def\arraystretch{1.6}
- \begin{tabular}{c@ {\hspace{2mm}}l}
- & \only<1>{\smath{\lambda a.\lambda b. (b\,M_6)\;} and \smath{\;\lambda a.\lambda a. (a\,M_7)}}%
- \only<2->{\smath{\abst{a}{\abst{b}{\pair{b}{M_6}}} \;\eqprob\; \abst{a}{\abst{a}{\pair{a}{M_7}}}}}\\
-
- \onslide<3->{\smath{\redu{\id}}} &
- \only<3->{\smath{\abst{b}{\pair{b}{M_6}} \eqprob \abst{a}{\pair{a}{M_7}}}}\\
-
- \onslide<4->{\smath{\redu{\id}}} &
- \only<4->{\smath{\pair{b}{M_6} \eqprob \pair{b}{\swap{b}{a}\act M_7}\;,\;b\freshprob\pair{a}{M_7}}}\\
-
- \onslide<5->{\smath{\redu{\id}}} &
- \only<5->{\smath{b\eqprob b\;,\; M_6 \eqprob \swap{b}{a}\act M_7\;,\;%
- b\freshprob \pair{a}{M_7}}}\\
-
- \onslide<6->{\smath{\redu{\id}}} &
- \only<6->{\smath{M_6 \eqprob \swap{b}{a}\act M_7\;,\;%
- b\freshprob \pair{a}{M_7}}}\\
-
- \onslide<7->{\makebox[0mm]{\smath{\redu{[M_6:=\swap{b}{a}\act M_7]}}}} &
- \only<7->{\smath{\qquad b\freshprob \pair{a}{M_7}}}\\
-
- \onslide<8->{\smath{\redu{\varnothing}}} &
- \only<8->{\smath{b\freshprob a\;,\;b\freshprob M_7}}\\
-
- \onslide<9->{\smath{\redu{\varnothing}}} &
- \only<9->{\smath{b\freshprob M_7}}\\
-
- \onslide<10->{\makebox[0mm]{\smath{\redu{\{b\fresh M_7\}}}}} &
- \only<10->{\smath{\;\;\varnothing}}\\
-
- \end{tabular}
-
- \only<10>{
- \begin{textblock}{6}(6,9)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=3mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\color{darkgray}
- \begin{minipage}{7cm}\raggedright
- \smath{\lambda a.\lambda b. (b\,M_6)\;} \smath{=_\alpha}
- \smath{\;\lambda a.\lambda a. (a\,M_7)}\\[2mm]
- we can take \smath{M_7} to be any $\lambda$-term that does not
- contain free occurrences of \smath{b}, so long as we take \smath{M_6} to
- be the result of swapping all occurrences of \smath{b} and \smath{a}
- throughout \smath{M_7}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{Properties}
-
- \begin{itemize}
- \item An interesting feature of nominal unification is that it
- does not need to create new atoms.\bigskip
-
- \begin{center}\small
- \colorbox{cream}{
- \smath{\{a.t \eqprob b.t'\}\cup P \redu{\id} \{t \eqprob \swap{a}{b}\act t', a \freshprob t'\} \cup P}}
- \end{center}\bigskip\bigskip
- \pause
-
- \item The alternative rule
-
- \begin{center}\small
- \colorbox{cream}{
- \begin{tabular}{@ {}l@ {}}
- \smath{\{a.t \eqprob b.t'\}\cup P \redu{\id}}\\
- \mbox{}\hspace{2cm}\smath{\{\swap{a}{c}\act t \eqprob
- \swap{b}{c}\act t', c \freshprob t, c \freshprob t'\} \cup P}
- \end{tabular}}
- \end{center}
-
- leads to a more complicated notion of mgu.\medskip\pause
-
- \footnotesize
- \smath{\{a.X \eqprob b.Y\} \redu{} (\{a\fresh Y, c\fresh Y\}, [X:=\swap{a}{c}\swap{b}{c}\act Y])}
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>
- \frametitle{Is it Useful?}
-
- Yes. $\alpha$Prolog by James Cheney (main developer)\bigskip\bigskip
-
- \color{darkgray}
- \begin{tabular}{@ {}l}
- type (Gamma, var(X), T) :- member (X,T) Gamma.\smallskip\medskip\\
-
- type (Gamma, app(M, N), T') :-\\
- \hspace{3cm}type (Gamma, M, arrow(T, T')),\\
- \hspace{3cm}type (Gamma, N, T).\smallskip\medskip\\
-
- type (Gamma, lam(\alert{x.M}), arrow(T, T')) / \alert{x \# Gamma} :-\\
- \hspace{3cm}type ((x, T)::Gamma, M, T').\smallskip\medskip\\
-
- member X X::Tail.\\
- member X Y::Tail :- member X Tail.\\
- \end{tabular}
-
- \only<2->{
- \begin{textblock}{6}(1.5,0.5)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=3mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\color{darkgray}
- \begin{minipage}{9cm}\raggedright
- {\bf One problem:} If we ask whether
-
- \begin{center}
- ?- type ([(x, T')], lam(x.Var(x)), T)
- \end{center}
-
- is typable, we expect an answer for T.\bigskip
-
- \onslide<3>{Solution: Before back-chaining freshen all variables and atoms
- in a program (clause).}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{Equivariant Unification}
-
- James Cheney proposed
-
- \begin{center}
- \colorbox{cream}{
- \smath{t \eqprob t' \redu{\nabla, \sigma, \pi}
- \nabla \vdash \sigma(t) \approx \pi \act \sigma(t')}}
- \end{center}\bigskip\bigskip
- \pause
-
- But he also showed this problem is undecidable\\ in general. :(
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{Taking Atoms as Variables}
-
- Instead of \smath{a.X}, have \smath{A.X}.\bigskip
- \pause
-
- Unfortunately this breaks the mgu-property:
-
- \begin{center}
- \smath{a.Z \eqprob X.Y.v(a)}
- \end{center}
-
- can be solved by
-
- \begin{center}
- \smath{[X:=a, Z:=Y.v(a)]} and
- \smath{[Y:=a, Z:=Y.v(Y)]}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{HOPU vs. NOMU}
-
- \begin{itemize}
- \item James Cheney showed\bigskip
- \begin{center}
- \colorbox{cream}{\smath{HOPU \Rightarrow NOMU}}
- \end{center}\bigskip
-
- \item Jordi Levy and Mateu Villaret established\bigskip
- \begin{center}
- \colorbox{cream}{\smath{HOPU \Leftarrow NOMU}}
- \end{center}\bigskip
- \end{itemize}
-
- The translations `explode' the problems quadratically.
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>
- \small\tt
-
- \begin{minipage}{13cm}
- \begin{tabular}{@ {\hspace{-2mm}}p{11.5cm}}
- \\
- From: Zhenyu Qian <zhqian@microsoft.com>\\
- To: Christian Urban <urbanc@in.tum.de>\\
- Subject: RE: Linear Higher-Order Pattern Unification\\
- Date: Mon, 14 Apr 2008 09:56:47 +0800\\
- \\
- Hi Christian,\\
- \\
- Thanks for your interests and asking. I know that that paper is complex. As
- I told Tobias when we met last time, I have raised the question to myself
- many times whether the proof could have some flaws, and so making it through
- a theorem prover would definitely bring piece to my mind (no matter what
- the result would be). The only problem for me is the time.\\
- \ldots\\
-
- Thanks/Zhenyu
- \end{tabular}
- \end{minipage}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>
- \frametitle{Complexity}
-
- \begin{itemize}
- \item Christiopher Calves and Maribel Fernandez showed first that
- it is polynomial and then also quadratic
-
- \item Jordi Levy and Mateu Villaret showed that it is quadratic
- by a translation into a subset of NOMU and using ideas from
- Martelli/Montenari.
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Conclusion}
-
- \begin{itemize}
- \item Nominal Unification is a completely first-order
- language, but implements unification modulo $\alpha$.
- \textcolor{gray}{(verification\ldots Ramana Kumar and Michael Norrish)}
- \medskip\pause
-
- \item NOMU has been applied in term-rewriting and
- logic programming. \textcolor{gray}{(Maribel Fernandez et
- al has a KB-completion procedure.)}
- I hope it will also be used in typing
- systems.\medskip\pause
-
- \item NOMU and HOPU are `equivalent' (it took a long time
- and considerable research to find this out).\medskip\pause
-
- \item The question about complexity is still an ongoing
- story.\medskip
- \end{itemize}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{
- \begin{tabular}{c}
- \mbox{}\\[23mm]
- \alert{\LARGE Thank you very much!}\\
- \alert{\Large Questions?}
- \end{tabular}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>
- \frametitle{Most General Unifiers}
-
- \underline{Definition}: For a unification problem
- \smath{P}, a solution \smath{(\sigma_1,\nabla_1)} is
- \alert{more general} than another solution
- \smath{(\sigma_2,\nabla_2)}, iff~there exists a substitution
- \smath{\tau} with
-
- \begin{center}
- \begin{tabular}{ll}
- \pgfuseshading{smallbluesphere} &
- \alt<2>{\smath{\alert{\nabla_2\vdash\tau(\nabla_1)}}}
- {\smath{\nabla_2\vdash\tau(\nabla_1)}}\\
- \pgfuseshading{smallbluesphere} &
- \alt<3>{\smath{\alert{\nabla_2\vdash\sigma_2\approx \tau\circ\sigma_1}}}
- {\smath{\nabla_2\vdash\sigma_2\approx \tau\circ\sigma_1}}
- \end{tabular}
- \end{center}
-
- \only<2>{
- \begin{textblock}{13}(1.5,10.5)
- \smath{\nabla_2\vdash a\fresh \sigma(X)} holds for all
- \smath{(a\fresh X)\in\nabla_1}
- \end{textblock}}
-
- \only<3>{
- \begin{textblock}{11}(1.5,10.5)
- \smath{\nabla_2\vdash \sigma_2(X)\approx
- \sigma(\sigma_1(X))}
- holds for all
- \smath{X\in\text{dom}(\sigma_2)\cup\text{dom}(\sigma\circ\sigma_1)}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-(*<*)
-end
-(*>*)
\ No newline at end of file
--- a/Slides/Slides4.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1135 +0,0 @@
-(*<*)
-theory Slides4
-imports "~~/src/HOL/Library/LaTeXsugar" "Nominal"
-begin
-
-declare [[show_question_marks = false]]
-
-notation (latex output)
- set ("_") and
- Cons ("_::/_" [66,65] 65)
-
-(*>*)
-
-text_raw {*
- \renewcommand{\slidecaption}{Nanjing, 31.~August 2010}
-
- \newcommand{\abst}[2]{#1.#2}% atom-abstraction
- \newcommand{\pair}[2]{\langle #1,#2\rangle} % pairing
- \newcommand{\susp}{{\boldsymbol{\cdot}}}% for suspensions
- \newcommand{\unit}{\langle\rangle}% unit
- \newcommand{\app}[2]{#1\,#2}% application
- \newcommand{\eqprob}{\mathrel{{\approx}?}}
- \newcommand{\freshprob}{\mathrel{\#?}}
- \newcommand{\redu}[1]{\stackrel{#1}{\Longrightarrow}}% reduction
- \newcommand{\id}{\varepsilon}% identity substitution
-
- \newcommand{\bl}[1]{\textcolor{blue}{#1}}
- \newcommand{\gr}[1]{\textcolor{gray}{#1}}
- \newcommand{\rd}[1]{\textcolor{red}{#1}}
-
- \newcommand{\ok}{\includegraphics[scale=0.07]{ok.png}}
- \newcommand{\notok}{\includegraphics[scale=0.07]{notok.png}}
- \newcommand{\largenotok}{\includegraphics[scale=1]{notok.png}}
-
- \renewcommand{\Huge}{\fontsize{61.92}{77}\selectfont}
- \newcommand{\veryHuge}{\fontsize{74.3}{93}\selectfont}
- \newcommand{\VeryHuge}{\fontsize{89.16}{112}\selectfont}
- \newcommand{\VERYHuge}{\fontsize{107}{134}\selectfont}
-
- \newcommand{\LL}{$\mathbb{L}\,$}
-
-
- \pgfdeclareradialshading{smallbluesphere}{\pgfpoint{0.5mm}{0.5mm}}%
- {rgb(0mm)=(0,0,0.9);
- rgb(0.9mm)=(0,0,0.7);
- rgb(1.3mm)=(0,0,0.5);
- rgb(1.4mm)=(1,1,1)}
-
- \def\myitemi{\begin{pgfpicture}{-1ex}{-0.55ex}{1ex}{1ex}
- \usebeamercolor[fg]{subitem projected}
- {\pgftransformscale{0.8}\pgftext{\normalsize\pgfuseshading{bigsphere}}}
- \pgftext{%
- \usebeamerfont*{subitem projected}}
- \end{pgfpicture}}
-
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[t]
- \frametitle{%
- \begin{tabular}{@ {\hspace{-3mm}}c@ {}}
- \\
- \huge Error-Free Programming\\[-1mm]
- \huge with Theorem Provers\\[5mm]
- \end{tabular}}
- \begin{center}
- Christian Urban
- \end{center}
- \begin{center}
- \small Technical University of Munich, Germany\\[7mm]
-
- \small in Nanjing on the kind invitation of\\ Professor Xingyuan Zhang
- and his group
- \end{center}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{My Background}
-
- \begin{itemize}
- \item researcher in Theoretical Computer Science\medskip
-
- \item programmer on a \alert<2->{software system} with 800 kloc (though I am
- responsible only for 35 kloc)
- \end{itemize}
-
- \only<2->{
- \begin{textblock}{6}(2,11)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=2mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\color{darkgray}
- \begin{minipage}{4cm}\raggedright
- A theorem prover called {\bf Isabelle}.
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
-
- \only<3>{
- \begin{textblock}{6}(9,11)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=2mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\color{darkgray}
- \begin{minipage}{4cm}\raggedright
- Like every other code, this code is very hard to
- get correct.
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Regular Expressions}
-
- An example many (should) know about:\\
- \rd{\bf Regular Expressions:}
-
- \only<2>{
- \begin{center}
- \bl{[] $\;\;\;|\;\;\;$ c $\;\;\;|\;\;\;$ r$_1$$|$r$_2$ $\;\;\;|\;\;\;$
- r$_1$$\cdot$r$_2$ $\;\;\;|\;\;\;$ r$^*$}
- \end{center}}
- \only<3->{
- \begin{center}
- \begin{tabular}{@ {}rrll@ {}}
- \bl{r} & \bl{$::=$} & \bl{NULL} & \gr{(matches no string)}\\
- & \bl{$\mid$} & \bl{EMPTY} & \gr{(matches the empty string, [])}\\
- & \bl{$\mid$} & \bl{CHR c} & \gr{(matches the character c)}\\
- & \bl{$\mid$} & \bl{ALT r$_1$ r$_2$} & \gr{(alternative, r$_1 |\,$r$_2$)}\\
- & \bl{$\mid$} & \bl{SEQ r$_1$ r$_2$} & \gr{(sequential, r$_1\cdot\,$r$_2$)}\\
- & \bl{$\mid$} & \bl{STAR r} & \gr{(repeat, r$^*$)}\\
- \end{tabular}
- \end{center}\medskip}
-
- \small
- \begin{textblock}{14.5}(1,12.5)
- \only<2->{\gr{(a$\cdot$b)$^*$ \hspace{3mm}$\mapsto$\hspace{3mm} \{[], ab, abab, ababab, \ldots\}}\\}
- \only<2->{\gr{x$\cdot$(0 $|$ 1 $|$ 2 \ldots 8 $|$ 9)$^*$ \hspace{3mm}$\mapsto$\hspace{3mm}
- \{x, x0, x1, \ldots, x00, \ldots, x123, \ldots\}}}
- \end{textblock}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{RegExp Matcher}
-
- Let's implement a regular expression matcher:\bigskip
-
- \begin{center}
- \begin{tikzpicture}
- %%\draw[help lines, black] (-3,0) grid (6,3);
-
- \draw[line width=1mm, red] (0.0,0.0) rectangle (4,2.3);
- \node[anchor=base] at (2,1)
- {\small\begin{tabular}{@ {}c@ {}}\Large\bf Regular\\
- \Large\bf Expression \\
- \Large\bf Matcher\end{tabular}};
-
- \coordinate (m1) at (0,1.5);
- \draw (-2,2) node (m2) {\small\begin{tabular}{c}\bl{regular}\\[-1mm] \bl{expression}\end{tabular}};
- \path[overlay, ->, line width = 1mm, shorten <=-3mm] (m2) edge (m1);
-
- \coordinate (s1) at (0,0.5);
- \draw (-1.8,-0) node (s2) {\small\begin{tabular}{c}\bl{string}\end{tabular}};
- \path[overlay, ->, line width = 1mm, shorten <=-3mm] (s2) edge (s1);
-
- \coordinate (r1) at (4,1.2);
- \draw (6,1.2) node (r2) {\small\begin{tabular}{c}\bl{true}, \bl{false}\end{tabular}};
- \path[overlay, ->, line width = 1mm, shorten >=-3mm] (r1) edge (r2);
-
- \end{tikzpicture}
- \end{center}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{RegExp Matcher}
- \small
-
- {\bf input:} a \underline{list} of RegExps and a string \hspace{6mm}{\bf output:} true or false
-
- \only<2->{
- \begin{center}
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {}}
- \bl{match [] []} & \bl{$=$} & \bl{true}\\
- \bl{match [] \_} & \bl{$=$} & \bl{false}\\
- \bl{match (NULL::rs) s} & \bl{$=$} & \bl{false}\\
- \bl{match (EMPTY::rs) s} & \bl{$=$} & \bl{match rs s}\\
- \bl{match (CHR c::rs) (c::s)} & \bl{$=$} & \bl{match rs s}\\
- \bl{match (CHR c::rs) \_} & \bl{$=$} & \bl{false}\\
- \bl{match (ALT r$_1$ r$_2$::rs) s} & \bl{$=$} & \bl{match (r$_1$::rs) s}\\
- & & \bl{\;\;\;\;orelse match (r$_2$::rs) s}\\
- \bl{match (SEQ r$_1$ r$_2$::rs) s} & \bl{$=$} & \bl{match (r$_1$::r$_2$::rs) s}\\
- \bl{match (STAR r::rs) s} & \bl{$=$} & \bl{match rs s}\\
- & & \bl{\;\;\;\;orelse match (r::STAR r::rs) s}\\
- \end{tabular}
- \end{center}}
-
- \onslide<3->{we start the program with\\
- \hspace{6mm}\bl{matches r s $=$ match [r] s}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{Program in Scala}
-
- \bl{\footnotesize
- \begin{tabular}{l}
- sealed abstract class Rexp\\
- sealed case class Null extends Rexp\\
- sealed case class Empty extends Rexp\\
- sealed case class Chr(c: Char) extends Rexp\\
- sealed case class Alt(r1 : Rexp, r2 : Rexp) extends Rexp\\
- sealed case class Seq(r1 : Rexp, r2 : Rexp) extends Rexp\\
- sealed case class Star(r : Rexp) extends Rexp\medskip\\
- def match1 (rs : List[Rexp], s : List[Char]) : Boolean = rs match \{\\
- \hspace{3mm}case Nil @{text "\<Rightarrow>"} if (s == Nil) true else false\\
- \hspace{3mm}case (Null()::rs) @{text "\<Rightarrow>"} false\\
- \hspace{3mm}case (Empty()::rs) @{text "\<Rightarrow>"} match1 (rs, s)\\
- \hspace{3mm}case (Chr(c)::rs) @{text "\<Rightarrow>"} s match \\
- \hspace{6mm}\{ case Nil @{text "\<Rightarrow>"} false\\
- \hspace{8mm}case (d::s) @{text "\<Rightarrow>"} if (c==d) match1 (rs,s) else false \}\\
- \hspace{3mm}case (Alt (r1, r2)::rs) @{text "\<Rightarrow>"} match1 (r1::rs, s) || match1 (r2::rs, s)\\
- \hspace{3mm}case (Seq (r1, r2)::rs) @{text "\<Rightarrow>"} match1 (r1::r2::rs, s) \\
- \hspace{3mm}case (Star (r)::rs) @{text "\<Rightarrow>"} match1 (r::rs, s) || match1 (r::Star (r)::rs, s)\\
- \}
- \end{tabular}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Testing}
-
- \small
- Every good programmer should do thourough tests:
-
-
- \begin{center}
- \begin{tabular}{@ {\hspace{-20mm}}lcl}
- \bl{matches (a$\cdot$b)$^*\;$ []} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches (a$\cdot$b)$^*\;$ ab} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches (a$\cdot$b)$^*\;$ aba} & \bl{$\mapsto$} & \bl{false}\\
- \bl{matches (a$\cdot$b)$^*\;$ abab} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches (a$\cdot$b)$^*\;$ abaa} & \bl{$\mapsto$} & \bl{false}\medskip\\
- \onslide<2->{\bl{matches x$\cdot$(0$|$1)$^*\;$ x} & \bl{$\mapsto$} & \bl{true}}\\
- \onslide<2->{\bl{matches x$\cdot$(0$|$1)$^*\;$ x0} & \bl{$\mapsto$} & \bl{true}}\\
- \onslide<2->{\bl{matches x$\cdot$(0$|$1)$^*\;$ x3} & \bl{$\mapsto$} & \bl{false}}
- \end{tabular}
- \end{center}
-
- \onslide<3->
- {looks OK \ldots let's ship it to customers\hspace{5mm}
- \raisebox{-5mm}{\includegraphics[scale=0.05]{sun.png}}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Testing}
-
- \begin{itemize}
- \item While testing is an important part in the process of programming development\pause
-
- \item we can only test a {\bf finite} amount of examples\bigskip\pause
-
- \begin{center}
- \colorbox{cream}
- {\gr{\begin{minipage}{10cm}
- ``Testing can only show the presence of errors, never their
- absence'' (Edsger W.~Dijkstra)
- \end{minipage}}}
- \end{center}\bigskip\pause
-
- \item In a theorem prover we can establish properties that apply to
- {\bf all} input and {\bf all} output.\pause
-
- \item For example we can establish that the matcher terminates
- on all input.
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{RegExp Matcher}
-
- \small
- We need to find a measure that gets smaller in each recursive call.\bigskip
-
- \onslide<1->{
- \begin{center}
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-9mm}}l@ {}}
- \bl{match [] []} & \bl{$=$} & \bl{true} & \onslide<2->{\ok}\\
- \bl{match [] \_} & \bl{$=$} & \bl{false} & \onslide<2->{\ok}\\
- \bl{match (NULL::rs) s} & \bl{$=$} & \bl{false} & \onslide<2->{\ok}\\
- \bl{match (EMPTY::rs) s} & \bl{$=$} & \bl{match rs s} & \onslide<3->{\ok}\\
- \bl{match (CHR c::rs) (c::s)} & \bl{$=$} & \bl{match rs s} & \onslide<4->{\ok}\\
- \bl{match (CHR c::rs) \_} & \bl{$=$} & \bl{false} & \onslide<2->{\ok}\\
- \bl{match (ALT r$_1$ r$_2$::rs) s} & \bl{$=$} & \bl{match (r$_1$::rs) s} & \onslide<5->{\ok}\\
- & & \bl{\;\;\;\;orelse match (r$_2$::rs) s}\\
- \bl{match (SEQ r$_1$ r$_2$::rs) s} & \bl{$=$} & \bl{match (r$_1$::r$_2$::rs) s} & \onslide<6->{\ok}\\
- \bl{match (STAR r::rs) s} & \bl{$=$} & \bl{match rs s} & \onslide<7->{\notok}\\
- & & \bl{\;\;\;\;orelse match (r::STAR r::rs) s}\\
- \end{tabular}
- \end{center}}
-
-
- \begin{textblock}{5}(4,4)
- \begin{tikzpicture}
- %%\draw[help lines, black] (-3,0) grid (6,3);
-
- \coordinate (m1) at (-2,0);
- \coordinate (m2) at ( 2,0);
- \path[overlay, ->, line width = 0.6mm, color = red] (m1) edge (m2);
- \draw (0,0) node[above=-1mm] {\footnotesize\rd{needs to get smaller}};
- \end{tikzpicture}
- \end{textblock}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Bug Hunting}
-
- \only<1>{Several hours later\ldots}\pause
-
-
- \begin{center}
- \begin{tabular}{@ {\hspace{-20mm}}lcl}
- \bl{matches (STAR (EMPTY)) s} & \bl{$\mapsto$} & loops\\
- \onslide<4->{\bl{matches (STAR (EMPTY $|$ \ldots)) s} & \bl{$\mapsto$} & loops\\}
- \end{tabular}
- \end{center}
-
- \small
- \onslide<3->{
- \begin{center}
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {}}
- \ldots\\
- \bl{match (EMPTY::rs) s} & \bl{$=$} & \bl{match rs s}\\
- \ldots\\
- \bl{match (STAR r::rs) s} & \bl{$=$} & \bl{match rs s}\\
- & & \bl{\;\;\;\;orelse match (r::STAR r::rs) s}\\
- \end{tabular}
- \end{center}}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{RegExp Matcher}
- \small
-
- \begin{center}
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {}}
- \bl{match [] []} & \bl{$=$} & \bl{true}\\
- \bl{match [] \_} & \bl{$=$} & \bl{false}\\
- \bl{match (NULL::rs) s} & \bl{$=$} & \bl{false}\\
- \bl{match (EMPTY::rs) s} & \bl{$=$} & \bl{match rs s}\\
- \bl{match (CHR c::rs) (c::s)} & \bl{$=$} & \bl{match rs s}\\
- \bl{match (CHR c::rs) \_} & \bl{$=$} & \bl{false}\\
- \bl{match (ALT r$_1$ r$_2$::rs) s} & \bl{$=$} & \bl{match (r$_1$::rs) s}\\
- & & \bl{\;\;\;\;orelse match (r$_2$::rs) s}\\
- \bl{match (SEQ r$_1$ r$_2$::rs) s} & \bl{$=$} & \bl{match (r$_1$::r$_2$::rs) s}\\
- \bl{match (STAR r::rs) s} & \bl{$=$} & \bl{match rs s}\\
- & & \bl{\;\;\;\;orelse match (r::STAR r::rs) s}\\
- \end{tabular}
- \end{center}
-
- \only<1>{
- \begin{textblock}{5}(4,4)
- \largenotok
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Second Attempt}
-
- Can a regular expression match the empty string?
-
- \small
- \begin{center}
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}ll@ {}}
- \bl{nullable (NULL)} & \bl{$=$} & \bl{false} & \onslide<2->{\ok}\\
- \bl{nullable (EMPTY)} & \bl{$=$} & \bl{true} & \onslide<2->{\ok}\\
- \bl{nullable (CHR c)} & \bl{$=$} & \bl{false} & \onslide<2->{\ok}\\
- \bl{nullable (ALT r$_1$ r$_2$)} & \bl{$=$} & \bl{(nullable r$_1$) orelse (nullable r$_2$)}
- & \onslide<2->{\ok}\\
- \bl{nullable (SEQ r$_1$ r$_2$)} & \bl{$=$} & \bl{(nullable r$_1$) andalso (nullable r$_2$)}
- & \onslide<2->{\ok}\\
- \bl{nullable (STAR r)} & \bl{$=$} & \bl{true} & \onslide<2->{\ok}\\
- \end{tabular}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{RegExp Matcher 2}
-
- If \bl{r} matches \bl{c::s}, which regular expression can match the string \bl{s}?
-
- \small
- \begin{center}
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-10mm}}l@ {}}
- \bl{der c (NULL)} & \bl{$=$} & \bl{NULL} & \onslide<3->{\ok}\\
- \bl{der c (EMPTY)} & \bl{$=$} & \bl{NULL} & \onslide<3->{\ok}\\
- \bl{der c (CHR d)} & \bl{$=$} & \bl{if c=d then EMPTY else NULL} & \onslide<3->{\ok}\\
- \bl{der c (ALT r$_1$ r$_2$)} & \bl{$=$} & \bl{ALT (der c r$_1$) (der c r$_2$)} & \onslide<3->{\ok}\\
- \bl{der c (SEQ r$_1$ r$_2$)} & \bl{$=$} & \bl{ALT (SEQ (der c r$_1$) r$_2$)} & \onslide<3->{\ok}\\
- & & \bl{\phantom{ALT} (if nullable r$_1$ then der c r$_2$ else NULL)}\\
- \bl{der c (STAR r)} & \bl{$=$} & \bl{SEQ (der c r) (STAR r)} & \onslide<3->{\ok}\medskip\\
- \pause
-
- \bl{derivative r []} & \bl{$=$} & \bl{r} & \onslide<3->{\ok}\\
- \bl{derivative r (c::s)} & \bl{$=$} & \bl{derivative (der c r) s} & \onslide<3->{\ok}\\
- \end{tabular}
- \end{center}
-
- we call the program with\\
- \bl{matches r s $=$ nullable (derivative r s)}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{But Now What?}
-
- \begin{center}
- {\usefont{T1}{ptm}{b}{N}\VERYHuge{\rd{?}}}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Testing}
-
- \small
-
- \begin{center}
- \begin{tabular}{@ {\hspace{-20mm}}lcl}
- \bl{matches []$^*$ []} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches ([]$|$a)$^*$ a} & \bl{$\mapsto$} & \bl{true}\medskip\\
-
- \bl{matches (a$\cdot$b)$^*\;$ []} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches (a$\cdot$b)$^*\;$ ab} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches (a$\cdot$b)$^*\;$ aba} & \bl{$\mapsto$} & \bl{false}\\
- \bl{matches (a$\cdot$b)$^*\;$ abab} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches (a$\cdot$b)$^*\;$ abaa} & \bl{$\mapsto$} & \bl{false}\medskip\\
-
- \bl{matches x$\cdot$(0$|$1)$^*\;$ x} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches x$\cdot$(0$|$1)$^*\;$ x0} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches x$\cdot$(0$|$1)$^*\;$ x3} & \bl{$\mapsto$} & \bl{false}
- \end{tabular}
- \end{center}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Specification}
-
- We have to specify what it means for a regular expression to match
- a string.
- %
- \only<2>{
- \mbox{}\\[8mm]
- \bl{(a$\cdot$b)$^*$}\\
- \hspace{7mm}\bl{$\mapsto$\hspace{3mm}\{[], ab, abab, ababab, \ldots\}}\bigskip\\
- \bl{x$\cdot$(0 $|$ 1 $|$ 2 \ldots 8 $|$ 9 )$^*$}\\
- \hspace{7mm}\bl{$\mapsto$\hspace{3mm}
- \{x, x0, x1, \ldots, x00, \ldots, x123, \ldots\}}}
- %
- \only<3->{
- \begin{center}
- \begin{tabular}{rcl}
- \bl{\LL (NULL)} & \bl{$\dn$} & \bl{\{\}}\\
- \bl{\LL (EMPTY)} & \bl{$\dn$} & \bl{\{[]\}}\\
- \bl{\LL (CHR c)} & \bl{$\dn$} & \bl{\{c\}}\\
- \bl{\LL (ALT r$_1$ r$_2$)} & \bl{$\dn$} & \onslide<4->{\bl{\LL (r$_1$) $\cup$ \LL (r$_2$)}}\\
- \bl{\LL (SEQ r$_1$ r$_2$)} & \bl{$\dn$} & \onslide<6->{\bl{\LL (r$_1$) ; \LL (r$_2$)}}\\
- \bl{\LL (STAR r)} & \bl{$\dn$} & \onslide<8->{\bl{(\LL (r))$^\star$}}\\
- \end{tabular}
- \end{center}}
-
- \only<5-6>{
- \begin{textblock}{6}(2.9,13.3)
- \colorbox{cream}{\bl{S$_1$ ; S$_2$ $\;\dn\;$ \{ s$_1$@s$_2$ $|$ s$_1$$\in$S$_1$ $\wedge$
- s$_2$$\in$S$_2$ \}}}
- \end{textblock}}
-
- \only<7->{
- \begin{textblock}{9}(4,13)
- \colorbox{cream}{\bl{$\infer{\mbox{[]} \in \mbox{S}^\star}{}$}}\hspace{3mm}
- \colorbox{cream}{\bl{$\infer{\mbox{s}_1\mbox{@}\mbox{s}_2 \in \mbox{S}^\star}
- {\mbox{s}_1 \in \mbox{S} & \mbox{s}_2 \in \mbox{S}^\star}$}}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Is the Matcher Error-Free?}
-
- We expect that
-
- \begin{center}
- \begin{tabular}{lcl}
- \bl{matches r s = true} & \only<1>{\rd{$\Longrightarrow\,\,$}}\only<2>{\rd{$\Longleftarrow\,\,$}}%
- \only<3->{\rd{$\Longleftrightarrow$}} & \bl{s $\in$ \LL(r)}\\
- \bl{matches r s = false} & \only<1>{\rd{$\Longrightarrow\,\,$}}\only<2>{\rd{$\Longleftarrow\,\,$}}%
- \only<3->{\rd{$\Longleftrightarrow$}} & \bl{s $\notin$ \LL(r)}\\
- \end{tabular}
- \end{center}
- \pause\pause\bigskip
- By \alert<4->{induction}, we can {\bf prove} these properties.\bigskip
-
- \begin{tabular}{lrcl}
- Lemmas: & \bl{nullable (r)} & \bl{$\Longleftrightarrow$} & \bl{[] $\in$ \LL (r)}\\
- & \bl{s $\in$ \LL (der c r)} & \bl{$\Longleftrightarrow$} & \bl{(c::s) $\in$ \LL (r)}\\
- \end{tabular}
-
- \only<4->{
- \begin{textblock}{3}(0.9,4.5)
- \rd{\huge$\forall$\large{}r s.}
- \end{textblock}}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
-
- \mbox{}\\[-2mm]
-
- \small
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}ll@ {}}
- \bl{nullable (NULL)} & \bl{$=$} & \bl{false} &\\
- \bl{nullable (EMPTY)} & \bl{$=$} & \bl{true} &\\
- \bl{nullable (CHR c)} & \bl{$=$} & \bl{false} &\\
- \bl{nullable (ALT r$_1$ r$_2$)} & \bl{$=$} & \bl{(nullable r$_1$) orelse (nullable r$_2$)} & \\
- \bl{nullable (SEQ r$_1$ r$_2$)} & \bl{$=$} & \bl{(nullable r$_1$) andalso (nullable r$_2$)} & \\
- \bl{nullable (STAR r)} & \bl{$=$} & \bl{true} & \\
- \end{tabular}\medskip
-
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-10mm}}l@ {}}
- \bl{der c (NULL)} & \bl{$=$} & \bl{NULL} & \\
- \bl{der c (EMPTY)} & \bl{$=$} & \bl{NULL} & \\
- \bl{der c (CHR d)} & \bl{$=$} & \bl{if c=d then EMPTY else NULL} & \\
- \bl{der c (ALT r$_1$ r$_2$)} & \bl{$=$} & \bl{ALT (der c r$_1$) (der c r$_2$)} & \\
- \bl{der c (SEQ r$_1$ r$_2$)} & \bl{$=$} & \bl{ALT (SEQ (der c r$_1$) r$_2$)} & \\
- & & \bl{\phantom{ALT} (if nullable r$_1$ then der c r$_2$ else NULL)}\\
- \bl{der c (STAR r)} & \bl{$=$} & \bl{SEQ (der c r) (STAR r)} &\smallskip\\
-
- \bl{derivative r []} & \bl{$=$} & \bl{r} & \\
- \bl{derivative r (c::s)} & \bl{$=$} & \bl{derivative (der c r) s} & \\
- \end{tabular}\medskip
-
- \bl{matches r s $=$ nullable (derivative r s)}
-
- \only<2>{
- \begin{textblock}{8}(1.5,4)
- \includegraphics[scale=0.3]{approved.png}
- \end{textblock}}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Interlude: TCB}
-
- \begin{itemize}
- \item The \alert{\bf Trusted Code Base} (TCB) is the code that can make your
- program behave in unintended ways (i.e.~cause bugs).\medskip
-
- \item Typically the TCB includes: CPUs, operating systems, C-libraries,
- device drivers, (J)VMs\ldots\bigskip
- \pause
-
- \item It also includes the compiler.
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
-
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>
- \frametitle{\LARGE\begin{tabular}{c}Hacking Compilers
- \end{tabular}}
-
- %Why is it so paramount to have a small trusted code base (TCB)?
- \bigskip\bigskip
-
- \begin{columns}
- \begin{column}{2.7cm}
- \begin{minipage}{2.5cm}%
- \begin{tabular}{c@ {}}
- \includegraphics[scale=0.2]{ken-thompson.jpg}\\[-1.8mm]
- \footnotesize Ken Thompson\\[-1.8mm]
- \footnotesize Turing Award, 1983\\
- \end{tabular}
- \end{minipage}
- \end{column}
- \begin{column}{9cm}
- \begin{tabular}{l@ {\hspace{1mm}}p{8cm}}
- \myitemi
- & Ken Thompson showed how to hide a Trojan Horse in a
- compiler \textcolor{red}{without} leaving any traces in the source code.\\[2mm]
- \myitemi
- & No amount of source level verification will protect
- you from such Thompson-hacks.\\[2mm]
-
- \myitemi
- & Therefore in safety-critical systems it is important to rely
- on only a very small TCB.
- \end{tabular}
- \end{column}
- \end{columns}
-
- \only<2>{
- \begin{textblock}{6}(4,2)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=3mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\normalsize
- \begin{minipage}{8cm}
- \begin{quote}
- \includegraphics[scale=0.05]{evil.png}
- \begin{enumerate}
- \item[1)] Assume you ship the compiler as binary and also with sources.
- \item[2)] Make the compiler aware when it compiles itself.
- \item[3)] Add the Trojan horse.
- \item[4)] Compile.
- \item[5)] Delete Trojan horse from the sources of the compiler.
- \item[6)] Go on holiday for the rest of your life. ;o)\\[-7mm]\mbox{}
- \end{enumerate}
- \end{quote}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}
- \frametitle{\LARGE\begin{tabular}{c}An Example: Small TCB for\\[-2mm]
- A Critical Infrastructure\end{tabular}}
- \mbox{}\\[-14mm]\mbox{}
-
- \begin{columns}
- \begin{column}{0.2\textwidth}
- \begin{tabular}{@ {} c@ {}}
- \includegraphics[scale=0.3]{appel.jpg}\\[-2mm]
- {\footnotesize Andrew Appel}\\[-2.5mm]
- {\footnotesize (Princeton)}
- \end{tabular}
- \end{column}
-
- \begin{column}{0.8\textwidth}
- \begin{textblock}{10}(4.5,3.95)
- \begin{block}{Proof-Carrying Code}
- \begin{center}
- \begin{tikzpicture}
- \draw[help lines,cream] (0,0.2) grid (8,4);
-
- \draw[line width=1mm, red] (5.5,0.6) rectangle (7.5,4);
- \node[anchor=base] at (6.5,2.8)
- {\small\begin{tabular}{@ {}p{1.9cm}@ {}}\centering user needs to run untrusted code\end{tabular}};
-
- \draw[line width=1mm, red] (0.5,0.6) rectangle (2.5,4);
- \node[anchor=base] at (1.5,2.3)
- {\small\begin{tabular}{@ {}p{1.9cm}@ {}}\centering code developer/ web server/ Apple
- Store\end{tabular}};
-
- \onslide<4->{
- \draw[line width=1mm, red, fill=red] (5.5,0.6) rectangle (7.5,1.8);
- \node[anchor=base,white] at (6.5,1.1)
- {\small\begin{tabular}{@ {}p{1.9cm}@ {}}\bf\centering proof- checker\end{tabular}};}
-
- \node at (3.8,3.0) [single arrow, fill=red,text=white, minimum height=3cm]{\bf code};
- \onslide<3->{
- \node at (3.8,1.3) [single arrow, fill=red,text=white, minimum height=3cm]{\bf LF proof};
- \node at (3.8,1.9) {\small certificate};
- }
-
- \onslide<2>{\node at (4.0,1.3) [text=red]{\begin{tabular}{c}\bf Highly\\\bf Dangerous!\end{tabular}};}
- % Code Developer
- % User (runs untrusted code)
- % transmits a proof that the code is safe
- %
- \end{tikzpicture}
- \end{center}
- \end{block}
- \end{textblock}
- \end{column}
- \end{columns}
-
- \small\mbox{}\\[2.5cm]
- \begin{itemize}
- \item<4-> TCB of the checker is $\sim$2700 lines of code (1865 loc of\\ LF definitions;
- 803 loc in C including 2 library functions)\\[-3mm]
- \item<5-> 167 loc in C implement a type-checker
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-
-
-text {*
- \tikzstyle{every node}=[node distance=25mm,text height=1.5ex, text depth=.25ex]
- \tikzstyle{node1}=[rectangle, minimum size=10mm, rounded corners=3mm, very thick,
- draw=black!50, top color=white, bottom color=black!20]
- \tikzstyle{node2}=[rectangle, minimum size=12mm, rounded corners=3mm, very thick,
- draw=red!70, top color=white, bottom color=red!50!black!20]
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[squeeze]
- \frametitle{Type-Checking in LF}
-
- \begin{columns}
- \begin{column}{0.2\textwidth}
- \begin{tabular}{@ {\hspace{-4mm}}c@ {}}
- \\[-4mm]
- \includegraphics[scale=0.1]{harper.jpg}\\[-2mm]
- {\footnotesize Bob Harper}\\[-2.5mm]
- {\footnotesize (CMU)}\\[2mm]
-
- \includegraphics[scale=0.3]{pfenning.jpg}\\[-2mm]
- {\footnotesize Frank Pfenning}\\[-2.5mm]
- {\footnotesize (CMU)}\\[2mm]
-
- \onslide<-6>{
- {\footnotesize 31 pages in }\\[-2.5mm]
- {\footnotesize ACM Transact.~on}\\[-2.5mm]
- {\footnotesize Comp.~Logic.,~2005}\\}
- \end{tabular}
- \end{column}
-
- \begin{column}{0.8\textwidth}
- \begin{textblock}{0}(3.1,2)
-
- \begin{tikzpicture}
- \matrix[ampersand replacement=\&,column sep=7mm, row sep=5mm]
- { \&[-10mm]
- \node (def1) [node1] {\large\hspace{1mm}Spec\hspace{1mm}\mbox{}}; \&
- \node (proof1) [node1] {\large Proof}; \&
- \node (alg1) [node1] {\large\hspace{1mm}Alg\hspace{1mm}\mbox{}}; \\
-
- \onslide<4->{\node {\begin{tabular}{c}\small 1st\\[-2.5mm] \footnotesize solution\end{tabular}};} \&
- \onslide<4->{\node (def2) [node2] {\large Spec$^\text{+ex}$};} \&
- \onslide<4->{\node (proof2) [node1] {\large Proof};} \&
- \onslide<4->{\node (alg2) [node1] {\large\hspace{1mm}Alg\hspace{1mm}\mbox{}};} \\
-
- \onslide<5->{\node {\begin{tabular}{c}\small 2nd\\[-2.5mm] \footnotesize solution\end{tabular}};} \&
- \onslide<5->{\node (def3) [node1] {\large\hspace{1mm}Spec\hspace{1mm}\mbox{}};} \&
- \onslide<5->{\node (proof3) [node1] {\large Proof};} \&
- \onslide<5->{\node (alg3) [node2] {\large Alg$^\text{-ex}$};} \\
-
- \onslide<6->{\node {\begin{tabular}{c}\small 3rd\\[-2.5mm] \footnotesize solution\end{tabular}};} \&
- \onslide<6->{\node (def4) [node1] {\large\hspace{1mm}Spec\hspace{1mm}\mbox{}};} \&
- \onslide<6->{\node (proof4) [node2] {\large\hspace{1mm}Proof\hspace{1mm}};} \&
- \onslide<6->{\node (alg4) [node1] {\large\hspace{1mm}Alg\hspace{1mm}\mbox{}};} \\
- };
-
- \draw[->,black!50,line width=2mm] (proof1) -- (def1);
- \draw[->,black!50,line width=2mm] (proof1) -- (alg1);
-
- \onslide<4->{\draw[->,black!50,line width=2mm] (proof2) -- (def2);}
- \onslide<4->{\draw[->,black!50,line width=2mm] (proof2) -- (alg2);}
-
- \onslide<5->{\draw[->,black!50,line width=2mm] (proof3) -- (def3);}
- \onslide<5->{\draw[->,black!50,line width=2mm] (proof3) -- (alg3);}
-
- \onslide<6->{\draw[->,black!50,line width=2mm] (proof4) -- (def4);}
- \onslide<6->{\draw[->,black!50,line width=2mm] (proof4) -- (alg4);}
-
- \onslide<3->{\draw[white,line width=1mm] (1.1,3.2) -- (0.9,2.85) -- (1.1,2.35) -- (0.9,2.0);}
- \end{tikzpicture}
-
- \end{textblock}
- \end{column}
- \end{columns}
-
- \only<2>{%
- \begin{textblock}{2}(.1,12.85)
- \begin{tikzpicture}
- \draw[line width=1mm, red] (0,0) ellipse (1.5cm and 0.88cm);
- \end{tikzpicture}
- \end{textblock}
- }
-
- \begin{textblock}{3}(14,3.6)
- \onslide<4->{
- \begin{tikzpicture}
- \node at (0,0) [single arrow, shape border rotate=270, fill=red,text=white]{2h};
- \end{tikzpicture}}
- \end{textblock}
-
- \only<7->{
- \begin{textblock}{14}(0.6,12.8)
- \begin{block}{}
- \small Each time one needs to check $\sim$31pp~of informal paper proofs.
- You have to be able to keep definitions and proofs consistent.
- \end{block}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{Theorem Provers}
-
- \begin{itemize}
- \item Theorem provers help with keeping large proofs consistent;
- make them modifiable.\medskip
-
- \item They can ensure that all cases are covered.\medskip
-
- \item Sometimes, tedious reasoning can be automated.
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{Theorem Provers}
-
- \begin{itemize}
- \item You also pay a (sometimes heavy) price: reasoning can be much, much harder.\medskip
-
- \item Depending on your domain, suitable reasoning infrastructure
- might not yet be available.
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{Theorem Provers}
-
- Recently impressive work has been accomplished proving the correctness
-
- \begin{itemize}
- \item of a compiler for C-light (compiled code has the same observable
- behaviour as the original source code),\medskip
-
- \item a mirco-kernel operating system (absence of certain
- bugs\ldots no nil pointers, no buffer overflows).
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{Trust in Theorem Provers}
-
- \begin{center}
- Why should we trust theorem provers?
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
-
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}
- \frametitle{Theorem Provers}
-
- \begin{itemize}
- \item Theorem provers are a \textcolor{red}{special kind} of software.
-
- \item We do \textcolor{red}{\bf not} need to trust them; we only need to trust:
- \end{itemize}
-
- \begin{quote}
- \begin{itemize}
- \item The logic they are based on \textcolor{gray}{(e.g.~HOL)}, and\smallskip
- \item a proof checker that checks the proofs
- \textcolor{gray}{(this can be a very small program)}.\smallskip\pause
- \item To a little extend, we also need to trust our definitions
- \textcolor{gray}{(this can be mitigated)}.
- \end{itemize}
- \end{quote}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
-
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}
- \frametitle{Isabelle}
-
- \begin{itemize}
- \item I am using the Isabelle theorem prover (development since 1990).\bigskip\bigskip\bigskip
-
- \item It follows the LCF-approach:
-
- \begin{itemize}
- \item Have a special abstract type \alert{\bf thm}.
- \item Make the constructors of this abstract type the inference rules
- of your logic.
- \item Implement the theorem prover in a strongly-typed language (e.g.~ML).
- \end{itemize}
-
- $\Rightarrow$ everything of type {\bf thm} has been proved (even if we do not
- have to explicitly generate proofs).
- \end{itemize}
-
- \only<1>{
- \begin{textblock}{5}(11,2.3)
- \begin{center}
- \includegraphics[scale=0.18]{robin-milner.jpg}\\[-0.8mm]
- \footnotesize Robin Milner\\[-0.8mm]
- \footnotesize Turing Award, 1991\\
- \end{center}
- \end{textblock}}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{
- \begin{tabular}{c}
- \mbox{}\\[23mm]
- \LARGE Demo
- \end{tabular}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Future Research}
-
- \begin{itemize}
- \item Make theorem provers more like a programming environment.\medskip\pause
-
- \item Use all the computational power we get from the hardware to
- automate reasoning (GPUs).\medskip\pause
-
- \item Provide a comprehensive reasoning infrastructure for many domains and
- design automated decision procedures.
- \end{itemize}\pause
-
-
- \begin{center}
- \colorbox{cream}{
- \begin{minipage}{10cm}
- \color{gray}
- \small
- ``Formal methods will never have a significant impact until
- they can be used by people that don't understand them.''\smallskip\\
- \mbox{}\footnotesize\hfill attributed to Tom Melham
- \end{minipage}}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Conclusion}
-
- \begin{itemize}
- \item The plan is to make this kind of programming the ``future''.\medskip\pause
-
- \item Though the technology is already there\\ (compiler + micro-kernel os).\medskip\pause
-
- \item Logic and reasoning (especially induction) are important skills for
- Computer Scientists.
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{
- \begin{tabular}{c}
- \mbox{}\\[23mm]
- \alert{\LARGE Thank you very much!}\\
- \alert{\Large Questions?}
- \end{tabular}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-
-(*<*)
-end
-(*>*)
\ No newline at end of file
--- a/Slides/Slides5.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,784 +0,0 @@
-(*<*)
-theory Slides5
-imports "~~/src/HOL/Library/LaTeXsugar" "Nominal"
-begin
-
-notation (latex output)
- set ("_") and
- Cons ("_::/_" [66,65] 65)
-
-(*>*)
-
-
-text_raw {*
- %% shallow, deep, and recursive binders
- %%
- %%\renewcommand{\slidecaption}{Cambridge, 8.~June 2010}
- %%\renewcommand{\slidecaption}{Uppsala, 3.~March 2011}
- \renewcommand{\slidecaption}{Saarbrücken, 31.~March 2011}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[t]
- \frametitle{%
- \begin{tabular}{@ {\hspace{-3mm}}c@ {}}
- \\
- \LARGE General Bindings and\\
- \LARGE Alpha-Equivalence\\
- \LARGE in Nominal Isabelle\\[3mm]
- \Large Or, Nominal Isabelle 2\\[-5mm]
- \end{tabular}}
- \begin{center}
- Christian Urban
- \end{center}
- \begin{center}
- joint work with {\bf Cezary Kaliszyk}\\[0mm]
- \end{center}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>
- \frametitle{\begin{tabular}{c}Binding in Old Nominal\end{tabular}}
- \mbox{}\\[-6mm]
-
- \begin{itemize}
- \item the old Nominal Isabelle provided a reasoning infrastructure for single binders\medskip
-
- \begin{center}
- Lam [a].(Var a)
- \end{center}\bigskip
-
- \item<2-> but representing
-
- \begin{center}
- $\forall\{a_1,\ldots,a_n\}.\; T$
- \end{center}\medskip
-
- with single binders and reasoning about it is a \alert{\bf major} pain;
- take my word for it!
- \end{itemize}
-
- \only<1>{
- \begin{textblock}{6}(1.5,11)
- \small
- for example\\
- \begin{tabular}{l@ {\hspace{2mm}}l}
- & a $\fresh$ Lam [a]. t\\
- & Lam [a]. (Var a) \alert{$=$} Lam [b]. (Var b)\\
- & Barendregt-style reasoning about bound variables\\
- \end{tabular}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-6>
- \frametitle{New Types in HOL}
-
- \begin{center}
- \begin{tikzpicture}[scale=1.5]
- %%%\draw[step=2mm] (-4,-1) grid (4,1);
-
- \onslide<2-4,6>{\draw[very thick] (0.7,0.4) circle (4.25mm);}
- \onslide<1-4,6>{\draw[rounded corners=1mm, very thick] ( 0.0,-0.8) rectangle ( 1.8, 0.9);}
- \onslide<3-5,6>{\draw[rounded corners=1mm, very thick] (-1.95,0.85) rectangle (-2.85,-0.05);}
-
- \onslide<3-4,6>{\draw (-2.0, 0.845) -- (0.7,0.845);}
- \onslide<3-4,6>{\draw (-2.0,-0.045) -- (0.7,-0.045);}
-
- \onslide<4-4,6>{\alert{\draw ( 0.7, 0.4) node {\footnotesize\begin{tabular}{@ {}c@ {}}$\alpha$-\\[-1mm]classes\end{tabular}};}}
- \onslide<4-5,6>{\alert{\draw (-2.4, 0.4) node {\footnotesize\begin{tabular}{@ {}c@ {}}$\alpha$-eq.\\[-1mm]terms\end{tabular}};}}
- \onslide<1-4,6>{\draw (1.8, 0.48) node[right=-0.1mm]
- {\footnotesize\begin{tabular}{@ {}l@ {}}existing\\[-1mm] type\\ \onslide<4-4,6>{\alert{(sets of raw terms)}}\end{tabular}};}
- \onslide<2-4,6>{\draw (0.9, -0.35) node {\footnotesize\begin{tabular}{@ {}l@ {}}non-empty\\[-1mm]subset\end{tabular}};}
- \onslide<3-5,6>{\draw (-3.25, 0.55) node {\footnotesize\begin{tabular}{@ {}l@ {}}new\\[-1mm]type\end{tabular}};}
-
- \onslide<3-4,6>{\draw[<->, very thick] (-1.8, 0.3) -- (-0.1,0.3);}
- \onslide<3-4,6>{\draw (-0.95, 0.3) node[above=0mm] {\footnotesize{}isomorphism};}
-
- \onslide<6>{\draw[->, line width=2mm, red] (-1.0,-0.4) -- (0.35,0.16);}
- \end{tikzpicture}
- \end{center}
-
- \begin{center}
- \textcolor{red}{\large\bf\onslide<6>{define $\alpha$-equivalence}}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-4>
- \frametitle{\begin{tabular}{c}Binding Sets of Names\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item binding sets of names has some interesting properties:\medskip
-
- \begin{center}
- \begin{tabular}{l}
- \textcolor{blue}{$\forall\{x, y\}.\, x \rightarrow y \;\;\approx_\alpha\;\; \forall\{y, x\}.\, y \rightarrow x$}
- \bigskip\smallskip\\
-
- \onslide<2->{%
- \textcolor{blue}{$\forall\{x, y\}.\, x \rightarrow y \;\;\not\approx_\alpha\;\; \forall\{z\}.\, z \rightarrow z$}
- }\bigskip\smallskip\\
-
- \onslide<3->{%
- \textcolor{blue}{$\forall\{x\}.\, x \rightarrow y \;\;\approx_\alpha\;\; \forall\{x, \alert{z}\}.\, x \rightarrow y$}
- }\medskip\\
- \onslide<3->{\hspace{4cm}\small provided $z$ is fresh for the type}
- \end{tabular}
- \end{center}
- \end{itemize}
-
- \begin{textblock}{8}(2,14.5)
- \footnotesize $^*$ $x$, $y$, $z$ are assumed to be distinct
- \end{textblock}
-
- \only<4>{
- \begin{textblock}{6}(2.5,4)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=3mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\normalsize\color{darkgray}
- \begin{minipage}{8cm}\raggedright
- For type-schemes the order of bound names does not matter, and
- $\alpha$-equivalence is preserved under \alert{vacuous} binders.
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>
- \frametitle{\begin{tabular}{c}Other Binding Modes\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item alpha-equivalence being preserved under vacuous binders is \underline{not} always
- wanted:\bigskip\bigskip\normalsize
-
- \textcolor{blue}{\begin{tabular}{@ {\hspace{-8mm}}l}
- $\text{let}\;x = 3\;\text{and}\;y = 2\;\text{in}\;x - y\;\text{end}$\medskip\\
- \onslide<2->{$\;\;\;\only<2>{\approx_\alpha}\only<3>{\alert{\not\approx_\alpha}}
- \text{let}\;y = 2\;\text{and}\;x = 3\only<3->{\alert{\;\text{and}
- \;z = \text{loop}}}\;\text{in}\;x - y\;\text{end}$}
- \end{tabular}}
-
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>
- \frametitle{\begin{tabular}{c}\LARGE{}Even Another Binding Mode\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item sometimes one wants to abstract more than one name, but the order \underline{does} matter\bigskip
-
- \begin{center}
- \textcolor{blue}{\begin{tabular}{@ {\hspace{-8mm}}l}
- $\text{let}\;(x, y) = (3, 2)\;\text{in}\;x - y\;\text{end}$\medskip\\
- $\;\;\;\not\approx_\alpha
- \text{let}\;(y, x) = (3, 2)\;\text{in}\;x - y\;\text{end}$
- \end{tabular}}
- \end{center}
-
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>
- \frametitle{\begin{tabular}{c}\LARGE{}Three Binding Modes\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item the order does not matter and alpha-equivelence is preserved under
- vacuous binders \textcolor{gray}{(restriction)}\medskip
-
- \item the order does not matter, but the cardinality of the binders
- must be the same \textcolor{gray}{(abstraction)}\medskip
-
- \item the order does matter \textcolor{gray}{(iterated single binders)}
- \end{itemize}
-
- \onslide<2->{
- \begin{center}
- \isacommand{bind (set+)}\hspace{6mm}
- \isacommand{bind (set)}\hspace{6mm}
- \isacommand{bind}
- \end{center}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>
- \frametitle{\begin{tabular}{c}Specification of Binding\end{tabular}}
- \mbox{}\\[-6mm]
-
- \mbox{}\hspace{10mm}
- \begin{tabular}{ll}
- \multicolumn{2}{l}{\isacommand{nominal\_datatype} trm $=$}\\
- \hspace{5mm}\phantom{$|$} Var name\\
- \hspace{5mm}$|$ App trm trm\\
- \hspace{5mm}$|$ Lam \only<2->{x::}name \only<2->{t::}trm
- & \onslide<2->{\isacommand{bind} x \isacommand{in} t}\\
- \hspace{5mm}$|$ Let \only<2->{as::}assns \only<2->{t::}trm
- & \onslide<2->{\isacommand{bind} bn(as) \isacommand{in} t}\\
- \multicolumn{2}{l}{\isacommand{and} assns $=$}\\
- \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} ANil}\\
- \multicolumn{2}{l}{\hspace{5mm}$|$ ACons name trm assns}\\
- \multicolumn{2}{l}{\onslide<3->{\isacommand{binder} bn \isacommand{where}}}\\
- \multicolumn{2}{l}{\onslide<3->{\hspace{5mm}\phantom{$|$} bn(ANil) $=$ []}}\\
- \multicolumn{2}{l}{\onslide<3->{\hspace{5mm}$|$ bn(ACons a t as) $=$ [a] @ bn(as)}}\\
- \end{tabular}
-
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-8>
- \frametitle{\begin{tabular}{c}Alpha-Equivalence\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item lets first look at pairs\bigskip\medskip
-
- \textcolor{blue}{\begin{tabular}{@ {\hspace{1cm}}l}
- $(as, x) \onslide<2->{\approx\!}\makebox[5mm][l]{\only<2-6>{${}_{\text{set}}$}%
- \only<7>{${}_{\text{\alert{list}}}$}%
- \only<8>{${}_{\text{\alert{set+}}}$}}%
- \,\onslide<2->{(bs,y)}$
- \end{tabular}}\bigskip
- \end{itemize}
-
- \only<1>{
- \begin{textblock}{8}(3,8.5)
- \begin{tabular}{l@ {\hspace{2mm}}p{8cm}}
- & \textcolor{blue}{$as$} is a set of names\ldots the binders\\
- & \textcolor{blue}{$x$} is the body (might be a tuple)\\
- & \textcolor{blue}{$\approx_{\text{set}}$} is where the cardinality
- of the binders has to be the same\\
- \end{tabular}
- \end{textblock}}
-
- \only<4->{
- \begin{textblock}{12}(5,8)
- \textcolor{blue}{
- \begin{tabular}{ll@ {\hspace{1mm}}l}
- $\dn$ & \onslide<5->{$\exists \pi.\,$} & $\text{fv}(x) - as = \text{fv}(y) - bs$\\[1mm]
- & \onslide<5->{$\;\;\;\wedge$} & \onslide<5->{$\text{fv}(x) - as \fresh^* \pi$}\\[1mm]
- & \onslide<5->{$\;\;\;\wedge$} & \onslide<5->{$(\pi \act x) = y$}\\[1mm]
- & \only<6-7>{$\;\;\;\wedge$}\only<8>{\textcolor{gray}{\xout{$\;\;\;\wedge$}}} &
- \only<6-7>{$\pi \act as = bs$}\only<8>{\textcolor{gray}{\xout{$\pi \act as = bs$}}}\\
- \end{tabular}}
- \end{textblock}}
-
- \only<7>{
- \begin{textblock}{7}(3,13.8)
- \footnotesize $^*$ $as$ and $bs$ are \alert{lists} of names
- \end{textblock}}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-3>
- \frametitle{\begin{tabular}{c}Examples\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{itemize}
- \item lets look at type-schemes:\medskip\medskip
-
- \begin{center}
- \textcolor{blue}{$(as, x) \approx\!\makebox[5mm][l]{${}_{\text{set}}$} (bs, y)$}
- \end{center}\medskip
-
- \onslide<2->{
- \begin{center}
- \textcolor{blue}{
- \begin{tabular}{l}
- $\text{fv}(x) = \{x\}$\\[1mm]
- $\text{fv}(T_1 \rightarrow T_2) = \text{fv}(T_1) \cup \text{fv}(T_2)$\\
- \end{tabular}}
- \end{center}}
- \end{itemize}
-
-
- \only<3->{
- \begin{textblock}{4}(0.3,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{set+:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- \\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
- \only<3->{
- \begin{textblock}{4}(5.2,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{set:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- $\wedge$ & $\pi \cdot as = bs$\\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
- \only<3->{
- \begin{textblock}{4}(10.2,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{list:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- $\wedge$ & $\pi \cdot as = bs$\\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>
- \frametitle{\begin{tabular}{c}Examples\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{center}
- \textcolor{blue}{
- \only<1>{$(\{x, y\}, x \rightarrow y) \approx_? (\{x, y\}, y \rightarrow x)$}
- \only<2>{$([x, y], x \rightarrow y) \approx_? ([x, y], y \rightarrow x)$}}
- \end{center}
-
- \begin{itemize}
- \item \textcolor{blue}{$\approx_{\text{set+}}$, $\approx_{\text{set}}$%
- \only<2>{, \alert{$\not\approx_{\text{list}}$}}}
- \end{itemize}
-
-
- \only<1->{
- \begin{textblock}{4}(0.3,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{set+:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- \\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
- \only<1->{
- \begin{textblock}{4}(5.2,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{set:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- $\wedge$ & $\pi \cdot as = bs$\\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
- \only<1->{
- \begin{textblock}{4}(10.2,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{list:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- $\wedge$ & $\pi \cdot as = bs$\\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>
- \frametitle{\begin{tabular}{c}Examples\end{tabular}}
- \mbox{}\\[-3mm]
-
- \begin{center}
- \textcolor{blue}{\only<1>{$(\{x\}, x) \approx_? (\{x, y\}, x)$}}
- \end{center}
-
- \begin{itemize}
- \item \textcolor{blue}{$\approx_{\text{set+}}$, $\not\approx_{\text{set}}$,
- $\not\approx_{\text{list}}$}
- \end{itemize}
-
-
- \only<1->{
- \begin{textblock}{4}(0.3,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{set+:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- \\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
- \only<1->{
- \begin{textblock}{4}(5.2,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{set:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- $\wedge$ & $\pi \cdot as = bs$\\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
- \only<1->{
- \begin{textblock}{4}(10.2,12)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\tiny\color{darkgray}
- \begin{minipage}{3.4cm}\raggedright
- \begin{tabular}{r@ {\hspace{1mm}}l}
- \multicolumn{2}{@ {}l}{list:}\\
- $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
- $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
- $\wedge$ & $\pi \cdot x = y$\\
- $\wedge$ & $\pi \cdot as = bs$\\
- \end{tabular}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \only<2>{
- \begin{textblock}{6}(2.5,4)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=5mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\normalsize
- \begin{minipage}{8cm}\raggedright
- \begin{itemize}
- \item \color{darkgray}$\alpha$-equivalences coincide when a single name is
- abstracted
- \item \color{darkgray}in that case they are equivalent to ``old-fashioned'' definitions of $\alpha$
- \end{itemize}
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{\begin{tabular}{c}Our Specifications\end{tabular}}
- \mbox{}\\[-6mm]
-
- \mbox{}\hspace{10mm}
- \begin{tabular}{ll}
- \multicolumn{2}{l}{\isacommand{nominal\_datatype} trm $=$}\\
- \hspace{5mm}\phantom{$|$} Var name\\
- \hspace{5mm}$|$ App trm trm\\
- \hspace{5mm}$|$ Lam x::name t::trm
- & \isacommand{bind} x \isacommand{in} t\\
- \hspace{5mm}$|$ Let as::assns t::trm
- & \isacommand{bind} bn(as) \isacommand{in} t\\
- \multicolumn{2}{l}{\isacommand{and} assns $=$}\\
- \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} ANil}\\
- \multicolumn{2}{l}{\hspace{5mm}$|$ ACons name trm assns}\\
- \multicolumn{2}{l}{\isacommand{binder} bn \isacommand{where}}\\
- \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} bn(ANil) $=$ $[]$}\\
- \multicolumn{2}{l}{\hspace{5mm}$|$ bn(ACons a t as) $=$ $[$a$]$ @ bn(as)}\\
- \end{tabular}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{\begin{tabular}{c}Binding Functions\end{tabular}}
-
- \begin{center}
- \begin{tikzpicture}
- \node (A) at (-0.5,1) {Foo $(\lambda y. \lambda x. t)$};
- \node (B) at ( 1.5,1) {$s$};
- \onslide<1>{\node (C) at (0.5,-0.5) {$\{y, x\}$};}
- \onslide<1>{\draw[->,red,line width=1mm] (A) -- (C);}
- \onslide<1>{\draw[->,red,line width=1mm] (C) -- (B);}
- \end{tikzpicture}
- \end{center}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{\begin{tabular}{c}Binder Clauses\end{tabular}}
-
- \begin{itemize}
- \item We need for a bound variable to have a `clear scope', and bound
- variables should not be free and bound at the same time.\bigskip
- \end{itemize}
-
- \begin{center}
- \only<1>{
- \begin{tabular}{@ {\hspace{-5mm}}l}
- \alert{\bf shallow binders}\\
- \hspace{4mm}Lam x::name t::trm\hspace{4mm} \isacommand{bind} x \isacommand{in} t\\
- \hspace{4mm}All xs::name set T::ty\hspace{4mm} \isacommand{bind} xs \isacommand{in} T\\
- \hspace{4mm}Foo x::name t$_1$::trm t$_2$::trm\hspace{4mm}
- \isacommand{bind} x \isacommand{in} t$_1$, \isacommand{bind} x \isacommand{in} t$_2$\\
- \hspace{4mm}Bar x::name t$_1$::trm t$_2$::trm\hspace{4mm}
- \isacommand{bind} x \isacommand{in} t$_1$ t$_2$\\
- \end{tabular}}
- \only<2>{
- \begin{tabular}{@ {\hspace{-5mm}}l}
- \alert{\bf deep binders} \\
- \hspace{4mm}Let as::assns t::trm\hspace{4mm} \isacommand{bind} bn(as) \isacommand{in} t\\
- \hspace{4mm}Foo as::assns t$_1$::trm t$_2$::trm\\
- \hspace{20mm}\isacommand{bind} bn(as) \isacommand{in} t$_1$, \isacommand{bind} bn(as) \isacommand{in} t$_2$\\[4mm]
- \makebox[0mm][l]{\alert{$\times$}}\hspace{4mm}Bar as::assns t$_1$::trm t$_2$::trm\\
- \hspace{20mm}\isacommand{bind} bn$_1$(as) \isacommand{in} t$_1$, \isacommand{bind} bn$_2$(as) \isacommand{in} t$_2$\\
- \end{tabular}}
- \only<3>{
- \begin{tabular}{@ {\hspace{-5mm}}l}
- {\bf deep \alert{recursive} binders} \\
- \hspace{4mm}Let\_rec as::assns t::trm\hspace{4mm} \isacommand{bind} bn(as) \isacommand{in} t as\\[4mm]
-
- \makebox[0mm][l]{\alert{$\times$}}\hspace{4mm}Foo\_rec as::assns t$_1$::trm t$_2$::trm\hspace{4mm}\\
- \hspace{20mm}\isacommand{bind} bn(as) \isacommand{in} t$_1$ as, \isacommand{bind} bn(as) \isacommand{in} t$_2$\\
-
- \end{tabular}}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<2-5>
- \frametitle{\begin{tabular}{c}Our Work\end{tabular}}
- \mbox{}\\[-6mm]
-
- \begin{center}
- \begin{tikzpicture}[scale=1.5]
- %%%\draw[step=2mm] (-4,-1) grid (4,1);
-
- \onslide<1>{\draw[very thick] (0.7,0.4) circle (4.25mm);}
- \onslide<1>{\draw[rounded corners=1mm, very thick] ( 0.0,-0.8) rectangle ( 1.8, 0.9);}
- \onslide<1->{\draw[rounded corners=1mm, very thick] (-1.95,0.85) rectangle (-2.85,-0.05);}
-
- \onslide<1>{\draw (-2.0, 0.845) -- (0.7,0.845);}
- \onslide<1>{\draw (-2.0,-0.045) -- (0.7,-0.045);}
-
- \onslide<1>{\alert{\draw ( 0.7, 0.4) node {\footnotesize\begin{tabular}{@ {}c@ {}}$\alpha$-\\[-1mm]classes\end{tabular}};}}
- \onslide<1->{\alert{\draw (-2.4, 0.4) node {\footnotesize\begin{tabular}{@ {}c@ {}}$\alpha$-eq.\\[-1mm]terms\end{tabular}};}}
- \onslide<1>{\draw (1.8, 0.48) node[right=-0.1mm]
- {\footnotesize\begin{tabular}{@ {}l@ {}}existing\\[-1mm] type\\ \onslide<1>{\alert{(sets of raw terms)}}\end{tabular}};}
- \onslide<1>{\draw (0.9, -0.35) node {\footnotesize\begin{tabular}{@ {}l@ {}}non-empty\\[-1mm]subset\end{tabular}};}
- \onslide<1->{\draw (-3.25, 0.55) node {\footnotesize\begin{tabular}{@ {}l@ {}}new\\[-1mm]type\end{tabular}};}
-
- \onslide<1>{\draw[<->, very thick] (-1.8, 0.3) -- (-0.1,0.3);}
- \onslide<1>{\draw (-0.95, 0.3) node[above=0mm] {\footnotesize{}isomorphism};}
-
- \onslide<1>{\draw[->, line width=2mm, red] (-1.0,-0.4) -- (0.35,0.16);}
- \end{tikzpicture}
- \end{center}
-
- \begin{textblock}{9.5}(6,3.5)
- \begin{itemize}
- \item<1-> defined fv and $\alpha$
- \item<3-> derived a reasoning infrastructure ($\fresh$, distinctness, injectivity, cases,\ldots)
- \item<4-> a (weak) induction principle
- \item<5-> derive a {\bf stronger} induction principle (Barendregt variable convention built in)\\
- \begin{center}
- \textcolor{blue}{Foo ($\lambda x. \lambda y. t$) ($\lambda u. \lambda v. s$)}
- \end{center}
- \end{itemize}
- \end{textblock}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->
- \frametitle{\begin{tabular}{c}Conclusion\end{tabular}}
- \mbox{}\\[-6mm]
-
- \begin{itemize}
- \item the user does not see anything of the raw level\medskip
- \only<1>{\begin{center}
- Lam a (Var a) \alert{$=$} Lam b (Var b)
- \end{center}\bigskip}
-
- \item<2-> it took quite some time to get here, but it seems worthwhile
- (Barendregt's variable convention is unsound in general,
- found bugs in two paper proofs)\bigskip\medskip
-
- \item<3-> \textcolor{blue}{http://isabelle.in.tum.de/nominal/}
- \end{itemize}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{\begin{tabular}{c}Questions?\end{tabular}}
- \mbox{}\\[-6mm]
-
- \begin{center}
- \alert{\huge{Thanks!}}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1-2>[c]
- \frametitle{\begin{tabular}{c}Examples\end{tabular}}
- \mbox{}\\[-6mm]
-
- \textcolor{blue}{
- \begin{center}
- $(\{a,b\}, a \rightarrow b) \approx_\alpha (\{a, b\}, a \rightarrow b)$
- $(\{a,b\}, a \rightarrow b) \approx_\alpha (\{a, b\}, b \rightarrow a)$
- \end{center}}
-
- \textcolor{blue}{
- \begin{center}
- $(\{a,b\}, (a \rightarrow b, a \rightarrow b))$\\
- \hspace{17mm}$\not\approx_\alpha (\{a, b\}, (a \rightarrow b, b \rightarrow a))$
- \end{center}}
-
- \onslide<2->
- {1.) \hspace{3mm}\isacommand{bind (set)} as \isacommand{in} $\tau_1$,
- \isacommand{bind (set)} as \isacommand{in} $\tau_2$\medskip
-
- 2.) \hspace{3mm}\isacommand{bind (set)} as \isacommand{in} $\tau_1$ $\tau_2$
- }
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-
-(*<*)
-end
-(*>*)
\ No newline at end of file
--- a/Slides/Slides6.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1606 +0,0 @@
-(*<*)
-theory Slides6
-imports "~~/src/HOL/Library/LaTeXsugar" "Nominal"
-begin
-
-declare [[show_question_marks = false]]
-
-notation (latex output)
- set ("_") and
- Cons ("_::/_" [66,65] 65)
-
-(*>*)
-
-text_raw {*
- \renewcommand{\slidecaption}{Hefei, 15.~April 2011}
-
- \newcommand{\abst}[2]{#1.#2}% atom-abstraction
- \newcommand{\pair}[2]{\langle #1,#2\rangle} % pairing
- \newcommand{\susp}{{\boldsymbol{\cdot}}}% for suspensions
- \newcommand{\unit}{\langle\rangle}% unit
- \newcommand{\app}[2]{#1\,#2}% application
- \newcommand{\eqprob}{\mathrel{{\approx}?}}
- \newcommand{\freshprob}{\mathrel{\#?}}
- \newcommand{\redu}[1]{\stackrel{#1}{\Longrightarrow}}% reduction
- \newcommand{\id}{\varepsilon}% identity substitution
-
- \newcommand{\bl}[1]{\textcolor{blue}{#1}}
- \newcommand{\gr}[1]{\textcolor{gray}{#1}}
- \newcommand{\rd}[1]{\textcolor{red}{#1}}
-
- \newcommand{\ok}{\includegraphics[scale=0.07]{ok.png}}
- \newcommand{\notok}{\includegraphics[scale=0.07]{notok.png}}
- \newcommand{\largenotok}{\includegraphics[scale=1]{notok.png}}
-
- \renewcommand{\Huge}{\fontsize{61.92}{77}\selectfont}
- \newcommand{\veryHuge}{\fontsize{74.3}{93}\selectfont}
- \newcommand{\VeryHuge}{\fontsize{89.16}{112}\selectfont}
- \newcommand{\VERYHuge}{\fontsize{107}{134}\selectfont}
-
- \newcommand{\LL}{$\mathbb{L}\,$}
-
-
- \pgfdeclareradialshading{smallbluesphere}{\pgfpoint{0.5mm}{0.5mm}}%
- {rgb(0mm)=(0,0,0.9);
- rgb(0.9mm)=(0,0,0.7);
- rgb(1.3mm)=(0,0,0.5);
- rgb(1.4mm)=(1,1,1)}
-
- \def\myitemi{\begin{pgfpicture}{-1ex}{-0.55ex}{1ex}{1ex}
- \usebeamercolor[fg]{subitem projected}
- {\pgftransformscale{0.8}\pgftext{\normalsize\pgfuseshading{bigsphere}}}
- \pgftext{%
- \usebeamerfont*{subitem projected}}
- \end{pgfpicture}}
-
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[t]
- \frametitle{%
- \begin{tabular}{@ {\hspace{-3mm}}c@ {}}
- \\
- \LARGE Verifying a Regular Expression\\[-1mm]
- \LARGE Matcher and Formal Language\\[-1mm]
- \LARGE Theory\\[5mm]
- \end{tabular}}
- \begin{center}
- Christian Urban\\
- \small Technical University of Munich, Germany
- \end{center}
-
-
- \begin{center}
- \small joint work with Chunhan Wu and Xingyuan Zhang from the PLA
- University of Science and Technology in Nanjing
- \end{center}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{My Background}
-
- \mbox{}\\[-10mm]
- \begin{itemize}
- \item My background is in theory and programming languages.\bigskip
- \pause
-
- \item But I am also a programmer with a \alert<2>{software system} of around 800 kloc
- (though I am responsible for only appr.~35 kloc),
-
- \item and I write papers.
- \end{itemize}
-
- \only<2>{
- \begin{textblock}{6}(6.5,11.5)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=2mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\color{darkgray}
- \begin{minipage}{6.5cm}\raggedright
- \begin{tabular}[b]{@ {}p{4.5cm}c@ {}}
- \raggedright
- The software is a theorem prover, called {\bf Isabelle}.
- & \mbox{}\hspace{-5mm}\raisebox{-14mm}{\includegraphics[scale=0.28]{isabelle1.png}}
- \end{tabular}%
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \only<4>{
- \begin{textblock}{6}(3,11.5)
- \begin{tikzpicture}
- \draw (0,0) node[inner sep=2mm,fill=cream, ultra thick, draw=red, rounded corners=2mm]
- {\color{darkgray}
- \begin{minipage}{9.6cm}\raggedright
- So I can experience every day that writing error-free code is {\bf very, very hard}
- and that papers are also {\bf hard} to get correct.
- \end{minipage}};
- \end{tikzpicture}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{3 Points}
- \large
- \begin{itemize}
- \item It is easy to make mistakes.\bigskip
- \item Theorem provers can prevent mistakes, {\bf if} the problem
- is formulated so that it is suitable for theorem provers.\bigskip
- \item This re-formulation can be done, even in domains where
- we do not expect it.
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
-
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}
- \frametitle{Getting Papers Correct}
-
- \begin{minipage}{1.1\textwidth}
- My work over the last 5 years.\\
- {\small (in the fields of programming languages, logic and lambda-calculi)}
- \end{minipage}\bigskip
-
- \only<1>{
- \mbox{}\\[15mm]
- \begin{center}
- \begin{tikzpicture}[node distance=0.5mm]
- \node at (-1.0,-0.3) (proof) [double arrow, fill=gray,text=white, minimum height=2cm]{\bf Proof};
- \node [left=of proof]{\Large\bf Specification};
- \node [right=of proof]{\Large\bf Code};
- \end{tikzpicture}
- \end{center}
- }
- \pause
-
- \begin{tabular}{c@ {\hspace{2mm}}c}
- \begin{tabular}{c}
- \includegraphics[scale=0.09]{harper.jpg}\\[-2mm]
- {\footnotesize Bob Harper}\\[-2.5mm]
- {\footnotesize (CMU)}
- \end{tabular}
- \begin{tabular}{c}
- \includegraphics[scale=0.31]{pfenning.jpg}\\[-2mm]
- {\footnotesize Frank Pfenning}\\[-2.5mm]
- {\footnotesize (CMU)}
- \end{tabular} &
-
- \begin{tabular}{p{6cm}}
- \raggedright\small
- \color{gray}{published a proof in ACM Transactions on Computational Logic (2005),
- $\sim$31pp}
- \end{tabular}\\
-
- \\[-4mm]
-
- \begin{tabular}{c}
- \includegraphics[scale=0.3]{appel.jpg}\\[-2mm]
- {\footnotesize Andrew Appel}\\[-2.5mm]
- {\footnotesize (Princeton)}
- \end{tabular} &
-
- \begin{tabular}{p{6cm}}
- \raggedright\small
- \color{gray}{relied on their proof in a safety critical system (proof carrying code)}
- \end{tabular}
-
- \end{tabular}\medskip
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-
-text_raw {*
-
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}
- \frametitle{Proof-Carrying Code}
-
- \begin{textblock}{10}(2.5,2.2)
- \begin{block}{Idea:}
- \begin{center}
- \begin{tikzpicture}
- \draw[help lines,cream] (0,0.2) grid (8,4);
-
- \draw[line width=1mm, red] (5.5,0.6) rectangle (7.5,4);
- \node[anchor=base] at (6.5,2.8)
- {\small\begin{tabular}{@ {}p{1.9cm}@ {}}\centering user needs to run untrusted code\end{tabular}};
-
- \draw[line width=1mm, red] (0.5,0.6) rectangle (2.5,4);
- \node[anchor=base] at (1.5,2.3)
- {\small\begin{tabular}{@ {}p{1.9cm}@ {}}\centering code developer/ web server/ Apple
- Store\end{tabular}};
-
- \onslide<4->{
- \draw[line width=1mm, red, fill=red] (5.5,0.6) rectangle (7.5,1.8);
- \node[anchor=base,white] at (6.5,1.1)
- {\small\begin{tabular}{@ {}p{1.9cm}@ {}}\bf\centering proof- checker\end{tabular}};}
-
- \node at (3.8,3.0) [single arrow, fill=red,text=white, minimum height=3cm]{\bf code};
- \onslide<3->{
- \node at (3.8,1.3) [single arrow, fill=red,text=white, minimum height=3cm]{\bf LF proof};
- \node at (3.8,1.9) {\small certificate};
- }
-
- \onslide<2>{\node at (4.0,1.3) [text=red]{\begin{tabular}{c}\bf Highly\\\bf Dangerous!\end{tabular}};}
-
- \end{tikzpicture}
- \end{center}
- \end{block}
- \end{textblock}
-
- \begin{textblock}{15}(2,12)
- \small
- \begin{itemize}
- \item<4-> Appel's checker is $\sim$2700 lines of code (1865 loc of\\ LF definitions;
- 803 loc in C including 2 library functions)\\[-3mm]
- \item<5-> 167 loc in C implement a type-checker
- \end{itemize}
- \end{textblock}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text {*
- \tikzstyle{every node}=[node distance=25mm,text height=1.5ex, text depth=.25ex]
- \tikzstyle{node1}=[rectangle, minimum size=10mm, rounded corners=3mm, very thick,
- draw=black!50, top color=white, bottom color=black!20]
- \tikzstyle{node2}=[rectangle, minimum size=12mm, rounded corners=3mm, very thick,
- draw=red!70, top color=white, bottom color=red!50!black!20]
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<2->[squeeze]
- \frametitle{Type-Checking in LF}
-
- \begin{columns}
-
- \begin{column}{0.8\textwidth}
- \begin{textblock}{0}(1,2)
-
- \begin{tikzpicture}
- \matrix[ampersand replacement=\&,column sep=7mm, row sep=5mm]
- { \&[-10mm]
- \node (def1) [node1] {\large\hspace{1mm}Spec\hspace{1mm}\mbox{}}; \&
- \node (proof1) [node1] {\large Proof}; \&
- \node (alg1) [node1] {\large\hspace{1mm}Alg\hspace{1mm}\mbox{}}; \\
-
- \onslide<4->{\node {\begin{tabular}{c}\small 1st\\[-2.5mm] \footnotesize solution\end{tabular}};} \&
- \onslide<4->{\node (def2) [node2] {\large Spec$^\text{+ex}$};} \&
- \onslide<4->{\node (proof2) [node1] {\large Proof};} \&
- \onslide<4->{\node (alg2) [node1] {\large\hspace{1mm}Alg\hspace{1mm}\mbox{}};} \\
-
- \onslide<5->{\node {\begin{tabular}{c}\small 2nd\\[-2.5mm] \footnotesize solution\end{tabular}};} \&
- \onslide<5->{\node (def3) [node1] {\large\hspace{1mm}Spec\hspace{1mm}\mbox{}};} \&
- \onslide<5->{\node (proof3) [node1] {\large Proof};} \&
- \onslide<5->{\node (alg3) [node2] {\large Alg$^\text{-ex}$};} \\
-
- \onslide<6->{\node {\begin{tabular}{c}\small 3rd\\[-2.5mm] \footnotesize solution\end{tabular}};} \&
- \onslide<6->{\node (def4) [node1] {\large\hspace{1mm}Spec\hspace{1mm}\mbox{}};} \&
- \onslide<6->{\node (proof4) [node2] {\large\hspace{1mm}Proof\hspace{1mm}};} \&
- \onslide<6->{\node (alg4) [node1] {\large\hspace{1mm}Alg\hspace{1mm}\mbox{}};} \\
- };
-
- \draw[->,black!50,line width=2mm] (proof1) -- (def1);
- \draw[->,black!50,line width=2mm] (proof1) -- (alg1);
-
- \onslide<4->{\draw[->,black!50,line width=2mm] (proof2) -- (def2);}
- \onslide<4->{\draw[->,black!50,line width=2mm] (proof2) -- (alg2);}
-
- \onslide<5->{\draw[->,black!50,line width=2mm] (proof3) -- (def3);}
- \onslide<5->{\draw[->,black!50,line width=2mm] (proof3) -- (alg3);}
-
- \onslide<6->{\draw[->,black!50,line width=2mm] (proof4) -- (def4);}
- \onslide<6->{\draw[->,black!50,line width=2mm] (proof4) -- (alg4);}
-
- \onslide<3->{\draw[white,line width=1mm] (1.1,3.2) -- (0.9,2.85) -- (1.1,2.35) -- (0.9,2.0);}
- \end{tikzpicture}
-
- \end{textblock}
- \end{column}
- \end{columns}
-
-
- \begin{textblock}{3}(12,3.6)
- \onslide<4->{
- \begin{tikzpicture}
- \node at (0,0) [single arrow, shape border rotate=270, fill=red,text=white]{2h};
- \end{tikzpicture}}
- \end{textblock}
-
- \only<7->{
- \begin{textblock}{14}(0.6,12.8)
- \begin{block}{}
- \small Each time one needs to check $\sim$31pp~of informal paper proofs.
- You have to be able to keep definitions and proofs consistent.
- \end{block}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Lessons Learned}
-
- \begin{itemize}
- \item Theorem provers help with keeping large proofs consistent;
- make them modifiable.\medskip
-
- \item They can ensure that all cases are covered.\medskip
-
- \item Some reasoning can be automated.
- \end{itemize}\bigskip\pause
-
- \begin{minipage}{1.1\textwidth}
- Formal reasoning needs to be ``smooth''.\\
- {\small (ideally as close as possible to reasoning with ``pen-and-paper'')}
- \end{minipage}
-
- \only<2->{
- \begin{textblock}{3}(0.1,9.9)
- \begin{tikzpicture}
- \node at (0,0) [single arrow, shape border rotate=0, fill=red,text=red]{a};
- \end{tikzpicture}
- \end{textblock}}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-(*<*)
-atom_decl name
-
-nominal_datatype lam =
- Var "name"
- | App "lam" "lam"
- | Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
-
-nominal_primrec
- subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]")
-where
- "(Var x)[y::=s] = (if x=y then s else (Var x))"
-| "(App t\<^isub>1 t\<^isub>2)[y::=s] = App (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])"
-| "x\<sharp>(y,s) \<Longrightarrow> (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])"
-apply(finite_guess)+
-apply(rule TrueI)+
-apply(simp add: abs_fresh)
-apply(fresh_guess)+
-done
-
-lemma subst_eqvt[eqvt]:
- fixes pi::"name prm"
- shows "pi\<bullet>(t1[x::=t2]) = (pi\<bullet>t1)[(pi\<bullet>x)::=(pi\<bullet>t2)]"
-by (nominal_induct t1 avoiding: x t2 rule: lam.strong_induct)
- (auto simp add: perm_bij fresh_atm fresh_bij)
-
-lemma fresh_fact:
- fixes z::"name"
- shows "\<lbrakk>z\<sharp>s; (z=y \<or> z\<sharp>t)\<rbrakk> \<Longrightarrow> z\<sharp>t[y::=s]"
-by (nominal_induct t avoiding: z y s rule: lam.strong_induct)
- (auto simp add: abs_fresh fresh_prod fresh_atm)
-
-lemma forget:
- assumes asm: "x\<sharp>L"
- shows "L[x::=P] = L"
- using asm
-by (nominal_induct L avoiding: x P rule: lam.strong_induct)
- (auto simp add: abs_fresh fresh_atm)
-(*>*)
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}
-
- \begin{textblock}{16}(1,1)
- \renewcommand{\isasymbullet}{$\cdot$}
- \tiny\color{black}
-*}
-lemma substitution_lemma_not_to_be_tried_at_home:
- assumes asm: "x\<noteq>y" "x\<sharp>L"
- shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
-using asm
-proof (induct M arbitrary: x y N L rule: lam.induct)
- case (Lam z M1)
- have ih: "\<And>x y N L. \<lbrakk>x\<noteq>y; x\<sharp>L\<rbrakk> \<Longrightarrow> M1[x::=N][y::=L] = M1[y::=L][x::=N[y::=L]]" by fact
- have "x\<noteq>y" by fact
- have "x\<sharp>L" by fact
- obtain z'::"name" where fc: "z'\<sharp>(x,y,z,M1,N,L)" by (rule exists_fresh) (auto simp add: fs_name1)
- have eq: "Lam [z'].([(z',z)]\<bullet>M1) = Lam [z].M1" using fc
- by (auto simp add: lam.inject alpha fresh_prod fresh_atm)
- have fc': "z'\<sharp>N[y::=L]" using fc by (simp add: fresh_fact fresh_prod)
- have "([(z',z)]\<bullet>x) \<noteq> ([(z',z)]\<bullet>y)" using `x\<noteq>y` by (auto simp add: calc_atm)
- moreover
- have "([(z',z)]\<bullet>x)\<sharp>([(z',z)]\<bullet>L)" using `x\<sharp>L` by (simp add: fresh_bij)
- ultimately
- have "M1[([(z',z)]\<bullet>x)::=([(z',z)]\<bullet>N)][([(z',z)]\<bullet>y)::=([(z',z)]\<bullet>L)]
- = M1[([(z',z)]\<bullet>y)::=([(z',z)]\<bullet>L)][([(z',z)]\<bullet>x)::=([(z',z)]\<bullet>N)[([(z',z)]\<bullet>y)::=([(z',z)]\<bullet>L)]]"
- using ih by simp
- then have "[(z',z)]\<bullet>(M1[([(z',z)]\<bullet>x)::=([(z',z)]\<bullet>N)][([(z',z)]\<bullet>y)::=([(z',z)]\<bullet>L)]
- = M1[([(z',z)]\<bullet>y)::=([(z',z)]\<bullet>L)][([(z',z)]\<bullet>x)::=([(z',z)]\<bullet>N)[([(z',z)]\<bullet>y)::=([(z',z)]\<bullet>L)]])"
- by (simp add: perm_bool)
- then have ih': "([(z',z)]\<bullet>M1)[x::=N][y::=L] = ([(z',z)]\<bullet>M1)[y::=L][x::=N[y::=L]]"
- by (simp add: eqvts perm_swap)
- show "(Lam [z].M1)[x::=N][y::=L] = (Lam [z].M1)[y::=L][x::=N[y::=L]]" (is "?LHS=?RHS")
- proof -
- have "?LHS = (Lam [z'].([(z',z)]\<bullet>M1))[x::=N][y::=L]" using eq by simp
- also have "\<dots> = Lam [z'].(([(z',z)]\<bullet>M1)[x::=N][y::=L])" using fc by (simp add: fresh_prod)
- also from ih have "\<dots> = Lam [z'].(([(z',z)]\<bullet>M1)[y::=L][x::=N[y::=L]])" sorry
- also have "\<dots> = (Lam [z'].([(z',z)]\<bullet>M1))[y::=L][x::=N[y::=L]]" using fc fc' by (simp add: fresh_prod)
- also have "\<dots> = ?RHS" using eq by simp
- finally show "?LHS = ?RHS" .
- qed
-qed (auto simp add: forget)
-text_raw {*
- \end{textblock}
- \mbox{}
-
- \only<2->{
- \begin{textblock}{11.5}(4,2.3)
- \begin{minipage}{9.3cm}
- \begin{block}{}\footnotesize
-*}
-lemma substitution_lemma\<iota>:
- assumes asm: "x \<noteq> y" "x \<sharp> L"
- shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
- using asm
-by (nominal_induct M avoiding: x y N L rule: lam.strong_induct)
- (auto simp add: forget fresh_fact)
-text_raw {*
- \end{block}
- \end{minipage}
- \end{textblock}}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{Getting Programs Correct}
-
- \begin{center}
- \begin{tikzpicture}[node distance=0.5mm]
- \node at (-1.0,-0.3) (proof) [double arrow, fill=gray,text=white, minimum height=2cm]{\bf Proof};
- \node [left=of proof]{\Large\bf Specification};
- \node [right=of proof]{\Large\bf Code};
- \end{tikzpicture}
- \end{center}
-
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Regular Expressions}
-
- \begin{textblock}{6}(2,4)
- \begin{tabular}{@ {}rrl}
- \bl{r} & \bl{$::=$} & \bl{$\varnothing$}\\
- & \bl{$\mid$} & \bl{[]}\\
- & \bl{$\mid$} & \bl{c}\\
- & \bl{$\mid$} & \bl{r$_1$ + r$_2$}\\
- & \bl{$\mid$} & \bl{r$_1$ $\cdot$ r$_2$}\\
- & \bl{$\mid$} & \bl{r$^*$}\\
- \end{tabular}
- \end{textblock}
-
- \begin{textblock}{6}(8,3.5)
- \includegraphics[scale=0.35]{Screen1.png}
- \end{textblock}
-
- \begin{textblock}{6}(10.2,2.8)
- \footnotesize Isabelle:
- \end{textblock}
-
- \only<2>{
- \begin{textblock}{9}(3.6,11.8)
- \bl{matches r s $\;\Longrightarrow\;$ true $\vee$ false}\\[3.5mm]
-
- \hspace{10mm}\begin{tikzpicture}
- \coordinate (m1) at (0.4,1);
- \draw (0,0.3) node (m2) {\small\color{gray}rexp};
- \path[overlay, ->, line width = 0.5mm, shorten <=-1mm, draw = gray] (m2) edge (m1);
-
- \coordinate (s1) at (0.81,1);
- \draw (1.3,0.3) node (s2) {\small\color{gray} string};
- \path[overlay, ->, line width = 0.5mm, shorten <=-1mm, draw = gray] (s2) edge (s1);
- \end{tikzpicture}
- \end{textblock}}
-
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Specification}
-
- \small
- \begin{textblock}{6}(0,3.5)
- \begin{tabular}{r@ {\hspace{0.5mm}}r@ {\hspace{1.5mm}}c@ {\hspace{1.5mm}}l}
- \multicolumn{4}{c}{rexp $\Rightarrow$ set of strings}\bigskip\\
- &\bl{\LL ($\varnothing$)} & \bl{$\dn$} & \bl{$\varnothing$}\\
- &\bl{\LL ([])} & \bl{$\dn$} & \bl{\{[]\}}\\
- &\bl{\LL (c)} & \bl{$\dn$} & \bl{\{c\}}\\
- &\bl{\LL (r$_1$ + r$_2$)} & \bl{$\dn$} & \bl{\LL (r$_1$) $\cup$ \LL (r$_2$)}\\
- \rd{$\Rightarrow$} &\bl{\LL (r$_1$ $\cdot$ r$_2$)} & \bl{$\dn$} & \bl{\LL (r$_1$) ;; \LL (r$_2$)}\\
- \rd{$\Rightarrow$} &\bl{\LL (r$^*$)} & \bl{$\dn$} & \bl{(\LL (r))$^\star$}\\
- \end{tabular}
- \end{textblock}
-
- \begin{textblock}{9}(7.3,3)
- {\mbox{}\hspace{2cm}\footnotesize Isabelle:\smallskip}
- \includegraphics[scale=0.325]{Screen3.png}
- \end{textblock}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Version 1}
- \small
- \mbox{}\\[-8mm]\mbox{}
-
- \begin{center}\def\arraystretch{1.05}
- \begin{tabular}{@ {\hspace{-5mm}}l@ {\hspace{2.5mm}}c@ {\hspace{2.5mm}}l@ {}}
- \bl{match [] []} & \bl{$=$} & \bl{true}\\
- \bl{match [] (c::s)} & \bl{$=$} & \bl{false}\\
- \bl{match ($\varnothing$::rs) s} & \bl{$=$} & \bl{false}\\
- \bl{match ([]::rs) s} & \bl{$=$} & \bl{match rs s}\\
- \bl{match (c::rs) []} & \bl{$=$} & \bl{false}\\
- \bl{match (c::rs) (d::s)} & \bl{$=$} & \bl{if c = d then match rs s else false}\\
- \bl{match (r$_1$ + r$_2$::rs) s} & \bl{$=$} & \bl{match (r$_1$::rs) s $\vee$ match (r$_2$::rs) s}\\
- \bl{match (r$_1$ $\cdot$ r$_2$::rs) s} & \bl{$=$} & \bl{match (r$_1$::r$_2$::rs) s}\\
- \bl{match (r$^*$::rs) s} & \bl{$=$} & \bl{match rs s $\vee$ match (r::r$^*$::rs) s}\\
- \end{tabular}
- \end{center}
-
- \begin{textblock}{9}(0.2,1.6)
- \hspace{10mm}\begin{tikzpicture}
- \coordinate (m1) at (0.44,-0.5);
- \draw (0,0.3) node (m2) {\small\color{gray}\mbox{}\hspace{-9mm}list of rexps};
- \path[overlay, ->, line width = 0.5mm, shorten <=-1mm, draw = gray] (m2) edge (m1);
-
- \coordinate (s1) at (0.86,-0.5);
- \draw (1.5,0.3) node (s2) {\small\color{gray} string};
- \path[overlay, ->, line width = 0.5mm, shorten <=-1mm, draw = gray] (s2) edge (s1);
- \end{tikzpicture}
- \end{textblock}
-
- \begin{textblock}{9}(2.8,11.8)
- \bl{matches$_1$ r s $\;=\;$ match [r] s}
- \end{textblock}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Testing}
-
- \small
- Every good programmer should do thourough tests:
-
- \begin{center}
- \begin{tabular}{@ {\hspace{-20mm}}lcl}
- \bl{matches (a$\cdot$b)$^*\;$ []} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches (a$\cdot$b)$^*\;$ ab} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches (a$\cdot$b)$^*\;$ aba} & \bl{$\mapsto$} & \bl{false}\\
- \bl{matches (a$\cdot$b)$^*\;$ abab} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches (a$\cdot$b)$^*\;$ abaa} & \bl{$\mapsto$} & \bl{false}\medskip\\
- \onslide<2->{\bl{matches x$\cdot$(0$|$1)$^*\;$ x} & \bl{$\mapsto$} & \bl{true}}\\
- \onslide<2->{\bl{matches x$\cdot$(0$|$1)$^*\;$ x0} & \bl{$\mapsto$} & \bl{true}}\\
- \onslide<2->{\bl{matches x$\cdot$(0$|$1)$^*\;$ x3} & \bl{$\mapsto$} & \bl{false}}
- \end{tabular}
- \end{center}
-
- \onslide<3->
- {looks OK \ldots let's ship it to customers\hspace{5mm}
- \raisebox{-5mm}{\includegraphics[scale=0.05]{sun.png}}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Version 1}
-
- \only<1->{Several hours later\ldots}\pause
-
-
- \begin{center}
- \begin{tabular}{@ {\hspace{0mm}}lcl}
- \bl{matches$_1$ []$^*$ s} & \bl{$\mapsto$} & loops\\
- \onslide<4->{\bl{matches$_1$ ([] + \ldots)$^*$ s} & \bl{$\mapsto$} & loops\\}
- \end{tabular}
- \end{center}
-
- \small
- \onslide<3->{
- \begin{center}
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {}}
- \ldots\\
- \bl{match ([]::rs) s} & \bl{$=$} & \bl{match rs s}\\
- \ldots\\
- \bl{match (r$^*$::rs) s} & \bl{$=$} & \bl{match rs s $\vee$ match (r::r$^*$::rs) s}\\
- \end{tabular}
- \end{center}}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Testing}
-
- \begin{itemize}
- \item While testing is an important part in the process of programming development\pause
-
- \item We can only test a {\bf finite} amount of examples.\bigskip\pause
-
- \begin{center}
- \colorbox{cream}
- {\gr{\begin{minipage}{10cm}
- ``Testing can only show the presence of errors, never their
- absence'' (Edsger W.~Dijkstra)
- \end{minipage}}}
- \end{center}\bigskip\pause
-
- \item In a theorem prover we can establish properties that apply to
- {\bf all} input and {\bf all} output.\pause
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Version 2}
- \mbox{}\\[-14mm]\mbox{}
-
- \small
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}ll@ {}}
- \bl{nullable ($\varnothing$)} & \bl{$=$} & \bl{false} &\\
- \bl{nullable ([])} & \bl{$=$} & \bl{true} &\\
- \bl{nullable (c)} & \bl{$=$} & \bl{false} &\\
- \bl{nullable (r$_1$ + r$_2$)} & \bl{$=$} & \bl{nullable r$_1$ $\vee$ nullable r$_2$} & \\
- \bl{nullable (r$_1$ $\cdot$ r$_2$)} & \bl{$=$} & \bl{nullable r$_1$ $\wedge$ nullable r$_2$} & \\
- \bl{nullable (r$^*$)} & \bl{$=$} & \bl{true} & \\
- \end{tabular}\medskip
-
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-10mm}}l@ {}}
- \bl{der c ($\varnothing$)} & \bl{$=$} & \bl{$\varnothing$} & \\
- \bl{der c ([])} & \bl{$=$} & \bl{$\varnothing$} & \\
- \bl{der c (d)} & \bl{$=$} & \bl{if c = d then [] else $\varnothing$} & \\
- \bl{der c (r$_1$ + r$_2$)} & \bl{$=$} & \bl{(der c r$_1$) + (der c r$_2$)} & \\
- \bl{der c (r$_1$ $\cdot$ r$_2$)} & \bl{$=$} & \bl{((der c r$_1$) $\cdot$ r$_2$)} & \\
- & & \bl{\;\;+ (if nullable r$_1$ then der c r$_2$ else $\varnothing$)}\\
- \bl{der c (r$^*$)} & \bl{$=$} & \bl{(der c r) $\cdot$ r$^*$} &\smallskip\\
-
- \bl{derivative r []} & \bl{$=$} & \bl{r} & \\
- \bl{derivative r (c::s)} & \bl{$=$} & \bl{derivative (der c r) s} & \\
- \end{tabular}\medskip
-
- \bl{matches$_2$ r s $=$ nullable (derivative r s)}
-
- \begin{textblock}{6}(9.5,0.9)
- \begin{flushright}
- \color{gray}``if r matches []''
- \end{flushright}
- \end{textblock}
-
- \begin{textblock}{6}(9.5,6.18)
- \begin{flushright}
- \color{gray}``derivative w.r.t.~a char''
- \end{flushright}
- \end{textblock}
-
- \begin{textblock}{6}(9.5,12.1)
- \begin{flushright}
- \color{gray}``deriv.~w.r.t.~a string''
- \end{flushright}
- \end{textblock}
-
- \begin{textblock}{6}(9.5,13.98)
- \begin{flushright}
- \color{gray}``main''
- \end{flushright}
- \end{textblock}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Is the Matcher Error-Free?}
-
- We expect that
-
- \begin{center}
- \begin{tabular}{lcl}
- \bl{matches$_2$ r s = true} & \only<1>{\rd{$\Longrightarrow\,\,$}}\only<2>{\rd{$\Longleftarrow\,\,$}}%
- \only<3->{\rd{$\Longleftrightarrow$}} & \bl{s $\in$ \LL(r)}\\
- \bl{matches$_2$ r s = false} & \only<1>{\rd{$\Longrightarrow\,\,$}}\only<2>{\rd{$\Longleftarrow\,\,$}}%
- \only<3->{\rd{$\Longleftrightarrow$}} & \bl{s $\notin$ \LL(r)}\\
- \end{tabular}
- \end{center}
- \pause\pause\bigskip
- By \alert<4->{induction}, we can {\bf prove} these properties.\bigskip
-
- \begin{tabular}{lrcl}
- Lemmas: & \bl{nullable (r)} & \bl{$\Longleftrightarrow$} & \bl{[] $\in$ \LL (r)}\\
- & \bl{s $\in$ \LL (der c r)} & \bl{$\Longleftrightarrow$} & \bl{(c::s) $\in$ \LL (r)}\\
- \end{tabular}
-
- \only<4->{
- \begin{textblock}{3}(0.9,4.5)
- \rd{\huge$\forall$\large{}r s.}
- \end{textblock}}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{
- \begin{tabular}{c}
- \mbox{}\\[23mm]
- \LARGE Demo
- \end{tabular}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
-
- \mbox{}\\[-2mm]
-
- \small
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}ll@ {}}
- \bl{nullable (NULL)} & \bl{$=$} & \bl{false} &\\
- \bl{nullable (EMPTY)} & \bl{$=$} & \bl{true} &\\
- \bl{nullable (CHR c)} & \bl{$=$} & \bl{false} &\\
- \bl{nullable (ALT r$_1$ r$_2$)} & \bl{$=$} & \bl{(nullable r$_1$) orelse (nullable r$_2$)} & \\
- \bl{nullable (SEQ r$_1$ r$_2$)} & \bl{$=$} & \bl{(nullable r$_1$) andalso (nullable r$_2$)} & \\
- \bl{nullable (STAR r)} & \bl{$=$} & \bl{true} & \\
- \end{tabular}\medskip
-
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-10mm}}l@ {}}
- \bl{der c (NULL)} & \bl{$=$} & \bl{NULL} & \\
- \bl{der c (EMPTY)} & \bl{$=$} & \bl{NULL} & \\
- \bl{der c (CHR d)} & \bl{$=$} & \bl{if c=d then EMPTY else NULL} & \\
- \bl{der c (ALT r$_1$ r$_2$)} & \bl{$=$} & \bl{ALT (der c r$_1$) (der c r$_2$)} & \\
- \bl{der c (SEQ r$_1$ r$_2$)} & \bl{$=$} & \bl{ALT (SEQ (der c r$_1$) r$_2$)} & \\
- & & \bl{\phantom{ALT} (if nullable r$_1$ then der c r$_2$ else NULL)}\\
- \bl{der c (STAR r)} & \bl{$=$} & \bl{SEQ (der c r) (STAR r)} &\smallskip\\
-
- \bl{derivative r []} & \bl{$=$} & \bl{r} & \\
- \bl{derivative r (c::s)} & \bl{$=$} & \bl{derivative (der c r) s} & \\
- \end{tabular}\medskip
-
- \bl{matches r s $=$ nullable (derivative r s)}
-
- \only<2>{
- \begin{textblock}{8}(1.5,4)
- \includegraphics[scale=0.3]{approved.png}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{No Automata?}
-
- You might be wondering why I did not use any automata:
-
- \begin{itemize}
- \item A \alert{regular language} is one where there is a DFA that
- recognises it.\bigskip\pause
- \end{itemize}
-
-
- There are many reasons why this is a good definition:\medskip
- \begin{itemize}
- \item pumping lemma
- \item closure properties of regular languages\\ (e.g.~closure under complement)
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[t]
- \frametitle{Really Bad News!}
-
- DFAs are bad news for formalisations in theorem provers. They might
- be represented as:
-
- \begin{itemize}
- \item graphs
- \item matrices
- \item partial functions
- \end{itemize}
-
- All constructions are messy to reason about.\bigskip\bigskip
- \pause
-
- \small
- \only<2>{
- Constable et al needed (on and off) 18 months for a 3-person team
- to formalise automata theory in Nuprl including Myhill-Nerode. There is
- only very little other formalised work on regular languages I know of
- in Coq, Isabelle and HOL.}
- \only<3>{typical textbook reasoning goes like: ``\ldots if \smath{M} and \smath{N} are any two
- automata with no inaccessible states \ldots''
- }
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{}
- \large
- \begin{center}
- \begin{tabular}{p{9cm}}
- My point:\bigskip\\
-
- The theory about regular languages can be reformulated
- to be more suitable for theorem proving.
- \end{tabular}
- \end{center}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE The Myhill-Nerode Theorem}
-
- \begin{itemize}
- \item provides necessary and suf\!ficient conditions for a language
- being regular (pumping lemma only necessary)\medskip
-
- \item will help with closure properties of regular languages\bigskip\pause
-
- \item key is the equivalence relation:\smallskip
- \begin{center}
- \smath{x \approx_{L} y \,\dn\, \forall z.\; x @ z \in L \Leftrightarrow y @ z \in L}
- \end{center}
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE The Myhill-Nerode Theorem}
-
- \mbox{}\\[5cm]
-
- \begin{itemize}
- \item \smath{\text{finite}\, (U\!N\!IV /\!/ \approx_L) \;\Leftrightarrow\; L\; \text{is regular}}
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Equivalence Classes}
-
- \begin{itemize}
- \item \smath{L = []}
- \begin{center}
- \smath{\Big\{\{[]\},\; U\!N\!IV - \{[]\}\Big\}}
- \end{center}\bigskip\bigskip
-
- \item \smath{L = [c]}
- \begin{center}
- \smath{\Big\{\{[]\},\; \{[c]\},\; U\!N\!IV - \{[], [c]\}\Big\}}
- \end{center}\bigskip\bigskip
-
- \item \smath{L = \varnothing}
- \begin{center}
- \smath{\Big\{U\!N\!IV\Big\}}
- \end{center}
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Regular Languages}
-
- \begin{itemize}
- \item \smath{L} is regular \smath{\dn} if there is an automaton \smath{M}
- such that \smath{\mathbb{L}(M) = L}\\[1.5cm]
-
- \item Myhill-Nerode:
-
- \begin{center}
- \begin{tabular}{l}
- finite $\Rightarrow$ regular\\
- \;\;\;\smath{\text{finite}\,(U\!N\!IV /\!/ \approx_L) \Rightarrow \exists r. L = \mathbb{L}(r)}\\[3mm]
- regular $\Rightarrow$ finite\\
- \;\;\;\smath{\text{finite}\, (U\!N\!IV /\!/ \approx_{\mathbb{L}(r)})}
- \end{tabular}
- \end{center}
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Final States}
-
- \mbox{}\\[3cm]
-
- \begin{itemize}
- \item \smath{\text{final}_L\,X \dn}\\
- \smath{\hspace{6mm}X \in (U\!N\!IV /\!/\approx_L) \;\wedge\; \forall s \in X.\; s \in L}
- \smallskip
- \item we can prove: \smath{L = \bigcup \{X.\;\text{final}_L\,X\}}
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Transitions between\\[-3mm] Equivalence Classes}
-
- \smath{L = \{[c]\}}
-
- \begin{tabular}{@ {\hspace{-7mm}}cc}
- \begin{tabular}{c}
- \begin{tikzpicture}[shorten >=1pt,node distance=2cm,auto, ultra thick]
- \tikzstyle{state}=[circle,thick,draw=blue!75,fill=blue!20,minimum size=0mm]
-
- %\draw[help lines] (0,0) grid (3,2);
-
- \node[state,initial] (q_0) {$R_1$};
- \node[state,accepting] (q_1) [above right of=q_0] {$R_2$};
- \node[state] (q_2) [below right of=q_0] {$R_3$};
-
- \path[->] (q_0) edge node {c} (q_1)
- edge node [swap] {$\Sigma-{c}$} (q_2)
- (q_2) edge [loop below] node {$\Sigma$} ()
- (q_1) edge node {$\Sigma$} (q_2);
- \end{tikzpicture}
- \end{tabular}
- &
- \begin{tabular}[t]{ll}
- \\[-20mm]
- \multicolumn{2}{l}{\smath{U\!N\!IV /\!/\approx_L} produces}\\[4mm]
-
- \smath{R_1}: & \smath{\{[]\}}\\
- \smath{R_2}: & \smath{\{[c]\}}\\
- \smath{R_3}: & \smath{U\!N\!IV - \{[], [c]\}}\\[6mm]
- \multicolumn{2}{l}{\onslide<2->{\smath{X \stackrel{c}{\longrightarrow} Y \dn X ;; [c] \subseteq Y}}}
- \end{tabular}
-
- \end{tabular}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Systems of Equations}
-
- Inspired by a method of Brzozowski\;'64, we can build an equational system
- characterising the equivalence classes:
-
- \begin{center}
- \begin{tabular}{@ {\hspace{-20mm}}c}
- \\[-13mm]
- \begin{tikzpicture}[shorten >=1pt,node distance=2cm,auto, ultra thick]
- \tikzstyle{state}=[circle,thick,draw=blue!75,fill=blue!20,minimum size=0mm]
-
- %\draw[help lines] (0,0) grid (3,2);
-
- \node[state,initial] (p_0) {$R_1$};
- \node[state,accepting] (p_1) [right of=q_0] {$R_2$};
-
- \path[->] (p_0) edge [bend left] node {a} (p_1)
- edge [loop above] node {b} ()
- (p_1) edge [loop above] node {a} ()
- edge [bend left] node {b} (p_0);
- \end{tikzpicture}\\
- \\[-13mm]
- \end{tabular}
- \end{center}
-
- \begin{center}
- \begin{tabular}{@ {\hspace{-6mm}}ll@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
- & \smath{R_1} & \smath{\equiv} & \smath{R_1;b + R_2;b \onslide<2->{\alert<2>{+ \lambda;[]}}}\\
- & \smath{R_2} & \smath{\equiv} & \smath{R_1;a + R_2;a}\medskip\\
- \onslide<3->{we can prove}
- & \onslide<3->{\smath{R_1}} & \onslide<3->{\smath{=}}
- & \onslide<3->{\smath{R_1; \mathbb{L}(b) \,\cup\, R_2;\mathbb{L}(b) \,\cup\, \{[]\};\{[]\}}}\\
- & \onslide<3->{\smath{R_2}} & \onslide<3->{\smath{=}}
- & \onslide<3->{\smath{R_1; \mathbb{L}(a) \,\cup\, R_2;\mathbb{L}(a)}}\\
- \end{tabular}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[t]
- \small
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}ll}
- \onslide<1->{\smath{R_1}} & \onslide<1->{\smath{=}}
- & \onslide<1->{\smath{R_1; b + R_2; b + \lambda;[]}}\\
- \onslide<1->{\smath{R_2}} & \onslide<1->{\smath{=}}
- & \onslide<1->{\smath{R_1; a + R_2; a}}\\
-
- & & & \onslide<2->{by Arden}\\
-
- \onslide<2->{\smath{R_1}} & \onslide<2->{\smath{=}}
- & \onslide<2->{\smath{R_1; b + R_2; b + \lambda;[]}}\\
- \onslide<2->{\smath{R_2}} & \onslide<2->{\smath{=}}
- & \only<2>{\smath{R_1; a + R_2; a}}%
- \only<3->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<4->{by Arden}\\
-
- \onslide<4->{\smath{R_1}} & \onslide<4->{\smath{=}}
- & \onslide<4->{\smath{R_2; b \cdot b^\star+ \lambda;b^\star}}\\
- \onslide<4->{\smath{R_2}} & \onslide<4->{\smath{=}}
- & \onslide<4->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<5->{by substitution}\\
-
- \onslide<5->{\smath{R_1}} & \onslide<5->{\smath{=}}
- & \onslide<5->{\smath{R_1; a\cdot a^\star \cdot b \cdot b^\star+ \lambda;b^\star}}\\
- \onslide<5->{\smath{R_2}} & \onslide<5->{\smath{=}}
- & \onslide<5->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<6->{by Arden}\\
-
- \onslide<6->{\smath{R_1}} & \onslide<6->{\smath{=}}
- & \onslide<6->{\smath{\lambda;b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star}}\\
- \onslide<6->{\smath{R_2}} & \onslide<6->{\smath{=}}
- & \onslide<6->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<7->{by substitution}\\
-
- \onslide<7->{\smath{R_1}} & \onslide<7->{\smath{=}}
- & \onslide<7->{\smath{\lambda;b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star}}\\
- \onslide<7->{\smath{R_2}} & \onslide<7->{\smath{=}}
- & \onslide<7->{\smath{\lambda; b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star
- \cdot a\cdot a^\star}}\\
- \end{tabular}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE A Variant of Arden's Lemma}
-
- {\bf Arden's Lemma:}\smallskip
-
- If \smath{[] \not\in A} then
- \begin{center}
- \smath{X = X; A + \text{something}}
- \end{center}
- has the (unique) solution
- \begin{center}
- \smath{X = \text{something} ; A^\star}
- \end{center}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \small
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}ll}
- \onslide<1->{\smath{R_1}} & \onslide<1->{\smath{=}}
- & \onslide<1->{\smath{R_1; b + R_2; b + \lambda;[]}}\\
- \onslide<1->{\smath{R_2}} & \onslide<1->{\smath{=}}
- & \onslide<1->{\smath{R_1; a + R_2; a}}\\
-
- & & & \onslide<2->{by Arden}\\
-
- \onslide<2->{\smath{R_1}} & \onslide<2->{\smath{=}}
- & \onslide<2->{\smath{R_1; b + R_2; b + \lambda;[]}}\\
- \onslide<2->{\smath{R_2}} & \onslide<2->{\smath{=}}
- & \only<2>{\smath{R_1; a + R_2; a}}%
- \only<3->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<4->{by Arden}\\
-
- \onslide<4->{\smath{R_1}} & \onslide<4->{\smath{=}}
- & \onslide<4->{\smath{R_2; b \cdot b^\star+ \lambda;b^\star}}\\
- \onslide<4->{\smath{R_2}} & \onslide<4->{\smath{=}}
- & \onslide<4->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<5->{by substitution}\\
-
- \onslide<5->{\smath{R_1}} & \onslide<5->{\smath{=}}
- & \onslide<5->{\smath{R_1; a\cdot a^\star \cdot b \cdot b^\star+ \lambda;b^\star}}\\
- \onslide<5->{\smath{R_2}} & \onslide<5->{\smath{=}}
- & \onslide<5->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<6->{by Arden}\\
-
- \onslide<6->{\smath{R_1}} & \onslide<6->{\smath{=}}
- & \onslide<6->{\smath{\lambda;b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star}}\\
- \onslide<6->{\smath{R_2}} & \onslide<6->{\smath{=}}
- & \onslide<6->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<7->{by substitution}\\
-
- \onslide<7->{\smath{R_1}} & \onslide<7->{\smath{=}}
- & \onslide<7->{\smath{\lambda;b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star}}\\
- \onslide<7->{\smath{R_2}} & \onslide<7->{\smath{=}}
- & \onslide<7->{\smath{\lambda; b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star
- \cdot a\cdot a^\star}}\\
- \end{tabular}
- \end{center}
-
- \only<8->{
- \begin{textblock}{6}(2.5,4)
- \begin{block}{}
- \begin{minipage}{8cm}\raggedright
-
- \begin{tikzpicture}[shorten >=1pt,node distance=2cm,auto, ultra thick, inner sep=1mm]
- \tikzstyle{state}=[circle,thick,draw=blue!75,fill=blue!20,minimum size=0mm]
-
- %\draw[help lines] (0,0) grid (3,2);
-
- \node[state,initial] (p_0) {$R_1$};
- \node[state,accepting] (p_1) [right of=q_0] {$R_2$};
-
- \path[->] (p_0) edge [bend left] node {a} (p_1)
- edge [loop above] node {b} ()
- (p_1) edge [loop above] node {a} ()
- edge [bend left] node {b} (p_0);
- \end{tikzpicture}
-
- \end{minipage}
- \end{block}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE The Equ's Solving Algorithm}
-
- \begin{itemize}
- \item The algorithm must terminate: Arden makes one equation smaller;
- substitution deletes one variable from the right-hand sides.\bigskip
-
- \item We need to maintain the invariant that Arden is applicable
- (if \smath{[] \not\in A} then \ldots):\medskip
-
- \begin{center}\small
- \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}ll}
- \smath{R_1} & \smath{=} & \smath{R_1; b + R_2; b + \lambda;[]}\\
- \smath{R_2} & \smath{=} & \smath{R_1; a + R_2; a}\\
-
- & & & by Arden\\
-
- \smath{R_1} & \smath{=} & \smath{R_1; b + R_2; b + \lambda;[]}\\
- \smath{R_2} & \smath{=} & \smath{R_1; a\cdot a^\star}\\
- \end{tabular}
- \end{center}
-
- \end{itemize}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Other Direction}
-
- One has to prove
-
- \begin{center}
- \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r)})}
- \end{center}
-
- by induction on \smath{r}. Not trivial, but after a bit
- of thinking, one can prove that if
-
- \begin{center}
- \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r_1)})}\hspace{5mm}
- \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r_2)})}
- \end{center}
-
- then
-
- \begin{center}
- \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r_1) \,\cup\, \mathbb{L}(r_2)})}
- \end{center}
-
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE What Have We Achieved?}
-
- \begin{itemize}
- \item \smath{\text{finite}\, (U\!N\!IV /\!/ \approx_L) \;\Leftrightarrow\; L\; \text{is regular}}
- \bigskip\pause
- \item regular languages are closed under complementation; this is now easy\medskip
- \begin{center}
- \smath{U\!N\!IV /\!/ \approx_L \;\;=\;\; U\!N\!IV /\!/ \approx_{-L}}
- \end{center}
- \end{itemize}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Examples}
-
- \begin{itemize}
- \item \smath{L \equiv \Sigma^\star 0 \Sigma} is regular
- \begin{quote}\small
- \begin{tabular}{lcl}
- \smath{A_1} & \smath{=} & \smath{\Sigma^\star 00}\\
- \smath{A_2} & \smath{=} & \smath{\Sigma^\star 01}\\
- \smath{A_3} & \smath{=} & \smath{\Sigma^\star 10 \cup \{0\}}\\
- \smath{A_4} & \smath{=} & \smath{\Sigma^\star 11 \cup \{1\} \cup \{[]\}}\\
- \end{tabular}
- \end{quote}
-
- \item \smath{L \equiv \{ 0^n 1^n \,|\, n \ge 0\}} is not regular
- \begin{quote}\small
- \begin{tabular}{lcl}
- \smath{B_0} & \smath{=} & \smath{\{0^n 1^n \,|\, n \ge 0\}}\\
- \smath{B_1} & \smath{=} & \smath{\{0^n 1^{(n-1)} \,|\, n \ge 1\}}\\
- \smath{B_2} & \smath{=} & \smath{\{0^n 1^{(n-2)} \,|\, n \ge 2\}}\\
- \smath{B_3} & \smath{=} & \smath{\{0^n 1^{(n-3)} \,|\, n \ge 3\}}\\
- & \smath{\vdots} &\\
- \end{tabular}
- \end{quote}
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE What We Have Not Achieved}
-
- \begin{itemize}
- \item regular expressions are not good if you look for a minimal
- one for a language (DFAs have this notion)\pause\bigskip
-
- \item Is there anything to be said about context free languages:\medskip
-
- \begin{quote}
- A context free language is where every string can be recognised by
- a pushdown automaton.\bigskip
- \end{quote}
- \end{itemize}
-
- \textcolor{gray}{\footnotesize Yes. Derivatives also work for c-f grammars. Ongoing work.}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Conclusion}
-
- \begin{itemize}
- \item We formalised the Myhill-Nerode theorem based on
- regular expressions only (DFAs are difficult to deal with in a theorem prover).\smallskip
-
- \item Seems to be a common theme: algorithms need to be reformulated
- to better suit formal treatment.\smallskip
-
- \item The most interesting aspect is that we are able to
- implement the matcher directly inside the theorem prover
- (ongoing work).\smallskip
-
- \item Parsing is a vast field which seem to offer new results.
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[b]
- \frametitle{
- \begin{tabular}{c}
- \mbox{}\\[13mm]
- \alert{\LARGE Thank you very much!}\\
- \alert{\Large Questions?}
- \end{tabular}}
-
- \begin{center}
- \bf \underline{Short Bio:}
- \end{center}
- \mbox{}\\[-17mm]\mbox{}\small
- \begin{itemize}
- \item PhD in Cambridge
- \item Emmy-Noether Research Fellowship at the TU Munich
- \item talks at: CMU, Yale, Princeton, MIT,$\ldots$
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{Future Research}
-
- My existing strengths:\bigskip
-
- \begin{itemize}
- \item Isabelle (implementation)\bigskip
- \item background in logic, programming languages, formal methods
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{Future Research}
-
- I want to have a single logic framework in which I can
- write programs and prove their correctness.\bigskip
-
- \begin{itemize}
- \item extensions of HOL (IO, modules, advanced types)
- \item high-level programming languages
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{Future Research}
-
- Compilers\bigskip
-
- \begin{itemize}
- \item the high-level language needs to be compiled to correct machine
- code
- \item compiler verification, machine code verification
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{Future Research}
-
- Stronger type-systems\bigskip
-
- \begin{itemize}
- \item ``correct by construction''
- \item GADTs, dependent types
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{Future Research}
-
- Proof automation\bigskip
-
- \begin{itemize}
- \item external tools generate ``proof-certificates''
- \item certificates are imported into Isabelle
- \item GPU based external provers
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{Future Research}
-
- Large-scale applications\bigskip
-
- \begin{itemize}
- \item verification of Java-Script, Scala,$\ldots$
- \item interesting code (INTEL in Shanghai)
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-(*<*)
-end
-(*>*)
\ No newline at end of file
--- a/Slides/Slides7.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1246 +0,0 @@
-(*<*)
-theory Slides7
-imports "~~/src/HOL/Library/LaTeXsugar" "Main"
-begin
-
-declare [[show_question_marks = false]]
-
-notation (latex output)
- set ("_") and
- Cons ("_::/_" [66,65] 65)
-
-(*>*)
-
-text_raw {*
- \renewcommand{\slidecaption}{Beijing, 29.~April 2011}
-
- \newcommand{\abst}[2]{#1.#2}% atom-abstraction
- \newcommand{\pair}[2]{\langle #1,#2\rangle} % pairing
- \newcommand{\susp}{{\boldsymbol{\cdot}}}% for suspensions
- \newcommand{\unit}{\langle\rangle}% unit
- \newcommand{\app}[2]{#1\,#2}% application
- \newcommand{\eqprob}{\mathrel{{\approx}?}}
- \newcommand{\freshprob}{\mathrel{\#?}}
- \newcommand{\redu}[1]{\stackrel{#1}{\Longrightarrow}}% reduction
- \newcommand{\id}{\varepsilon}% identity substitution
-
- \newcommand{\bl}[1]{\textcolor{blue}{#1}}
- \newcommand{\gr}[1]{\textcolor{gray}{#1}}
- \newcommand{\rd}[1]{\textcolor{red}{#1}}
-
- \newcommand{\ok}{\includegraphics[scale=0.07]{ok.png}}
- \newcommand{\notok}{\includegraphics[scale=0.07]{notok.png}}
- \newcommand{\largenotok}{\includegraphics[scale=1]{notok.png}}
-
- \renewcommand{\Huge}{\fontsize{61.92}{77}\selectfont}
- \newcommand{\veryHuge}{\fontsize{74.3}{93}\selectfont}
- \newcommand{\VeryHuge}{\fontsize{89.16}{112}\selectfont}
- \newcommand{\VERYHuge}{\fontsize{107}{134}\selectfont}
-
- \newcommand{\LL}{$\mathbb{L}\,$}
-
-
- \pgfdeclareradialshading{smallbluesphere}{\pgfpoint{0.5mm}{0.5mm}}%
- {rgb(0mm)=(0,0,0.9);
- rgb(0.9mm)=(0,0,0.7);
- rgb(1.3mm)=(0,0,0.5);
- rgb(1.4mm)=(1,1,1)}
-
- \def\myitemi{\begin{pgfpicture}{-1ex}{-0.55ex}{1ex}{1ex}
- \usebeamercolor[fg]{subitem projected}
- {\pgftransformscale{0.8}\pgftext{\normalsize\pgfuseshading{bigsphere}}}
- \pgftext{%
- \usebeamerfont*{subitem projected}}
- \end{pgfpicture}}
-
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[t]
- \frametitle{%
- \begin{tabular}{@ {\hspace{-3mm}}c@ {}}
- \\
- \LARGE Verifying a Regular Expression\\[-1mm]
- \LARGE Matcher and Formal Language\\[-1mm]
- \LARGE Theory\\[5mm]
- \end{tabular}}
- \begin{center}
- Christian Urban\\
- \small Technical University of Munich, Germany
- \end{center}
-
-
- \begin{center}
- \small joint work with Chunhan Wu and Xingyuan Zhang from the PLA
- University of Science and Technology in Nanjing
- \end{center}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{This Talk: 3 Points}
- \large
- \begin{itemize}
- \item It is easy to make mistakes.\bigskip
- \item Theorem provers can prevent mistakes, {\bf if} the problem
- is formulated so that it is suitable for theorem provers.\bigskip
- \item This re-formulation can be done, even in domains where
- we least expect it.
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Regular Expressions}
-
- \begin{textblock}{6}(2,4)
- \begin{tabular}{@ {}rrl}
- \bl{r} & \bl{$::=$} & \bl{$\varnothing$}\\
- & \bl{$\mid$} & \bl{[]}\\
- & \bl{$\mid$} & \bl{c}\\
- & \bl{$\mid$} & \bl{r$_1$ + r$_2$}\\
- & \bl{$\mid$} & \bl{r$_1$ $\cdot$ r$_2$}\\
- & \bl{$\mid$} & \bl{r$^*$}\\
- \end{tabular}
- \end{textblock}
-
- \begin{textblock}{6}(8,3.5)
- \includegraphics[scale=0.35]{Screen1.png}
- \end{textblock}
-
- \begin{textblock}{6}(10.2,2.8)
- \footnotesize Isabelle:
- \end{textblock}
-
- \only<2>{
- \begin{textblock}{9}(3.6,11.8)
- \bl{matches r s $\;\Longrightarrow\;$ true $\vee$ false}\\[3.5mm]
-
- \hspace{10mm}\begin{tikzpicture}
- \coordinate (m1) at (0.4,1);
- \draw (0,0.3) node (m2) {\small\color{gray}rexp};
- \path[overlay, ->, line width = 0.5mm, shorten <=-1mm, draw = gray] (m2) edge (m1);
-
- \coordinate (s1) at (0.81,1);
- \draw (1.3,0.3) node (s2) {\small\color{gray} string};
- \path[overlay, ->, line width = 0.5mm, shorten <=-1mm, draw = gray] (s2) edge (s1);
- \end{tikzpicture}
- \end{textblock}}
-
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Specification}
-
- \small
- \begin{textblock}{6}(0,3.5)
- \begin{tabular}{r@ {\hspace{0.5mm}}r@ {\hspace{1.5mm}}c@ {\hspace{1.5mm}}l}
- \multicolumn{4}{c}{rexp $\Rightarrow$ set of strings}\bigskip\\
- &\bl{\LL ($\varnothing$)} & \bl{$\dn$} & \bl{$\varnothing$}\\
- &\bl{\LL ([])} & \bl{$\dn$} & \bl{\{[]\}}\\
- &\bl{\LL (c)} & \bl{$\dn$} & \bl{\{c\}}\\
- &\bl{\LL (r$_1$ + r$_2$)} & \bl{$\dn$} & \bl{\LL (r$_1$) $\cup$ \LL (r$_2$)}\\
- \rd{$\Rightarrow$} &\bl{\LL (r$_1$ $\cdot$ r$_2$)} & \bl{$\dn$} & \bl{\LL (r$_1$) ;; \LL (r$_2$)}\\
- \rd{$\Rightarrow$} &\bl{\LL (r$^*$)} & \bl{$\dn$} & \bl{(\LL (r))$^\star$}\\
- \end{tabular}
- \end{textblock}
-
- \begin{textblock}{9}(7.3,3)
- {\mbox{}\hspace{2cm}\footnotesize Isabelle:\smallskip}
- \includegraphics[scale=0.325]{Screen3.png}
- \end{textblock}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Version 1}
- \small
- \mbox{}\\[-8mm]\mbox{}
-
- \begin{center}\def\arraystretch{1.05}
- \begin{tabular}{@ {\hspace{-5mm}}l@ {\hspace{2.5mm}}c@ {\hspace{2.5mm}}l@ {}}
- \bl{match [] []} & \bl{$=$} & \bl{true}\\
- \bl{match [] (c::s)} & \bl{$=$} & \bl{false}\\
- \bl{match ($\varnothing$::rs) s} & \bl{$=$} & \bl{false}\\
- \bl{match ([]::rs) s} & \bl{$=$} & \bl{match rs s}\\
- \bl{match (c::rs) []} & \bl{$=$} & \bl{false}\\
- \bl{match (c::rs) (d::s)} & \bl{$=$} & \bl{if c = d then match rs s else false}\\
- \bl{match (r$_1$ + r$_2$::rs) s} & \bl{$=$} & \bl{match (r$_1$::rs) s $\vee$ match (r$_2$::rs) s}\\
- \bl{match (r$_1$ $\cdot$ r$_2$::rs) s} & \bl{$=$} & \bl{match (r$_1$::r$_2$::rs) s}\\
- \bl{match (r$^*$::rs) s} & \bl{$=$} & \bl{match rs s $\vee$ match (r::r$^*$::rs) s}\\
- \end{tabular}
- \end{center}
-
- \begin{textblock}{9}(0.2,1.6)
- \hspace{10mm}\begin{tikzpicture}
- \coordinate (m1) at (0.44,-0.5);
- \draw (0,0.3) node (m2) {\small\color{gray}\mbox{}\hspace{-9mm}list of rexps};
- \path[overlay, ->, line width = 0.5mm, shorten <=-1mm, draw = gray] (m2) edge (m1);
-
- \coordinate (s1) at (0.86,-0.5);
- \draw (1.5,0.3) node (s2) {\small\color{gray} string};
- \path[overlay, ->, line width = 0.5mm, shorten <=-1mm, draw = gray] (s2) edge (s1);
- \end{tikzpicture}
- \end{textblock}
-
- \begin{textblock}{9}(2.8,11.8)
- \bl{matches$_1$ r s $\;=\;$ match [r] s}
- \end{textblock}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Testing}
-
- \small
- Every good programmer should do thourough tests:
-
- \begin{center}
- \begin{tabular}{@ {\hspace{-20mm}}lcl}
- \bl{matches$_1$ (a$\cdot$b)$^*\;$ []} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches$_1$ (a$\cdot$b)$^*\;$ ab} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches$_1$ (a$\cdot$b)$^*\;$ aba} & \bl{$\mapsto$} & \bl{false}\\
- \bl{matches$_1$ (a$\cdot$b)$^*\;$ abab} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches$_1$ (a$\cdot$b)$^*\;$ abaa} & \bl{$\mapsto$} & \bl{false}\medskip\\
- \onslide<2->{\bl{matches$_1$ x$\cdot$(0$|$1)$^*\;$ x} & \bl{$\mapsto$} & \bl{true}}\\
- \onslide<2->{\bl{matches$_1$ x$\cdot$(0$|$1)$^*\;$ x0} & \bl{$\mapsto$} & \bl{true}}\\
- \onslide<2->{\bl{matches$_1$ x$\cdot$(0$|$1)$^*\;$ x3} & \bl{$\mapsto$} & \bl{false}}
- \end{tabular}
- \end{center}
-
- \onslide<3->
- {Looks OK \ldots let's ship it to customers\hspace{5mm}
- \raisebox{-5mm}{\includegraphics[scale=0.05]{sun.png}}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Version 1}
-
- \only<1->{Several hours later\ldots}\pause
-
-
- \begin{center}
- \begin{tabular}{@ {\hspace{0mm}}lcl}
- \bl{matches$_1$ []$^*$ s} & \bl{$\mapsto$} & loops\\
- \onslide<4->{\bl{matches$_1$ ([] + \ldots)$^*$ s} & \bl{$\mapsto$} & loops\\}
- \end{tabular}
- \end{center}
-
- \small
- \onslide<3->{
- \begin{center}
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {}}
- \ldots\\
- \bl{match ([]::rs) s} & \bl{$=$} & \bl{match rs s}\\
- \ldots\\
- \bl{match (r$^*$::rs) s} & \bl{$=$} & \bl{match rs s $\vee$ match (r::r$^*$::rs) s}\\
- \end{tabular}
- \end{center}}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Testing}
-
- \begin{itemize}
- \item While testing is an important part in the process of programming development\pause\ldots
-
- \item we can only test a {\bf finite} amount of examples.\bigskip\pause
-
- \begin{center}
- \colorbox{cream}
- {\gr{\begin{minipage}{10cm}
- ``Testing can only show the presence of errors, never their
- absence.'' (Edsger W.~Dijkstra)
- \end{minipage}}}
- \end{center}\bigskip\pause
-
- \item In a theorem prover we can establish properties that apply to
- {\bf all} input and {\bf all} output.
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Version 2}
- \mbox{}\\[-14mm]\mbox{}
-
- \small
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}ll@ {}}
- \bl{nullable ($\varnothing$)} & \bl{$=$} & \bl{false} &\\
- \bl{nullable ([])} & \bl{$=$} & \bl{true} &\\
- \bl{nullable (c)} & \bl{$=$} & \bl{false} &\\
- \bl{nullable (r$_1$ + r$_2$)} & \bl{$=$} & \bl{nullable r$_1$ $\vee$ nullable r$_2$} & \\
- \bl{nullable (r$_1$ $\cdot$ r$_2$)} & \bl{$=$} & \bl{nullable r$_1$ $\wedge$ nullable r$_2$} & \\
- \bl{nullable (r$^*$)} & \bl{$=$} & \bl{true} & \\
- \end{tabular}\medskip
-
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-10mm}}l@ {}}
- \bl{der c ($\varnothing$)} & \bl{$=$} & \bl{$\varnothing$} & \\
- \bl{der c ([])} & \bl{$=$} & \bl{$\varnothing$} & \\
- \bl{der c (d)} & \bl{$=$} & \bl{if c = d then [] else $\varnothing$} & \\
- \bl{der c (r$_1$ + r$_2$)} & \bl{$=$} & \bl{(der c r$_1$) + (der c r$_2$)} & \\
- \bl{der c (r$_1$ $\cdot$ r$_2$)} & \bl{$=$} & \bl{((der c r$_1$) $\cdot$ r$_2$)} & \\
- & & \bl{\;\;\;\;+ (if nullable r$_1$ then der c r$_2$ else $\varnothing$)}\\
- \bl{der c (r$^*$)} & \bl{$=$} & \bl{(der c r) $\cdot$ r$^*$} &\smallskip\\
-
- \bl{derivative r []} & \bl{$=$} & \bl{r} & \\
- \bl{derivative r (c::s)} & \bl{$=$} & \bl{derivative (der c r) s} & \\
- \end{tabular}\medskip
-
- \bl{matches$_2$ r s $=$ nullable (derivative r s)}
-
- \begin{textblock}{6}(9.5,0.9)
- \begin{flushright}
- \color{gray}``if r matches []''
- \end{flushright}
- \end{textblock}
-
- \begin{textblock}{6}(9.5,6.18)
- \begin{flushright}
- \color{gray}``derivative w.r.t.~a char''
- \end{flushright}
- \end{textblock}
-
- \begin{textblock}{6}(9.5,12.1)
- \begin{flushright}
- \color{gray}``deriv.~w.r.t.~a string''
- \end{flushright}
- \end{textblock}
-
- \begin{textblock}{6}(9.5,13.98)
- \begin{flushright}
- \color{gray}``main''
- \end{flushright}
- \end{textblock}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Is the Matcher Error-Free?}
-
- We expect that
-
- \begin{center}
- \begin{tabular}{lcl}
- \bl{matches$_2$ r s = true} & \only<1>{\rd{$\Longrightarrow\,\,$}}\only<2>{\rd{$\Longleftarrow\,\,$}}%
- \only<3->{\rd{$\Longleftrightarrow$}} & \bl{s $\in$ \LL(r)}\\
- \bl{matches$_2$ r s = false} & \only<1>{\rd{$\Longrightarrow\,\,$}}\only<2>{\rd{$\Longleftarrow\,\,$}}%
- \only<3->{\rd{$\Longleftrightarrow$}} & \bl{s $\notin$ \LL(r)}\\
- \end{tabular}
- \end{center}
- \pause\pause\bigskip
- By \alert<4->{induction}, we can {\bf prove} these properties.\bigskip
-
- \begin{tabular}{lrcl}
- Lemmas: & \bl{nullable (r)} & \bl{$\Longleftrightarrow$} & \bl{[] $\in$ \LL (r)}\\
- & \bl{s $\in$ \LL (der c r)} & \bl{$\Longleftrightarrow$} & \bl{(c::s) $\in$ \LL (r)}\\
- \end{tabular}
-
- \only<4->{
- \begin{textblock}{3}(0.9,4.5)
- \rd{\huge$\forall$\large{}r s.}
- \end{textblock}}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{
- \begin{tabular}{c}
- \mbox{}\\[23mm]
- \LARGE Demo
- \end{tabular}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
-
- \mbox{}\\[-2mm]
-
- \small
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}ll@ {}}
- \bl{nullable (NULL)} & \bl{$=$} & \bl{false} &\\
- \bl{nullable (EMPTY)} & \bl{$=$} & \bl{true} &\\
- \bl{nullable (CHR c)} & \bl{$=$} & \bl{false} &\\
- \bl{nullable (ALT r$_1$ r$_2$)} & \bl{$=$} & \bl{(nullable r$_1$) orelse (nullable r$_2$)} & \\
- \bl{nullable (SEQ r$_1$ r$_2$)} & \bl{$=$} & \bl{(nullable r$_1$) andalso (nullable r$_2$)} & \\
- \bl{nullable (STAR r)} & \bl{$=$} & \bl{true} & \\
- \end{tabular}\medskip
-
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-10mm}}l@ {}}
- \bl{der c (NULL)} & \bl{$=$} & \bl{NULL} & \\
- \bl{der c (EMPTY)} & \bl{$=$} & \bl{NULL} & \\
- \bl{der c (CHR d)} & \bl{$=$} & \bl{if c=d then EMPTY else NULL} & \\
- \bl{der c (ALT r$_1$ r$_2$)} & \bl{$=$} & \bl{ALT (der c r$_1$) (der c r$_2$)} & \\
- \bl{der c (SEQ r$_1$ r$_2$)} & \bl{$=$} & \bl{ALT (SEQ (der c r$_1$) r$_2$)} & \\
- & & \bl{\phantom{ALT} (if nullable r$_1$ then der c r$_2$ else NULL)}\\
- \bl{der c (STAR r)} & \bl{$=$} & \bl{SEQ (der c r) (STAR r)} &\smallskip\\
-
- \bl{derivative r []} & \bl{$=$} & \bl{r} & \\
- \bl{derivative r (c::s)} & \bl{$=$} & \bl{derivative (der c r) s} & \\
- \end{tabular}\medskip
-
- \bl{matches r s $=$ nullable (derivative r s)}
-
- \only<2>{
- \begin{textblock}{8}(1.5,4)
- \includegraphics[scale=0.3]{approved.png}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{No Automata?}
-
- You might be wondering why I did not use any automata?
-
- \begin{itemize}
- \item {\bf Def.:} A \alert{regular language} is one where there is a DFA that
- recognises it.\bigskip\pause
- \end{itemize}
-
-
- There are many reasons why this is a good definition:\medskip
- \begin{itemize}
- \item pumping lemma
- \item closure properties of regular languages\\ (e.g.~closure under complement)
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[t]
- \frametitle{Really Bad News!}
-
- DFAs are bad news for formalisations in theorem provers. They might
- be represented as:
-
- \begin{itemize}
- \item graphs
- \item matrices
- \item partial functions
- \end{itemize}
-
- All constructions are messy to reason about.\bigskip\bigskip
- \pause
-
- \small
- \only<2>{
- Constable et al needed (on and off) 18 months for a 3-person team
- to formalise automata theory in Nuprl including Myhill-Nerode. There is
- only very little other formalised work on regular languages I know of
- in Coq, Isabelle and HOL.}
- \only<3>{Typical textbook reasoning goes like: ``\ldots if \smath{M} and \smath{N} are any two
- automata with no inaccessible states \ldots''
- }
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{}
- \large
- \begin{center}
- \begin{tabular}{p{9cm}}
- My point:\bigskip\\
-
- The theory about regular languages can be reformulated
- to be more\\ suitable for theorem proving.
- \end{tabular}
- \end{center}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE The Myhill-Nerode Theorem}
-
- \begin{itemize}
- \item provides necessary and suf\!ficient conditions for a language
- being regular (pumping lemma only necessary)\medskip
-
- \item will help with closure properties of regular languages\bigskip\pause
-
- \item key is the equivalence relation:\smallskip
- \begin{center}
- \smath{x \approx_{L} y \,\dn\, \forall z.\; x @ z \in L \Leftrightarrow y @ z \in L}
- \end{center}
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE The Myhill-Nerode Theorem}
-
- \mbox{}\\[5cm]
-
- \begin{itemize}
- \item \smath{\text{finite}\, (U\!N\!IV /\!/ \approx_L) \;\Leftrightarrow\; L\; \text{is regular}}
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Equivalence Classes}
-
- \begin{itemize}
- \item \smath{L = []}
- \begin{center}
- \smath{\Big\{\{[]\},\; U\!N\!IV - \{[]\}\Big\}}
- \end{center}\bigskip\bigskip
-
- \item \smath{L = [c]}
- \begin{center}
- \smath{\Big\{\{[]\},\; \{[c]\},\; U\!N\!IV - \{[], [c]\}\Big\}}
- \end{center}\bigskip\bigskip
-
- \item \smath{L = \varnothing}
- \begin{center}
- \smath{\Big\{U\!N\!IV\Big\}}
- \end{center}
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Regular Languages}
-
- \begin{itemize}
- \item \smath{L} is regular \smath{\dn} if there is an automaton \smath{M}
- such that \smath{\mathbb{L}(M) = L}\\[1.5cm]
-
- \item Myhill-Nerode:
-
- \begin{center}
- \begin{tabular}{l}
- finite $\Rightarrow$ regular\\
- \;\;\;\smath{\text{finite}\,(U\!N\!IV /\!/ \approx_L) \Rightarrow \exists r.\; L = \mathbb{L}(r)}\\[3mm]
- regular $\Rightarrow$ finite\\
- \;\;\;\smath{\text{finite}\, (U\!N\!IV /\!/ \approx_{\mathbb{L}(r)})}
- \end{tabular}
- \end{center}
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Final Equiv.~Classes}
-
- \mbox{}\\[3cm]
-
- \begin{itemize}
- \item \smath{\text{finals}\,L \dn
- \{{\lbrack\mkern-2mu\lbrack{s}\rbrack\mkern-2mu\rbrack}_\approx\;|\; s \in L\}}\\
- \medskip
-
- \item we can prove: \smath{L = \bigcup (\text{finals}\,L)}
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Transitions between ECs}
-
- \smath{L = \{[c]\}}
-
- \begin{tabular}{@ {\hspace{-7mm}}cc}
- \begin{tabular}{c}
- \begin{tikzpicture}[shorten >=1pt,node distance=2cm,auto, ultra thick]
- \tikzstyle{state}=[circle,thick,draw=blue!75,fill=blue!20,minimum size=0mm]
-
- %\draw[help lines] (0,0) grid (3,2);
-
- \node[state,initial] (q_0) {$R_1$};
- \node[state,accepting] (q_1) [above right of=q_0] {$R_2$};
- \node[state] (q_2) [below right of=q_0] {$R_3$};
-
- \path[->] (q_0) edge node {c} (q_1)
- edge node [swap] {$\Sigma-{c}$} (q_2)
- (q_2) edge [loop below] node {$\Sigma$} ()
- (q_1) edge node {$\Sigma$} (q_2);
- \end{tikzpicture}
- \end{tabular}
- &
- \begin{tabular}[t]{ll}
- \\[-20mm]
- \multicolumn{2}{l}{\smath{U\!N\!IV /\!/\approx_L} produces}\\[4mm]
-
- \smath{R_1}: & \smath{\{[]\}}\\
- \smath{R_2}: & \smath{\{[c]\}}\\
- \smath{R_3}: & \smath{U\!N\!IV - \{[], [c]\}}\\[6mm]
- \multicolumn{2}{l}{\onslide<2->{\smath{X \stackrel{c}{\longrightarrow} Y \dn X ;; [c] \subseteq Y}}}
- \end{tabular}
-
- \end{tabular}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Systems of Equations}
-
- Inspired by a method of Brzozowski\;'64, we can build an equational system
- characterising the equivalence classes:
-
- \begin{center}
- \begin{tabular}{@ {\hspace{-20mm}}c}
- \\[-13mm]
- \begin{tikzpicture}[shorten >=1pt,node distance=2cm,auto, ultra thick]
- \tikzstyle{state}=[circle,thick,draw=blue!75,fill=blue!20,minimum size=0mm]
-
- %\draw[help lines] (0,0) grid (3,2);
-
- \node[state,initial] (p_0) {$R_1$};
- \node[state,accepting] (p_1) [right of=q_0] {$R_2$};
-
- \path[->] (p_0) edge [bend left] node {a} (p_1)
- edge [loop above] node {b} ()
- (p_1) edge [loop above] node {a} ()
- edge [bend left] node {b} (p_0);
- \end{tikzpicture}\\
- \\[-13mm]
- \end{tabular}
- \end{center}
-
- \begin{center}
- \begin{tabular}{@ {\hspace{-6mm}}ll@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
- & \smath{R_1} & \smath{\equiv} & \smath{R_1;b + R_2;b \onslide<2->{\alert<2>{+ \lambda;[]}}}\\
- & \smath{R_2} & \smath{\equiv} & \smath{R_1;a + R_2;a}\medskip\\
- \onslide<3->{we can prove}
- & \onslide<3->{\smath{R_1}} & \onslide<3->{\smath{=}}
- & \onslide<3->{\smath{R_1;; \mathbb{L}(b) \,\cup\, R_2;;\mathbb{L}(b) \,\cup\, \{[]\}}}\\
- & \onslide<3->{\smath{R_2}} & \onslide<3->{\smath{=}}
- & \onslide<3->{\smath{R_1;; \mathbb{L}(a) \,\cup\, R_2;;\mathbb{L}(a)}}\\
- \end{tabular}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[t]
- \small
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}ll}
- \onslide<1->{\smath{R_1}} & \onslide<1->{\smath{=}}
- & \onslide<1->{\smath{R_1; b + R_2; b + \lambda;[]}}\\
- \onslide<1->{\smath{R_2}} & \onslide<1->{\smath{=}}
- & \onslide<1->{\smath{R_1; a + R_2; a}}\\
-
- & & & \onslide<2->{by Arden}\\
-
- \onslide<2->{\smath{R_1}} & \onslide<2->{\smath{=}}
- & \onslide<2->{\smath{R_1; b + R_2; b + \lambda;[]}}\\
- \onslide<2->{\smath{R_2}} & \onslide<2->{\smath{=}}
- & \only<2>{\smath{R_1; a + R_2; a}}%
- \only<3->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<4->{by Arden}\\
-
- \onslide<4->{\smath{R_1}} & \onslide<4->{\smath{=}}
- & \onslide<4->{\smath{R_2; b \cdot b^\star+ \lambda;b^\star}}\\
- \onslide<4->{\smath{R_2}} & \onslide<4->{\smath{=}}
- & \onslide<4->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<5->{by substitution}\\
-
- \onslide<5->{\smath{R_1}} & \onslide<5->{\smath{=}}
- & \onslide<5->{\smath{R_1; a\cdot a^\star \cdot b \cdot b^\star+ \lambda;b^\star}}\\
- \onslide<5->{\smath{R_2}} & \onslide<5->{\smath{=}}
- & \onslide<5->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<6->{by Arden}\\
-
- \onslide<6->{\smath{R_1}} & \onslide<6->{\smath{=}}
- & \onslide<6->{\smath{\lambda;b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star}}\\
- \onslide<6->{\smath{R_2}} & \onslide<6->{\smath{=}}
- & \onslide<6->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<7->{by substitution}\\
-
- \onslide<7->{\smath{R_1}} & \onslide<7->{\smath{=}}
- & \onslide<7->{\smath{\lambda;b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star}}\\
- \onslide<7->{\smath{R_2}} & \onslide<7->{\smath{=}}
- & \onslide<7->{\smath{\lambda; b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star
- \cdot a\cdot a^\star}}\\
- \end{tabular}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE A Variant of Arden's Lemma}
-
- {\bf Arden's Lemma:}\smallskip
-
- If \smath{[] \not\in A} then
- \begin{center}
- \smath{X = X; A + \text{something}}
- \end{center}
- has the (unique) solution
- \begin{center}
- \smath{X = \text{something} ; A^\star}
- \end{center}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \small
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}ll}
- \onslide<1->{\smath{R_1}} & \onslide<1->{\smath{=}}
- & \onslide<1->{\smath{R_1; b + R_2; b + \lambda;[]}}\\
- \onslide<1->{\smath{R_2}} & \onslide<1->{\smath{=}}
- & \onslide<1->{\smath{R_1; a + R_2; a}}\\
-
- & & & \onslide<2->{by Arden}\\
-
- \onslide<2->{\smath{R_1}} & \onslide<2->{\smath{=}}
- & \onslide<2->{\smath{R_1; b + R_2; b + \lambda;[]}}\\
- \onslide<2->{\smath{R_2}} & \onslide<2->{\smath{=}}
- & \only<2>{\smath{R_1; a + R_2; a}}%
- \only<3->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<4->{by Arden}\\
-
- \onslide<4->{\smath{R_1}} & \onslide<4->{\smath{=}}
- & \onslide<4->{\smath{R_2; b \cdot b^\star+ \lambda;b^\star}}\\
- \onslide<4->{\smath{R_2}} & \onslide<4->{\smath{=}}
- & \onslide<4->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<5->{by substitution}\\
-
- \onslide<5->{\smath{R_1}} & \onslide<5->{\smath{=}}
- & \onslide<5->{\smath{R_1; a\cdot a^\star \cdot b \cdot b^\star+ \lambda;b^\star}}\\
- \onslide<5->{\smath{R_2}} & \onslide<5->{\smath{=}}
- & \onslide<5->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<6->{by Arden}\\
-
- \onslide<6->{\smath{R_1}} & \onslide<6->{\smath{=}}
- & \onslide<6->{\smath{\lambda;b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star}}\\
- \onslide<6->{\smath{R_2}} & \onslide<6->{\smath{=}}
- & \onslide<6->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<7->{by substitution}\\
-
- \onslide<7->{\smath{R_1}} & \onslide<7->{\smath{=}}
- & \onslide<7->{\smath{\lambda;b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star}}\\
- \onslide<7->{\smath{R_2}} & \onslide<7->{\smath{=}}
- & \onslide<7->{\smath{\lambda; b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star
- \cdot a\cdot a^\star}}\\
- \end{tabular}
- \end{center}
-
- \only<8->{
- \begin{textblock}{6}(2.5,4)
- \begin{block}{}
- \begin{minipage}{8cm}\raggedright
-
- \begin{tikzpicture}[shorten >=1pt,node distance=2cm,auto, ultra thick, inner sep=1mm]
- \tikzstyle{state}=[circle,thick,draw=blue!75,fill=blue!20,minimum size=0mm]
-
- %\draw[help lines] (0,0) grid (3,2);
-
- \node[state,initial] (p_0) {$R_1$};
- \node[state,accepting] (p_1) [right of=q_0] {$R_2$};
-
- \path[->] (p_0) edge [bend left] node {a} (p_1)
- edge [loop above] node {b} ()
- (p_1) edge [loop above] node {a} ()
- edge [bend left] node {b} (p_0);
- \end{tikzpicture}
-
- \end{minipage}
- \end{block}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE The Equ's Solving Algorithm}
-
- \begin{itemize}
- \item The algorithm must terminate: Arden makes one equation smaller;
- substitution deletes one variable from the right-hand sides.\bigskip
-
- \item We need to maintain the invariant that Arden is applicable
- (if \smath{[] \not\in A} then \ldots):\medskip
-
- \begin{center}\small
- \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}ll}
- \smath{R_1} & \smath{=} & \smath{R_1; b + R_2; b + \lambda;[]}\\
- \smath{R_2} & \smath{=} & \smath{R_1; a + R_2; a}\\
-
- & & & by Arden\\
-
- \smath{R_1} & \smath{=} & \smath{R_1; b + R_2; b + \lambda;[]}\\
- \smath{R_2} & \smath{=} & \smath{R_1; a\cdot a^\star}\\
- \end{tabular}
- \end{center}
-
- \end{itemize}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE The Other Direction}
-
- One has to prove
-
- \begin{center}
- \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r)})}
- \end{center}
-
- by induction on \smath{r}. This is straightforward for \\the base cases:\small
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{1mm}}l}
- \smath{U\!N\!IV /\!/ \!\approx_{\emptyset}} & \smath{= \{U\!N\!IV\}}\smallskip\\
- \smath{U\!N\!IV /\!/ \!\approx_{\{[]\}}} & \smath{\subseteq \{\{[]\}, U\!N\!IV - \{[]\}\}}\smallskip\\
- \smath{U\!N\!IV /\!/ \!\approx_{\{[c]\}}} & \smath{\subseteq \{\{[]\}, \{[c]\}, U\!N\!IV - \{[], [c]\}\}}
- \end{tabular}
- \end{center}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[t]
- \frametitle{\LARGE The Other Direction}
-
- More complicated are the inductive cases:\\ one needs to prove that if
-
- \begin{center}
- \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r_1)})}\hspace{3mm}
- \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r_2)})}
- \end{center}
-
- then
-
- \begin{center}
- \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r_1) \,\cup\, \mathbb{L}(r_2)})}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[t]
- \frametitle{\LARGE Helper Lemma}
-
- \begin{center}
- \begin{tabular}{p{10cm}}
- %If \smath{\text{finite} (f\;' A)} and \smath{f} is injective
- %(on \smath{A}),\\ then \smath{\text{finite}\,A}.
- Given two equivalence relations \smath{R_1} and \smath{R_2} with
- \smath{R_1} refining \smath{R_2} (\smath{R_1 \subseteq R_2}).\\
- Then\medskip\\
- \smath{\;\;\text{finite} (U\!N\!IV /\!/ R_1) \Rightarrow \text{finite} (U\!N\!IV /\!/ R_2)}
- \end{tabular}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\Large Derivatives and Left-Quotients}
- \small
- Work by Brozowski ('64) and Antimirov ('96):\pause\smallskip
-
-
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-10mm}}l@ {}}
- \multicolumn{4}{@ {}l}{Left-Quotient:}\\
- \multicolumn{4}{@ {}l}{\bl{$\text{Ders}\;\text{s}\,A \dn \{\text{s'} \;|\; \text{s @ s'} \in A\}$}}\bigskip\\
-
- \multicolumn{4}{@ {}l}{Derivative:}\\
- \bl{der c ($\varnothing$)} & \bl{$=$} & \bl{$\varnothing$} & \\
- \bl{der c ([])} & \bl{$=$} & \bl{$\varnothing$} & \\
- \bl{der c (d)} & \bl{$=$} & \bl{if c = d then [] else $\varnothing$} & \\
- \bl{der c (r$_1$ + r$_2$)} & \bl{$=$} & \bl{(der c r$_1$) + (der c r$_2$)} & \\
- \bl{der c (r$_1$ $\cdot$ r$_2$)} & \bl{$=$} & \bl{((der c r$_1$) $\cdot$ r$_2$)} & \\
- & & \bl{\;\;\;\;+ (if nullable r$_1$ then der c r$_2$ else $\varnothing$)}\\
- \bl{der c (r$^*$)} & \bl{$=$} & \bl{(der c r) $\cdot$ r$^*$} &\smallskip\\
-
- \bl{ders [] r} & \bl{$=$} & \bl{r} & \\
- \bl{ders (s @ [c]) r} & \bl{$=$} & \bl{der c (ders s r)} & \\
- \end{tabular}\pause
-
- \begin{center}
- \alert{$\Rightarrow$}\smath{\;\;\text{Ders}\,\text{s}\,(\mathbb{L}(\text{r})) = \mathbb{L} (\text{ders s r})}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Left-Quotients and MN-Rels}
-
- \begin{itemize}
- \item \smath{x \approx_{A} y \,\dn\, \forall z.\; x @ z \in A \Leftrightarrow y @ z \in A}\medskip
- \item \bl{$\text{Ders}\;s\,A \dn \{s' \;|\; s @ s' \in A\}$}
- \end{itemize}\bigskip
-
- \begin{center}
- \smath{x \approx_A y \Longleftrightarrow \text{Ders}\;x\;A = \text{Ders}\;y\;A}
- \end{center}\bigskip\pause\small
-
- which means
-
- \begin{center}
- \smath{x \approx_{\mathbb{L}(r)} y \Longleftrightarrow
- \mathbb{L}(\text{ders}\;x\;r) = \mathbb{L}(\text{ders}\;y\;r)}
- \end{center}\pause
-
- \hspace{8.8mm}or
- \smath{\;x \approx_{\mathbb{L}(r)} y \Longleftarrow
- \text{ders}\;x\;r = \text{ders}\;y\;r}
-
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Partial Derivatives}
-
- Antimirov: \bl{pder : rexp $\Rightarrow$ rexp set}\bigskip
-
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-10mm}}l@ {}}
- \bl{pder c ($\varnothing$)} & \bl{$=$} & \bl{\{$\varnothing$\}} & \\
- \bl{pder c ([])} & \bl{$=$} & \bl{\{$\varnothing$\}} & \\
- \bl{pder c (d)} & \bl{$=$} & \bl{if c = d then \{[]\} else \{$\varnothing$\}} & \\
- \bl{pder c (r$_1$ + r$_2$)} & \bl{$=$} & \bl{(pder c r$_1$) $\cup$ (pder c r$_2$)} & \\
- \bl{pder c (r$_1$ $\cdot$ r$_2$)} & \bl{$=$} & \bl{(pder c r$_1$) $\odot$ r$_2$} & \\
- & & \bl{\hspace{-10mm}$\cup$ (if nullable r$_1$ then pder c r$_2$ else $\varnothing$)}\\
- \bl{pder c (r$^*$)} & \bl{$=$} & \bl{(pder c r) $\odot$ r$^*$} &\smallskip\\
- \end{tabular}
-
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-10mm}}l@ {}}
- \bl{pders [] r} & \bl{$=$} & \bl{r} & \\
- \bl{pders (s @ [c]) r} & \bl{$=$} & \bl{pder c (pders s r)} & \\
- \end{tabular}\pause
-
- \begin{center}
- \alert{$\Rightarrow$}\smath{\;\;\text{Ders}\,\text{s}\,(\mathbb{L}(\text{r})) = \bigcup (\mathbb{L}\;`\; (\text{pders s r}))}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[t]
- \frametitle{\LARGE Final Result}
-
- \mbox{}\\[7mm]
-
- \begin{itemize}
- \item \alt<1>{\smath{\text{pders x r \mbox{$=$} pders y r}}}
- {\smath{\underbrace{\text{pders x r \mbox{$=$} pders y r}}_{R_1}}}
- refines \bl{x $\approx_{\mathbb{L}(\text{r})}$ y}\pause
- \item \smath{\text{finite} (U\!N\!IV /\!/ R_1)} \bigskip\pause
- \item Therefore \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r)})}. Qed.
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE What Have We Achieved?}
-
- \begin{itemize}
- \item \smath{\text{finite}\, (U\!N\!IV /\!/ \approx_L) \;\Leftrightarrow\; L\; \text{is regular}}
- \bigskip\pause
- \item regular languages are closed under complementation; this is now easy\medskip
- \begin{center}
- \smath{U\!N\!IV /\!/ \approx_L \;\;=\;\; U\!N\!IV /\!/ \approx_{-L}}
- \end{center}
- \end{itemize}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Examples}
-
- \begin{itemize}
- \item \smath{L \equiv \Sigma^\star 0 \Sigma} is regular
- \begin{quote}\small
- \begin{tabular}{lcl}
- \smath{A_1} & \smath{=} & \smath{\Sigma^\star 00}\\
- \smath{A_2} & \smath{=} & \smath{\Sigma^\star 01}\\
- \smath{A_3} & \smath{=} & \smath{\Sigma^\star 10 \cup \{0\}}\\
- \smath{A_4} & \smath{=} & \smath{\Sigma^\star 11 \cup \{1\} \cup \{[]\}}\\
- \end{tabular}
- \end{quote}
-
- \item \smath{L \equiv \{ 0^n 1^n \,|\, n \ge 0\}} is not regular
- \begin{quote}\small
- \begin{tabular}{lcl}
- \smath{B_0} & \smath{=} & \smath{\{0^n 1^n \,|\, n \ge 0\}}\\
- \smath{B_1} & \smath{=} & \smath{\{0^n 1^{(n-1)} \,|\, n \ge 1\}}\\
- \smath{B_2} & \smath{=} & \smath{\{0^n 1^{(n-2)} \,|\, n \ge 2\}}\\
- \smath{B_3} & \smath{=} & \smath{\{0^n 1^{(n-3)} \,|\, n \ge 3\}}\\
- & \smath{\vdots} &\\
- \end{tabular}
- \end{quote}
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE What We Have Not Achieved}
-
- \begin{itemize}
- \item regular expressions are not good if you look for a minimal
- one for a language (DFAs have this notion)\pause\bigskip
-
- \item Is there anything to be said about context free languages:\medskip
-
- \begin{quote}
- A context free language is where every string can be recognised by
- a pushdown automaton.\bigskip
- \end{quote}
- \end{itemize}
-
- \textcolor{gray}{\footnotesize Yes. Derivatives also work for c-f grammars. Ongoing work.}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Conclusion}
-
- \begin{itemize}
- \item We formalised the Myhill-Nerode theorem based on
- regular expressions only (DFAs are difficult to deal with in a theorem prover).\smallskip
-
- \item Seems to be a common theme: algorithms need to be reformulated
- to better suit formal treatment.\smallskip
-
- \item The most interesting aspect is that we are able to
- implement the matcher directly inside the theorem prover
- (ongoing work).\smallskip
-
- \item Parsing is a vast field which seem to offer new results.
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[b]
- \frametitle{
- \begin{tabular}{c}
- \mbox{}\\[13mm]
- \alert{\LARGE Thank you very much!}\\
- \alert{\Large Questions?}
- \end{tabular}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-
-(*<*)
-end
-(*>*)
\ No newline at end of file
--- a/Slides/Slides8.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1419 +0,0 @@
-(*<*)
-theory Slides8
-imports "~~/src/HOL/Library/LaTeXsugar" "Nominal"
-begin
-
-declare [[show_question_marks = false]]
-
-notation (latex output)
- set ("_") and
- Cons ("_::/_" [66,65] 65)
-
-(*>*)
-
-text_raw {*
- \renewcommand{\slidecaption}{Copenhagen, 23rd~May 2011}
-
- \newcommand{\abst}[2]{#1.#2}% atom-abstraction
- \newcommand{\pair}[2]{\langle #1,#2\rangle} % pairing
- \newcommand{\susp}{{\boldsymbol{\cdot}}}% for suspensions
- \newcommand{\unit}{\langle\rangle}% unit
- \newcommand{\app}[2]{#1\,#2}% application
- \newcommand{\eqprob}{\mathrel{{\approx}?}}
- \newcommand{\freshprob}{\mathrel{\#?}}
- \newcommand{\redu}[1]{\stackrel{#1}{\Longrightarrow}}% reduction
- \newcommand{\id}{\varepsilon}% identity substitution
-
- \newcommand{\bl}[1]{\textcolor{blue}{#1}}
- \newcommand{\gr}[1]{\textcolor{gray}{#1}}
- \newcommand{\rd}[1]{\textcolor{red}{#1}}
-
- \newcommand{\ok}{\includegraphics[scale=0.07]{ok.png}}
- \newcommand{\notok}{\includegraphics[scale=0.07]{notok.png}}
- \newcommand{\largenotok}{\includegraphics[scale=1]{notok.png}}
-
- \renewcommand{\Huge}{\fontsize{61.92}{77}\selectfont}
- \newcommand{\veryHuge}{\fontsize{74.3}{93}\selectfont}
- \newcommand{\VeryHuge}{\fontsize{89.16}{112}\selectfont}
- \newcommand{\VERYHuge}{\fontsize{107}{134}\selectfont}
-
- \newcommand{\LL}{$\mathbb{L}\,$}
-
-
- \pgfdeclareradialshading{smallbluesphere}{\pgfpoint{0.5mm}{0.5mm}}%
- {rgb(0mm)=(0,0,0.9);
- rgb(0.9mm)=(0,0,0.7);
- rgb(1.3mm)=(0,0,0.5);
- rgb(1.4mm)=(1,1,1)}
-
- \def\myitemi{\begin{pgfpicture}{-1ex}{-0.55ex}{1ex}{1ex}
- \usebeamercolor[fg]{subitem projected}
- {\pgftransformscale{0.8}\pgftext{\normalsize\pgfuseshading{bigsphere}}}
- \pgftext{%
- \usebeamerfont*{subitem projected}}
- \end{pgfpicture}}
-
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[t]
- \frametitle{%
- \begin{tabular}{@ {\hspace{-3mm}}c@ {}}
- \\
- \LARGE Verifying a Regular Expression\\[-1mm]
- \LARGE Matcher and Formal Language\\[-1mm]
- \LARGE Theory\\[5mm]
- \end{tabular}}
- \begin{center}
- Christian Urban\\
- \small Technical University of Munich, Germany
- \end{center}
-
-
- \begin{center}
- \small joint work with Chunhan Wu and Xingyuan Zhang from the PLA
- University of Science and Technology in Nanjing
- \end{center}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{This Talk: 4 Points}
- \large
- \begin{itemize}
- \item It is easy to make mistakes.\medskip
- \item Theorem provers can prevent mistakes, {\bf if} the problem
- is formulated so that it is suitable for theorem provers.\medskip
- \item This re-formulation can be done, even in domains where
- we least expect it.\medskip
- \item Where theorem provers are superior to the {\color{gray}{(best)}} human reasoners. ;o)
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
-
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{}
-
- \begin{tabular}{c@ {\hspace{2mm}}c}
- \\[6mm]
- \begin{tabular}{c}
- \includegraphics[scale=0.12]{harper.jpg}\\[-2mm]
- {\footnotesize Bob Harper}\\[-2.5mm]
- {\footnotesize (CMU)}
- \end{tabular}
- \begin{tabular}{c}
- \includegraphics[scale=0.36]{pfenning.jpg}\\[-2mm]
- {\footnotesize Frank Pfenning}\\[-2.5mm]
- {\footnotesize (CMU)}
- \end{tabular} &
-
- \begin{tabular}{p{6cm}}
- \raggedright
- \color{gray}{published a proof in\\ {\bf ACM Transactions on Computational Logic} (2005),
- $\sim$31pp}
- \end{tabular}\\
-
- \pause
- \\[0mm]
-
- \begin{tabular}{c}
- \includegraphics[scale=0.36]{appel.jpg}\\[-2mm]
- {\footnotesize Andrew Appel}\\[-2.5mm]
- {\footnotesize (Princeton)}
- \end{tabular} &
-
- \begin{tabular}{p{6cm}}
- \raggedright
- \color{gray}{relied on their proof in a\\ {\bf security} critical application}
- \end{tabular}
- \end{tabular}
-
-
-
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
-
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}
- \frametitle{Proof-Carrying Code}
-
- \begin{textblock}{10}(2.5,2.2)
- \begin{block}{Idea:}
- \begin{center}
- \begin{tikzpicture}
- \draw[help lines,cream] (0,0.2) grid (8,4);
-
- \draw[line width=1mm, red] (5.5,0.6) rectangle (7.5,4);
- \node[anchor=base] at (6.5,2.8)
- {\small\begin{tabular}{@ {}p{1.9cm}@ {}}\centering user: untrusted code\end{tabular}};
-
- \draw[line width=1mm, red] (0.5,0.6) rectangle (2.5,4);
- \node[anchor=base] at (1.5,2.3)
- {\small\begin{tabular}{@ {}p{1.9cm}@ {}}\centering developer ---\\ web server\end{tabular}};
-
- \onslide<3->{
- \draw[line width=1mm, red, fill=red] (5.5,0.6) rectangle (7.5,1.8);
- \node[anchor=base,white] at (6.5,1.1)
- {\small\begin{tabular}{@ {}p{1.9cm}@ {}}\bf\centering proof- checker\end{tabular}};}
-
- \node at (3.8,3.0) [single arrow, fill=red,text=white, minimum height=3cm]{\bf code};
- \onslide<2->{
- \node at (3.8,1.3) [single arrow, fill=red,text=white, minimum height=3cm]{\bf certificate};
- \node at (3.8,1.9) {\small\color{gray}{\mbox{}\hspace{-1mm}a proof in LF}};
- }
-
-
- \end{tikzpicture}
- \end{center}
- \end{block}
- \end{textblock}
-
- %\begin{textblock}{15}(2,12)
- %\small
- %\begin{itemize}
- %\item<4-> Appel's checker is $\sim$2700 lines of code (1865 loc of\\ LF definitions;
- %803 loc in C including 2 library functions)\\[-3mm]
- %\item<5-> 167 loc in C implement a type-checker
- %\end{itemize}
- %\end{textblock}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text {*
- \tikzstyle{every node}=[node distance=25mm,text height=1.5ex, text depth=.25ex]
- \tikzstyle{node1}=[rectangle, minimum size=10mm, rounded corners=3mm, very thick,
- draw=black!50, top color=white, bottom color=black!20]
- \tikzstyle{node2}=[rectangle, minimum size=12mm, rounded corners=3mm, very thick,
- draw=red!70, top color=white, bottom color=red!50!black!20]
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<2->[squeeze]
- \frametitle{}
-
- \begin{columns}
-
- \begin{column}{0.8\textwidth}
- \begin{textblock}{0}(1,2)
-
- \begin{tikzpicture}
- \matrix[ampersand replacement=\&,column sep=7mm, row sep=5mm]
- { \&[-10mm]
- \node (def1) [node1] {\large\hspace{1mm}Spec\hspace{1mm}\mbox{}}; \&
- \node (proof1) [node1] {\large Proof}; \&
- \node (alg1) [node1] {\large\hspace{1mm}Alg\hspace{1mm}\mbox{}}; \\
-
- \onslide<4->{\node {\begin{tabular}{c}\small 1st\\[-2.5mm] \footnotesize solution\end{tabular}};} \&
- \onslide<4->{\node (def2) [node2] {\large Spec$^\text{+ex}$};} \&
- \onslide<4->{\node (proof2) [node1] {\large Proof};} \&
- \onslide<4->{\node (alg2) [node1] {\large\hspace{1mm}Alg\hspace{1mm}\mbox{}};} \\
-
- \onslide<5->{\node {\begin{tabular}{c}\small 2nd\\[-2.5mm] \footnotesize solution\end{tabular}};} \&
- \onslide<5->{\node (def3) [node1] {\large\hspace{1mm}Spec\hspace{1mm}\mbox{}};} \&
- \onslide<5->{\node (proof3) [node1] {\large Proof};} \&
- \onslide<5->{\node (alg3) [node2] {\large Alg$^\text{-ex}$};} \\
-
- \onslide<6->{\node {\begin{tabular}{c}\small 3rd\\[-2.5mm] \footnotesize solution\end{tabular}};} \&
- \onslide<6->{\node (def4) [node1] {\large\hspace{1mm}Spec\hspace{1mm}\mbox{}};} \&
- \onslide<6->{\node (proof4) [node2] {\large\hspace{1mm}Proof\hspace{1mm}};} \&
- \onslide<6->{\node (alg4) [node1] {\large\hspace{1mm}Alg\hspace{1mm}\mbox{}};} \\
- };
-
- \draw[->,black!50,line width=2mm] (proof1) -- (def1);
- \draw[->,black!50,line width=2mm] (proof1) -- (alg1);
-
- \onslide<4->{\draw[->,black!50,line width=2mm] (proof2) -- (def2);}
- \onslide<4->{\draw[->,black!50,line width=2mm] (proof2) -- (alg2);}
-
- \onslide<5->{\draw[->,black!50,line width=2mm] (proof3) -- (def3);}
- \onslide<5->{\draw[->,black!50,line width=2mm] (proof3) -- (alg3);}
-
- \onslide<6->{\draw[->,black!50,line width=2mm] (proof4) -- (def4);}
- \onslide<6->{\draw[->,black!50,line width=2mm] (proof4) -- (alg4);}
-
- \onslide<3->{\draw[white,line width=1mm] (1.1,3.2) -- (0.9,2.85) -- (1.1,2.35) -- (0.9,2.0);}
- \end{tikzpicture}
-
- \end{textblock}
- \end{column}
- \end{columns}
-
-
- \begin{textblock}{3}(12,3.6)
- \onslide<4->{
- \begin{tikzpicture}
- \node at (0,0) [single arrow, shape border rotate=270, fill=red,text=white]{2h};
- \end{tikzpicture}}
- \end{textblock}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-
-(*<*)
-atom_decl name
-
-nominal_datatype lam =
- Var "name"
- | App "lam" "lam"
- | Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
-
-nominal_primrec
- subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]")
-where
- "(Var x)[y::=s] = (if x=y then s else (Var x))"
-| "(App t\<^isub>1 t\<^isub>2)[y::=s] = App (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])"
-| "x\<sharp>(y,s) \<Longrightarrow> (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])"
-apply(finite_guess)+
-apply(rule TrueI)+
-apply(simp add: abs_fresh)
-apply(fresh_guess)+
-done
-
-lemma subst_eqvt[eqvt]:
- fixes pi::"name prm"
- shows "pi\<bullet>(t1[x::=t2]) = (pi\<bullet>t1)[(pi\<bullet>x)::=(pi\<bullet>t2)]"
-by (nominal_induct t1 avoiding: x t2 rule: lam.strong_induct)
- (auto simp add: perm_bij fresh_atm fresh_bij)
-
-lemma fresh_fact:
- fixes z::"name"
- shows "\<lbrakk>z\<sharp>s; (z=y \<or> z\<sharp>t)\<rbrakk> \<Longrightarrow> z\<sharp>t[y::=s]"
-by (nominal_induct t avoiding: z y s rule: lam.strong_induct)
- (auto simp add: abs_fresh fresh_prod fresh_atm)
-
-lemma forget:
- assumes asm: "x\<sharp>L"
- shows "L[x::=P] = L"
- using asm
-by (nominal_induct L avoiding: x P rule: lam.strong_induct)
- (auto simp add: abs_fresh fresh_atm)
-(*>*)
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}
-
- \begin{textblock}{16}(1,1)
- \renewcommand{\isasymbullet}{$\cdot$}
- \tiny\color{black}
-*}
-lemma substitution_lemma_not_to_be_tried_at_home:
- assumes asm: "x\<noteq>y" "x\<sharp>L"
- shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
-using asm
-proof (induct M arbitrary: x y N L rule: lam.induct)
- case (Lam z M1)
- have ih: "\<And>x y N L. \<lbrakk>x\<noteq>y; x\<sharp>L\<rbrakk> \<Longrightarrow> M1[x::=N][y::=L] = M1[y::=L][x::=N[y::=L]]" by fact
- have "x\<noteq>y" by fact
- have "x\<sharp>L" by fact
- obtain z'::"name" where fc: "z'\<sharp>(x,y,z,M1,N,L)" by (rule exists_fresh) (auto simp add: fs_name1)
- have eq: "Lam [z'].([(z',z)]\<bullet>M1) = Lam [z].M1" using fc
- by (auto simp add: lam.inject alpha fresh_prod fresh_atm)
- have fc': "z'\<sharp>N[y::=L]" using fc by (simp add: fresh_fact fresh_prod)
- have "([(z',z)]\<bullet>x) \<noteq> ([(z',z)]\<bullet>y)" using `x\<noteq>y` by (auto simp add: calc_atm)
- moreover
- have "([(z',z)]\<bullet>x)\<sharp>([(z',z)]\<bullet>L)" using `x\<sharp>L` by (simp add: fresh_bij)
- ultimately
- have "M1[([(z',z)]\<bullet>x)::=([(z',z)]\<bullet>N)][([(z',z)]\<bullet>y)::=([(z',z)]\<bullet>L)]
- = M1[([(z',z)]\<bullet>y)::=([(z',z)]\<bullet>L)][([(z',z)]\<bullet>x)::=([(z',z)]\<bullet>N)[([(z',z)]\<bullet>y)::=([(z',z)]\<bullet>L)]]"
- using ih by simp
- then have "[(z',z)]\<bullet>(M1[([(z',z)]\<bullet>x)::=([(z',z)]\<bullet>N)][([(z',z)]\<bullet>y)::=([(z',z)]\<bullet>L)]
- = M1[([(z',z)]\<bullet>y)::=([(z',z)]\<bullet>L)][([(z',z)]\<bullet>x)::=([(z',z)]\<bullet>N)[([(z',z)]\<bullet>y)::=([(z',z)]\<bullet>L)]])"
- by (simp add: perm_bool)
- then have ih': "([(z',z)]\<bullet>M1)[x::=N][y::=L] = ([(z',z)]\<bullet>M1)[y::=L][x::=N[y::=L]]"
- by (simp add: eqvts perm_swap)
- show "(Lam [z].M1)[x::=N][y::=L] = (Lam [z].M1)[y::=L][x::=N[y::=L]]" (is "?LHS=?RHS")
- proof -
- have "?LHS = (Lam [z'].([(z',z)]\<bullet>M1))[x::=N][y::=L]" using eq by simp
- also have "\<dots> = Lam [z'].(([(z',z)]\<bullet>M1)[x::=N][y::=L])" using fc by (simp add: fresh_prod)
- also from ih have "\<dots> = Lam [z'].(([(z',z)]\<bullet>M1)[y::=L][x::=N[y::=L]])" sorry
- also have "\<dots> = (Lam [z'].([(z',z)]\<bullet>M1))[y::=L][x::=N[y::=L]]" using fc fc' by (simp add: fresh_prod)
- also have "\<dots> = ?RHS" using eq by simp
- finally show "?LHS = ?RHS" .
- qed
-qed (auto simp add: forget)
-text_raw {*
- \end{textblock}
- \mbox{}
-
- \only<2->{
- \begin{textblock}{11.5}(4,2.3)
- \begin{minipage}{9.3cm}
- \begin{block}{}\footnotesize
-*}
-lemma substitution_lemma\<iota>:
- assumes asm: "x \<noteq> y" "x \<sharp> L"
- shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
- using asm
-by (nominal_induct M avoiding: x y N L rule: lam.strong_induct)
- (auto simp add: forget fresh_fact)
-text_raw {*
- \end{block}
- \end{minipage}
- \end{textblock}}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Lesson Learned}
-
- \begin{textblock}{11.5}(1.2,5)
- \begin{minipage}{10.5cm}
- \begin{block}{}
- Theorem provers can keep large proofs and definitions consistent and
- make them modifiable.
- \end{block}
- \end{minipage}
- \end{textblock}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}
- \frametitle{}
-
- \begin{textblock}{11.5}(0.8,2.3)
- \begin{minipage}{11.2cm}
- In most papers/books:
- \begin{block}{}
- \color{darkgray}
- ``\ldots this necessary hygienic discipline is somewhat swept under the carpet via
- the so-called `{\bf variable convention}' \ldots
- The {\color{black}{\bf belief}} that this is {\bf sound} came from the calculus
- with nameless binders in de Bruijn''
- \end{block}\medskip
- \end{minipage}
- \end{textblock}
-
- \begin{textblock}{11.5}(0.8,10)
- \includegraphics[scale=0.25]{LambdaBook.jpg}\hspace{-3mm}\includegraphics[scale=0.3]{barendregt.jpg}
- \end{textblock}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Regular Expressions}
-
- \begin{textblock}{6}(2,4)
- \begin{tabular}{@ {}rrl}
- \bl{r} & \bl{$::=$} & \bl{$\varnothing$}\\
- & \bl{$\mid$} & \bl{[]}\\
- & \bl{$\mid$} & \bl{c}\\
- & \bl{$\mid$} & \bl{r$_1$ + r$_2$}\\
- & \bl{$\mid$} & \bl{r$_1$ $\cdot$ r$_2$}\\
- & \bl{$\mid$} & \bl{r$^*$}\\
- \end{tabular}
- \end{textblock}
-
- \begin{textblock}{6}(8,3.5)
- \includegraphics[scale=0.35]{Screen1.png}
- \end{textblock}
-
- \begin{textblock}{6}(10.2,2.8)
- \footnotesize Isabelle:
- \end{textblock}
-
- \only<2>{
- \begin{textblock}{9}(3.6,11.8)
- \bl{matches r s $\;\Longrightarrow\;$ true $\vee$ false}\\[3.5mm]
-
- \hspace{10mm}\begin{tikzpicture}
- \coordinate (m1) at (0.4,1);
- \draw (0,0.3) node (m2) {\small\color{gray}rexp};
- \path[overlay, ->, line width = 0.5mm, shorten <=-1mm, draw = gray] (m2) edge (m1);
-
- \coordinate (s1) at (0.81,1);
- \draw (1.3,0.3) node (s2) {\small\color{gray} string};
- \path[overlay, ->, line width = 0.5mm, shorten <=-1mm, draw = gray] (s2) edge (s1);
- \end{tikzpicture}
- \end{textblock}}
-
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Specification}
-
- \small
- \begin{textblock}{6}(0,3.5)
- \begin{tabular}{r@ {\hspace{0.5mm}}r@ {\hspace{1.5mm}}c@ {\hspace{1.5mm}}l}
- \multicolumn{4}{c}{rexp $\Rightarrow$ set of strings}\bigskip\\
- &\bl{\LL ($\varnothing$)} & \bl{$\dn$} & \bl{$\varnothing$}\\
- &\bl{\LL ([])} & \bl{$\dn$} & \bl{\{[]\}}\\
- &\bl{\LL (c)} & \bl{$\dn$} & \bl{\{c\}}\\
- &\bl{\LL (r$_1$ + r$_2$)} & \bl{$\dn$} & \bl{\LL (r$_1$) $\cup$ \LL (r$_2$)}\\
- \rd{$\Rightarrow$} &\bl{\LL (r$_1$ $\cdot$ r$_2$)} & \bl{$\dn$} & \bl{\LL (r$_1$) ;; \LL (r$_2$)}\\
- \rd{$\Rightarrow$} &\bl{\LL (r$^*$)} & \bl{$\dn$} & \bl{(\LL (r))$^\star$}\\
- \end{tabular}
- \end{textblock}
-
- \begin{textblock}{9}(7.3,3)
- {\mbox{}\hspace{2cm}\footnotesize Isabelle:\smallskip}
- \includegraphics[scale=0.325]{Screen3.png}
- \end{textblock}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Version 1}
- \small
- \mbox{}\\[-8mm]\mbox{}
-
- \begin{center}\def\arraystretch{1.05}
- \begin{tabular}{@ {\hspace{-5mm}}l@ {\hspace{2.5mm}}c@ {\hspace{2.5mm}}l@ {}}
- \bl{match [] []} & \bl{$=$} & \bl{true}\\
- \bl{match [] (c::s)} & \bl{$=$} & \bl{false}\\
- \bl{match ($\varnothing$::rs) s} & \bl{$=$} & \bl{false}\\
- \bl{match ([]::rs) s} & \bl{$=$} & \bl{match rs s}\\
- \bl{match (c::rs) []} & \bl{$=$} & \bl{false}\\
- \bl{match (c::rs) (d::s)} & \bl{$=$} & \bl{if c = d then match rs s else false}\\
- \bl{match (r$_1$ + r$_2$::rs) s} & \bl{$=$} & \bl{match (r$_1$::rs) s $\vee$ match (r$_2$::rs) s}\\
- \bl{match (r$_1$ $\cdot$ r$_2$::rs) s} & \bl{$=$} & \bl{match (r$_1$::r$_2$::rs) s}\\
- \bl{match (r$^*$::rs) s} & \bl{$=$} & \bl{match rs s $\vee$ match (r::r$^*$::rs) s}\\
- \end{tabular}
- \end{center}
-
- \begin{textblock}{9}(0.2,1.6)
- \hspace{10mm}\begin{tikzpicture}
- \coordinate (m1) at (0.44,-0.5);
- \draw (0,0.3) node (m2) {\small\color{gray}\mbox{}\hspace{-9mm}list of rexps};
- \path[overlay, ->, line width = 0.5mm, shorten <=-1mm, draw = gray] (m2) edge (m1);
-
- \coordinate (s1) at (0.86,-0.5);
- \draw (1.5,0.3) node (s2) {\small\color{gray} string};
- \path[overlay, ->, line width = 0.5mm, shorten <=-1mm, draw = gray] (s2) edge (s1);
- \end{tikzpicture}
- \end{textblock}
-
- \begin{textblock}{9}(2.8,11.8)
- \bl{matches$_1$ r s $\;=\;$ match [r] s}
- \end{textblock}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Testing}
-
- \small
- Every good programmer should do thourough tests:
-
- \begin{center}
- \begin{tabular}{@ {\hspace{-20mm}}lcl}
- \bl{matches$_1$ (a$\cdot$b)$^*\;$ []} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches$_1$ (a$\cdot$b)$^*\;$ ab} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches$_1$ (a$\cdot$b)$^*\;$ aba} & \bl{$\mapsto$} & \bl{false}\\
- \bl{matches$_1$ (a$\cdot$b)$^*\;$ abab} & \bl{$\mapsto$} & \bl{true}\\
- \bl{matches$_1$ (a$\cdot$b)$^*\;$ abaa} & \bl{$\mapsto$} & \bl{false}\medskip\\
- \onslide<2->{\bl{matches$_1$ x$\cdot$(0$|$1)$^*\;$ x} & \bl{$\mapsto$} & \bl{true}}\\
- \onslide<2->{\bl{matches$_1$ x$\cdot$(0$|$1)$^*\;$ x0} & \bl{$\mapsto$} & \bl{true}}\\
- \onslide<2->{\bl{matches$_1$ x$\cdot$(0$|$1)$^*\;$ x3} & \bl{$\mapsto$} & \bl{false}}
- \end{tabular}
- \end{center}
-
- \onslide<3->
- {Looks OK \ldots let's ship it to customers\hspace{5mm}
- \raisebox{-5mm}{\includegraphics[scale=0.05]{sun.png}}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Version 1}
-
- \only<1->{Several hours later\ldots}\pause
-
-
- \begin{center}
- \begin{tabular}{@ {\hspace{0mm}}lcl}
- \bl{matches$_1$ []$^*$ s} & \bl{$\mapsto$} & loops\\
- \onslide<4->{\bl{matches$_1$ ([] + \ldots)$^*$ s} & \bl{$\mapsto$} & loops\\}
- \end{tabular}
- \end{center}
-
- \small
- \onslide<3->{
- \begin{center}
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {}}
- \ldots\\
- \bl{match ([]::rs) s} & \bl{$=$} & \bl{match rs s}\\
- \ldots\\
- \bl{match (r$^*$::rs) s} & \bl{$=$} & \bl{match rs s $\vee$ match (r::r$^*$::rs) s}\\
- \end{tabular}
- \end{center}}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[c]
- \frametitle{Testing}
-
- \begin{itemize}
- \item We can only test a {\bf finite} amount of examples:\bigskip
-
- \begin{center}
- \colorbox{cream}
- {\gr{\begin{minipage}{10cm}
- ``Testing can only show the presence of errors, never their
- absence.'' (Edsger W.~Dijkstra)
- \end{minipage}}}
- \end{center}\bigskip\pause
-
- \item In a theorem prover we can establish properties that apply to
- {\bf all} input and {\bf all} output.
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Version 2}
- \mbox{}\\[-14mm]\mbox{}
-
- \small
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}ll@ {}}
- \bl{nullable ($\varnothing$)} & \bl{$=$} & \bl{false} &\\
- \bl{nullable ([])} & \bl{$=$} & \bl{true} &\\
- \bl{nullable (c)} & \bl{$=$} & \bl{false} &\\
- \bl{nullable (r$_1$ + r$_2$)} & \bl{$=$} & \bl{nullable r$_1$ $\vee$ nullable r$_2$} & \\
- \bl{nullable (r$_1$ $\cdot$ r$_2$)} & \bl{$=$} & \bl{nullable r$_1$ $\wedge$ nullable r$_2$} & \\
- \bl{nullable (r$^*$)} & \bl{$=$} & \bl{true} & \\
- \end{tabular}\medskip
-
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-10mm}}l@ {}}
- \bl{der c ($\varnothing$)} & \bl{$=$} & \bl{$\varnothing$} & \\
- \bl{der c ([])} & \bl{$=$} & \bl{$\varnothing$} & \\
- \bl{der c (d)} & \bl{$=$} & \bl{if c = d then [] else $\varnothing$} & \\
- \bl{der c (r$_1$ + r$_2$)} & \bl{$=$} & \bl{(der c r$_1$) + (der c r$_2$)} & \\
- \bl{der c (r$_1$ $\cdot$ r$_2$)} & \bl{$=$} & \bl{((der c r$_1$) $\cdot$ r$_2$)} & \\
- & & \bl{\;\;\;\;+ (if nullable r$_1$ then der c r$_2$ else $\varnothing$)}\\
- \bl{der c (r$^*$)} & \bl{$=$} & \bl{(der c r) $\cdot$ r$^*$} &\smallskip\\
-
- \bl{derivative r []} & \bl{$=$} & \bl{r} & \\
- \bl{derivative r (c::s)} & \bl{$=$} & \bl{derivative (der c r) s} & \\
- \end{tabular}\medskip
-
- \bl{matches$_2$ r s $=$ nullable (derivative r s)}
-
- \begin{textblock}{6}(9.5,0.9)
- \begin{flushright}
- \color{gray}``if r matches []''
- \end{flushright}
- \end{textblock}
-
- \begin{textblock}{6}(9.5,6.18)
- \begin{flushright}
- \color{gray}``derivative w.r.t.~a char''
- \end{flushright}
- \end{textblock}
-
- \begin{textblock}{6}(9.5,12.1)
- \begin{flushright}
- \color{gray}``deriv.~w.r.t.~a string''
- \end{flushright}
- \end{textblock}
-
- \begin{textblock}{6}(9.5,13.98)
- \begin{flushright}
- \color{gray}``main''
- \end{flushright}
- \end{textblock}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \frametitle{Is the Matcher Error-Free?}
-
- We expect that
-
- \begin{center}
- \begin{tabular}{lcl}
- \bl{matches$_2$ r s = true} & \only<1>{\rd{$\Longrightarrow\,\,$}}\only<2>{\rd{$\Longleftarrow\,\,$}}%
- \only<3->{\rd{$\Longleftrightarrow$}} & \bl{s $\in$ \LL(r)}\\
- \bl{matches$_2$ r s = false} & \only<1>{\rd{$\Longrightarrow\,\,$}}\only<2>{\rd{$\Longleftarrow\,\,$}}%
- \only<3->{\rd{$\Longleftrightarrow$}} & \bl{s $\notin$ \LL(r)}\\
- \end{tabular}
- \end{center}
- \pause\pause\bigskip
- By \alert<4->{induction}, we can {\bf prove} these properties.\bigskip
-
- \begin{tabular}{lrcl}
- Lemmas: & \bl{nullable (r)} & \bl{$\Longleftrightarrow$} & \bl{[] $\in$ \LL (r)}\\
- & \bl{s $\in$ \LL (der c r)} & \bl{$\Longleftrightarrow$} & \bl{(c::s) $\in$ \LL (r)}\\
- \end{tabular}
-
- \only<4->{
- \begin{textblock}{3}(0.9,4.5)
- \rd{\huge$\forall$\large{}r s.}
- \end{textblock}}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[c]
- \frametitle{
- \begin{tabular}{c}
- \mbox{}\\[23mm]
- \LARGE Demo
- \end{tabular}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
-
- \mbox{}\\[-2mm]
-
- \small
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}ll@ {}}
- \bl{nullable (NULL)} & \bl{$=$} & \bl{false} &\\
- \bl{nullable (EMPTY)} & \bl{$=$} & \bl{true} &\\
- \bl{nullable (CHR c)} & \bl{$=$} & \bl{false} &\\
- \bl{nullable (ALT r$_1$ r$_2$)} & \bl{$=$} & \bl{(nullable r$_1$) orelse (nullable r$_2$)} & \\
- \bl{nullable (SEQ r$_1$ r$_2$)} & \bl{$=$} & \bl{(nullable r$_1$) andalso (nullable r$_2$)} & \\
- \bl{nullable (STAR r)} & \bl{$=$} & \bl{true} & \\
- \end{tabular}\medskip
-
- \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-10mm}}l@ {}}
- \bl{der c (NULL)} & \bl{$=$} & \bl{NULL} & \\
- \bl{der c (EMPTY)} & \bl{$=$} & \bl{NULL} & \\
- \bl{der c (CHR d)} & \bl{$=$} & \bl{if c=d then EMPTY else NULL} & \\
- \bl{der c (ALT r$_1$ r$_2$)} & \bl{$=$} & \bl{ALT (der c r$_1$) (der c r$_2$)} & \\
- \bl{der c (SEQ r$_1$ r$_2$)} & \bl{$=$} & \bl{ALT (SEQ (der c r$_1$) r$_2$)} & \\
- & & \bl{\phantom{ALT} (if nullable r$_1$ then der c r$_2$ else NULL)}\\
- \bl{der c (STAR r)} & \bl{$=$} & \bl{SEQ (der c r) (STAR r)} &\smallskip\\
-
- \bl{derivative r []} & \bl{$=$} & \bl{r} & \\
- \bl{derivative r (c::s)} & \bl{$=$} & \bl{derivative (der c r) s} & \\
- \end{tabular}\medskip
-
- \bl{matches r s $=$ nullable (derivative r s)}
-
- \only<2>{
- \begin{textblock}{8}(1.5,4)
- \includegraphics[scale=0.3]{approved.png}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{No Automata?}
-
- You might be wondering why I did not use any automata?
-
- \begin{itemize}
- \item {\bf Def.:} A \alert{regular language} is one where there is a DFA that
- recognises it.\bigskip\pause
- \end{itemize}
-
-
- There are many reasons why this is a good definition:\medskip
- \begin{itemize}
- \item pumping lemma
- \item closure properties of regular languages\\ (e.g.~closure under complement)
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[t]
- \frametitle{Really Bad News!}
-
- DFAs are bad news for formalisations in theorem provers. They might
- be represented as:
-
- \begin{itemize}
- \item graphs
- \item matrices
- \item partial functions
- \end{itemize}
-
- All constructions are messy to reason about.\bigskip\bigskip
- \pause
-
- \small
- \only<2>{
- Constable et al needed (on and off) 18 months for a 3-person team
- to formalise automata theory in Nuprl including Myhill-Nerode. There is
- only very little other formalised work on regular languages I know of
- in Coq, Isabelle and HOL.}
- \only<3>{Typical textbook reasoning goes like: ``\ldots if \smath{M} and \smath{N} are any two
- automata with no inaccessible states \ldots''
- }
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{}
- \large
- \begin{center}
- \begin{tabular}{p{9cm}}
- My point:\bigskip\\
-
- The theory about regular languages can be reformulated
- to be more\\ suitable for theorem proving.
- \end{tabular}
- \end{center}
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE The Myhill-Nerode Theorem}
-
- \begin{itemize}
- \item provides necessary and suf\!ficient conditions for a language
- being regular (pumping lemma only necessary)\medskip
-
- \item will help with closure properties of regular languages\bigskip\pause
-
- \item key is the equivalence relation:\smallskip
- \begin{center}
- \smath{x \approx_{L} y \,\dn\, \forall z.\; x @ z \in L \Leftrightarrow y @ z \in L}
- \end{center}
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE The Myhill-Nerode Theorem}
-
- \mbox{}\\[5cm]
-
- \begin{itemize}
- \item \smath{\text{finite}\, (U\!N\!IV /\!/ \approx_L) \;\Leftrightarrow\; L\; \text{is regular}}
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Equivalence Classes}
-
- \begin{itemize}
- \item \smath{L = []}
- \begin{center}
- \smath{\Big\{\{[]\},\; U\!N\!IV - \{[]\}\Big\}}
- \end{center}\bigskip\bigskip
-
- \item \smath{L = [c]}
- \begin{center}
- \smath{\Big\{\{[]\},\; \{[c]\},\; U\!N\!IV - \{[], [c]\}\Big\}}
- \end{center}\bigskip\bigskip
-
- \item \smath{L = \varnothing}
- \begin{center}
- \smath{\Big\{U\!N\!IV\Big\}}
- \end{center}
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Regular Languages}
-
- \begin{itemize}
- \item \smath{L} is regular \smath{\dn} if there is an automaton \smath{M}
- such that \smath{\mathbb{L}(M) = L}\\[1.5cm]
-
- \item Myhill-Nerode:
-
- \begin{center}
- \begin{tabular}{l}
- finite $\Rightarrow$ regular\\
- \;\;\;\smath{\text{finite}\,(U\!N\!IV /\!/ \approx_L) \Rightarrow \exists r.\; L = \mathbb{L}(r)}\\[3mm]
- regular $\Rightarrow$ finite\\
- \;\;\;\smath{\text{finite}\, (U\!N\!IV /\!/ \approx_{\mathbb{L}(r)})}
- \end{tabular}
- \end{center}
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Final Equiv.~Classes}
-
- \mbox{}\\[3cm]
-
- \begin{itemize}
- \item \smath{\text{finals}\,L \dn
- \{{\lbrack\mkern-2mu\lbrack{s}\rbrack\mkern-2mu\rbrack}_\approx\;|\; s \in L\}}\\
- \medskip
-
- \item we can prove: \smath{L = \bigcup (\text{finals}\,L)}
-
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Transitions between ECs}
-
- \smath{L = \{[c]\}}
-
- \begin{tabular}{@ {\hspace{-7mm}}cc}
- \begin{tabular}{c}
- \begin{tikzpicture}[shorten >=1pt,node distance=2cm,auto, ultra thick]
- \tikzstyle{state}=[circle,thick,draw=blue!75,fill=blue!20,minimum size=0mm]
-
- %\draw[help lines] (0,0) grid (3,2);
-
- \node[state,initial] (q_0) {$R_1$};
- \node[state,accepting] (q_1) [above right of=q_0] {$R_2$};
- \node[state] (q_2) [below right of=q_0] {$R_3$};
-
- \path[->] (q_0) edge node {c} (q_1)
- edge node [swap] {$\Sigma-{c}$} (q_2)
- (q_2) edge [loop below] node {$\Sigma$} ()
- (q_1) edge node {$\Sigma$} (q_2);
- \end{tikzpicture}
- \end{tabular}
- &
- \begin{tabular}[t]{ll}
- \\[-20mm]
- \multicolumn{2}{l}{\smath{U\!N\!IV /\!/\approx_L} produces}\\[4mm]
-
- \smath{R_1}: & \smath{\{[]\}}\\
- \smath{R_2}: & \smath{\{[c]\}}\\
- \smath{R_3}: & \smath{U\!N\!IV - \{[], [c]\}}\\[6mm]
- \multicolumn{2}{l}{\onslide<2->{\smath{X \stackrel{c}{\longrightarrow} Y \dn X ;; [c] \subseteq Y}}}
- \end{tabular}
-
- \end{tabular}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Systems of Equations}
-
- Inspired by a method of Brzozowski\;'64, we can build an equational system
- characterising the equivalence classes:
-
- \begin{center}
- \begin{tabular}{@ {\hspace{-20mm}}c}
- \\[-13mm]
- \begin{tikzpicture}[shorten >=1pt,node distance=2cm,auto, ultra thick]
- \tikzstyle{state}=[circle,thick,draw=blue!75,fill=blue!20,minimum size=0mm]
-
- %\draw[help lines] (0,0) grid (3,2);
-
- \node[state,initial] (p_0) {$R_1$};
- \node[state,accepting] (p_1) [right of=q_0] {$R_2$};
-
- \path[->] (p_0) edge [bend left] node {a} (p_1)
- edge [loop above] node {b} ()
- (p_1) edge [loop above] node {a} ()
- edge [bend left] node {b} (p_0);
- \end{tikzpicture}\\
- \\[-13mm]
- \end{tabular}
- \end{center}
-
- \begin{center}
- \begin{tabular}{@ {\hspace{-6mm}}ll@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
- & \smath{R_1} & \smath{\equiv} & \smath{R_1;b + R_2;b \onslide<2->{\alert<2>{+ \lambda;[]}}}\\
- & \smath{R_2} & \smath{\equiv} & \smath{R_1;a + R_2;a}\medskip\\
- \onslide<3->{we can prove}
- & \onslide<3->{\smath{R_1}} & \onslide<3->{\smath{=}}
- & \onslide<3->{\smath{R_1;; \mathbb{L}(b) \,\cup\, R_2;;\mathbb{L}(b) \,\cup\, \{[]\}}}\\
- & \onslide<3->{\smath{R_2}} & \onslide<3->{\smath{=}}
- & \onslide<3->{\smath{R_1;; \mathbb{L}(a) \,\cup\, R_2;;\mathbb{L}(a)}}\\
- \end{tabular}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[t]
- \small
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}ll}
- \onslide<1->{\smath{R_1}} & \onslide<1->{\smath{=}}
- & \onslide<1->{\smath{R_1; b + R_2; b + \lambda;[]}}\\
- \onslide<1->{\smath{R_2}} & \onslide<1->{\smath{=}}
- & \onslide<1->{\smath{R_1; a + R_2; a}}\\
-
- & & & \onslide<2->{by Arden}\\
-
- \onslide<2->{\smath{R_1}} & \onslide<2->{\smath{=}}
- & \onslide<2->{\smath{R_1; b + R_2; b + \lambda;[]}}\\
- \onslide<2->{\smath{R_2}} & \onslide<2->{\smath{=}}
- & \only<2>{\smath{R_1; a + R_2; a}}%
- \only<3->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<4->{by Arden}\\
-
- \onslide<4->{\smath{R_1}} & \onslide<4->{\smath{=}}
- & \onslide<4->{\smath{R_2; b \cdot b^\star+ \lambda;b^\star}}\\
- \onslide<4->{\smath{R_2}} & \onslide<4->{\smath{=}}
- & \onslide<4->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<5->{by substitution}\\
-
- \onslide<5->{\smath{R_1}} & \onslide<5->{\smath{=}}
- & \onslide<5->{\smath{R_1; a\cdot a^\star \cdot b \cdot b^\star+ \lambda;b^\star}}\\
- \onslide<5->{\smath{R_2}} & \onslide<5->{\smath{=}}
- & \onslide<5->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<6->{by Arden}\\
-
- \onslide<6->{\smath{R_1}} & \onslide<6->{\smath{=}}
- & \onslide<6->{\smath{\lambda;b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star}}\\
- \onslide<6->{\smath{R_2}} & \onslide<6->{\smath{=}}
- & \onslide<6->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<7->{by substitution}\\
-
- \onslide<7->{\smath{R_1}} & \onslide<7->{\smath{=}}
- & \onslide<7->{\smath{\lambda;b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star}}\\
- \onslide<7->{\smath{R_2}} & \onslide<7->{\smath{=}}
- & \onslide<7->{\smath{\lambda; b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star
- \cdot a\cdot a^\star}}\\
- \end{tabular}
- \end{center}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE A Variant of Arden's Lemma}
-
- {\bf Arden's Lemma:}\smallskip
-
- If \smath{[] \not\in A} then
- \begin{center}
- \smath{X = X; A + \text{something}}
- \end{center}
- has the (unique) solution
- \begin{center}
- \smath{X = \text{something} ; A^\star}
- \end{center}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1->[t]
- \small
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}ll}
- \onslide<1->{\smath{R_1}} & \onslide<1->{\smath{=}}
- & \onslide<1->{\smath{R_1; b + R_2; b + \lambda;[]}}\\
- \onslide<1->{\smath{R_2}} & \onslide<1->{\smath{=}}
- & \onslide<1->{\smath{R_1; a + R_2; a}}\\
-
- & & & \onslide<2->{by Arden}\\
-
- \onslide<2->{\smath{R_1}} & \onslide<2->{\smath{=}}
- & \onslide<2->{\smath{R_1; b + R_2; b + \lambda;[]}}\\
- \onslide<2->{\smath{R_2}} & \onslide<2->{\smath{=}}
- & \only<2>{\smath{R_1; a + R_2; a}}%
- \only<3->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<4->{by Arden}\\
-
- \onslide<4->{\smath{R_1}} & \onslide<4->{\smath{=}}
- & \onslide<4->{\smath{R_2; b \cdot b^\star+ \lambda;b^\star}}\\
- \onslide<4->{\smath{R_2}} & \onslide<4->{\smath{=}}
- & \onslide<4->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<5->{by substitution}\\
-
- \onslide<5->{\smath{R_1}} & \onslide<5->{\smath{=}}
- & \onslide<5->{\smath{R_1; a\cdot a^\star \cdot b \cdot b^\star+ \lambda;b^\star}}\\
- \onslide<5->{\smath{R_2}} & \onslide<5->{\smath{=}}
- & \onslide<5->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<6->{by Arden}\\
-
- \onslide<6->{\smath{R_1}} & \onslide<6->{\smath{=}}
- & \onslide<6->{\smath{\lambda;b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star}}\\
- \onslide<6->{\smath{R_2}} & \onslide<6->{\smath{=}}
- & \onslide<6->{\smath{R_1; a\cdot a^\star}}\\
-
- & & & \onslide<7->{by substitution}\\
-
- \onslide<7->{\smath{R_1}} & \onslide<7->{\smath{=}}
- & \onslide<7->{\smath{\lambda;b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star}}\\
- \onslide<7->{\smath{R_2}} & \onslide<7->{\smath{=}}
- & \onslide<7->{\smath{\lambda; b^\star\cdot (a\cdot a^\star \cdot b \cdot b^\star)^\star
- \cdot a\cdot a^\star}}\\
- \end{tabular}
- \end{center}
-
- \only<8->{
- \begin{textblock}{6}(2.5,4)
- \begin{block}{}
- \begin{minipage}{8cm}\raggedright
-
- \begin{tikzpicture}[shorten >=1pt,node distance=2cm,auto, ultra thick, inner sep=1mm]
- \tikzstyle{state}=[circle,thick,draw=blue!75,fill=blue!20,minimum size=0mm]
-
- %\draw[help lines] (0,0) grid (3,2);
-
- \node[state,initial] (p_0) {$R_1$};
- \node[state,accepting] (p_1) [right of=q_0] {$R_2$};
-
- \path[->] (p_0) edge [bend left] node {a} (p_1)
- edge [loop above] node {b} ()
- (p_1) edge [loop above] node {a} ()
- edge [bend left] node {b} (p_0);
- \end{tikzpicture}
-
- \end{minipage}
- \end{block}
- \end{textblock}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE The Equ's Solving Algorithm}
-
- \begin{itemize}
- \item The algorithm must terminate: Arden makes one equation smaller;
- substitution deletes one variable from the right-hand sides.\bigskip
-
- \item We need to maintain the invariant that Arden is applicable
- (if \smath{[] \not\in A} then \ldots):\medskip
-
- \begin{center}\small
- \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}ll}
- \smath{R_1} & \smath{=} & \smath{R_1; b + R_2; b + \lambda;[]}\\
- \smath{R_2} & \smath{=} & \smath{R_1; a + R_2; a}\\
-
- & & & by Arden\\
-
- \smath{R_1} & \smath{=} & \smath{R_1; b + R_2; b + \lambda;[]}\\
- \smath{R_2} & \smath{=} & \smath{R_1; a\cdot a^\star}\\
- \end{tabular}
- \end{center}
-
- \end{itemize}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Other Direction}
-
- One has to prove
-
- \begin{center}
- \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r)})}
- \end{center}
-
- by induction on \smath{r}. Not trivial, but after a bit
- of thinking, one can prove that if
-
- \begin{center}
- \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r_1)})}\hspace{5mm}
- \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r_2)})}
- \end{center}
-
- then
-
- \begin{center}
- \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r_1) \,\cup\, \mathbb{L}(r_2)})}
- \end{center}
-
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE What Have We Achieved?}
-
- \begin{itemize}
- \item \smath{\text{finite}\, (U\!N\!IV /\!/ \approx_L) \;\Leftrightarrow\; L\; \text{is regular}}
- \bigskip\pause
- \item regular languages are closed under complementation; this is now easy\medskip
- \begin{center}
- \smath{U\!N\!IV /\!/ \approx_L \;\;=\;\; U\!N\!IV /\!/ \approx_{-L}}
- \end{center}
- \end{itemize}
-
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Examples}
-
- \begin{itemize}
- \item \smath{L \equiv \Sigma^\star 0 \Sigma} is regular
- \begin{quote}\small
- \begin{tabular}{lcl}
- \smath{A_1} & \smath{=} & \smath{\Sigma^\star 00}\\
- \smath{A_2} & \smath{=} & \smath{\Sigma^\star 01}\\
- \smath{A_3} & \smath{=} & \smath{\Sigma^\star 10 \cup \{0\}}\\
- \smath{A_4} & \smath{=} & \smath{\Sigma^\star 11 \cup \{1\} \cup \{[]\}}\\
- \end{tabular}
- \end{quote}
-
- \item \smath{L \equiv \{ 0^n 1^n \,|\, n \ge 0\}} is not regular
- \begin{quote}\small
- \begin{tabular}{lcl}
- \smath{B_0} & \smath{=} & \smath{\{0^n 1^n \,|\, n \ge 0\}}\\
- \smath{B_1} & \smath{=} & \smath{\{0^n 1^{(n-1)} \,|\, n \ge 1\}}\\
- \smath{B_2} & \smath{=} & \smath{\{0^n 1^{(n-2)} \,|\, n \ge 2\}}\\
- \smath{B_3} & \smath{=} & \smath{\{0^n 1^{(n-3)} \,|\, n \ge 3\}}\\
- & \smath{\vdots} &\\
- \end{tabular}
- \end{quote}
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE What We Have Not Achieved}
-
- \begin{itemize}
- \item regular expressions are not good if you look for a minimal
- one for a language (DFAs have this notion)\pause\bigskip
-
- \item Is there anything to be said about context free languages:\medskip
-
- \begin{quote}
- A context free language is where every string can be recognised by
- a pushdown automaton.\bigskip
- \end{quote}
- \end{itemize}
-
- \textcolor{gray}{\footnotesize Yes. Derivatives also work for c-f grammars. Ongoing work.}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}[c]
- \frametitle{\LARGE Conclusion}
-
- \begin{itemize}
- \item We formalised the Myhill-Nerode theorem based on
- regular expressions only (DFAs are difficult to deal with in a theorem prover).\smallskip
-
- \item Seems to be a common theme: algorithms need to be reformulated
- to better suit formal treatment.\smallskip
-
- \item The most interesting aspect is that we are able to
- implement the matcher directly inside the theorem prover
- (ongoing work).\smallskip
-
- \item Parsing is a vast field which seem to offer new results.
- \end{itemize}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-text_raw {*
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
- \mode<presentation>{
- \begin{frame}<1>[b]
- \frametitle{
- \begin{tabular}{c}
- \mbox{}\\[13mm]
- \alert{\LARGE Thank you very much!}\\
- \alert{\Large Questions?}
- \end{tabular}}
-
- \end{frame}}
- %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-*}
-
-
-
-(*<*)
-end
-(*>*)
\ No newline at end of file
Binary file Slides/document/Screen1.png has changed
Binary file Slides/document/Screen2.png has changed
Binary file Slides/document/Screen3.png has changed
Binary file Slides/document/appel.jpg has changed
Binary file Slides/document/approved.png has changed
--- a/Slides/document/beamerthemeplaincudark.sty Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,171 +0,0 @@
-\ProvidesPackage{beamerthemeplaincu}[2003/11/07 ver 0.93]
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-
-% Copyright 2003 by Till Tantau <tantau@cs.tu-berlin.de>.
-%
-% This program can be redistributed and/or modified under the terms
-% of the LaTeX Project Public License Distributed from CTAN
-% archives in directory macros/latex/base/lppl.txt.
-
-\newcommand{\slidecaption}{}
-
-\mode<presentation>
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% comic fonts fonts
-\DeclareFontFamily{T1}{comic}{}%
-\DeclareFontShape{T1}{comic}{m}{n}{<->s*[.9]comic8t}{}%
-\DeclareFontShape{T1}{comic}{m}{it}{<->s*[.9]comic8t}{}%
-\DeclareFontShape{T1}{comic}{m}{sc}{<->s*[.9]comic8t}{}%
-\DeclareFontShape{T1}{comic}{b}{n}{<->s*[.9]comicbd8t}{}%
-\DeclareFontShape{T1}{comic}{b}{it}{<->s*[.9]comicbd8t}{}%
-\DeclareFontShape{T1}{comic}{m}{sl}{<->ssub * comic/m/it}{}%
-\DeclareFontShape{T1}{comic}{b}{sc}{<->sub * comic/m/sc}{}%
-\DeclareFontShape{T1}{comic}{b}{sl}{<->ssub * comic/b/it}{}%
-\DeclareFontShape{T1}{comic}{bx}{n}{<->ssub * comic/b/n}{}%
-\DeclareFontShape{T1}{comic}{bx}{it}{<->ssub * comic/b/it}{}%
-\DeclareFontShape{T1}{comic}{bx}{sc}{<->sub * comic/m/sc}{}%
-\DeclareFontShape{T1}{comic}{bx}{sl}{<->ssub * comic/b/it}{}%
-%
-\renewcommand{\rmdefault}{comic}%
-\renewcommand{\sfdefault}{comic}%
-\renewcommand{\mathfamilydefault}{cmr}% mathfont should be still the old one
-%
-\DeclareMathVersion{bold}% mathfont needs to be bold
-\DeclareSymbolFont{operators}{OT1}{cmr}{b}{n}%
-\SetSymbolFont{operators}{bold}{OT1}{cmr}{b}{n}%
-\DeclareSymbolFont{letters}{OML}{cmm}{b}{it}%
-\SetSymbolFont{letters}{bold}{OML}{cmm}{b}{it}%
-\DeclareSymbolFont{symbols}{OMS}{cmsy}{b}{n}%
-\SetSymbolFont{symbols}{bold}{OMS}{cmsy}{b}{n}%
-\DeclareSymbolFont{largesymbols}{OMX}{cmex}{b}{n}%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% colours
-\setbeamercolor{normal text}{fg=white}
-%%%\setbeamercolor{background canvas}{bg=black}
-\setbeamertemplate{background canvas}[vertical shading]
- [bottom=white!5!black,top=white!5!black,middle=white!25!black]
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Frametitles
-\setbeamerfont{frametitle}{size={\LARGE}}
-\setbeamerfont{frametitle}{family={\usefont{T1}{ptm}{b}{n}}}
-\setbeamercolor{frametitle}{fg=white!80!black}
-
-\setbeamertemplate{frametitle}{%
-\vskip 2mm % distance from the top margin
-\hskip -3mm % distance from left margin
-\vbox{%
-\begin{minipage}{1.05\textwidth}%
-\centering%
-\insertframetitle%
-\end{minipage}}%
-}
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Foot
-%
-\setbeamertemplate{navigation symbols}{}
-\usefoottemplate{%
-\vbox{%
- \tinyline{%
- \tiny\hfill\textcolor{gray!50}{\slidecaption{} --
- p.~\insertframenumber/\inserttotalframenumber}}}%
-}
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%%\beamertemplateballitem
-\setbeamertemplate{itemize items}[ball]
-\setbeamercolor{item}{fg=red}
-
-\pgfdeclareradialshading[bg,parent.bg]{bigsphere}{\pgfpoint{-0.1849315ex}{.2260273ex}}%
-{%
- color(0cm)=(bg!15);
- color(0.1643835ex)=(bg!75);
- color(0.3287671ex)=(bg!70!black);
- color(0.4520547ex)=(bg!50!black)}
-
-\pgfdeclareradialshading{bigspherered}{\pgfpoint{-0.1849315ex}{.2260273ex}}%
-{%
- color(0cm)=(red!15);
- color(0.1643835ex)=(red!75);
- color(0.3287671ex)=(red!70!black);
- color(0.4520547ex)=(red!50!black)}
-
-\pgfdeclareradialshading[bg,parent.bg]{smallsphere}{\pgfpoint{-0.1479452ex}{0.18287671ex}}%
-{%
- color(0cm)=(bg!15);
- color(0.1315068ex)=(bg!75);
- color(0.2630136ex)=(bg!70!black);
- color(0.36164383ex)=(bg!50!black)}
-
-\setlength\leftmargini{2mm}
-\setlength\leftmarginii{0.6cm}
-\setlength\leftmarginiii{1.5cm}
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% blocks
-%\definecolor{cream}{rgb}{1,1,.65}
-\definecolor{cream}{rgb}{1,1,.8}
-
-\setbeamerfont{block title}{size=\normalsize}
-\setbeamercolor{block title}{fg=black,bg=cream}
-\setbeamercolor{block body}{fg=black,bg=cream}
-\setbeamertemplate{blocks}[rounded][shadow=true]
-\setbeamercolor{boxcolor}{fg=red,bg=cream}
-
-
-%
-% Boxed environment with semi-transparent shadow.
-%
-\newlength{\boxw}
-\newlength{\boxh}
-\newlength{\shadowsize}
-\newlength{\boxroundness}
-\newlength{\tmpa}
-\newsavebox{\shadowblockbox}
-
-\setlength{\shadowsize}{6pt}
-\setlength{\boxroundness}{3pt}
-
-\newenvironment{shadowblock}[1]%
-{\begin{lrbox}{\shadowblockbox}\begin{minipage}{#1}}%
-{\end{minipage}\end{lrbox}%
-\settowidth{\boxw}{\usebox{\shadowblockbox}}%
-\settodepth{\tmpa}{\usebox{\shadowblockbox}}%
-\settoheight{\boxh}{\usebox{\shadowblockbox}}%
-\addtolength{\boxh}{\tmpa}%
-\begin{tikzpicture}
- \addtolength{\boxw}{\boxroundness * 2}
- \addtolength{\boxh}{\boxroundness * 2}
-
- \foreach \x in {0,.05,...,1}
- {
- \setlength{\tmpa}{\shadowsize * \real{\x}}
- \fill[xshift=\shadowsize - 1pt,yshift=-\shadowsize + 1pt,
- black,opacity=.04,rounded corners=\boxroundness]
- (\tmpa, \tmpa) rectangle +(\boxw - \tmpa - \tmpa, \boxh - \tmpa - \tmpa);
- }
-
- \filldraw[fill=cream, draw=black, rounded corners=\boxroundness]
- (0, 0) rectangle (\boxw, \boxh);
- \draw node[xshift=\boxroundness,yshift=\boxroundness,inner sep=0pt,
- outer sep=0pt,anchor=south west] (0,0) {\usebox{\shadowblockbox}};
-\end{tikzpicture}}
-
-
-\mode
-<all>
-
-
-
-
-
-
--- a/Slides/document/beamerthemeplainculight.sty Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,101 +0,0 @@
-%%\Providespackage{beamerthemeplainculight}[2003/11/07 ver 0.93]
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-
-% Copyright 2003 by Till Tantau <tantau@cs.tu-berlin.de>.
-%
-% This program can be redistributed and/or modified under the terms
-% of the LaTeX Project Public License Distributed from CTAN
-% archives in directory macros/latex/base/lppl.txt.
-
-\newcommand{\slidecaption}{}
-
-\mode<presentation>
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% comic fonts fonts
-\DeclareFontFamily{T1}{comic}{}%
-\DeclareFontShape{T1}{comic}{m}{n}{<->s*[.9]comic8t}{}%
-\DeclareFontShape{T1}{comic}{m}{it}{<->s*[.9]comic8t}{}%
-\DeclareFontShape{T1}{comic}{m}{sc}{<->s*[.9]comic8t}{}%
-\DeclareFontShape{T1}{comic}{b}{n}{<->s*[.9]comicbd8t}{}%
-\DeclareFontShape{T1}{comic}{b}{it}{<->s*[.9]comicbd8t}{}%
-\DeclareFontShape{T1}{comic}{m}{sl}{<->ssub * comic/m/it}{}%
-\DeclareFontShape{T1}{comic}{b}{sc}{<->sub * comic/m/sc}{}%
-\DeclareFontShape{T1}{comic}{b}{sl}{<->ssub * comic/b/it}{}%
-\DeclareFontShape{T1}{comic}{bx}{n}{<->ssub * comic/b/n}{}%
-\DeclareFontShape{T1}{comic}{bx}{it}{<->ssub * comic/b/it}{}%
-\DeclareFontShape{T1}{comic}{bx}{sc}{<->sub * comic/m/sc}{}%
-\DeclareFontShape{T1}{comic}{bx}{sl}{<->ssub * comic/b/it}{}%
-%
-\renewcommand{\rmdefault}{comic}%
-\renewcommand{\sfdefault}{comic}%
-\renewcommand{\mathfamilydefault}{cmr}% mathfont should be still the old one
-%
-\DeclareMathVersion{bold}% mathfont needs to be bold
-\DeclareSymbolFont{operators}{OT1}{cmr}{b}{n}%
-\SetSymbolFont{operators}{bold}{OT1}{cmr}{b}{n}%
-\DeclareSymbolFont{letters}{OML}{cmm}{b}{it}%
-\SetSymbolFont{letters}{bold}{OML}{cmm}{b}{it}%
-\DeclareSymbolFont{symbols}{OMS}{cmsy}{b}{n}%
-\SetSymbolFont{symbols}{bold}{OMS}{cmsy}{b}{n}%
-\DeclareSymbolFont{largesymbols}{OMX}{cmex}{b}{n}%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Frametitles
-\setbeamerfont{frametitle}{size={\LARGE}}
-\setbeamerfont{frametitle}{family={\usefont{T1}{ptm}{b}{n}}}
-\setbeamercolor{frametitle}{fg=gray,bg=white}
-
-\setbeamertemplate{frametitle}{%
-\vskip 2mm % distance from the top margin
-\hskip -3mm % distance from left margin
-\vbox{%
-\begin{minipage}{1.05\textwidth}%
-\centering%
-\insertframetitle%
-\end{minipage}}%
-}
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Foot
-%
-\setbeamertemplate{navigation symbols}{}
-\usefoottemplate{%
-\vbox{%
- \tinyline{%
- \tiny\hfill\textcolor{gray!50}{\slidecaption{} --
- p.~\insertframenumber/\inserttotalframenumber}}}%
-}
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-\beamertemplateballitem
-\setlength\leftmargini{2mm}
-\setlength\leftmarginii{0.6cm}
-\setlength\leftmarginiii{1.5cm}
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% blocks
-%\definecolor{cream}{rgb}{1,1,.65}
-\definecolor{cream}{rgb}{1,1,.8}
-\setbeamerfont{block title}{size=\normalsize}
-\setbeamercolor{block title}{fg=black,bg=cream}
-\setbeamercolor{block body}{fg=black,bg=cream}
-
-\setbeamertemplate{blocks}[rounded][shadow=true]
-
-\setbeamercolor{boxcolor}{fg=black,bg=cream}
-
-\mode
-<all>
-
-
-
-
-
-
Binary file Slides/document/evil.png has changed
Binary file Slides/document/harper.jpg has changed
Binary file Slides/document/isabelle1.png has changed
Binary file Slides/document/ken-thompson.jpg has changed
Binary file Slides/document/notok.png has changed
Binary file Slides/document/ok.png has changed
Binary file Slides/document/pfenning.jpg has changed
Binary file Slides/document/robin-milner.jpg has changed
--- a/Slides/document/root.beamer.tex Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,12 +0,0 @@
-\documentclass[14pt,t]{beamer}
-%%%\usepackage{pstricks}
-
-\input{root.tex}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% TeX-command-default: "Slides"
-%%% TeX-view-style: (("." "kghostview --landscape --scale 0.45 --geometry 605x505 %f"))
-%%% End:
-
--- a/Slides/document/root.notes.tex Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,18 +0,0 @@
-\documentclass[11pt]{article}
-%%\usepackage{pstricks}
-\usepackage{dina4}
-\usepackage{beamerarticle}
-\usepackage{times}
-\usepackage{hyperref}
-\usepackage{pgf}
-\usepackage{amssymb}
-\setjobnamebeamerversion{root.beamer}
-\input{root.tex}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% TeX-command-default: "Slides"
-%%% TeX-view-style: (("." "kghostview --landscape --scale 0.45 --geometry 605x505 %f"))
-%%% End:
-
--- a/Slides/document/root.tex Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,167 +0,0 @@
-\usepackage{beamerthemeplainculight}
-\usepackage[T1]{fontenc}
-\usepackage{proof}
-\usepackage{german}
-\usepackage[latin1]{inputenc}
-\usepackage{isabelle}
-\usepackage{isabellesym}
-%%\usepackage{mathpartir}
-\usepackage[absolute,overlay]{textpos}
-\usepackage{proof}
-\usepackage{ifthen}
-%%\usepackage{animate}
-\usepackage{tikz}
-\usepackage{pgf}
-\usepackage{calc}
-\usepackage{ulem}
-%%%\newcommand{\uline}[1]{}
-\usetikzlibrary{arrows}
-\usetikzlibrary{automata}
-\usetikzlibrary{shapes}
-\usetikzlibrary{shadows}
-\usetikzlibrary{positioning}
-%%%\usetikzlibrary{mindmap}
-
-\usepackage{graphicx}
-\usepackage{xcolor}
-
-% general math stuff
-\newcommand{\dn}{\stackrel{\mbox{\scriptsize def}}{=}}% for definitions
-\newcommand{\dnn}{\stackrel{\mbox{\Large def}}{=}}
-\renewcommand{\emptyset}{\varnothing}% nice round empty set
-\renewcommand{\Gamma}{\varGamma}
-\DeclareRobustCommand{\flqq}{\mbox{\guillemotleft}}
-\DeclareRobustCommand{\frqq}{\mbox{\guillemotright}}
-\newcommand{\smath}[1]{\textcolor{blue}{\ensuremath{#1}}}
-\newcommand{\fresh}{\mathrel{\#}}
-\newcommand{\act}{{\raisebox{-0.5mm}{\Large$\boldsymbol{\cdot}$}}}% swapping action
-\newcommand{\swap}[2]{(#1\,#2)}% swapping operation
-
-
-
-% Isabelle configuration
-%%\urlstyle{rm}
-\isabellestyle{rm}
-\renewcommand{\isastyle}{\rm}%
-\renewcommand{\isastyleminor}{\rm}%
-\renewcommand{\isastylescript}{\footnotesize\rm\slshape}%
-\renewcommand{\isatagproof}{}
-\renewcommand{\endisatagproof}{}
-\renewcommand{\isamarkupcmt}[1]{#1}
-
-% Isabelle characters
-\renewcommand{\isacharunderscore}{\_}
-\renewcommand{\isacharbar}{\isamath{\mid}}
-\renewcommand{\isasymiota}{}
-\renewcommand{\isacharbraceleft}{\{}
-\renewcommand{\isacharbraceright}{\}}
-\renewcommand{\isacharless}{$\langle$}
-\renewcommand{\isachargreater}{$\rangle$}
-\renewcommand{\isasymsharp}{\isamath{\#}}
-\renewcommand{\isasymdots}{\isamath{...}}
-\renewcommand{\isasymbullet}{\act}
-\renewcommand{\isasymequiv}{$\dn$}
-
-% mathpatir
-%%\mprset{sep=1em}
-
-
-% beamer stuff
-\renewcommand{\slidecaption}{Salvador, 26.~August 2008}
-
-
-% colours for Isar Code (in article mode everything is black and white)
-\mode<presentation>{
-\definecolor{isacol:brown}{rgb}{.823,.411,.117}
-\definecolor{isacol:lightblue}{rgb}{.274,.509,.705}
-\definecolor{isacol:green}{rgb}{0,.51,0.14}
-\definecolor{isacol:red}{rgb}{.803,0,0}
-\definecolor{isacol:blue}{rgb}{0,0,.803}
-\definecolor{isacol:darkred}{rgb}{.545,0,0}
-\definecolor{isacol:black}{rgb}{0,0,0}}
-\mode<article>{
-\definecolor{isacol:brown}{rgb}{0,0,0}
-\definecolor{isacol:lightblue}{rgb}{0,0,0}
-\definecolor{isacol:green}{rgb}{0,0,0}
-\definecolor{isacol:red}{rgb}{0,0,0}
-\definecolor{isacol:blue}{rgb}{0,0,0}
-\definecolor{isacol:darkred}{rgb}{0,0,0}
-\definecolor{isacol:black}{rgb}{0,0,0}
-}
-
-
-\newcommand{\strong}[1]{{\bfseries {#1}}}
-\newcommand{\bluecmd}[1]{{\color{isacol:lightblue}{\strong{#1}}}}
-\newcommand{\browncmd}[1]{{\color{isacol:brown}{\strong{#1}}}}
-\newcommand{\redcmd}[1]{{\color{isacol:red}{\strong{#1}}}}
-
-\renewcommand{\isakeyword}[1]{%
-\ifthenelse{\equal{#1}{show}}{\browncmd{#1}}{%
-\ifthenelse{\equal{#1}{case}}{\browncmd{#1}}{%
-\ifthenelse{\equal{#1}{assume}}{\browncmd{#1}}{%
-\ifthenelse{\equal{#1}{obtain}}{\browncmd{#1}}{%
-\ifthenelse{\equal{#1}{fix}}{\browncmd{#1}}{%
-\ifthenelse{\equal{#1}{oops}}{\redcmd{#1}}{%
-\ifthenelse{\equal{#1}{thm}}{\redcmd{#1}}{%
-{\bluecmd{#1}}}}}}}}}}%
-
-% inner syntax colour
-\chardef\isachardoublequoteopen=`\"%
-\chardef\isachardoublequoteclose=`\"%
-\chardef\isacharbackquoteopen=`\`%
-\chardef\isacharbackquoteclose=`\`%
-
-\let\oldisachardoublequoteopen=\isachardoublequoteopen
-\let\oldisachardoublequoteclose=\isachardoublequoteclose
-\let\oldisacharbackquoteopen=\isacharbackquoteopen
-\let\oldisacharbackquoteclose=\isacharbackquoteclose
-\newenvironment{innerdouble}%
- {\oldisachardoublequoteopen \color{isacol:green}}%
- {\color{isacol:black} \oldisachardoublequoteclose}
-\newenvironment{innersingle}%
-{\oldisacharbackquoteopen\color{isacol:green}}%
-{\color{isacol:black}\oldisacharbackquoteclose}
-
-\renewcommand{\isachardoublequoteopen}{\egroup\begin{innerdouble}\bgroup}
-\renewcommand{\isachardoublequoteclose}{\egroup\end{innerdouble}\bgroup}
-\renewcommand{\isacharbackquoteopen}{\egroup\begin{innersingle}\bgroup}
-\renewcommand{\isacharbackquoteclose}{\egroup\end{innersingle}\bgroup}
-
-%% misc
-\newcommand{\gb}[1]{\textcolor{isacol:green}{#1}}
-\newcommand{\rb}[1]{\textcolor{red}{#1}}
-
-%% animations
-\newcounter{growcnt}
-\newcommand{\grow}[2]
-{\begin{tikzpicture}[baseline=(n.base)]%
- \node[scale=(0.1 *#1 + 0.001),inner sep=0pt] (n) {#2};
- \end{tikzpicture}%
-}
-
-%% isatabbing
-%\renewcommand{\isamarkupcmt}[1]%
-%{\ifthenelse{\equal{TABSET}{#1}}{\=}%
-% {\ifthenelse{\equal{TAB}{#1}}{\>}%
-% {\ifthenelse{\equal{NEWLINE}{#1}}{\\}%
-% {\ifthenelse{\equal{DOTS}{#1}}{\ldots}{\isastylecmt--- {#1}}}%
-% }%
-% }%
-%}%
-
-
-\newenvironment{isatabbing}%
-{\renewcommand{\isanewline}{\\}\begin{tabbing}}%
-{\end{tabbing}}
-
-\begin{document}
-\input{session}
-\end{document}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% TeX-command-default: "Slides"
-%%% TeX-view-style: (("." "kghostview --landscape --scale 0.45 --geometry 605x505 %f"))
-%%% End:
-
Binary file Slides/document/sun.png has changed
--- a/TODO Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,49 +0,0 @@
-Function definitions
-
-- nominal_datatype should export proofs about alpha_bn: reflexivity, induction, ...?
-
-- nominal_datatype should declare xxx.fv_bn_eqvt as [eqvt]
-
-- equations that talk about True/False, for example
- | "simp p (App t1 t2) = (if True then (App (simp p t1) (simp p t2)) else t1)"
- do not work (since 'function' does not work)
-
-- equivariance for functions can be derived automatically (knowing termination and
- equivariance of the graph).
-
-- for multiple simultanously defined functions, compatibility cases should fold the
- definition and derive eqvt_at for the folded one
-
-- equivariance of the graph for multiple functions does not use equivariance
- infrastructure
-
-Nominal_Datatype Parser should check that:
-
-- types of bindings match types of binding functions
-- fsets are not bound in lst bindings
-- bound arguments are not datatypes
-- binder is referred to by name and not by type
-
-Smaller things:
-
-- maybe <t1_t2>.perm_simps should be called permute_<type>.simps;
- that would conform with the terminology in Nominal2
-
-- nominal_induct does not work without an induction rule
-
-- nominal_induct throws strange "THM" exceptions if not enough
- variables are given
-
-- nominal_induct cannot avoid a term (for example function applied
- to an argument).
-
-- .induct and .exhaust could be the default methods for induction/cases
- on nominal datatypes
-
-Other:
-
-- nested recursion, like types "trm list" in a constructor
-
-- store information about defined nominal datatypes, so that
- it can be used to define new types that depend on these
-
--- a/Tutorial/Lambda.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,173 +0,0 @@
-theory Lambda
-imports "../Nominal/Nominal2"
-begin
-
-section {* Definitions for Lambda Terms *}
-
-
-text {* type of variables *}
-
-atom_decl name
-
-
-subsection {* Alpha-Equated Lambda Terms *}
-
-nominal_datatype lam =
- Var "name"
-| App "lam" "lam"
-| Lam x::"name" l::"lam" bind x in l ("Lam [_]. _" [100, 100] 100)
-
-
-text {* some automatically derived theorems *}
-
-thm lam.distinct
-thm lam.eq_iff
-thm lam.fresh
-thm lam.size
-thm lam.exhaust
-thm lam.strong_exhaust
-thm lam.induct
-thm lam.strong_induct
-
-
-subsection {* Height Function *}
-
-nominal_primrec
- height :: "lam \<Rightarrow> int"
-where
- "height (Var x) = 1"
-| "height (App t1 t2) = max (height t1) (height t2) + 1"
-| "height (Lam [x].t) = height t + 1"
-apply(subgoal_tac "\<And>p x r. height_graph x r \<Longrightarrow> height_graph (p \<bullet> x) (p \<bullet> r)")
-unfolding eqvt_def
-apply(rule allI)
-apply(simp add: permute_fun_def)
-apply(rule ext)
-apply(rule ext)
-apply(simp add: permute_bool_def)
-apply(rule iffI)
-apply(drule_tac x="p" in meta_spec)
-apply(drule_tac x="- p \<bullet> x" in meta_spec)
-apply(drule_tac x="- p \<bullet> xa" in meta_spec)
-apply(simp)
-apply(drule_tac x="-p" in meta_spec)
-apply(drule_tac x="x" in meta_spec)
-apply(drule_tac x="xa" in meta_spec)
-apply(simp)
-apply(erule height_graph.induct)
-apply(perm_simp)
-apply(rule height_graph.intros)
-apply(perm_simp)
-apply(rule height_graph.intros)
-apply(assumption)
-apply(assumption)
-apply(perm_simp)
-apply(rule height_graph.intros)
-apply(assumption)
-apply(rule_tac y="x" in lam.exhaust)
-apply(auto simp add: lam.distinct lam.eq_iff)
-apply(simp add: Abs_eq_iff alphas)
-apply(clarify)
-apply(subst (4) supp_perm_eq[where p="p", symmetric])
-apply(simp add: pure_supp fresh_star_def)
-apply(simp add: eqvt_at_def)
-done
-
-termination
- by (relation "measure size") (simp_all add: lam.size)
-
-
-subsection {* Capture-Avoiding Substitution *}
-
-nominal_primrec
- subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_ [_ ::= _]" [90,90,90] 90)
-where
- "(Var x)[y ::= s] = (if x = y then s else (Var x))"
-| "(App t1 t2)[y ::= s] = App (t1[y ::= s]) (t2[y ::= s])"
-| "atom x \<sharp> (y, s) \<Longrightarrow> (Lam [x]. t)[y ::= s] = Lam [x].(t[y ::= s])"
-apply(subgoal_tac "\<And>p x r. subst_graph x r \<Longrightarrow> subst_graph (p \<bullet> x) (p \<bullet> r)")
-unfolding eqvt_def
-apply(rule allI)
-apply(simp add: permute_fun_def)
-apply(rule ext)
-apply(rule ext)
-apply(simp add: permute_bool_def)
-apply(rule iffI)
-apply(drule_tac x="p" in meta_spec)
-apply(drule_tac x="- p \<bullet> x" in meta_spec)
-apply(drule_tac x="- p \<bullet> xa" in meta_spec)
-apply(simp)
-apply(drule_tac x="-p" in meta_spec)
-apply(drule_tac x="x" in meta_spec)
-apply(drule_tac x="xa" in meta_spec)
-apply(simp)
-apply(erule subst_graph.induct)
-apply(perm_simp)
-apply(rule subst_graph.intros)
-apply(perm_simp)
-apply(rule subst_graph.intros)
-apply(assumption)
-apply(assumption)
-apply(perm_simp)
-apply(rule subst_graph.intros)
-apply(simp add: fresh_Pair)
-apply(assumption)
-apply(auto simp add: lam.distinct lam.eq_iff)
-apply(rule_tac y="a" and c="(aa, b)" in lam.strong_exhaust)
-apply(blast)+
-apply(simp add: fresh_star_def)
-apply(subgoal_tac "atom xa \<sharp> [[atom x]]lst. t \<and> atom x \<sharp> [[atom xa]]lst. ta")
-apply(subst (asm) Abs_eq_iff2)
-apply(simp add: alphas atom_eqvt)
-apply(clarify)
-apply(rule trans)
-apply(rule_tac p="p" in supp_perm_eq[symmetric])
-apply(rule fresh_star_supp_conv)
-apply(drule fresh_star_perm_set_conv)
-apply(simp add: finite_supp)
-apply(subgoal_tac "{atom (p \<bullet> x), atom x} \<sharp>* ([[atom x]]lst. subst_sumC (t, ya, sa))")
-apply(auto simp add: fresh_star_def)[1]
-apply(simp (no_asm) add: fresh_star_def)
-apply(rule conjI)
-apply(simp (no_asm) add: Abs_fresh_iff)
-apply(clarify)
-apply(drule_tac a="atom (p \<bullet> x)" in fresh_eqvt_at)
-apply(simp add: finite_supp)
-apply(simp (no_asm_use) add: fresh_Pair)
-apply(simp add: Abs_fresh_iff)
-apply(simp)
-apply(simp add: Abs_fresh_iff)
-apply(subgoal_tac "p \<bullet> ya = ya")
-apply(subgoal_tac "p \<bullet> sa = sa")
-apply(simp add: atom_eqvt eqvt_at_def)
-apply(rule perm_supp_eq)
-apply(auto simp add: fresh_star_def fresh_Pair)[1]
-apply(rule perm_supp_eq)
-apply(auto simp add: fresh_star_def fresh_Pair)[1]
-apply(rule conjI)
-apply(simp add: Abs_fresh_iff)
-apply(drule sym)
-apply(simp add: Abs_fresh_iff)
-done
-
-termination
- by (relation "measure (\<lambda>(t, _, _). size t)")
- (simp_all add: lam.size)
-
-lemma subst_eqvt[eqvt]:
- shows "(p \<bullet> t[x ::= s]) = (p \<bullet> t)[(p \<bullet> x) ::= (p \<bullet> s)]"
-by (induct t x s rule: subst.induct) (simp_all)
-
-lemma fresh_fact:
- assumes a: "atom z \<sharp> s"
- and b: "z = y \<or> atom z \<sharp> t"
- shows "atom z \<sharp> t[y ::= s]"
-using a b
-by (nominal_induct t avoiding: z y s rule: lam.strong_induct)
- (auto simp add: lam.fresh fresh_at_base)
-
-
-end
-
-
-
--- a/Tutorial/Minimal.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,18 +0,0 @@
-theory Minimal
-imports "Nominal2"
-begin
-
-atom_decl name
-
-nominal_datatype lam =
- Var "name"
-| App "lam" "lam"
-| Lam x::"name" l::"lam" bind x in l ("Lam [_]. _" [100, 100] 100)
-
-
-
-lemma alpha_test:
- shows "Lam [x]. (Var x) = Lam [y]. (Var y)"
- by (simp add: lam.eq_iff Abs1_eq_iff lam.fresh fresh_at_base)
-
-end
\ No newline at end of file
--- a/Tutorial/Tutorial1.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,738 +0,0 @@
-header {*
-
- Nominal Isabelle Tutorial at POPL'11
- ====================================
-
- Nominal Isabelle is a definitional extension of Isabelle/HOL, which
- means it does not add any new axioms to higher-order logic. It just
- adds new definitions and an infrastructure for 'nominal resoning'.
-
-
- The jEdit Interface
- -------------------
-
- The Isabelle theorem prover comes with an interface for jEdit.
- Unlike many other theorem prover interfaces (e.g. ProofGeneral) where you
- advance a 'checked' region in a proof script, this interface immediately
- checks the whole buffer. The text you type is also immediately checked.
- Malformed text or unfinished proofs are highlighted in red with a little
- red 'stop' signal on the left-hand side. If you drag the 'red-box' cursor
- over a line, the Output window gives further feedback.
-
- Note: If a section is not parsed correctly, the work-around is to cut it
- out and paste it back into the text (cut-out: Ctrl + X; paste in: Ctrl + V;
- Cmd is Ctrl on the Mac)
-
- Nominal Isabelle and jEdit can be started by typing on the command line
-
- isabelle jedit -l HOL-Nominal2
- isabelle jedit -l HOL-Nominal2 A.thy B.thy ...
-
-
- Symbols
- -------
-
- Symbols can considerably improve the readability of your statements and proofs.
- They can be input by just typing 'name-of-symbol' where 'name-of-symbol' is the
- usual latex name of that symbol. A little window will then appear in which
- you can select the symbol. With `Escape' you can ignore any suggestion.
-
- There are some handy short-cuts for frequently used symbols.
- For example
-
- short-cut symbol
-
- => \<Rightarrow>
- ==> \<Longrightarrow>
- --> \<longrightarrow>
- ! \<forall>
- ? \<exists>
- /\ \<and>
- \/ \<or>
- ~ \<not>
- ~= \<noteq>
- : \<in>
- ~: \<notin>
-
- For nominal the following two symbols have a special meaning
-
- \<sharp> sharp (freshness)
- \<bullet> bullet (permutation application)
-*}
-
-
-theory Tutorial1
-imports Lambda
-begin
-
-section {* Theories *}
-
-text {*
- All formal developments in Isabelle are part of a theory. A theory
- needs to have a name and must import some pre-existing theory. The
- imported theory will normally be Nominal2 (which provides many useful
- concepts like set-theory, lists, nominal things etc). For the purpose
- of this tutorial we import the theory Lambda.thy, which contains
- already some useful definitions for (alpha-equated) lambda terms.
-*}
-
-
-
-section {* Types *}
-
-text {*
- Isabelle is based on simple types including some polymorphism. It also
- includes overloading, which means that sometimes explicit type annotations
- need to be given.
-
- - Base types include: nat, bool, string, lam (defined in the Lambda theory)
-
- - Type formers include: 'a list, ('a \<times> 'b), 'c set
-
- - Type variables are written like in ML with an apostrophe: 'a, 'b, \<dots>
-
- Types known to Isabelle can be queried using the command "typ".
-*}
-
-typ nat
-typ bool
-typ string
-typ lam -- {* alpha-equated lambda terms defined in Lambda.thy *}
-typ name -- {* type of (object) variables defined in Lambda.thy *}
-typ "('a \<times> 'b)" -- {* pair type *}
-typ "'c set" -- {* set type *}
-typ "'a list" -- {* list type *}
-typ "lam \<Rightarrow> nat" -- {* type of functions from lambda terms to natural numbers *}
-
-
-text {* Some malformed types - note the "stop" signal on the left margin *}
-
-(*
-typ boolean -- {* undeclared type *}
-typ set -- {* type argument missing *}
-*)
-
-
-section {* Terms *}
-
-text {*
- Every term in Isabelle needs to be well-typed. However they can have
- polymorphic type. Whether a term is accepted can be queried using
- the command "term".
-*}
-
-term c -- {* a variable of polymorphic type *}
-term "1::nat" -- {* the constant 1 in natural numbers *}
-term 1 -- {* the constant 1 with polymorphic type *}
-term "{1, 2, 3::nat}" -- {* the set containing natural numbers 1, 2 and 3 *}
-term "[1, 2, 3]" -- {* the list containing the polymorphic numbers 1, 2 and 3 *}
-term "(True, ''c'')" -- {* a pair consisting of the boolean true and the string "c" *}
-term "Suc 0" -- {* successor of 0, in other words 1::nat *}
-term "Lam [x].Var x" -- {* a lambda-term *}
-term "App t1 t2" -- {* another lambda-term *}
-term "x::name" -- {* an (object) variable of type name *}
-term "atom (x::name)" -- {* atom is an overloaded function *}
-
-text {*
- Lam [x].Var is the syntax we made up for lambda abstractions. If you
- prefer, you can have your own syntax (but \<lambda> is already taken up for
- Isabelle's functions). We also came up with the type "name" for variables.
- You can introduce your own types of object variables using the
- command atom_decl:
-*}
-
-atom_decl ident
-atom_decl ty_var
-
-text {*
- Isabelle provides some useful colour feedback about its constants (black),
- free variables (blue) and bound variables (green).
-*}
-
-term "True" -- {* a constant defined somewhere...written in black *}
-term "true" -- {* not recognised as a constant, therefore it is interpreted
- as a free variable, written in blue *}
-term "\<forall>x. P x" -- {* x is bound (green), P is free (blue) *}
-
-
-text {* Formulae in Isabelle/HOL are terms of type bool *}
-
-term "True"
-term "True \<and> False"
-term "True \<or> B"
-term "{1,2,3} = {3,2,1}"
-term "\<forall>x. P x"
-term "A \<longrightarrow> B"
-term "atom a \<sharp> t" -- {* freshness in Nominal *}
-
-text {*
- When working with Isabelle, one deals with an object logic (that is HOL) and
- Isabelle's rule framework (called Pure). Occasionally one has to pay attention
- to this fact. But for the moment we ignore this completely.
-*}
-
-term "A \<longrightarrow> B" -- {* versus *}
-term "A \<Longrightarrow> B"
-
-term "\<forall>x. P x" -- {* versus *}
-term "\<And>x. P x"
-
-
-section {* Inductive Definitions: Evaluation Relation *}
-
-text {*
- In this section we want to define inductively the evaluation
- relation and for cbv-reduction relation.
-
- Inductive definitions in Isabelle start with the keyword "inductive"
- and a predicate name. One can optionally provide a type for the predicate.
- Clauses of the inductive predicate are introduced by "where" and more than
- two clauses need to be separated by "|". One can also give a name to each
- clause and indicate that it should be added to the hints database ("[intro]").
- A typical clause has some premises and a conclusion. This is written in
- Isabelle as:
-
- "premise \<Longrightarrow> conclusion"
- "premise1 \<Longrightarrow> premise2 \<Longrightarrow> \<dots> premisen \<Longrightarrow> conclusion"
-
- There is an alternative way of writing the latter clause as
-
- "\<lbrakk>premise1; premise2; \<dots> premisen\<rbrakk> \<Longrightarrow> conclusion"
-
- If no premise is present, then one just writes
-
- "conclusion"
-
- Below we give two definitions for the transitive closure
-*}
-
-inductive
- eval :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixr "\<Down>" 60)
-where
- e_Lam[intro]: "Lam [x].t \<Down> Lam [x].t"
-| e_App[intro]: "\<lbrakk>t1 \<Down> Lam [x].t; t2 \<Down> v'; t[x::=v'] \<Down> v\<rbrakk> \<Longrightarrow> App t1 t2 \<Down> v"
-
-text {*
- Values and cbv are also defined using inductive.
-*}
-
-inductive
- val :: "lam \<Rightarrow> bool"
-where
- v_Lam[intro]: "val (Lam [x].t)"
-
-section {* Theorems *}
-
-text {*
- A central concept in Isabelle is that of theorems. Isabelle's theorem
- database can be queried using
-*}
-
-thm e_App
-thm e_Lam
-thm conjI
-thm conjunct1
-
-text {*
- Notice that theorems usually contain schematic variables (e.g. ?P, ?Q, \<dots>).
- These schematic variables can be substituted with any term (of the right type
- of course).
-
- When defining the predicates beta_star, Isabelle provides us automatically
- with the following theorems that state how they can be introduced and what
- constitutes an induction over them.
-*}
-
-thm eval.intros
-thm eval.induct
-
-
-section {* Lemmas / Theorems / Corollaries *}
-
-text {*
- Whether to use lemma, theorem or corollary makes no semantic difference
- in Isabelle.
-
- A lemma starts with "lemma" and consists of a statement ("shows \<dots>") and
- optionally a lemma name, some type-information for variables ("fixes \<dots>")
- and some assumptions ("assumes \<dots>").
-
- Lemmas also need to have a proof, but ignore this 'detail' for the moment.
-
- Examples are
-*}
-
-lemma alpha_equ:
- shows "Lam [x].Var x = Lam [y].Var y"
- by (simp add: lam.eq_iff Abs1_eq_iff lam.fresh fresh_at_base)
-
-lemma Lam_freshness:
- assumes a: "atom y \<sharp> Lam [x].t"
- shows "(y = x) \<or> (y \<noteq> x \<and> atom y \<sharp> t)"
- using a by (auto simp add: lam.fresh Abs_fresh_iff)
-
-lemma neutral_element:
- fixes x::"nat"
- shows "x + 0 = x"
- by simp
-
-text {*
- Note that in the last statement, the explicit type annotation is important
- in order for Isabelle to be able to figure out what 0 stands for (e.g. a
- natural number, a vector, etc) and which lemmas to apply.
-*}
-
-
-section {* Isar Proofs *}
-
-text {*
- Isar is a language for writing formal proofs that can be understood by humans
- and by Isabelle. An Isar proof can be thought of as a sequence of 'stepping stones'
- that start with some assumptions and lead to the goal to be established. Every such
- stepping stone is introduced by "have" followed by the statement of the stepping
- stone. An exception is the goal to be proved, which need to be introduced with "show".
-
- have "statement" \<dots>
- show "goal_to_be_proved" \<dots>
-
- Since proofs usually do not proceed in a linear fashion, labels can be given
- to every stepping stone and they can be used later to explicitly refer to this
- corresponding stepping stone ("using").
-
- have label: "statement1" \<dots>
- \<dots>
- have "later_statement" using label \<dots>
-
- Each stepping stone (or have-statement) needs to have a justification. The
- simplest justification is "sorry" which admits any stepping stone, even false
- ones (this is good during the development of proofs).
-
- have "outrageous_false_statement" sorry
-
- Assumptions can be 'justified' using "by fact".
-
- have "assumption" by fact
-
- Derived facts can be justified using
-
- have "statement" by simp -- simplification
- have "statement" by auto -- proof search and simplification
- have "statement" by blast -- only proof search
-
- If facts or lemmas are needed in order to justify a have-statement, then
- one can feed these facts into the proof by using "using label \<dots>" or
- "using theorem-name \<dots>". More than one label at a time is allowed. For
- example
-
- have "statement" using label1 label2 theorem_name by auto
-
- Induction proofs in Isar are set up by indicating over which predicate(s)
- the induction proceeds ("using a b") followed by the command "proof (induct)".
- In this way, Isabelle uses defaults for which induction should be performed.
- These defaults can be overridden by giving more information, like the variable
- over which a structural induction should proceed, or a specific induction principle,
- such as well-founded inductions.
-
- After the induction is set up, the proof proceeds by cases. In Isar these
- cases can be given in any order. Most commonly they are started with "case" and the
- name of the case, and optionally some legible names for the variables used inside
- the case.
-
- In each "case", we need to establish a statement introduced by "show". Once
- this has been done, the next case can be started using "next". When all cases
- are completed, the proof can be finished using "qed".
-
- This means a typical induction proof has the following pattern
-
- proof (induct)
- case \<dots>
- \<dots>
- show \<dots>
- next
- case \<dots>
- \<dots>
- show \<dots>
- \<dots>
- qed
-
- Two very simple example proofs are as follows.
-*}
-
-subsection {* EXERCISE 1 *}
-
-lemma eval_val:
- assumes a: "val t"
- shows "t \<Down> t"
-using a
-proof (induct)
- case (v_Lam x t)
- show "Lam [x]. t \<Down> Lam [x]. t" sorry
-qed
-
-subsection {* EXERCISE 2 *}
-
-text {* Fill the gaps in the proof below. *}
-
-lemma eval_to_val:
- assumes a: "t \<Down> t'"
- shows "val t'"
-using a
-proof (induct)
- case (e_Lam x t)
- show "val (Lam [x].t)" sorry
-next
- case (e_App t1 x t t2 v v')
- -- {* all assumptions in this case *}
- have "t1 \<Down> Lam [x].t" by fact
- have ih1: "val (Lam [x]. t)" by fact
- have "t2 \<Down> v" by fact
- have ih2: "val v" by fact
- have "t [x ::= v] \<Down> v'" by fact
- have ih3: "val v'" by fact
-
- show "val v'" sorry
-qed
-
-
-section {* Datatypes: Evaluation Contexts*}
-
-text {*
- Datatypes can be defined in Isabelle as follows: we have to provide the name
- of the datatype and a list its type-constructors. Each type-constructor needs
- to have the information about the types of its arguments, and optionally
- can also contain some information about pretty syntax. For example, we like
- to write "\<box>" for holes.
-*}
-
-datatype ctx =
- Hole ("\<box>")
- | CAppL "ctx" "lam"
- | CAppR "lam" "ctx"
-
-text {* Now Isabelle knows about: *}
-
-typ ctx
-term "\<box>"
-term "CAppL"
-term "CAppL \<box> (Var x)"
-
-subsection {* MINI EXERCISE *}
-
-text {*
- Try and see what happens if you apply a Hole to a Hole?
-*}
-
-
-section {* Machines for Evaluation *}
-
-type_synonym ctxs = "ctx list"
-
-inductive
- machine :: "lam \<Rightarrow> ctxs \<Rightarrow>lam \<Rightarrow> ctxs \<Rightarrow> bool" ("<_,_> \<mapsto> <_,_>" [60, 60, 60, 60] 60)
-where
- m1: "<App t1 t2, Es> \<mapsto> <t1, (CAppL \<box> t2) # Es>"
-| m2: "val v \<Longrightarrow> <v, (CAppL \<box> t2) # Es> \<mapsto> <t2, (CAppR v \<box>) # Es>"
-| m3: "val v \<Longrightarrow> <v, (CAppR (Lam [x].t) \<box>) # Es> \<mapsto> <t[x ::= v], Es>"
-
-text {*
- Since the machine defined above only performs a single reduction,
- we need to define the transitive closure of this machine. *}
-
-inductive
- machines :: "lam \<Rightarrow> ctxs \<Rightarrow> lam \<Rightarrow> ctxs \<Rightarrow> bool" ("<_,_> \<mapsto>* <_,_>" [60, 60, 60, 60] 60)
-where
- ms1: "<t,Es> \<mapsto>* <t,Es>"
-| ms2: "\<lbrakk><t1,Es1> \<mapsto> <t2,Es2>; <t2,Es2> \<mapsto>* <t3,Es3>\<rbrakk> \<Longrightarrow> <t1,Es1> \<mapsto>* <t3,Es3>"
-
-declare machine.intros[intro] machines.intros[intro]
-
-section {* EXERCISE 3 *}
-
-text {*
- We need to show that the machines-relation is transitive.
- Fill in the gaps below.
-*}
-
-lemma ms3[intro]:
- assumes a: "<e1, Es1> \<mapsto>* <e2, Es2>"
- and b: "<e2, Es2> \<mapsto>* <e3, Es3>"
- shows "<e1, Es1> \<mapsto>* <e3, Es3>"
-using a b
-proof(induct)
- case (ms1 e1 Es1)
- have c: "<e1, Es1> \<mapsto>* <e3, Es3>" by fact
-
- show "<e1, Es1> \<mapsto>* <e3, Es3>" sorry
-next
- case (ms2 e1 Es1 e2 Es2 e2' Es2')
- have ih: "<e2', Es2'> \<mapsto>* <e3, Es3> \<Longrightarrow> <e2, Es2> \<mapsto>* <e3, Es3>" by fact
- have d1: "<e2', Es2'> \<mapsto>* <e3, Es3>" by fact
- have d2: "<e1, Es1> \<mapsto> <e2, Es2>" by fact
-
- show "<e1, Es1> \<mapsto>* <e3, Es3>" sorry
-qed
-
-text {*
- Just like gotos in the Basic programming language, labels often reduce
- the readability of proofs. Therefore one can use in Isar the notation
- "then have" in order to feed a have-statement to the proof of
- the next have-statement. This is used in the second case below.
-*}
-
-lemma
- assumes a: "<e1, Es1> \<mapsto>* <e2, Es2>"
- and b: "<e2, Es2> \<mapsto>* <e3, Es3>"
- shows "<e1, Es1> \<mapsto>* <e3, Es3>"
-using a b
-proof(induct)
- case (ms1 e1 Es1)
- show "<e1, Es1> \<mapsto>* <e3, Es3>" by fact
-next
- case (ms2 e1 Es1 e2 Es2 e2' Es2')
- have ih: "<e2', Es2'> \<mapsto>* <e3, Es3> \<Longrightarrow> <e2, Es2> \<mapsto>* <e3, Es3>" by fact
- have "<e2', Es2'> \<mapsto>* <e3, Es3>" by fact
- then have d3: "<e2, Es2> \<mapsto>* <e3, Es3>" using ih by simp
- have d2: "<e1, Es1> \<mapsto> <e2, Es2>" by fact
- show "<e1, Es1> \<mapsto>* <e3, Es3>" using d2 d3 by auto
-qed
-
-text {*
- The label ih2 cannot be got rid of in this way, because it is used
- two lines below and we cannot rearrange them. We can still avoid the
- label by feeding a sequence of facts into a proof using the
- "moreover"-chaining mechanism:
-
- have "statement_1" \<dots>
- moreover
- have "statement_2" \<dots>
- \<dots>
- moreover
- have "statement_n" \<dots>
- ultimately have "statement" \<dots>
-
- In this chain, all "statement_i" can be used in the proof of the final
- "statement". With this we can simplify our proof further to:
-*}
-
-lemma
- assumes a: "<e1, Es1> \<mapsto>* <e2, Es2>"
- and b: "<e2, Es2> \<mapsto>* <e3, Es3>"
- shows "<e1, Es1> \<mapsto>* <e3, Es3>"
-using a b
-proof(induct)
- case (ms1 e1 Es1)
- show "<e1, Es1> \<mapsto>* <e3, Es3>" by fact
-next
- case (ms2 e1 Es1 e2 Es2 e2' Es2')
- have ih: "<e2', Es2'> \<mapsto>* <e3, Es3> \<Longrightarrow> <e2, Es2> \<mapsto>* <e3, Es3>" by fact
- have "<e2', Es2'> \<mapsto>* <e3, Es3>" by fact
- then have "<e2, Es2> \<mapsto>* <e3, Es3>" using ih by simp
- moreover
- have "<e1, Es1> \<mapsto> <e2, Es2>" by fact
- ultimately show "<e1, Es1> \<mapsto>* <e3, Es3>" by auto
-qed
-
-
-text {*
- While automatic proof procedures in Isabelle are not able to prove statements
- like "P = NP" assuming usual definitions for P and NP, they can automatically
- discharge the lemmas we just proved. For this we only have to set up the induction
- and auto will take care of the rest. This means we can write:
-*}
-
-lemma
- assumes a: "val t"
- shows "t \<Down> t"
-using a by (induct) (auto)
-
-lemma
- assumes a: "t \<Down> t'"
- shows "val t'"
-using a by (induct) (auto)
-
-lemma
- assumes a: "<e1, Es1> \<mapsto>* <e2, Es2>"
- and b: "<e2, Es2> \<mapsto>* <e3, Es3>"
- shows "<e1, Es1> \<mapsto>* <e3, Es3>"
-using a b by (induct) (auto)
-
-
-section {* EXERCISE 4 *}
-
-text {*
- The point of the machine is that it simulates the evaluation
- relation. Therefore we like to prove the following:
-*}
-
-theorem
- assumes a: "t \<Down> t'"
- shows "<t, []> \<mapsto>* <t', []>"
-using a
-proof (induct)
- case (e_Lam x t)
- -- {* no assumptions *}
- show "<Lam [x].t, []> \<mapsto>* <Lam [x].t, []>" by auto
-next
- case (e_App t1 x t t2 v' v)
- -- {* all assumptions in this case *}
- have a1: "t1 \<Down> Lam [x].t" by fact
- have ih1: "<t1, []> \<mapsto>* <Lam [x].t, []>" by fact
- have a2: "t2 \<Down> v'" by fact
- have ih2: "<t2, []> \<mapsto>* <v', []>" by fact
- have a3: "t[x::=v'] \<Down> v" by fact
- have ih3: "<t[x::=v'], []> \<mapsto>* <v, []>" by fact
-
- -- {* your reasoning *}
- show "<App t1 t2, []> \<mapsto>* <v, []>" sorry
-qed
-
-
-text {*
- Again the automatic tools in Isabelle can discharge automatically
- of the routine work in these proofs. We can write:
-*}
-
-theorem eval_implies_machines_ctx:
- assumes a: "t \<Down> t'"
- shows "<t, Es> \<mapsto>* <t', Es>"
-using a
-by (induct arbitrary: Es)
- (metis eval_to_val machine.intros ms1 ms2 ms3 v_Lam)+
-
-corollary eval_implies_machines:
- assumes a: "t \<Down> t'"
- shows "<t, []> \<mapsto>* <t', []>"
-using a eval_implies_machines_ctx by simp
-
-
-section {* Function Definitions: Filling a Lambda-Term into a Context *}
-
-text {*
- Many functions over datatypes can be defined by recursion on the
- structure. For this purpose, Isabelle provides "fun". To use it one needs
- to give a name for the function, its type, optionally some pretty-syntax
- and then some equations defining the function. Like in "inductive",
- "fun" expects that more than one such equation is separated by "|".
-*}
-
-fun
- filling :: "ctx \<Rightarrow> lam \<Rightarrow> lam" ("_\<lbrakk>_\<rbrakk>" [100, 100] 100)
-where
- "\<box>\<lbrakk>t\<rbrakk> = t"
-| "(CAppL E t')\<lbrakk>t\<rbrakk> = App (E\<lbrakk>t\<rbrakk>) t'"
-| "(CAppR t' E)\<lbrakk>t\<rbrakk> = App t' (E\<lbrakk>t\<rbrakk>)"
-
-
-text {*
- After this definition Isabelle will be able to simplify
- statements like:
-*}
-
-lemma
- shows "(CAppL \<box> (Var x))\<lbrakk>Var y\<rbrakk> = App (Var y) (Var x)"
- by simp
-
-fun
- ctx_compose :: "ctx \<Rightarrow> ctx \<Rightarrow> ctx" (infixr "\<odot>" 99)
-where
- "\<box> \<odot> E' = E'"
-| "(CAppL E t') \<odot> E' = CAppL (E \<odot> E') t'"
-| "(CAppR t' E) \<odot> E' = CAppR t' (E \<odot> E')"
-
-fun
- ctx_composes :: "ctxs \<Rightarrow> ctx" ("_\<down>" [110] 110)
-where
- "[]\<down> = \<box>"
- | "(E # Es)\<down> = (Es\<down>) \<odot> E"
-
-text {*
- Notice that we not just have given a pretty syntax for the functions, but
- also some precedences. The numbers inside the [\<dots>] stand for the precedences
- of the arguments; the one next to it the precedence of the whole term.
-
- This means we have to write (Es1 \<odot> Es2) \<odot> Es3 otherwise Es1 \<odot> Es2 \<odot> Es3 is
- interpreted as Es1 \<odot> (Es2 \<odot> Es3).
-*}
-
-section {* Structural Inductions over Contexts *}
-
-text {*
- So far we have had a look at an induction over an inductive predicate.
- Another important induction principle is structural inductions for
- datatypes. To illustrate structural inductions we prove some facts
- about context composition, some of which we will need later on.
-*}
-
-subsection {* EXERCISE 5 *}
-
-text {* Complete the proof and remove the sorries. *}
-
-lemma ctx_compose:
- shows "(E1 \<odot> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>"
-proof (induct E1)
- case Hole
- show "(\<box> \<odot> E2)\<lbrakk>t\<rbrakk> = \<box>\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" sorry
-next
- case (CAppL E1 t')
- have ih: "(E1 \<odot> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" by fact
- show "((CAppL E1 t') \<odot> E2)\<lbrakk>t\<rbrakk> = (CAppL E1 t')\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" sorry
-next
- case (CAppR t' E1)
- have ih: "(E1 \<odot> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" by fact
- show "((CAppR t' E1) \<odot> E2)\<lbrakk>t\<rbrakk> = (CAppR t' E1)\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" sorry
-qed
-
-
-subsection {* EXERCISE 6 *}
-
-text {*
- Remove the sorries in the ctx_append proof below. You can make
- use of the following two properties.
-*}
-
-lemma neut_hole:
- shows "E \<odot> \<box> = E"
-by (induct E) (simp_all)
-
-lemma compose_assoc:
- shows "(E1 \<odot> E2) \<odot> E3 = E1 \<odot> (E2 \<odot> E3)"
-by (induct E1) (simp_all)
-
-lemma
- shows "(Es1 @ Es2)\<down> = (Es2\<down>) \<odot> (Es1\<down>)"
-proof (induct Es1)
- case Nil
- show "([] @ Es2)\<down> = Es2\<down> \<odot> []\<down>" sorry
-next
- case (Cons E Es1)
- have ih: "(Es1 @ Es2)\<down> = Es2\<down> \<odot> Es1\<down>" by fact
-
- show "((E # Es1) @ Es2)\<down> = Es2\<down> \<odot> (E # Es1)\<down>" sorry
-qed
-
-text {*
- The last proof involves several steps of equational reasoning.
- Isar provides some convenient means to express such equational
- reasoning in a much cleaner fashion using the "also have"
- construction.
-*}
-
-lemma
- shows "(Es1 @ Es2)\<down> = (Es2\<down>) \<odot> (Es1\<down>)"
-proof (induct Es1)
- case Nil
- show "([] @ Es2)\<down> = Es2\<down> \<odot> []\<down>" using neut_hole by simp
-next
- case (Cons E Es1)
- have ih: "(Es1 @ Es2)\<down> = Es2\<down> \<odot> Es1\<down>" by fact
- have "((E # Es1) @ Es2)\<down> = (E # (Es1 @ Es2))\<down>" by simp
- also have "\<dots> = (Es1 @ Es2)\<down> \<odot> E" by simp
- also have "\<dots> = (Es2\<down> \<odot> Es1\<down>) \<odot> E" using ih by simp
- also have "\<dots> = Es2\<down> \<odot> (Es1\<down> \<odot> E)" using compose_assoc by simp
- also have "\<dots> = Es2\<down> \<odot> (E # Es1)\<down>" by simp
- finally show "((E # Es1) @ Es2)\<down> = Es2\<down> \<odot> (E # Es1)\<down>" by simp
-qed
-
-
-end
-
--- a/Tutorial/Tutorial1s.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,752 +0,0 @@
-
-header {*
-
- Nominal Isabelle Tutorial at POPL'11
- ====================================
-
- Nominal Isabelle is a definitional extension of Isabelle/HOL, which
- means it does not add any new axioms to higher-order logic. It just
- adds new definitions and an infrastructure for 'nominal resoning'.
-
-
- The jEdit Interface
- -------------------
-
- The Isabelle theorem prover comes with an interface for jEdit.
- Unlike many other theorem prover interfaces (e.g. ProofGeneral) where you
- advance a 'checked' region in a proof script, this interface immediately
- checks the whole buffer. The text you type is also immediately checked.
- Malformed text or unfinished proofs are highlighted in red with a little
- red 'stop' signal on the left-hand side. If you drag the 'red-box' cursor
- over a line, the Output window gives further feedback.
-
- Note: If a section is not parsed correctly, the work-around is to cut it
- out and paste it back into the text (cut-out: Ctrl + X; paste in: Ctrl + V;
- Cmd is Ctrl on the Mac)
-
- Nominal Isabelle and jEdit can be started by typing on the command line
-
- isabelle jedit -l HOL-Nominal2
- isabelle jedit -l HOL-Nominal2 A.thy B.thy ...
-
-
- Symbols
- -------
-
- Symbols can considerably improve the readability of your statements and proofs.
- They can be input by just typing 'name-of-symbol' where 'name-of-symbol' is the
- usual latex name of that symbol. A little window will then appear in which
- you can select the symbol. With `Escape' you can ignore any suggestion.
-
- There are some handy short-cuts for frequently used symbols.
- For example
-
- short-cut symbol
-
- => \<Rightarrow>
- ==> \<Longrightarrow>
- --> \<longrightarrow>
- ! \<forall>
- ? \<exists>
- /\ \<and>
- \/ \<or>
- ~ \<not>
- ~= \<noteq>
- : \<in>
- ~: \<notin>
-
- For nominal the following two symbols have a special meaning
-
- \<sharp> sharp (freshness)
- \<bullet> bullet (permutation application)
-*}
-
-theory Tutorial1s
-imports Lambda
-begin
-
-section {* Theories *}
-
-text {*
- All formal developments in Isabelle are part of a theory. A theory
- needs to have a name and must import some pre-existing theory. The
- imported theory will normally be Nominal2 (which provides many useful
- concepts like set-theory, lists, nominal things etc). For the purpose
- of this tutorial we import the theory Lambda.thy, which contains
- already some useful definitions for (alpha-equated) lambda terms.
-*}
-
-
-
-section {* Types *}
-
-text {*
- Isabelle is based on simple types including some polymorphism. It also
- includes overloading, which means that sometimes explicit type annotations
- need to be given.
-
- - Base types include: nat, bool, string, lam (defined in the Lambda theory)
-
- - Type formers include: 'a list, ('a \<times> 'b), 'c set
-
- - Type variables are written like in ML with an apostrophe: 'a, 'b, \<dots>
-
- Types known to Isabelle can be queried using the command "typ".
-*}
-
-typ nat
-typ bool
-typ string
-typ lam -- {* alpha-equated lambda terms defined in Lambda.thy *}
-typ name -- {* type of (object) variables defined in Lambda.thy *}
-typ "('a \<times> 'b)" -- {* pair type *}
-typ "'c set" -- {* set type *}
-typ "'a list" -- {* list type *}
-typ "lam \<Rightarrow> nat" -- {* type of functions from lambda terms to natural numbers *}
-
-
-text {* Some malformed types - note the "stop" signal on the left margin *}
-
-(*
-typ boolean -- {* undeclared type *}
-typ set -- {* type argument missing *}
-*)
-
-
-section {* Terms *}
-
-text {*
- Every term in Isabelle needs to be well-typed. However they can have
- polymorphic type. Whether a term is accepted can be queried using
- the command "term".
-*}
-
-term c -- {* a variable of polymorphic type *}
-term "1::nat" -- {* the constant 1 in natural numbers *}
-term 1 -- {* the constant 1 with polymorphic type *}
-term "{1, 2, 3::nat}" -- {* the set containing natural numbers 1, 2 and 3 *}
-term "[1, 2, 3]" -- {* the list containing the polymorphic numbers 1, 2 and 3 *}
-term "(True, ''c'')" -- {* a pair consisting of the boolean true and the string "c" *}
-term "Suc 0" -- {* successor of 0, in other words 1::nat *}
-term "Lam [x].Var x" -- {* a lambda-term *}
-term "App t1 t2" -- {* another lambda-term *}
-term "x::name" -- {* an (object) variable of type name *}
-term "atom (x::name)" -- {* atom is an overloded function *}
-
-text {*
- Lam [x].Var is the syntax we made up for lambda abstractions. If you
- prefer, you can have your own syntax (but \<lambda> is already taken up for
- Isabelle's functions). We also came up with the type "name" for variables.
- You can introduce your own types of object variables using the
- command atom_decl:
-*}
-
-atom_decl ident
-atom_decl ty_var
-
-text {*
- Isabelle provides some useful colour feedback about its constants (black),
- free variables (blue) and bound variables (green).
-*}
-
-term "True" -- {* a constant defined somewhere...written in black *}
-term "true" -- {* not recognised as a constant, therefore it is interpreted
- as a free variable, written in blue *}
-term "\<forall>x. P x" -- {* x is bound (green), P is free (blue) *}
-
-
-text {* Formulae in Isabelle/HOL are terms of type bool *}
-
-term "True"
-term "True \<and> False"
-term "True \<or> B"
-term "{1,2,3} = {3,2,1}"
-term "\<forall>x. P x"
-term "A \<longrightarrow> B"
-term "atom a \<sharp> t" -- {* freshness in Nominal *}
-
-text {*
- When working with Isabelle, one deals with an object logic (that is HOL) and
- Isabelle's rule framework (called Pure). Occasionally one has to pay attention
- to this fact. But for the moment we ignore this completely.
-*}
-
-term "A \<longrightarrow> B" -- {* versus *}
-term "A \<Longrightarrow> B"
-
-term "\<forall>x. P x" -- {* versus *}
-term "\<And>x. P x"
-
-
-section {* Inductive Definitions: Evaluation Relation *}
-
-text {*
- In this section we want to define inductively the evaluation
- relation and for cbv-reduction relation.
-
- Inductive definitions in Isabelle start with the keyword "inductive"
- and a predicate name. One can optionally provide a type for the predicate.
- Clauses of the inductive predicate are introduced by "where" and more than
- two clauses need to be separated by "|". One can also give a name to each
- clause and indicate that it should be added to the hints database ("[intro]").
- A typical clause has some premises and a conclusion. This is written in
- Isabelle as:
-
- "premise \<Longrightarrow> conclusion"
- "premise1 \<Longrightarrow> premise2 \<Longrightarrow> \<dots> premisen \<Longrightarrow> conclusion"
-
- There is an alternative way of writing the latter clause as
-
- "\<lbrakk>premise1; premise2; \<dots> premisen\<rbrakk> \<Longrightarrow> conclusion"
-
- If no premise is present, then one just writes
-
- "conclusion"
-
- Below we give two definitions for the transitive closure
-*}
-
-inductive
- eval :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<Down> _" [60, 60] 60)
-where
- e_Lam[intro]: "Lam [x].t \<Down> Lam [x].t"
-| e_App[intro]: "\<lbrakk>t1 \<Down> Lam [x].t; t2 \<Down> v'; t[x::=v'] \<Down> v\<rbrakk> \<Longrightarrow> App t1 t2 \<Down> v"
-
-text {*
- Values and cbv are also defined using inductive.
-*}
-
-inductive
- val :: "lam \<Rightarrow> bool"
-where
- v_Lam[intro]: "val (Lam [x].t)"
-
-section {* Theorems *}
-
-text {*
- A central concept in Isabelle is that of theorems. Isabelle's theorem
- database can be queried using
-*}
-
-thm e_App
-thm e_Lam
-thm conjI
-thm conjunct1
-
-text {*
- Notice that theorems usually contain schematic variables (e.g. ?P, ?Q, \<dots>).
- These schematic variables can be substituted with any term (of the right type
- of course).
-
- When defining the predicates beta_star, Isabelle provides us automatically
- with the following theorems that state how they can be introduced and what
- constitutes an induction over them.
-*}
-
-thm eval.intros
-thm eval.induct
-
-
-section {* Lemmas / Theorems / Corollaries *}
-
-text {*
- Whether to use lemma, theorem or corollary makes no semantic difference
- in Isabelle.
-
- A lemma starts with "lemma" and consists of a statement ("shows \<dots>") and
- optionally a lemma name, some type-information for variables ("fixes \<dots>")
- and some assumptions ("assumes \<dots>").
-
- Lemmas also need to have a proof, but ignore this 'detail' for the moment.
-
- Examples are
-*}
-
-lemma alpha_equ:
- shows "Lam [x].Var x = Lam [y].Var y"
- by (simp add: lam.eq_iff Abs1_eq_iff lam.fresh fresh_at_base)
-
-lemma Lam_freshness:
- assumes a: "atom y \<sharp> Lam [x].t"
- shows "(y = x) \<or> (y \<noteq> x \<and> atom y \<sharp> t)"
- using a by (auto simp add: lam.fresh Abs_fresh_iff)
-
-lemma neutral_element:
- fixes x::"nat"
- shows "x + 0 = x"
- by simp
-
-text {*
- Note that in the last statement, the explicit type annotation is important
- in order for Isabelle to be able to figure out what 0 stands for (e.g. a
- natural number, a vector, etc) and which lemmas to apply.
-*}
-
-
-section {* Isar Proofs *}
-
-text {*
- Isar is a language for writing formal proofs that can be understood by humans
- and by Isabelle. An Isar proof can be thought of as a sequence of 'stepping stones'
- that start with some assumptions and lead to the goal to be established. Every such
- stepping stone is introduced by "have" followed by the statement of the stepping
- stone. An exception is the goal to be proved, which need to be introduced with "show".
-
- have "statement" \<dots>
- show "goal_to_be_proved" \<dots>
-
- Since proofs usually do not proceed in a linear fashion, labels can be given
- to every stepping stone and they can be used later to explicitly refer to this
- corresponding stepping stone ("using").
-
- have label: "statement1" \<dots>
- \<dots>
- have "later_statement" using label \<dots>
-
- Each stepping stone (or have-statement) needs to have a justification. The
- simplest justification is "sorry" which admits any stepping stone, even false
- ones (this is good during the development of proofs).
-
- have "outrageous_false_statement" sorry
-
- Assumptions can be 'justified' using "by fact".
-
- have "assumption" by fact
-
- Derived facts can be justified using
-
- have "statement" by simp -- simplification
- have "statement" by auto -- proof search and simplification
- have "statement" by blast -- only proof search
-
- If facts or lemmas are needed in order to justify a have-statement, then
- one can feed these facts into the proof by using "using label \<dots>" or
- "using theorem-name \<dots>". More than one label at a time is allowed. For
- example
-
- have "statement" using label1 label2 theorem_name by auto
-
- Induction proofs in Isar are set up by indicating over which predicate(s)
- the induction proceeds ("using a b") followed by the command "proof (induct)".
- In this way, Isabelle uses defaults for which induction should be performed.
- These defaults can be overridden by giving more information, like the variable
- over which a structural induction should proceed, or a specific induction principle,
- such as well-founded inductions.
-
- After the induction is set up, the proof proceeds by cases. In Isar these
- cases can be given in any order. Most commonly they are started with "case" and the
- name of the case, and optionally some legible names for the variables used inside
- the case.
-
- In each "case", we need to establish a statement introduced by "show". Once
- this has been done, the next case can be started using "next". When all cases
- are completed, the proof can be finished using "qed".
-
- This means a typical induction proof has the following pattern
-
- proof (induct)
- case \<dots>
- \<dots>
- show \<dots>
- next
- case \<dots>
- \<dots>
- show \<dots>
- \<dots>
- qed
-
- Two very simple example proofs are as follows.
-*}
-
-subsection {* EXERCISE 1 *}
-
-lemma eval_val:
- assumes a: "val t"
- shows "t \<Down> t"
-using a
-proof (induct)
- case (v_Lam x t)
- show "Lam [x]. t \<Down> Lam [x]. t" by auto
-qed
-
-subsection {* EXERCISE 2 *}
-
-text {* Fill the gaps in the proof below. *}
-
-lemma eval_to_val:
- assumes a: "t \<Down> t'"
- shows "val t'"
-using a
-proof (induct)
- case (e_Lam x t)
- show "val (Lam [x].t)" by auto
-next
- case (e_App t1 x t t2 v v')
- -- {* all assumptions in this case *}
- have "t1 \<Down> Lam [x].t" by fact
- have ih1: "val (Lam [x]. t)" by fact
- have "t2 \<Down> v" by fact
- have ih2: "val v" by fact
- have "t [x ::= v] \<Down> v'" by fact
- have ih3: "val v'" by fact
-
- show "val v'" using ih3 by auto
-qed
-
-
-section {* Datatypes: Evaluation Contexts*}
-
-text {*
- Datatypes can be defined in Isabelle as follows: we have to provide the name
- of the datatype and a list its type-constructors. Each type-constructor needs
- to have the information about the types of its arguments, and optionally
- can also contain some information about pretty syntax. For example, we like
- to write "\<box>" for holes.
-*}
-
-datatype ctx =
- Hole ("\<box>")
- | CAppL "ctx" "lam"
- | CAppR "lam" "ctx"
-
-text {* Now Isabelle knows about: *}
-
-typ ctx
-term "\<box>"
-term "CAppL"
-term "CAppL \<box> (Var x)"
-
-subsection {* MINI EXERCISE *}
-
-text {*
- Try and see what happens if you apply a Hole to a Hole?
-*}
-
-
-section {* Machines for Evaluation *}
-
-type_synonym ctxs = "ctx list"
-
-inductive
- machine :: "lam \<Rightarrow> ctxs \<Rightarrow>lam \<Rightarrow> ctxs \<Rightarrow> bool" ("<_,_> \<mapsto> <_,_>" [60, 60, 60, 60] 60)
-where
- m1: "<App t1 t2, Es> \<mapsto> <t1, (CAppL \<box> t2) # Es>"
-| m2: "val v \<Longrightarrow> <v, (CAppL \<box> t2) # Es> \<mapsto> <t2, (CAppR v \<box>) # Es>"
-| m3: "val v \<Longrightarrow> <v, (CAppR (Lam [x].t) \<box>) # Es> \<mapsto> <t[x ::= v], Es>"
-
-text {*
- Since the machine defined above only performs a single reduction,
- we need to define the transitive closure of this machine. *}
-
-inductive
- machines :: "lam \<Rightarrow> ctxs \<Rightarrow> lam \<Rightarrow> ctxs \<Rightarrow> bool" ("<_,_> \<mapsto>* <_,_>" [60, 60, 60, 60] 60)
-where
- ms1: "<t,Es> \<mapsto>* <t,Es>"
-| ms2: "\<lbrakk><t1,Es1> \<mapsto> <t2,Es2>; <t2,Es2> \<mapsto>* <t3,Es3>\<rbrakk> \<Longrightarrow> <t1,Es1> \<mapsto>* <t3,Es3>"
-
-declare machine.intros[intro] machines.intros[intro]
-
-section {* EXERCISE 3 *}
-
-text {*
- We need to show that the machines-relation is transitive.
- Fill in the gaps below.
-*}
-
-lemma ms3[intro]:
- assumes a: "<e1, Es1> \<mapsto>* <e2, Es2>"
- and b: "<e2, Es2> \<mapsto>* <e3, Es3>"
- shows "<e1, Es1> \<mapsto>* <e3, Es3>"
-using a b
-proof(induct)
- case (ms1 e1 Es1)
- have c: "<e1, Es1> \<mapsto>* <e3, Es3>" by fact
- then show "<e1, Es1> \<mapsto>* <e3, Es3>" by simp
-next
- case (ms2 e1 Es1 e2 Es2 e2' Es2')
- have ih: "<e2', Es2'> \<mapsto>* <e3, Es3> \<Longrightarrow> <e2, Es2> \<mapsto>* <e3, Es3>" by fact
- have d1: "<e2', Es2'> \<mapsto>* <e3, Es3>" by fact
- have d2: "<e1, Es1> \<mapsto> <e2, Es2>" by fact
- have d3: "<e2, Es2> \<mapsto>* <e3, Es3>" using ih d1 by auto
- show "<e1, Es1> \<mapsto>* <e3, Es3>" using d2 d3 by auto
-qed
-
-text {*
- Just like gotos in the Basic programming language, labels often reduce
- the readability of proofs. Therefore one can use in Isar the notation
- "then have" in order to feed a have-statement to the proof of
- the next have-statement. This is used in teh second case below.
-*}
-
-lemma
- assumes a: "<e1, Es1> \<mapsto>* <e2, Es2>"
- and b: "<e2, Es2> \<mapsto>* <e3, Es3>"
- shows "<e1, Es1> \<mapsto>* <e3, Es3>"
-using a b
-proof(induct)
- case (ms1 e1 Es1)
- show "<e1, Es1> \<mapsto>* <e3, Es3>" by fact
-next
- case (ms2 e1 Es1 e2 Es2 e2' Es2')
- have ih: "<e2', Es2'> \<mapsto>* <e3, Es3> \<Longrightarrow> <e2, Es2> \<mapsto>* <e3, Es3>" by fact
- have "<e2', Es2'> \<mapsto>* <e3, Es3>" by fact
- then have d3: "<e2, Es2> \<mapsto>* <e3, Es3>" using ih by simp
- have d2: "<e1, Es1> \<mapsto> <e2, Es2>" by fact
- show "<e1, Es1> \<mapsto>* <e3, Es3>" using d2 d3 by auto
-qed
-
-text {*
- The label ih2 cannot be got rid of in this way, because it is used
- two lines below and we cannot rearange them. We can still avoid the
- label by feeding a sequence of facts into a proof using the
- "moreover"-chaining mechanism:
-
- have "statement_1" \<dots>
- moreover
- have "statement_2" \<dots>
- \<dots>
- moreover
- have "statement_n" \<dots>
- ultimately have "statement" \<dots>
-
- In this chain, all "statement_i" can be used in the proof of the final
- "statement". With this we can simplify our proof further to:
-*}
-
-lemma
- assumes a: "<e1, Es1> \<mapsto>* <e2, Es2>"
- and b: "<e2, Es2> \<mapsto>* <e3, Es3>"
- shows "<e1, Es1> \<mapsto>* <e3, Es3>"
-using a b
-proof(induct)
- case (ms1 e1 Es1)
- show "<e1, Es1> \<mapsto>* <e3, Es3>" by fact
-next
- case (ms2 e1 Es1 e2 Es2 e2' Es2')
- have ih: "<e2', Es2'> \<mapsto>* <e3, Es3> \<Longrightarrow> <e2, Es2> \<mapsto>* <e3, Es3>" by fact
- have "<e2', Es2'> \<mapsto>* <e3, Es3>" by fact
- then have "<e2, Es2> \<mapsto>* <e3, Es3>" using ih by simp
- moreover
- have "<e1, Es1> \<mapsto> <e2, Es2>" by fact
- ultimately show "<e1, Es1> \<mapsto>* <e3, Es3>" by auto
-qed
-
-
-text {*
- While automatic proof procedures in Isabelle are not able to prove statements
- like "P = NP" assuming usual definitions for P and NP, they can automatically
- discharge the lemmas we just proved. For this we only have to set up the induction
- and auto will take care of the rest. This means we can write:
-*}
-
-lemma
- assumes a: "val t"
- shows "t \<Down> t"
-using a by (induct) (auto)
-
-lemma
- assumes a: "t \<Down> t'"
- shows "val t'"
-using a by (induct) (auto)
-
-lemma
- assumes a: "<e1, Es1> \<mapsto>* <e2, Es2>"
- and b: "<e2, Es2> \<mapsto>* <e3, Es3>"
- shows "<e1, Es1> \<mapsto>* <e3, Es3>"
-using a b by (induct) (auto)
-
-
-section {* EXERCISE 4 *}
-
-text {*
- The point of the machine is that it simulates the evaluation
- relation. Therefore we like to prove the following:
-*}
-
-theorem
- assumes a: "t \<Down> t'"
- shows "<t, Es> \<mapsto>* <t', Es>"
-using a
-proof (induct arbitrary: Es)
- case (e_Lam x t Es)
- -- {* no assumptions *}
- show "<Lam [x].t, Es> \<mapsto>* <Lam [x].t, Es>" by auto
-next
- case (e_App t1 x t t2 v' v Es)
- -- {* all assumptions in this case *}
- have a1: "t1 \<Down> Lam [x].t" by fact
- have ih1: "\<And>Es. <t1, Es> \<mapsto>* <Lam [x].t, Es>" by fact
- have a2: "t2 \<Down> v'" by fact
- have ih2: "\<And>Es. <t2, Es> \<mapsto>* <v', Es>" by fact
- have a3: "t[x::=v'] \<Down> v" by fact
- have ih3: "\<And>Es. <t[x::=v'], Es> \<mapsto>* <v, Es>" by fact
- -- {* your reasoning *}
- have "<App t1 t2, Es> \<mapsto>* <t1, CAppL \<box> t2 # Es>" by auto
- moreover
- have "<t1, CAppL \<box> t2 # Es> \<mapsto>* <Lam [x].t, CAppL \<box> t2 # Es>" using ih1 by auto
- moreover
- have "<Lam [x].t, CAppL \<box> t2 # Es> \<mapsto>* <t2, CAppR (Lam [x].t) \<box> # Es>" by auto
- moreover
- have "<t2, CAppR (Lam [x].t) \<box> # Es> \<mapsto>* <v', CAppR (Lam [x].t) \<box> # Es>"
- using ih2 by auto
- moreover
- have "val v'" using a2 eval_to_val by auto
- then have "<v', CAppR (Lam [x].t) \<box> # Es> \<mapsto>* <t[x::=v'], Es>" by auto
- moreover
- have "<t[x::=v'], Es> \<mapsto>* <v, Es>" using ih3 by auto
- ultimately show "<App t1 t2, Es> \<mapsto>* <v, Es>" by blast
-qed
-
-
-text {*
- Again the automatic tools in Isabelle can discharge automatically
- of the routine work in these proofs. We can write:
-*}
-
-theorem eval_implies_machines_ctx:
- assumes a: "t \<Down> t'"
- shows "<t, Es> \<mapsto>* <t', Es>"
-using a
-by (induct arbitrary: Es)
- (metis eval_to_val machine.intros ms1 ms2 ms3 v_Lam)+
-
-corollary eval_implies_machines:
- assumes a: "t \<Down> t'"
- shows "<t, []> \<mapsto>* <t', []>"
-using a eval_implies_machines_ctx by simp
-
-
-section {* Function Definitions: Filling a Lambda-Term into a Context *}
-
-text {*
- Many functions over datatypes can be defined by recursion on the
- structure. For this purpose, Isabelle provides "fun". To use it one needs
- to give a name for the function, its type, optionally some pretty-syntax
- and then some equations defining the function. Like in "inductive",
- "fun" expects that more than one such equation is separated by "|".
-*}
-
-fun
- filling :: "ctx \<Rightarrow> lam \<Rightarrow> lam" ("_\<lbrakk>_\<rbrakk>" [100, 100] 100)
-where
- "\<box>\<lbrakk>t\<rbrakk> = t"
-| "(CAppL E t')\<lbrakk>t\<rbrakk> = App (E\<lbrakk>t\<rbrakk>) t'"
-| "(CAppR t' E)\<lbrakk>t\<rbrakk> = App t' (E\<lbrakk>t\<rbrakk>)"
-
-
-text {*
- After this definition Isabelle will be able to simplify
- statements like:
-*}
-
-lemma
- shows "(CAppL \<box> (Var x))\<lbrakk>Var y\<rbrakk> = App (Var y) (Var x)"
- by simp
-
-fun
- ctx_compose :: "ctx \<Rightarrow> ctx \<Rightarrow> ctx" (infixr "\<odot>" 99)
-where
- "\<box> \<odot> E' = E'"
-| "(CAppL E t') \<odot> E' = CAppL (E \<odot> E') t'"
-| "(CAppR t' E) \<odot> E' = CAppR t' (E \<odot> E')"
-
-fun
- ctx_composes :: "ctxs \<Rightarrow> ctx" ("_\<down>" [110] 110)
-where
- "[]\<down> = \<box>"
- | "(E # Es)\<down> = (Es\<down>) \<odot> E"
-
-text {*
- Notice that we not just have given a pretty syntax for the functions, but
- also some precedences. The numbers inside the [\<dots>] stand for the precedences
- of the arguments; the one next to it the precedence of the whole term.
-
- This means we have to write (Es1 \<odot> Es2) \<odot> Es3 otherwise Es1 \<odot> Es2 \<odot> Es3 is
- interpreted as Es1 \<odot> (Es2 \<odot> Es3).
-*}
-
-section {* Structural Inductions over Contexts *}
-
-text {*
- So far we have had a look at an induction over an inductive predicate.
- Another important induction principle is structural inductions for
- datatypes. To illustrate structural inductions we prove some facts
- about context composition, some of which we will need later on.
-*}
-
-subsection {* EXERCISE 5 *}
-
-text {* Complete the proof and remove the sorries. *}
-
-lemma
- shows "(E1 \<odot> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>"
-proof (induct E1)
- case Hole
- show "(\<box> \<odot> E2)\<lbrakk>t\<rbrakk> = \<box>\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" by simp
-next
- case (CAppL E1 t')
- have ih: "(E1 \<odot> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" by fact
- show "((CAppL E1 t') \<odot> E2)\<lbrakk>t\<rbrakk> = (CAppL E1 t')\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" using ih by simp
-next
- case (CAppR t' E1)
- have ih: "(E1 \<odot> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" by fact
- show "((CAppR t' E1) \<odot> E2)\<lbrakk>t\<rbrakk> = (CAppR t' E1)\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" using ih by simp
-qed
-
-subsection {* EXERCISE 6 *}
-
-text {*
- Remove the sorries in the ctx_append proof below. You can make
- use of the following two properties.
-*}
-
-lemma neut_hole:
- shows "E \<odot> \<box> = E"
-by (induct E) (simp_all)
-
-lemma compose_assoc:
- shows "(E1 \<odot> E2) \<odot> E3 = E1 \<odot> (E2 \<odot> E3)"
-by (induct E1) (simp_all)
-
-lemma
- shows "(Es1 @ Es2)\<down> = (Es2\<down>) \<odot> (Es1\<down>)"
-proof (induct Es1)
- case Nil
- show "([] @ Es2)\<down> = Es2\<down> \<odot> []\<down>" using neut_hole by simp
-next
- case (Cons E Es1)
- have ih: "(Es1 @ Es2)\<down> = Es2\<down> \<odot> Es1\<down>" by fact
- have eq1: "((E # Es1) @ Es2)\<down> = (E # (Es1 @ Es2))\<down>" by simp
- have eq2: "(E # (Es1 @ Es2))\<down> = (Es1 @ Es2)\<down> \<odot> E" by simp
- have eq3: "(Es1 @ Es2)\<down> \<odot> E = (Es2\<down> \<odot> Es1\<down>) \<odot> E" using ih by simp
- have eq4: "(Es2\<down> \<odot> Es1\<down>) \<odot> E = Es2\<down> \<odot> (Es1\<down> \<odot> E)" using compose_assoc by simp
- have eq5: "Es2\<down> \<odot> (Es1\<down> \<odot> E) = Es2\<down> \<odot> (E # Es1)\<down>" by simp
- show "((E # Es1) @ Es2)\<down> = Es2\<down> \<odot> (E # Es1)\<down>" using eq1 eq2 eq3 eq4 eq5 by (simp only:)
-qed
-
-text {*
- The last proof involves several steps of equational reasoning.
- Isar provides some convenient means to express such equational
- reasoning in a much cleaner fashion using the "also have"
- construction.
-*}
-
-lemma
- shows "(Es1 @ Es2)\<down> = (Es2\<down>) \<odot> (Es1\<down>)"
-proof (induct Es1)
- case Nil
- show "([] @ Es2)\<down> = Es2\<down> \<odot> []\<down>" using neut_hole by simp
-next
- case (Cons E Es1)
- have ih: "(Es1 @ Es2)\<down> = Es2\<down> \<odot> Es1\<down>" by fact
- have "((E # Es1) @ Es2)\<down> = (E # (Es1 @ Es2))\<down>" by simp
- also have "\<dots> = (Es1 @ Es2)\<down> \<odot> E" by simp
- also have "\<dots> = (Es2\<down> \<odot> Es1\<down>) \<odot> E" using ih by simp
- also have "\<dots> = Es2\<down> \<odot> (Es1\<down> \<odot> E)" using compose_assoc by simp
- also have "\<dots> = Es2\<down> \<odot> (E # Es1)\<down>" by simp
- finally show "((E # Es1) @ Es2)\<down> = Es2\<down> \<odot> (E # Es1)\<down>" by simp
-qed
-
-
-end
-
--- a/Tutorial/Tutorial2.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,384 +0,0 @@
-
-theory Tutorial2
-imports Lambda
-begin
-
-section {* Height of a Lambda-Term *}
-
-text {*
- The theory Lambda defines the notions of capture-avoiding
- substitution and the height of lambda terms.
-*}
-
-thm height.simps
-thm subst.simps
-
-subsection {* EXERCISE 7 *}
-
-text {*
- Lets prove a property about the height of substitutions.
-
- Assume that the height of a lambda-term is always greater
- or equal to 1.
-*}
-
-lemma height_ge_one:
- shows "1 \<le> height t"
-by (induct t rule: lam.induct) (auto)
-
-text {* Remove the sorries *}
-
-theorem height_subst:
- shows "height (t[x ::= t']) \<le> height t - 1 + height t'"
-proof (induct t rule: lam.induct)
- case (Var y)
- show "height (Var y[x ::= t']) \<le> height (Var y) - 1 + height t'" sorry
-next
- case (App t1 t2)
- have ih1: "height (t1[x::=t']) \<le> (height t1) - 1 + height t'" by fact
- have ih2: "height (t2[x::=t']) \<le> (height t2) - 1 + height t'" by fact
-
- show "height ((App t1 t2)[x ::= t']) \<le> height (App t1 t2) - 1 + height t'" sorry
-next
- case (Lam y t1)
- have ih: "height (t1[x::=t']) \<le> height t1 - 1 + height t'" by fact
- -- {* the definition of capture-avoiding substitution *}
- thm subst.simps
-
- show "height ((Lam [y].t1)[x ::= t']) \<le> height (Lam [y].t1) - 1 + height t'" sorry
-qed
-
-
-text {*
- The point is that substitutions can only be moved under a binder
- provided certain freshness conditions are satisfied. The structural
- induction above does not say anything about such freshness conditions.
-
- Fortunately, Nominal derives automatically some stronger induction
- principle for lambda terms which has the usual variable convention
- build in.
-
- In the proof below, we use this stronger induction principle. The
- variable and application case are as before.
-*}
-
-
-theorem
- shows "height (t[x ::= t']) \<le> height t - 1 + height t'"
-proof (nominal_induct t avoiding: x t' rule: lam.strong_induct)
- case (Var y)
- have "1 \<le> height t'" using height_ge_one by simp
- then show "height (Var y[x ::= t']) \<le> height (Var y) - 1 + height t'" by simp
-next
- case (App t1 t2)
- have ih1: "height (t1[x::=t']) \<le> (height t1) - 1 + height t'"
- and ih2: "height (t2[x::=t']) \<le> (height t2) - 1 + height t'" by fact+
- then show "height ((App t1 t2)[x ::= t']) \<le> height (App t1 t2) - 1 + height t'" by simp
-next
- case (Lam y t1)
- have ih: "height (t1[x::=t']) \<le> height t1 - 1 + height t'" by fact
- have vc: "atom y \<sharp> x" "atom y \<sharp> t'" by fact+ -- {* usual variable convention *}
-
- show "height ((Lam [y].t1)[x ::= t']) \<le> height (Lam [y].t1) - 1 + height t'" sorry
-qed
-
-
-section {* Types and the Weakening Lemma *}
-
-nominal_datatype ty =
- tVar "string"
-| tArr "ty" "ty" (infixr "\<rightarrow>" 100)
-
-
-text {*
- Having defined them as nominal datatypes gives us additional
- definitions and theorems that come with types (see below).
-
- We next define the type of typing contexts, a predicate that
- defines valid contexts (i.e. lists that contain only unique
- variables) and the typing judgement.
-*}
-
-type_synonym ty_ctx = "(name \<times> ty) list"
-
-inductive
- valid :: "ty_ctx \<Rightarrow> bool"
-where
- v1[intro]: "valid []"
-| v2[intro]: "\<lbrakk>valid \<Gamma>; atom x \<sharp> \<Gamma>\<rbrakk>\<Longrightarrow> valid ((x, T) # \<Gamma>)"
-
-
-inductive
- typing :: "ty_ctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60, 60, 60] 60)
-where
- t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
-| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"
-| t_Lam[intro]: "\<lbrakk>atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T1 \<rightarrow> T2"
-
-
-text {*
- The predicate atom x \<sharp> \<Gamma>, read as x fresh for \<Gamma>, is defined by
- Nominal Isabelle. It is defined for lambda-terms, products, lists etc.
- For example we have:
-*}
-
-lemma
- fixes x::"name"
- shows "atom x \<sharp> Lam [x].t"
- and "atom x \<sharp> (t1, t2) \<Longrightarrow> atom x \<sharp> App t1 t2"
- and "atom x \<sharp> Var y \<Longrightarrow> atom x \<sharp> y"
- and "\<lbrakk>atom x \<sharp> t1; atom x \<sharp> t2\<rbrakk> \<Longrightarrow> atom x \<sharp> (t1, t2)"
- and "\<lbrakk>atom x \<sharp> l1; atom x \<sharp> l2\<rbrakk> \<Longrightarrow> atom x \<sharp> (l1 @ l2)"
- and "atom x \<sharp> y \<Longrightarrow> x \<noteq> y"
- by (simp_all add: lam.fresh fresh_append fresh_at_base)
-
-text {*
- We can also prove that every variable is fresh for a type.
-*}
-
-lemma ty_fresh:
- fixes x::"name"
- and T::"ty"
- shows "atom x \<sharp> T"
-by (induct T rule: ty.induct)
- (simp_all add: ty.fresh pure_fresh)
-
-text {*
- In order to state the weakening lemma in a convenient form, we
- using the following abbreviation. Abbreviations behave like
- definitions, except that they are always automatically folded
- and unfolded.
-*}
-
-abbreviation
- "sub_ty_ctx" :: "ty_ctx \<Rightarrow> ty_ctx \<Rightarrow> bool" ("_ \<sqsubseteq> _" [60, 60] 60)
-where
- "\<Gamma>1 \<sqsubseteq> \<Gamma>2 \<equiv> \<forall>x. x \<in> set \<Gamma>1 \<longrightarrow> x \<in> set \<Gamma>2"
-
-
-subsection {* EXERCISE 8 *}
-
-text {*
- Fill in the details and give a proof of the weakening lemma.
-*}
-
-lemma
- assumes a: "\<Gamma>1 \<turnstile> t : T"
- and b: "valid \<Gamma>2"
- and c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
- shows "\<Gamma>2 \<turnstile> t : T"
-using a b c
-proof (induct arbitrary: \<Gamma>2)
- case (t_Var \<Gamma>1 x T)
- have a1: "valid \<Gamma>1" by fact
- have a2: "(x, T) \<in> set \<Gamma>1" by fact
- have a3: "valid \<Gamma>2" by fact
- have a4: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
-
- show "\<Gamma>2 \<turnstile> Var x : T" sorry
-next
- case (t_Lam x \<Gamma>1 T1 t T2)
- have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x, T1) # \<Gamma>1 \<sqsubseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact
- have a0: "atom x \<sharp> \<Gamma>1" by fact
- have a1: "valid \<Gamma>2" by fact
- have a2: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
-
- show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" sorry
-qed (auto) -- {* the application case is automatic*}
-
-
-text {*
- Despite the frequent claim that the weakening lemma is trivial,
- routine or obvious, the proof in the lambda-case does not go
- through smoothly. Painful variable renamings seem to be necessary.
- But the proof using renamings is horribly complicated (see below).
-*}
-
-equivariance valid
-equivariance typing
-
-lemma weakening_NOT_TO_BE_TRIED_AT_HOME:
- assumes a: "\<Gamma>1 \<turnstile> t : T"
- and b: "valid \<Gamma>2"
- and c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
- shows "\<Gamma>2 \<turnstile> t : T"
-using a b c
-proof (induct arbitrary: \<Gamma>2)
- -- {* lambda case *}
- case (t_Lam x \<Gamma>1 T1 t T2)
- have fc0: "atom x \<sharp> \<Gamma>1" by fact
- have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x, T1) # \<Gamma>1 \<sqsubseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact
- -- {* we choose a fresh variable *}
- obtain c::"name" where fc1: "atom c \<sharp> (x, t, \<Gamma>1, \<Gamma>2)" by (rule obtain_fresh)
- -- {* we alpha-rename the abstraction *}
- have "Lam [c].((c \<leftrightarrow> x) \<bullet> t) = Lam [x].t" using fc1
- by (auto simp add: lam.eq_iff Abs1_eq_iff flip_def)
- moreover
- -- {* we can then alpha rename the goal *}
- have "\<Gamma>2 \<turnstile> Lam [c].((c \<leftrightarrow> x) \<bullet> t) : T1 \<rightarrow> T2"
- proof -
- -- {* we need to establish *}
- -- {* (1) (x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) *}
- -- {* (2) valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)) *}
- have "(1)": "(x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)"
- proof -
- have "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
- then have "(c \<leftrightarrow> x) \<bullet> ((x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))" using fc0 fc1
- by (perm_simp) (simp add: flip_fresh_fresh)
- then show "(x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)" by (rule permute_boolE)
- qed
- moreover
- have "(2)": "valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))"
- proof -
- have "valid \<Gamma>2" by fact
- then show "valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))" using fc1
- by (auto simp add: fresh_permute_left atom_eqvt valid.eqvt)
- qed
- -- {* these two facts give us by induction hypothesis the following *}
- ultimately have "(x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) \<turnstile> t : T2" using ih by simp
- -- {* we now apply renamings to get to our goal *}
- then have "(c \<leftrightarrow> x) \<bullet> ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) \<turnstile> t : T2)" by (rule permute_boolI)
- then have "(c, T1) # \<Gamma>2 \<turnstile> ((c \<leftrightarrow> x) \<bullet> t) : T2" using fc1
- by (perm_simp) (simp add: flip_fresh_fresh ty_fresh)
- then show "\<Gamma>2 \<turnstile> Lam [c].((c \<leftrightarrow> x) \<bullet> t) : T1 \<rightarrow> T2" using fc1 by auto
- qed
- ultimately show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" by simp
-qed (auto) -- {* var and app cases, luckily, are automatic *}
-
-
-text {*
- The remedy is to use again a stronger induction principle that
- has the usual "variable convention" already build in. The
- following command derives this induction principle for the typing
- relation. (We shall explain what happens here later.)
-*}
-
-nominal_inductive typing
- avoids t_Lam: "x"
- unfolding fresh_star_def
- by (simp_all add: fresh_Pair lam.fresh ty_fresh)
-
-text {* Compare the two induction principles *}
-
-thm typing.induct
-thm typing.strong_induct -- {* note the additional assumption {atom x} \<sharp> c *}
-
-text {*
- We can use the stronger induction principle by replacing
- the line
-
- proof (induct arbitrary: \<Gamma>2)
-
- with
-
- proof (nominal_induct avoiding: \<Gamma>2 rule: typing.strong_induct)
-
- Try now the proof.
-*}
-
-subsection {* EXERCISE 9 *}
-
-lemma weakening:
- assumes a: "\<Gamma>1 \<turnstile> t : T"
- and b: "valid \<Gamma>2"
- and c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
- shows "\<Gamma>2 \<turnstile> t : T"
-using a b c
-proof (nominal_induct avoiding: \<Gamma>2 rule: typing.strong_induct)
- case (t_Var \<Gamma>1 x T) -- {* again the variable case is as before *}
- have "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
- moreover
- have "valid \<Gamma>2" by fact
- moreover
- have "(x, T)\<in> set \<Gamma>1" by fact
- ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto
-next
- case (t_Lam x \<Gamma>1 T1 t T2)
- have vc: "atom x \<sharp> \<Gamma>2" by fact -- {* additional fact afforded by the strong induction *}
- have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x, T1) # \<Gamma>1 \<sqsubseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact
- have a0: "atom x \<sharp> \<Gamma>1" by fact
- have a1: "valid \<Gamma>2" by fact
- have a2: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
-
- show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" sorry
-qed (auto) -- {* app case *}
-
-
-
-section {* Unbind and an Inconsistency Example *}
-
-text {*
- You might wonder why we had to discharge some proof
- obligations in order to obtain the stronger induction
- principle for the typing relation. (Remember for
- lambda terms this stronger induction principle is
- derived automatically.)
-
- This proof obligation is really needed, because if we
- assume universally a stronger induction principle, then
- in some cases we can derive false. This is "shown" below.
-*}
-
-
-inductive
- unbind :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixr "\<mapsto>" 60)
-where
- u_Var[intro]: "Var x \<mapsto> Var x"
-| u_App[intro]: "App t1 t2 \<mapsto> App t1 t2"
-| u_Lam[intro]: "t \<mapsto> t' \<Longrightarrow> Lam [x].t \<mapsto> t'"
-
-text {* It is clear that the following judgement holds. *}
-
-lemma unbind_simple:
- shows "Lam [x].Var x \<mapsto> Var x"
- by auto
-
-text {*
- Now lets derive the strong induction principle for unbind.
- The point is that we cannot establish the proof obligations,
- therefore we force Isabelle to accept them by using "sorry".
-*}
-
-equivariance unbind
-nominal_inductive unbind
- avoids u_Lam: "x"
- sorry
-
-text {*
- Using the stronger induction principle, we can establish
- th follwoing bogus property.
-*}
-
-lemma unbind_fake:
- fixes y::"name"
- assumes a: "t \<mapsto> t'"
- and b: "atom y \<sharp> t"
- shows "atom y \<sharp> t'"
-using a b
-proof (nominal_induct avoiding: y rule: unbind.strong_induct)
- case (u_Lam t t' x y)
- have ih: "atom y \<sharp> t \<Longrightarrow> atom y \<sharp> t'" by fact
- have asm: "atom y \<sharp> Lam [x]. t" by fact
- have fc: "atom x \<sharp> y" by fact
- then have in_eq: "x \<noteq> y" by (simp add: fresh_at_base)
- then have "atom y \<sharp> t" using asm by (simp add: lam.fresh)
- then show "atom y \<sharp> t'" using ih by simp
-qed
-
-text {*
- And from this follows the inconsitency:
-*}
-
-lemma
- shows "False"
-proof -
- have "atom x \<sharp> Lam [x]. Var x" by (simp add: lam.fresh)
- moreover
- have "Lam [x]. Var x \<mapsto> Var x" using unbind_simple by auto
- ultimately
- have "atom x \<sharp> Var x" using unbind_fake by auto
- then have "x \<noteq> x" by (simp add: lam.fresh fresh_at_base)
- then show "False" by simp
-qed
-
-end
--- a/Tutorial/Tutorial2s.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,354 +0,0 @@
-
-theory Tutorial2s
-imports Lambda
-begin
-
-section {* Height of a Lambda-Term *}
-
-text {*
- The theory Lambda defines the notions of capture-avoiding
- substitution and the height of lambda terms.
-*}
-
-thm height.simps
-thm subst.simps
-
-subsection {* EXERCISE 7 *}
-
-text {*
- Lets prove a property about the height of substitutions.
-
- Assume that the height of a lambda-term is always greater
- or equal to 1.
-*}
-
-lemma height_ge_one:
- shows "1 \<le> height t"
-by (induct t rule: lam.induct) (auto)
-
-text {* Remove the sorries *}
-
-theorem
- shows "height (t[x ::= t']) \<le> height t - 1 + height t'"
-proof (nominal_induct t avoiding: x t' rule: lam.strong_induct)
- case (Var y)
- have "1 \<le> height t'" using height_ge_one by simp
- then show "height (Var y[x ::= t']) \<le> height (Var y) - 1 + height t'" by simp
-next
- case (App t1 t2)
- have ih1: "height (t1[x::=t']) \<le> (height t1) - 1 + height t'"
- and ih2: "height (t2[x::=t']) \<le> (height t2) - 1 + height t'" by fact+
- then show "height ((App t1 t2)[x ::= t']) \<le> height (App t1 t2) - 1 + height t'" by simp
-next
- case (Lam y t1)
- have ih: "height (t1[x::=t']) \<le> height t1 - 1 + height t'" by fact
- moreover
- have vc: "atom y \<sharp> x" "atom y \<sharp> t'" by fact+ -- {* usual variable convention *}
- ultimately
- show "height ((Lam [y].t1)[x ::= t']) \<le> height (Lam [y].t1) - 1 + height t'" by simp
-qed
-
-
-section {* Types and the Weakening Lemma *}
-
-nominal_datatype ty =
- tVar "string"
-| tArr "ty" "ty" (infixr "\<rightarrow>" 100)
-
-
-text {*
- Having defined them as nominal datatypes gives us additional
- definitions and theorems that come with types (see below).
-
- We next define the type of typing contexts, a predicate that
- defines valid contexts (i.e. lists that contain only unique
- variables) and the typing judgement.
-*}
-
-type_synonym ty_ctx = "(name \<times> ty) list"
-
-inductive
- valid :: "ty_ctx \<Rightarrow> bool"
-where
- v1[intro]: "valid []"
-| v2[intro]: "\<lbrakk>valid \<Gamma>; atom x \<sharp> \<Gamma>\<rbrakk>\<Longrightarrow> valid ((x, T) # \<Gamma>)"
-
-
-inductive
- typing :: "ty_ctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60, 60, 60] 60)
-where
- t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
-| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"
-| t_Lam[intro]: "\<lbrakk>atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T1 \<rightarrow> T2"
-
-
-text {*
- The predicate atom x \<sharp> \<Gamma>, read as x fresh for \<Gamma>, is defined by
- Nominal Isabelle. It is defined for lambda-terms, products, lists etc.
- For example we have:
-*}
-
-lemma
- fixes x::"name"
- shows "atom x \<sharp> Lam [x].t"
- and "atom x \<sharp> (t1, t2) \<Longrightarrow> atom x \<sharp> App t1 t2"
- and "atom x \<sharp> Var y \<Longrightarrow> atom x \<sharp> y"
- and "\<lbrakk>atom x \<sharp> t1; atom x \<sharp> t2\<rbrakk> \<Longrightarrow> atom x \<sharp> (t1, t2)"
- and "\<lbrakk>atom x \<sharp> l1; atom x \<sharp> l2\<rbrakk> \<Longrightarrow> atom x \<sharp> (l1 @ l2)"
- and "atom x \<sharp> y \<Longrightarrow> x \<noteq> y"
- by (simp_all add: lam.fresh fresh_append fresh_at_base)
-
-text {*
- We can also prove that every variable is fresh for a type.
-*}
-
-lemma ty_fresh:
- fixes x::"name"
- and T::"ty"
- shows "atom x \<sharp> T"
-by (induct T rule: ty.induct)
- (simp_all add: ty.fresh pure_fresh)
-
-text {*
- In order to state the weakening lemma in a convenient form, we
- using the following abbreviation. Abbreviations behave like
- definitions, except that they are always automatically folded
- and unfolded.
-*}
-
-abbreviation
- "sub_ty_ctx" :: "ty_ctx \<Rightarrow> ty_ctx \<Rightarrow> bool" ("_ \<sqsubseteq> _" [60, 60] 60)
-where
- "\<Gamma>1 \<sqsubseteq> \<Gamma>2 \<equiv> \<forall>x. x \<in> set \<Gamma>1 \<longrightarrow> x \<in> set \<Gamma>2"
-
-
-subsection {* EXERCISE 8 *}
-
-text {*
- Fill in the details and give a proof of the weakening lemma.
-*}
-
-lemma
- assumes a: "\<Gamma>1 \<turnstile> t : T"
- and b: "valid \<Gamma>2"
- and c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
- shows "\<Gamma>2 \<turnstile> t : T"
-using a b c
-proof (induct arbitrary: \<Gamma>2)
- case (t_Var \<Gamma>1 x T)
- have a1: "valid \<Gamma>1" by fact
- have a2: "(x, T) \<in> set \<Gamma>1" by fact
- have a3: "valid \<Gamma>2" by fact
- have a4: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
-
- show "\<Gamma>2 \<turnstile> Var x : T" sorry
-next
- case (t_Lam x \<Gamma>1 T1 t T2)
- have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x, T1) # \<Gamma>1 \<sqsubseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact
- have a0: "atom x \<sharp> \<Gamma>1" by fact
- have a1: "valid \<Gamma>2" by fact
- have a2: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
-
- show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" sorry
-qed (auto) -- {* the application case is automatic*}
-
-
-text {*
- Despite the frequent claim that the weakening lemma is trivial,
- routine or obvious, the proof in the lambda-case does not go
- through smoothly. Painful variable renamings seem to be necessary.
- But the proof using renamings is horribly complicated (see below).
-*}
-
-equivariance valid
-equivariance typing
-
-lemma weakening_NOT_TO_BE_TRIED_AT_HOME:
- assumes a: "\<Gamma>1 \<turnstile> t : T"
- and b: "valid \<Gamma>2"
- and c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
- shows "\<Gamma>2 \<turnstile> t : T"
-using a b c
-proof (induct arbitrary: \<Gamma>2)
- -- {* lambda case *}
- case (t_Lam x \<Gamma>1 T1 t T2)
- have fc0: "atom x \<sharp> \<Gamma>1" by fact
- have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x, T1) # \<Gamma>1 \<sqsubseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact
- -- {* we choose a fresh variable *}
- obtain c::"name" where fc1: "atom c \<sharp> (x, t, \<Gamma>1, \<Gamma>2)" by (rule obtain_fresh)
- -- {* we alpha-rename the abstraction *}
- have "Lam [c].((c \<leftrightarrow> x) \<bullet> t) = Lam [x].t" using fc1
- by (auto simp add: lam.eq_iff Abs1_eq_iff flip_def)
- moreover
- -- {* we can then alpha rename the goal *}
- have "\<Gamma>2 \<turnstile> Lam [c].((c \<leftrightarrow> x) \<bullet> t) : T1 \<rightarrow> T2"
- proof -
- -- {* we need to establish *}
- -- {* (1) (x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) *}
- -- {* (2) valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)) *}
- have "(1)": "(x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)"
- proof -
- have "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
- then have "(c \<leftrightarrow> x) \<bullet> ((x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))" using fc0 fc1
- by (perm_simp) (simp add: flip_fresh_fresh)
- then show "(x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)" by (rule permute_boolE)
- qed
- moreover
- have "(2)": "valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))"
- proof -
- have "valid \<Gamma>2" by fact
- then show "valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))" using fc1
- by (auto simp add: fresh_permute_left atom_eqvt valid.eqvt)
- qed
- -- {* these two facts give us by induction hypothesis the following *}
- ultimately have "(x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) \<turnstile> t : T2" using ih by simp
- -- {* we now apply renamings to get to our goal *}
- then have "(c \<leftrightarrow> x) \<bullet> ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) \<turnstile> t : T2)" by (rule permute_boolI)
- then have "(c, T1) # \<Gamma>2 \<turnstile> ((c \<leftrightarrow> x) \<bullet> t) : T2" using fc1
- by (perm_simp) (simp add: flip_fresh_fresh ty_fresh)
- then show "\<Gamma>2 \<turnstile> Lam [c].((c \<leftrightarrow> x) \<bullet> t) : T1 \<rightarrow> T2" using fc1 by auto
- qed
- ultimately show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" by simp
-qed (auto) -- {* var and app cases, luckily, are automatic *}
-
-
-text {*
- The remedy is to use again a stronger induction principle that
- has the usual "variable convention" already build in. The
- following command derives this induction principle for the typing
- relation. (We shall explain what happens here later.)
-*}
-
-nominal_inductive typing
- avoids t_Lam: "x"
- unfolding fresh_star_def
- by (simp_all add: fresh_Pair lam.fresh ty_fresh)
-
-text {* Compare the two induction principles *}
-
-thm typing.induct
-thm typing.strong_induct -- {* note the additional assumption {atom x} \<sharp> c *}
-
-text {*
- We can use the stronger induction principle by replacing
- the line
-
- proof (induct arbitrary: \<Gamma>2)
-
- with
-
- proof (nominal_induct avoiding: \<Gamma>2 rule: typing.strong_induct)
-
- Try now the proof.
-*}
-
-subsection {* EXERCISE 9 *}
-
-lemma weakening:
- assumes a: "\<Gamma>1 \<turnstile> t : T"
- and b: "valid \<Gamma>2"
- and c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
- shows "\<Gamma>2 \<turnstile> t : T"
-using a b c
-proof (nominal_induct avoiding: \<Gamma>2 rule: typing.strong_induct)
- case (t_Var \<Gamma>1 x T) -- {* variable case is as before *}
- have "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
- moreover
- have "valid \<Gamma>2" by fact
- moreover
- have "(x, T)\<in> set \<Gamma>1" by fact
- ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto
-next
- case (t_Lam x \<Gamma>1 T1 t T2)
- have vc: "atom x \<sharp> \<Gamma>2" by fact -- {* additional fact afforded by the strong induction *}
- have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x, T1) # \<Gamma>1 \<sqsubseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact
- have a0: "atom x \<sharp> \<Gamma>1" by fact
- have a1: "valid \<Gamma>2" by fact
- have a2: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
- have "valid ((x, T1) # \<Gamma>2)" using a1 vc by auto
- moreover
- have "(x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # \<Gamma>2" using a2 by auto
- ultimately
- have "(x, T1) # \<Gamma>2 \<turnstile> t : T2" using ih by simp
- then show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" using vc by auto
-qed (auto) -- {* app case *}
-
-
-
-section {* Unbind and an Inconsistency Example *}
-
-text {*
- You might wonder why we had to discharge some proof
- obligations in order to obtain the stronger induction
- principle for the typing relation. (Remember for
- lambda terms this stronger induction principle is
- derived automatically.)
-
- This proof obligation is really needed, because if we
- assume universally a stronger induction principle, then
- in some cases we can derive false. This is "shown" below.
-*}
-
-
-inductive
- unbind :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixr "\<mapsto>" 60)
-where
- u_Var[intro]: "Var x \<mapsto> Var x"
-| u_App[intro]: "App t1 t2 \<mapsto> App t1 t2"
-| u_Lam[intro]: "t \<mapsto> t' \<Longrightarrow> Lam [x].t \<mapsto> t'"
-
-text {* It is clear that the following judgement holds. *}
-
-lemma unbind_simple:
- shows "Lam [x].Var x \<mapsto> Var x"
- by auto
-
-text {*
- Now lets derive the strong induction principle for unbind.
- The point is that we cannot establish the proof obligations,
- therefore we force Isabelle to accept them by using "sorry".
-*}
-
-equivariance unbind
-nominal_inductive unbind
- avoids u_Lam: "x"
- sorry
-
-text {*
- Using the stronger induction principle, we can establish
- th follwoing bogus property.
-*}
-
-lemma unbind_fake:
- fixes y::"name"
- assumes a: "t \<mapsto> t'"
- and b: "atom y \<sharp> t"
- shows "atom y \<sharp> t'"
-using a b
-proof (nominal_induct avoiding: y rule: unbind.strong_induct)
- case (u_Lam t t' x y)
- have ih: "atom y \<sharp> t \<Longrightarrow> atom y \<sharp> t'" by fact
- have asm: "atom y \<sharp> Lam [x]. t" by fact
- have fc: "atom x \<sharp> y" by fact
- then have in_eq: "x \<noteq> y" by (simp add: fresh_at_base)
- then have "atom y \<sharp> t" using asm by (simp add: lam.fresh)
- then show "atom y \<sharp> t'" using ih by simp
-qed
-
-text {*
- And from this follows the inconsitency:
-*}
-
-lemma
- shows "False"
-proof -
- have "atom x \<sharp> Lam [x]. Var x" by (simp add: lam.fresh)
- moreover
- have "Lam [x]. Var x \<mapsto> Var x" using unbind_simple by auto
- ultimately
- have "atom x \<sharp> Var x" using unbind_fake by auto
- then have "x \<noteq> x" by (simp add: lam.fresh fresh_at_base)
- then show "False" by simp
-qed
-
-end
--- a/Tutorial/Tutorial3.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,157 +0,0 @@
-theory Tutorial3
-imports Lambda
-begin
-
-section {* Formalising Barendregt's Proof of the Substitution Lemma *}
-
-text {*
- The substitution lemma is another theorem where the variable
- convention plays a crucial role.
-
- Barendregt's proof of this lemma needs in the variable case a
- case distinction. One way to do this in Isar is to use blocks.
- A block consist of some assumptions and reasoning steps
- enclosed in curly braces, like
-
- { \<dots>
- have "statement"
- have "last_statement_in_the_block"
- }
-
- Such a block may contain local assumptions like
-
- { assume "A"
- assume "B"
- \<dots>
- have "C" by \<dots>
- }
-
- Where "C" is the last have-statement in this block. The behaviour
- of such a block to the 'outside' is the implication
-
- A \<Longrightarrow> B \<Longrightarrow> C
-
- Now if we want to prove a property "smth" using the case-distinctions
- P1, P2 and P3 then we can use the following reasoning:
-
- { assume "P1"
- \<dots>
- have "smth"
- }
- moreover
- { assume "P2"
- \<dots>
- have "smth"
- }
- moreover
- { assume "P3"
- \<dots>
- have "smth"
- }
- ultimately have "smth" by blast
-
- The blocks establish the implications
-
- P1 \<Longrightarrow> smth
- P2 \<Longrightarrow> smth
- P3 \<Longrightarrow> smth
-
- If we know that P1, P2 and P3 cover all the cases, that is P1 \<or> P2 \<or> P3
- holds, then we have 'ultimately' established the property "smth"
-
-*}
-
-subsection {* Two preliminary facts *}
-
-lemma forget:
- shows "atom x \<sharp> t \<Longrightarrow> t[x ::= s] = t"
-by (nominal_induct t avoiding: x s rule: lam.strong_induct)
- (auto simp add: lam.fresh fresh_at_base)
-
-lemma fresh_fact:
- assumes a: "atom z \<sharp> s"
- and b: "z = y \<or> atom z \<sharp> t"
- shows "atom z \<sharp> t[y ::= s]"
-using a b
-by (nominal_induct t avoiding: z y s rule: lam.strong_induct)
- (auto simp add: lam.fresh fresh_at_base)
-
-
-
-section {* EXERCISE 10 *}
-
-text {*
- Fill in the cases 1.2 and 1.3 and the equational reasoning
- in the lambda-case.
-*}
-
-lemma
- assumes a: "x \<noteq> y"
- and b: "atom x \<sharp> L"
- shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
-using a b
-proof (nominal_induct M avoiding: x y N L rule: lam.strong_induct)
- case (Var z)
- have a1: "x \<noteq> y" by fact
- have a2: "atom x \<sharp> L" by fact
- show "Var z[x::=N][y::=L] = Var z[y::=L][x::=N[y::=L]]" (is "?LHS = ?RHS")
- proof -
- { -- {* Case 1.1 *}
- assume c1: "z = x"
- have "(1)": "?LHS = N[y::=L]" using c1 by simp
- have "(2)": "?RHS = N[y::=L]" using c1 a1 by simp
- have "?LHS = ?RHS" using "(1)" "(2)" by simp
- }
- moreover
- { -- {* Case 1.2 *}
- assume c2: "z = y" "z \<noteq> x"
-
- have "?LHS = ?RHS" sorry
- }
- moreover
- { -- {* Case 1.3 *}
- assume c3: "z \<noteq> x" "z \<noteq> y"
-
- have "?LHS = ?RHS" sorry
- }
- ultimately show "?LHS = ?RHS" by blast
- qed
-next
- case (Lam z M1) -- {* case 2: lambdas *}
- have ih: "\<lbrakk>x \<noteq> y; atom x \<sharp> L\<rbrakk> \<Longrightarrow> M1[x ::= N][y ::= L] = M1[y ::= L][x ::= N[y ::= L]]" by fact
- have a1: "x \<noteq> y" by fact
- have a2: "atom x \<sharp> L" by fact
- have fs: "atom z \<sharp> x" "atom z \<sharp> y" "atom z \<sharp> N" "atom z \<sharp> L" by fact+ -- {* !! *}
- then have b: "atom z \<sharp> N[y::=L]" by (simp add: fresh_fact)
- show "(Lam [z].M1)[x ::= N][y ::= L] = (Lam [z].M1)[y ::= L][x ::= N[y ::= L]]" (is "?LHS=?RHS")
- proof -
- have "?LHS = \<dots>" sorry
-
- also have "\<dots> = ?RHS" sorry
- finally show "?LHS = ?RHS" by simp
- qed
-next
- case (App M1 M2) -- {* case 3: applications *}
- then show "(App M1 M2)[x::=N][y::=L] = (App M1 M2)[y::=L][x::=N[y::=L]]" by simp
-qed
-
-text {*
- Again the strong induction principle enables Isabelle to find
- the proof of the substitution lemma completely automatically.
-*}
-
-lemma substitution_lemma_version:
- assumes asm: "x \<noteq> y" "atom x \<sharp> L"
- shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
- using asm
-by (nominal_induct M avoiding: x y N L rule: lam.strong_induct)
- (auto simp add: fresh_fact forget)
-
-subsection {* MINI EXERCISE *}
-
-text {*
- Compare and contrast Barendregt's reasoning and the
- formalised proofs.
-*}
-
-end
--- a/Tutorial/Tutorial3s.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,162 +0,0 @@
-
-theory Tutorial3s
-imports Lambda
-begin
-
-section {* Formalising Barendregt's Proof of the Substitution Lemma *}
-
-text {*
- The substitution lemma is another theorem where the variable
- convention plays a crucial role.
-
- Barendregt's proof of this lemma needs in the variable case a
- case distinction. One way to do this in Isar is to use blocks.
- A block consist of some assumptions and reasoning steps
- enclosed in curly braces, like
-
- { \<dots>
- have "statement"
- have "last_statement_in_the_block"
- }
-
- Such a block may contain local assumptions like
-
- { assume "A"
- assume "B"
- \<dots>
- have "C" by \<dots>
- }
-
- Where "C" is the last have-statement in this block. The behaviour
- of such a block to the 'outside' is the implication
-
- A \<Longrightarrow> B \<Longrightarrow> C
-
- Now if we want to prove a property "smth" using the case-distinctions
- P1, P2 and P3 then we can use the following reasoning:
-
- { assume "P1"
- \<dots>
- have "smth"
- }
- moreover
- { assume "P2"
- \<dots>
- have "smth"
- }
- moreover
- { assume "P3"
- \<dots>
- have "smth"
- }
- ultimately have "smth" by blast
-
- The blocks establish the implications
-
- P1 \<Longrightarrow> smth
- P2 \<Longrightarrow> smth
- P3 \<Longrightarrow> smth
-
- If we know that P1, P2 and P3 cover all the cases, that is P1 \<or> P2 \<or> P3
- holds, then we have 'ultimately' established the property "smth"
-
-*}
-
-subsection {* Two preliminary facts *}
-
-lemma forget:
- shows "atom x \<sharp> t \<Longrightarrow> t[x ::= s] = t"
-by (nominal_induct t avoiding: x s rule: lam.strong_induct)
- (auto simp add: lam.fresh fresh_at_base)
-
-lemma fresh_fact:
- assumes a: "atom z \<sharp> s"
- and b: "z = y \<or> atom z \<sharp> t"
- shows "atom z \<sharp> t[y ::= s]"
-using a b
-by (nominal_induct t avoiding: z y s rule: lam.strong_induct)
- (auto simp add: lam.fresh fresh_at_base)
-
-
-
-section {* EXERCISE 10 *}
-
-text {*
- Fill in the cases 1.2 and 1.3 and the equational reasoning
- in the lambda-case.
-*}
-
-lemma
- assumes a: "x \<noteq> y"
- and b: "atom x \<sharp> L"
- shows "M[x ::= N][y ::= L] = M[y ::= L][x ::= N[y ::= L]]"
-using a b
-proof (nominal_induct M avoiding: x y N L rule: lam.strong_induct)
- case (Var z)
- have a1: "x \<noteq> y" by fact
- have a2: "atom x \<sharp> L" by fact
- show "Var z[x::=N][y::=L] = Var z[y::=L][x::=N[y::=L]]" (is "?LHS = ?RHS")
- proof -
- { -- {* Case 1.1 *}
- assume c1: "z = x"
- have "(1)": "?LHS = N[y::=L]" using c1 by simp
- have "(2)": "?RHS = N[y::=L]" using c1 a1 by simp
- have "?LHS = ?RHS" using "(1)" "(2)" by simp
- }
- moreover
- { -- {* Case 1.2 *}
- assume c2: "z = y" "z \<noteq> x"
- have "(1)": "?LHS = L" using c2 by simp
- have "(2)": "?RHS = L[x::=N[y::=L]]" using c2 by simp
- have "(3)": "L[x::=N[y::=L]] = L" using a2 forget by simp
- have "?LHS = ?RHS" using "(1)" "(2)" "(3)" by simp
- }
- moreover
- { -- {* Case 1.3 *}
- assume c3: "z \<noteq> x" "z \<noteq> y"
- have "(1)": "?LHS = Var z" using c3 by simp
- have "(2)": "?RHS = Var z" using c3 by simp
- have "?LHS = ?RHS" using "(1)" "(2)" by simp
- }
- ultimately show "?LHS = ?RHS" by blast
- qed
-next
- case (Lam z M1) -- {* case 2: lambdas *}
- have ih: "\<lbrakk>x \<noteq> y; atom x \<sharp> L\<rbrakk> \<Longrightarrow> M1[x ::= N][y ::= L] = M1[y ::= L][x ::= N[y ::= L]]" by fact
- have a1: "x \<noteq> y" by fact
- have a2: "atom x \<sharp> L" by fact
- have fs: "atom z \<sharp> x" "atom z \<sharp> y" "atom z \<sharp> N" "atom z \<sharp> L" by fact+ -- {* !! *}
- then have b: "atom z \<sharp> N[y::=L]" by (simp add: fresh_fact)
- show "(Lam [z].M1)[x ::= N][y ::= L] = (Lam [z].M1)[y ::= L][x ::= N[y ::= L]]" (is "?LHS=?RHS")
- proof -
- have "?LHS = Lam [z].(M1[x ::= N][y ::= L])" using fs by simp
- also have "\<dots> = Lam [z].(M1[y ::= L][x ::= N[y ::= L]])" using ih a1 a2 by simp
- also have "\<dots> = (Lam [z].(M1[y ::= L]))[x ::= N[y ::= L]]" using b fs by simp
- also have "\<dots> = ?RHS" using fs by simp
- finally show "?LHS = ?RHS" by simp
- qed
-next
- case (App M1 M2) -- {* case 3: applications *}
- then show "(App M1 M2)[x::=N][y::=L] = (App M1 M2)[y::=L][x::=N[y::=L]]" by simp
-qed
-
-text {*
- Again the strong induction principle enables Isabelle to find
- the proof of the substitution lemma completely automatically.
-*}
-
-lemma substitution_lemma_version:
- assumes asm: "x \<noteq> y" "atom x \<sharp> L"
- shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
- using asm
-by (nominal_induct M avoiding: x y N L rule: lam.strong_induct)
- (auto simp add: fresh_fact forget)
-
-subsection {* MINI EXERCISE *}
-
-text {*
- Compare and contrast Barendregt's reasoning and the
- formalised proofs.
-*}
-
-end
--- a/Tutorial/Tutorial4.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,270 +0,0 @@
-
-theory Tutorial4
-imports Tutorial1 Tutorial2
-begin
-
-section {* The CBV Reduction Relation (Small-Step Semantics) *}
-
-
-inductive
- cbv :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<longrightarrow>cbv _" [60, 60] 60)
-where
- cbv1: "\<lbrakk>val v; atom x \<sharp> v\<rbrakk> \<Longrightarrow> App (Lam [x].t) v \<longrightarrow>cbv t[x ::= v]"
-| cbv2: "t \<longrightarrow>cbv t' \<Longrightarrow> App t t2 \<longrightarrow>cbv App t' t2"
-| cbv3: "t \<longrightarrow>cbv t' \<Longrightarrow> App t2 t \<longrightarrow>cbv App t2 t'"
-
-inductive
- "cbv_star" :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>cbv* _" [60, 60] 60)
-where
- cbvs1: "e \<longrightarrow>cbv* e"
-| cbvs2: "\<lbrakk>e1\<longrightarrow>cbv e2; e2 \<longrightarrow>cbv* e3\<rbrakk> \<Longrightarrow> e1 \<longrightarrow>cbv* e3"
-
-declare cbv.intros[intro] cbv_star.intros[intro]
-
-subsection {* EXERCISE 11 *}
-
-text {*
- Show that cbv* is transitive, by filling the gaps in the
- proof below.
-*}
-
-lemma cbvs3 [intro]:
- assumes a: "e1 \<longrightarrow>cbv* e2" "e2 \<longrightarrow>cbv* e3"
- shows "e1 \<longrightarrow>cbv* e3"
-using a
-proof (induct)
- case (cbvs1 e1)
- have asm: "e1 \<longrightarrow>cbv* e3" by fact
- show "e1 \<longrightarrow>cbv* e3" sorry
-next
- case (cbvs2 e1 e2 e3')
- have "e1 \<longrightarrow>cbv e2" by fact
- have "e2 \<longrightarrow>cbv* e3'" by fact
- have "e3' \<longrightarrow>cbv* e3 \<Longrightarrow> e2 \<longrightarrow>cbv* e3" by fact
- have "e3' \<longrightarrow>cbv* e3" by fact
-
- show "e1 \<longrightarrow>cbv* e3" sorry
-qed
-
-
-text {*
- In order to help establishing the property that the machine
- calculates a normal form that corresponds to the evaluation
- relation, we introduce the call-by-value small-step semantics.
-*}
-
-
-equivariance val
-equivariance cbv
-nominal_inductive cbv
- avoids cbv1: "x"
- unfolding fresh_star_def
- by (simp_all add: lam.fresh Abs_fresh_iff fresh_Pair fresh_fact)
-
-text {*
- In order to satisfy the vc-condition we have to formulate
- this relation with the additional freshness constraint
- atom x \<sharp> v. Although this makes the definition vc-ompatible, it
- makes the definition less useful. We can with a little bit of
- pain show that the more restricted rule is equivalent to the
- usual rule.
-*}
-
-lemma subst_rename:
- assumes a: "atom y \<sharp> t"
- shows "t[x ::= s] = ((y \<leftrightarrow> x) \<bullet> t)[y ::= s]"
-using a
-by (nominal_induct t avoiding: x y s rule: lam.strong_induct)
- (auto simp add: lam.fresh fresh_at_base)
-
-
-lemma better_cbv1 [intro]:
- assumes a: "val v"
- shows "App (Lam [x].t) v \<longrightarrow>cbv t[x ::= v]"
-proof -
- obtain y::"name" where fs: "atom y \<sharp> (x, t, v)" by (rule obtain_fresh)
- have "App (Lam [x].t) v = App (Lam [y].((y \<leftrightarrow> x) \<bullet> t)) v" using fs
- by (auto simp add: lam.eq_iff Abs1_eq_iff' flip_def fresh_Pair fresh_at_base)
- also have "\<dots> \<longrightarrow>cbv ((y \<leftrightarrow> x) \<bullet> t)[y ::= v]" using fs a cbv1 by auto
- also have "\<dots> = t[x ::= v]" using fs subst_rename[symmetric] by simp
- finally show "App (Lam [x].t) v \<longrightarrow>cbv t[x ::= v]" by simp
-qed
-
-
-
-subsection {* EXERCISE 12 *}
-
-text {*
- If more simple exercises are needed, then complete the following proof.
-*}
-
-lemma cbv_in_ctx:
- assumes a: "t \<longrightarrow>cbv t'"
- shows "E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>"
-using a
-proof (induct E)
- case Hole
- have "t \<longrightarrow>cbv t'" by fact
- show "\<box>\<lbrakk>t\<rbrakk> \<longrightarrow>cbv \<box>\<lbrakk>t'\<rbrakk>" sorry
-next
- case (CAppL E s)
- have ih: "t \<longrightarrow>cbv t' \<Longrightarrow> E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" by fact
- have "t \<longrightarrow>cbv t'" by fact
-
- show "(CAppL E s)\<lbrakk>t\<rbrakk> \<longrightarrow>cbv (CAppL E s)\<lbrakk>t'\<rbrakk>" sorry
-next
- case (CAppR s E)
- have ih: "t \<longrightarrow>cbv t' \<Longrightarrow> E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" by fact
- have a: "t \<longrightarrow>cbv t'" by fact
-
- show "(CAppR s E)\<lbrakk>t\<rbrakk> \<longrightarrow>cbv (CAppR s E)\<lbrakk>t'\<rbrakk>" sorry
-qed
-
-section {* EXERCISE 13 *}
-
-text {*
- The point of the cbv-reduction was that we can easily relatively
- establish the following property:
-*}
-
-lemma machine_implies_cbvs_ctx:
- assumes a: "<e, Es> \<mapsto> <e', Es'>"
- shows "(Es\<down>)\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* (Es'\<down>)\<lbrakk>e'\<rbrakk>"
-using a
-proof (induct)
- case (m1 t1 t2 Es)
-
- show "Es\<down>\<lbrakk>App t1 t2\<rbrakk> \<longrightarrow>cbv* ((CAppL \<box> t2) # Es)\<down>\<lbrakk>t1\<rbrakk>" sorry
-next
- case (m2 v t2 Es)
- have "val v" by fact
-
- show "((CAppL \<box> t2) # Es)\<down>\<lbrakk>v\<rbrakk> \<longrightarrow>cbv* (CAppR v \<box> # Es)\<down>\<lbrakk>t2\<rbrakk>" sorry
-next
- case (m3 v x t Es)
- have "val v" by fact
-
- show "(((CAppR (Lam [x].t) \<box>) # Es)\<down>)\<lbrakk>v\<rbrakk> \<longrightarrow>cbv* (Es\<down>)\<lbrakk>(t[x ::= v])\<rbrakk>" sorry
-qed
-
-text {*
- It is not difficult to extend the lemma above to
- arbitrary reductions sequences of the machine.
-*}
-
-lemma machines_implies_cbvs_ctx:
- assumes a: "<e, Es> \<mapsto>* <e', Es'>"
- shows "(Es\<down>)\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* (Es'\<down>)\<lbrakk>e'\<rbrakk>"
-using a machine_implies_cbvs_ctx
-by (induct) (blast)+
-
-text {*
- So whenever we let the machine start in an initial
- state and it arrives at a final state, then there exists
- a corresponding cbv-reduction sequence.
-*}
-
-corollary machines_implies_cbvs:
- assumes a: "<e, []> \<mapsto>* <e', []>"
- shows "e \<longrightarrow>cbv* e'"
-proof -
- have "[]\<down>\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* []\<down>\<lbrakk>e'\<rbrakk>"
- using a machines_implies_cbvs_ctx by blast
- then show "e \<longrightarrow>cbv* e'" by simp
-qed
-
-text {*
- We now want to relate the cbv-reduction to the evaluation
- relation. For this we need one auxiliary lemma about
- inverting the e_App rule.
-*}
-
-
-lemma e_App_elim:
- assumes a: "App t1 t2 \<Down> v"
- obtains x t v' where "t1 \<Down> Lam [x].t" "t2 \<Down> v'" "t[x::=v'] \<Down> v"
-using a by (cases) (auto simp add: lam.eq_iff lam.distinct)
-
-
-subsection {* EXERCISE 13 *}
-
-text {*
- Complete the first and second case in the
- proof below.
-*}
-
-lemma cbv_eval:
- assumes a: "t1 \<longrightarrow>cbv t2" "t2 \<Down> t3"
- shows "t1 \<Down> t3"
-using a
-proof(induct arbitrary: t3)
- case (cbv1 v x t t3)
- have a1: "val v" by fact
- have a2: "t[x ::= v] \<Down> t3" by fact
-
- show "App (Lam [x].t) v \<Down> t3" sorry
-next
- case (cbv2 t t' t2 t3)
- have ih: "\<And>t3. t' \<Down> t3 \<Longrightarrow> t \<Down> t3" by fact
- have "App t' t2 \<Down> t3" by fact
- then obtain x t'' v'
- where a1: "t' \<Down> Lam [x].t''"
- and a2: "t2 \<Down> v'"
- and a3: "t''[x ::= v'] \<Down> t3" by (rule e_App_elim)
-
- show "App t t2 \<Down> t3" sorry
-qed (auto elim!: e_App_elim)
-
-
-text {*
- Next we extend the lemma above to arbitray initial
- sequences of cbv-reductions.
-*}
-
-lemma cbvs_eval:
- assumes a: "t1 \<longrightarrow>cbv* t2" "t2 \<Down> t3"
- shows "t1 \<Down> t3"
-using a by (induct) (auto intro: cbv_eval)
-
-text {*
- Finally, we can show that if from a term t we reach a value
- by a cbv-reduction sequence, then t evaluates to this value.
-
- This proof is not by induction. So we have to set up the proof
- with
-
- proof -
-
- in order to prevent Isabelle from applying a default introduction
- rule.
-*}
-
-lemma cbvs_implies_eval:
- assumes a: "t \<longrightarrow>cbv* v"
- and b: "val v"
- shows "t \<Down> v"
-proof -
- have "v \<Down> v" using b eval_val by simp
- then show "t \<Down> v" using a cbvs_eval by auto
-qed
-
-section {* EXERCISE 15 *}
-
-text {*
- All facts tied together give us the desired property
- about machines: we know that a machines transitions
- correspond to cbvs transitions, and with the lemma
- above they correspond to an eval judgement.
-*}
-
-theorem machines_implies_eval:
- assumes a: "<t1, []> \<mapsto>* <t2, []>"
- and b: "val t2"
- shows "t1 \<Down> t2"
-proof -
-
- show "t1 \<Down> t2" sorry
-qed
-
-end
-
--- a/Tutorial/Tutorial4s.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,244 +0,0 @@
-theory Tutorial4
-imports Tutorial1
-begin
-
-section {* The CBV Reduction Relation (Small-Step Semantics) *}
-
-text {*
- In order to help establishing the property that the CK Machine
- calculates a nomrmalform that corresponds to the evaluation
- relation, we introduce the call-by-value small-step semantics.
-*}
-
-inductive
- cbv :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<longrightarrow>cbv _" [60, 60] 60)
-where
- cbv1: "\<lbrakk>val v; atom x \<sharp> v\<rbrakk> \<Longrightarrow> App (Lam [x].t) v \<longrightarrow>cbv t[x ::= v]"
-| cbv2[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> App t t2 \<longrightarrow>cbv App t' t2"
-| cbv3[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> App t2 t \<longrightarrow>cbv App t2 t'"
-
-equivariance val
-equivariance cbv
-nominal_inductive cbv
- avoids cbv1: "x"
- unfolding fresh_star_def
- by (simp_all add: lam.fresh Abs_fresh_iff fresh_Pair fresh_fact)
-
-text {*
- In order to satisfy the vc-condition we have to formulate
- this relation with the additional freshness constraint
- atom x \<sharp> v. Although this makes the definition vc-ompatible, it
- makes the definition less useful. We can with a little bit of
- pain show that the more restricted rule is equivalent to the
- usual rule.
-*}
-
-lemma subst_rename:
- assumes a: "atom y \<sharp> t"
- shows "t[x ::= s] = ((y \<leftrightarrow> x) \<bullet> t)[y ::= s]"
-using a
-by (nominal_induct t avoiding: x y s rule: lam.strong_induct)
- (auto simp add: lam.fresh fresh_at_base)
-
-
-lemma better_cbv1 [intro]:
- assumes a: "val v"
- shows "App (Lam [x].t) v \<longrightarrow>cbv t[x::=v]"
-proof -
- obtain y::"name" where fs: "atom y \<sharp> (x, t, v)" by (rule obtain_fresh)
- have "App (Lam [x].t) v = App (Lam [y].((y \<leftrightarrow> x) \<bullet> t)) v" using fs
- by (auto simp add: lam.eq_iff Abs1_eq_iff' flip_def fresh_Pair fresh_at_base)
- also have "\<dots> \<longrightarrow>cbv ((y \<leftrightarrow> x) \<bullet> t)[y ::= v]" using fs a cbv1 by auto
- also have "\<dots> = t[x ::= v]" using fs subst_rename[symmetric] by simp
- finally show "App (Lam [x].t) v \<longrightarrow>cbv t[x ::= v]" by simp
-qed
-
-text {*
- The transitive closure of the cbv-reduction relation:
-*}
-
-inductive
- "cbvs" :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>cbv* _" [60, 60] 60)
-where
- cbvs1[intro]: "e \<longrightarrow>cbv* e"
-| cbvs2[intro]: "\<lbrakk>e1\<longrightarrow>cbv e2; e2 \<longrightarrow>cbv* e3\<rbrakk> \<Longrightarrow> e1 \<longrightarrow>cbv* e3"
-
-lemma cbvs3 [intro]:
- assumes a: "e1 \<longrightarrow>cbv* e2" "e2 \<longrightarrow>cbv* e3"
- shows "e1 \<longrightarrow>cbv* e3"
-using a by (induct) (auto)
-
-
-subsection {* EXERCISE 8 *}
-
-text {*
- If more simple exercises are needed, then complete the following proof.
-*}
-
-lemma cbv_in_ctx:
- assumes a: "t \<longrightarrow>cbv t'"
- shows "E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>"
-using a
-proof (induct E)
- case Hole
- have "t \<longrightarrow>cbv t'" by fact
- then show "\<box>\<lbrakk>t\<rbrakk> \<longrightarrow>cbv \<box>\<lbrakk>t'\<rbrakk>" by simp
-next
- case (CAppL E s)
- have ih: "t \<longrightarrow>cbv t' \<Longrightarrow> E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" by fact
- moreover
- have "t \<longrightarrow>cbv t'" by fact
- ultimately
- have "E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" by simp
- then show "(CAppL E s)\<lbrakk>t\<rbrakk> \<longrightarrow>cbv (CAppL E s)\<lbrakk>t'\<rbrakk>" by auto
-next
- case (CAppR s E)
- have ih: "t \<longrightarrow>cbv t' \<Longrightarrow> E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" by fact
- moreover
- have a: "t \<longrightarrow>cbv t'" by fact
- ultimately
- have "E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" by simp
- then show "(CAppR s E)\<lbrakk>t\<rbrakk> \<longrightarrow>cbv (CAppR s E)\<lbrakk>t'\<rbrakk>" by auto
-qed
-
-section {* EXERCISE 9 *}
-
-text {*
- The point of the cbv-reduction was that we can easily relatively
- establish the follwoing property:
-*}
-
-lemma machine_implies_cbvs_ctx:
- assumes a: "<e, Es> \<mapsto> <e', Es'>"
- shows "(Es\<down>)\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* (Es'\<down>)\<lbrakk>e'\<rbrakk>"
-using a
-proof (induct)
- case (m1 t1 t2 Es)
-thm machine.intros thm cbv2
- have "Es\<down>\<lbrakk>App t1 t2\<rbrakk> = (Es\<down> \<odot> CAppL \<box> t2)\<lbrakk>t1\<rbrakk>" using ctx_compose ctx_composes.simps filling.simps by simp
- then show "Es\<down>\<lbrakk>App t1 t2\<rbrakk> \<longrightarrow>cbv* ((CAppL \<box> t2) # Es)\<down>\<lbrakk>t1\<rbrakk>" using cbvs.intros by simp
-next
- case (m2 v t2 Es)
- have "val v" by fact
- have "((CAppL \<box> t2) # Es)\<down>\<lbrakk>v\<rbrakk> = (CAppR v \<box> # Es)\<down>\<lbrakk>t2\<rbrakk>" using ctx_compose ctx_composes.simps filling.simps by simp
- then show "((CAppL \<box> t2) # Es)\<down>\<lbrakk>v\<rbrakk> \<longrightarrow>cbv* (CAppR v \<box> # Es)\<down>\<lbrakk>t2\<rbrakk>" using cbvs.intros by simp
-next
- case (m3 v x t Es)
- have aa: "val v" by fact
- have "(((CAppR (Lam [x].t) \<box>) # Es)\<down>)\<lbrakk>v\<rbrakk> = Es\<down>\<lbrakk>App (Lam [x]. t) v\<rbrakk>" using ctx_compose ctx_composes.simps filling.simps by simp
- then have "(((CAppR (Lam [x].t) \<box>) # Es)\<down>)\<lbrakk>v\<rbrakk> \<longrightarrow>cbv (Es\<down>)\<lbrakk>(t[x ::= v])\<rbrakk>" using better_cbv1[OF aa] cbv_in_ctx by simp
- then show "(((CAppR (Lam [x].t) \<box>) # Es)\<down>)\<lbrakk>v\<rbrakk> \<longrightarrow>cbv* (Es\<down>)\<lbrakk>(t[x ::= v])\<rbrakk>" using cbvs.intros by blast
-qed
-
-text {*
- It is not difficult to extend the lemma above to
- arbitrary reductions sequences of the CK machine. *}
-
-lemma machines_implies_cbvs_ctx:
- assumes a: "<e, Es> \<mapsto>* <e', Es'>"
- shows "(Es\<down>)\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* (Es'\<down>)\<lbrakk>e'\<rbrakk>"
-using a machine_implies_cbvs_ctx
-by (induct) (blast)+
-
-text {*
- So whenever we let the CL machine start in an initial
- state and it arrives at a final state, then there exists
- a corresponding cbv-reduction sequence.
-*}
-
-corollary machines_implies_cbvs:
- assumes a: "<e, []> \<mapsto>* <e', []>"
- shows "e \<longrightarrow>cbv* e'"
-proof -
- have "[]\<down>\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* []\<down>\<lbrakk>e'\<rbrakk>"
- using a machines_implies_cbvs_ctx by blast
- then show "e \<longrightarrow>cbv* e'" by simp
-qed
-
-text {*
- We now want to relate the cbv-reduction to the evaluation
- relation. For this we need two auxiliary lemmas.
-*}
-
-lemma eval_val:
- assumes a: "val t"
- shows "t \<Down> t"
-using a by (induct) (auto)
-
-
-lemma e_App_elim:
- assumes a: "App t1 t2 \<Down> v"
- obtains x t v' where "t1 \<Down> Lam [x].t" "t2 \<Down> v'" "t[x::=v'] \<Down> v"
-using a by (cases) (auto simp add: lam.eq_iff lam.distinct)
-
-
-subsection {* EXERCISE *}
-
-text {*
- Complete the first and second case in the
- proof below.
-*}
-
-lemma cbv_eval:
- assumes a: "t1 \<longrightarrow>cbv t2" "t2 \<Down> t3"
- shows "t1 \<Down> t3"
-using a
-proof(induct arbitrary: t3)
- case (cbv1 v x t t3)
- have a1: "val v" by fact
- have a2: "t[x ::= v] \<Down> t3" by fact
- have a3: "Lam [x].t \<Down> Lam [x].t" by auto
- have a4: "v \<Down> v" using a1 eval_val by auto
- show "App (Lam [x].t) v \<Down> t3" using a3 a4 a2 by auto
-next
- case (cbv2 t t' t2 t3)
- have ih: "\<And>t3. t' \<Down> t3 \<Longrightarrow> t \<Down> t3" by fact
- have "App t' t2 \<Down> t3" by fact
- then obtain x t'' v'
- where a1: "t' \<Down> Lam [x].t''"
- and a2: "t2 \<Down> v'"
- and a3: "t''[x ::= v'] \<Down> t3" by (rule e_App_elim)
- have "t \<Down> Lam [x].t''" using ih a1 by auto
- then show "App t t2 \<Down> t3" using a2 a3 by auto
-qed (auto elim!: e_App_elim)
-
-
-text {*
- Next we extend the lemma above to arbitray initial
- sequences of cbv-reductions. *}
-
-lemma cbvs_eval:
- assumes a: "t1 \<longrightarrow>cbv* t2" "t2 \<Down> t3"
- shows "t1 \<Down> t3"
-using a by (induct) (auto intro: cbv_eval)
-
-text {*
- Finally, we can show that if from a term t we reach a value
- by a cbv-reduction sequence, then t evaluates to this value.
-*}
-
-lemma cbvs_implies_eval:
- assumes a: "t \<longrightarrow>cbv* v" "val v"
- shows "t \<Down> v"
-using a
-by (induct) (auto intro: eval_val cbvs_eval)
-
-text {*
- All facts tied together give us the desired property about
- machines.
-*}
-
-theorem machines_implies_eval:
- assumes a: "<t1, []> \<mapsto>* <t2, []>"
- and b: "val t2"
- shows "t1 \<Down> t2"
-proof -
- have "t1 \<longrightarrow>cbv* t2" using a machines_implies_cbvs by simp
- then show "t1 \<Down> t2" using b cbvs_implies_eval by simp
-qed
-
-
-
-
-end
-
--- a/Tutorial/Tutorial5.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,218 +0,0 @@
-
-
-theory Tutorial5
-imports Tutorial4
-begin
-
-section {* Type-Preservation and Progress Lemma*}
-
-text {*
- The point of this tutorial is to prove the
- type-preservation and progress lemma. Since
- we now know that \<Down>, \<longrightarrow>cbv* and the machine
- correspond to each other, we only need to
- prove this property for one of them. We chose
- \<longrightarrow>cbv.
-
- First we need to establish two elimination
- properties and two auxiliary lemmas about contexts.
-*}
-
-
-lemma valid_elim:
- assumes a: "valid ((x, T) # \<Gamma>)"
- shows "atom x \<sharp> \<Gamma> \<and> valid \<Gamma>"
-using a by (cases) (auto)
-
-lemma valid_insert:
- assumes a: "valid (\<Delta> @ [(x, T)] @ \<Gamma>)"
- shows "valid (\<Delta> @ \<Gamma>)"
-using a
-by (induct \<Delta>)
- (auto simp add: fresh_append fresh_Cons dest!: valid_elim)
-
-lemma fresh_list:
- shows "atom y \<sharp> xs = (\<forall>x \<in> set xs. atom y \<sharp> x)"
-by (induct xs) (simp_all add: fresh_Nil fresh_Cons)
-
-lemma context_unique:
- assumes a1: "valid \<Gamma>"
- and a2: "(x, T) \<in> set \<Gamma>"
- and a3: "(x, U) \<in> set \<Gamma>"
- shows "T = U"
-using a1 a2 a3
-by (induct) (auto simp add: fresh_list fresh_Pair fresh_at_base)
-
-
-section {* EXERCISE 16 *}
-
-text {*
- Next we want to show the type substitution lemma. Unfortunately,
- we have to prove a slightly more general version of it, where
- the variable being substituted occurs somewhere inside the
- context.
-*}
-
-lemma type_substitution_aux:
- assumes a: "\<Delta> @ [(x, T')] @ \<Gamma> \<turnstile> e : T"
- and b: "\<Gamma> \<turnstile> e' : T'"
- shows "\<Delta> @ \<Gamma> \<turnstile> e[x ::= e'] : T"
-using a b
-proof (nominal_induct \<Gamma>'\<equiv>"\<Delta> @ [(x, T')] @ \<Gamma>" e T avoiding: x e' \<Delta> rule: typing.strong_induct)
- case (t_Var y T x e' \<Delta>)
- have a1: "valid (\<Delta> @ [(x, T')] @ \<Gamma>)" by fact
- have a2: "(y,T) \<in> set (\<Delta> @ [(x, T')] @ \<Gamma>)" by fact
- have a3: "\<Gamma> \<turnstile> e' : T'" by fact
-
- from a1 have a4: "valid (\<Delta> @ \<Gamma>)" by (rule valid_insert)
- { assume eq: "x = y"
-
- have "\<Delta> @ \<Gamma> \<turnstile> Var y[x ::= e'] : T" sorry
- }
- moreover
- { assume ineq: "x \<noteq> y"
- from a2 have "(y, T) \<in> set (\<Delta> @ \<Gamma>)" using ineq by simp
- then have "\<Delta> @ \<Gamma> \<turnstile> Var y[x ::= e'] : T" using ineq a4 by auto
- }
- ultimately show "\<Delta> @ \<Gamma> \<turnstile> Var y[x::=e'] : T" by blast
-next
- case (t_Lam y T1 t T2 x e' \<Delta>)
- have a1: "atom y \<sharp> e'" by fact
- have a2: "atom y \<sharp> \<Delta> @ [(x, T')] @ \<Gamma>" by fact
- have a3: "\<Gamma> \<turnstile> e' : T'" by fact
- have ih: "\<Gamma> \<turnstile> e' : T' \<Longrightarrow> ((y, T1) # \<Delta>) @ \<Gamma> \<turnstile> t [x ::= e'] : T2"
- using t_Lam(6)[of "(y, T1) # \<Delta>"] by auto
-
-
- show "\<Delta> @ \<Gamma> \<turnstile> (Lam [y]. t)[x ::= e'] : T1 \<rightarrow> T2" sorry
-next
- case (t_App t1 T1 T2 t2 x e' \<Delta>)
- have ih1: "\<Gamma> \<turnstile> e' : T' \<Longrightarrow> \<Delta> @ \<Gamma> \<turnstile> t1 [x ::= e'] : T1 \<rightarrow> T2" using t_App(2) by auto
- have ih2: "\<Gamma> \<turnstile> e' : T' \<Longrightarrow> \<Delta> @ \<Gamma> \<turnstile> t2 [x ::= e'] : T1" using t_App(4) by auto
- have a: "\<Gamma> \<turnstile> e' : T'" by fact
-
- show "\<Delta> @ \<Gamma> \<turnstile> App t1 t2 [x ::= e'] : T2" sorry
-qed
-
-text {*
- From this we can derive the usual version of the substitution
- lemma.
-*}
-
-corollary type_substitution:
- assumes a: "(x, T') # \<Gamma> \<turnstile> e : T"
- and b: "\<Gamma> \<turnstile> e' : T'"
- shows "\<Gamma> \<turnstile> e[x ::= e'] : T"
-using a b type_substitution_aux[of "[]"]
-by auto
-
-
-section {* Type Preservation *}
-
-text {*
- Finally we are in a position to establish the type preservation
- property. We just need the following two inversion rules for
- particualr typing instances.
-*}
-
-lemma t_App_elim:
- assumes a: "\<Gamma> \<turnstile> App t1 t2 : T"
- obtains T' where "\<Gamma> \<turnstile> t1 : T' \<rightarrow> T" "\<Gamma> \<turnstile> t2 : T'"
-using a
-by (cases) (auto simp add: lam.eq_iff lam.distinct)
-
-text {* we have not yet generated strong elimination rules *}
-lemma t_Lam_elim:
- assumes ty: "\<Gamma> \<turnstile> Lam [x].t : T"
- and fc: "atom x \<sharp> \<Gamma>"
- obtains T1 T2 where "T = T1 \<rightarrow> T2" "(x, T1) # \<Gamma> \<turnstile> t : T2"
-using ty fc
-apply(cases)
-apply(auto simp add: lam.eq_iff lam.distinct ty.eq_iff)
-apply(auto simp add: Abs1_eq_iff)
-apply(rotate_tac 3)
-apply(drule_tac p="(x \<leftrightarrow> xa)" in permute_boolI)
-apply(perm_simp)
-apply(auto simp add: flip_def swap_fresh_fresh ty_fresh)
-done
-
-
-section {* EXERCISE 17 *}
-
-text {*
- Fill in the gaps in the t_Lam case. You will need
- the type substitution lemma proved above.
-*}
-
-theorem cbv_type_preservation:
- assumes a: "t \<longrightarrow>cbv t'"
- and b: "\<Gamma> \<turnstile> t : T"
- shows "\<Gamma> \<turnstile> t' : T"
-using a b
-proof (nominal_induct avoiding: \<Gamma> T rule: cbv.strong_induct)
- case (cbv1 v x t \<Gamma> T)
- have fc: "atom x \<sharp> \<Gamma>" by fact
- have "\<Gamma> \<turnstile> App (Lam [x]. t) v : T" by fact
- then obtain T' where
- *: "\<Gamma> \<turnstile> Lam [x]. t : T' \<rightarrow> T" and
- **: "\<Gamma> \<turnstile> v : T'" by (rule t_App_elim)
- have "(x, T') # \<Gamma> \<turnstile> t : T" using * fc by (rule t_Lam_elim) (simp add: ty.eq_iff)
-
- show "\<Gamma> \<turnstile> t [x ::= v] : T " sorry
-qed (auto elim!: t_App_elim)
-
-text {*
- We can easily extend this to sequences of cbv* reductions.
-*}
-
-corollary cbvs_type_preservation:
- assumes a: "t \<longrightarrow>cbv* t'"
- and b: "\<Gamma> \<turnstile> t : T"
- shows "\<Gamma> \<turnstile> t' : T"
-using a b
-by (induct) (auto intro: cbv_type_preservation)
-
-text {*
- The type-preservation property for the machine and
- evaluation relation.
-*}
-
-theorem machine_type_preservation:
- assumes a: "<t, []> \<mapsto>* <t', []>"
- and b: "\<Gamma> \<turnstile> t : T"
- shows "\<Gamma> \<turnstile> t' : T"
-proof -
- have "t \<longrightarrow>cbv* t'" using a machines_implies_cbvs by simp
- then show "\<Gamma> \<turnstile> t' : T" using b cbvs_type_preservation by simp
-qed
-
-theorem eval_type_preservation:
- assumes a: "t \<Down> t'"
- and b: "\<Gamma> \<turnstile> t : T"
- shows "\<Gamma> \<turnstile> t' : T"
-proof -
- have "<t, []> \<mapsto>* <t', []>" using a eval_implies_machines by simp
- then show "\<Gamma> \<turnstile> t' : T" using b machine_type_preservation by simp
-qed
-
-text {* The Progress Property *}
-
-lemma canonical_tArr:
- assumes a: "[] \<turnstile> t : T1 \<rightarrow> T2"
- and b: "val t"
- obtains x t' where "t = Lam [x].t'"
-using b a by (induct) (auto)
-
-theorem progress:
- assumes a: "[] \<turnstile> t : T"
- shows "(\<exists>t'. t \<longrightarrow>cbv t') \<or> (val t)"
-using a
-by (induct \<Gamma>\<equiv>"[]::ty_ctx" t T)
- (auto elim: canonical_tArr simp add: val.simps)
-
-text {*
- Done! Congratulations!
-*}
-
-end
-
--- a/Tutorial/Tutorial6.thy Sat Dec 17 16:58:11 2011 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,49 +0,0 @@
-theory Tutorial6
-imports "../Nominal/Nominal2"
-begin
-
-
-section {* Type Schemes *}
-
-text {*
- Nominal2 can deal also with more complicated binding
- structures; for example the finite set of binders found
- in type schemes.
-*}
-
-atom_decl name
-
-nominal_datatype ty =
- Var "name"
-| Fun "ty" "ty" (infixr "\<rightarrow>" 100)
-and tys =
- All as::"name fset" ty::"ty" bind (set+) as in ty ("All _. _" [100, 100] 100)
-
-
-text {* Some alpha-equivalences *}
-
-lemma
- shows "All {|a, b|}. (Var a) \<rightarrow> (Var b) = All {|b, a|}. (Var a) \<rightarrow> (Var b)"
- unfolding ty_tys.eq_iff Abs_eq_iff alphas fresh_star_def
- by (auto simp add: ty_tys.eq_iff ty_tys.supp supp_at_base fresh_star_def intro: permute_zero)
-
-lemma
- shows "All {|a, b|}. (Var a) \<rightarrow> (Var b) = All {|a, b|}. (Var b) \<rightarrow> (Var a)"
- unfolding ty_tys.eq_iff Abs_eq_iff alphas fresh_star_def
- by (auto simp add: ty_tys.eq_iff ty_tys.supp supp_at_base fresh_star_def swap_atom)
- (metis permute_flip_at)
-
-lemma
- shows "All {|a, b, c|}. (Var a) \<rightarrow> (Var b) = All {|a, b|}. (Var a) \<rightarrow> (Var b)"
- unfolding ty_tys.eq_iff Abs_eq_iff alphas fresh_star_def
- by (auto simp add: ty_tys.eq_iff ty_tys.supp supp_at_base fresh_star_def intro: permute_zero)
-
-lemma
- assumes a: "a \<noteq> b"
- shows "All {|a, b|}. (Var a) \<rightarrow> (Var b) \<noteq> All {|c|}. (Var c) \<rightarrow> (Var c)"
- using a
- unfolding ty_tys.eq_iff Abs_eq_iff alphas fresh_star_def
- by (auto simp add: ty_tys.eq_iff ty_tys.supp supp_at_base fresh_star_def)
-
-
-end
\ No newline at end of file
Binary file Tutorial/slides.pdf has changed