--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/Atoms.thy Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,46 @@
+
+theory Atoms = Main + Swap + Terms:
+
+consts atms :: "(string \<times> string) list \<Rightarrow> string set"
+primrec
+"atms [] = {}"
+"atms (x#xs) = ((atms xs) \<union> {fst(x),snd(x)})"
+
+lemma [simp]:
+ "atms (xs@ys) = atms xs \<union> atms ys"
+apply(induct_tac xs, auto)
+done
+
+lemma [simp]:
+ "atms (rev pi) = atms pi"
+apply(induct_tac pi, simp_all)
+done
+
+lemma a_not_in_atms:
+ "a\<notin>atms pi \<longrightarrow> a = swapas pi a"
+apply(induct_tac pi, auto)
+done
+
+lemma swapas_pi_ineq_a:
+ "swapas pi a \<noteq> a \<longrightarrow> a\<in>atms pi"
+apply(case_tac "a\<in>atms pi")
+apply(simp)+
+apply(drule a_not_in_atms[THEN mp])
+apply(simp)
+done
+
+lemma a_ineq_swapas_pi:
+ "a \<noteq> swapas pi a \<longrightarrow> a\<in>atms pi"
+apply(case_tac "a\<in>atms pi")
+apply(simp_all add: a_not_in_atms)
+done
+
+lemma swapas_pi_in_atms[THEN mp]:
+ "a\<in>atms pi \<longrightarrow> swapas pi a\<in>atms pi"
+apply(subgoal_tac "\<forall>pi1. (a\<in>atms pi1) \<longrightarrow> swapas pi a\<in>(atms pi1 \<union> atms pi)")
+apply(force)
+apply(induct_tac pi)
+apply(auto)
+done
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/Disagreement.thy Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,147 @@
+
+theory Disagreement = Main + Swap + Atoms:
+
+consts
+ ds :: "(string \<times> string) list \<Rightarrow> (string \<times> string) list \<Rightarrow> string set"
+defs
+ ds_def: "ds xs ys \<equiv> { a . a \<in> (atms xs \<union> atms ys) \<and> (swapas xs a \<noteq> swapas ys a) }"
+
+lemma
+ ds_elem: "\<lbrakk>swapas pi a\<noteq>a\<rbrakk>\<Longrightarrow>a\<in>ds [] pi"
+apply(simp add: ds_def)
+apply(auto simp add: swapas_pi_ineq_a)
+done
+
+lemma
+ elem_ds: "\<lbrakk>a\<in>ds [] pi\<rbrakk>\<Longrightarrow>a\<noteq>swapas pi a"
+apply(simp add: ds_def)
+done
+
+lemma
+ ds_sym: "ds pi1 pi2 = ds pi2 pi1"
+apply(simp only: ds_def)
+apply(auto)
+done
+
+lemma
+ ds_trans: "c\<in>ds pi1 pi3\<longrightarrow>(c\<in>ds pi1 pi2 \<or> c\<in>ds pi2 pi3)"
+apply(auto)
+apply(simp only: ds_def)
+apply(auto)
+apply(drule a_not_in_atms[THEN mp])+
+apply(simp)
+apply(drule a_not_in_atms[THEN mp])
+apply(simp)
+apply(drule swapas_pi_ineq_a[THEN mp])
+apply(assumption)
+done
+
+lemma ds_cancel_pi_left:
+ "(c\<in> ds (pi1@pi) (pi2@pi)) \<longrightarrow> (swapas pi c\<in> ds pi1 pi2)"
+apply(simp only: ds_def)
+apply(auto)
+apply(simp_all add: swapas_append)
+apply(rule a_ineq_swapas_pi[THEN mp], clarify, drule a_not_in_atms[THEN mp], simp)+
+done
+
+lemma ds_cancel_pi_right:
+ "(swapas pi c\<in> ds pi1 pi2) \<longrightarrow> (c\<in> ds (pi1@pi) (pi2@pi))"
+apply(simp only: ds_def)
+apply(auto)
+apply(simp_all add: swapas_append)
+apply(rule a_ineq_swapas_pi[THEN mp],clarify,
+ drule a_not_in_atms[THEN mp],drule a_not_in_atms[THEN mp],simp)+
+done
+
+lemma ds_equality:
+ "(ds [] pi)-{a,swapas pi a} = (ds [] ((a,swapas pi a)#pi))-{swapas pi a}"
+apply(simp only: ds_def)
+apply(auto)
+done
+
+lemma ds_7:
+ "\<lbrakk>b\<noteq> swapas pi b;a\<in>ds [] ((b,swapas pi b)#pi)\<rbrakk>\<Longrightarrow>a\<in>ds [] pi"
+apply(simp only: ds_def)
+apply(case_tac "b=a")
+apply(auto)
+apply(rule swapas_pi_in_atms)
+apply(drule a_ineq_swapas_pi[THEN mp])
+apply(assumption)
+apply(drule sym)
+apply(drule swapas_rev_pi_a)
+apply(simp)
+apply(case_tac "swapas pi b = a")
+apply(auto)
+apply(drule sym)
+apply(drule swapas_rev_pi_a)
+apply(simp)
+done
+
+lemma ds_cancel_pi_front:
+ "ds (pi@pi1) (pi@pi2) = ds pi1 pi2"
+apply(simp only: ds_def)
+apply(auto)
+apply(simp_all add: swapas_append)
+apply(rule swapas_pi_ineq_a[THEN mp], clarify, drule a_not_in_atms[THEN mp], simp)+
+apply(drule swapas_rev_pi_a, simp)+
+done
+
+lemma ds_rev_pi_pi:
+ "ds ((rev pi1)@pi1) pi2 = ds [] pi2"
+apply(simp only: ds_def)
+apply(auto)
+apply(simp_all add: swapas_append)
+apply(drule a_ineq_swapas_pi[THEN mp], assumption)+
+done
+
+lemma ds_rev:
+ "ds [] ((rev pi1)@pi2) = ds pi1 pi2"
+apply(subgoal_tac "ds pi1 pi2 = ds ((rev pi1)@pi1) ((rev pi1)@pi2)")
+apply(simp add: ds_rev_pi_pi)
+apply(simp only: ds_cancel_pi_front)
+done
+
+lemma ds_acabbc:
+ "\<lbrakk>a\<noteq>b;b\<noteq>c;c\<noteq>a\<rbrakk>\<Longrightarrow>ds [] [(a, c), (a, b), (b, c)] = {a, b}"
+apply(simp only: ds_def)
+apply(auto)
+done
+
+lemma ds_baab:
+ "\<lbrakk>a\<noteq>b\<rbrakk>\<Longrightarrow>ds [] [(b, a), (a, b)] = {}"
+apply(simp only: ds_def)
+apply(auto)
+done
+
+lemma ds_abab:
+ "\<lbrakk>a\<noteq>b\<rbrakk>\<Longrightarrow>ds [] [(a, b), (a, b)] = {}"
+apply(simp only: ds_def)
+apply(auto)
+done
+
+(* disagreement set as list *)
+
+consts flatten :: "(string \<times> string)list \<Rightarrow> string list"
+primrec
+"flatten [] = []"
+"flatten (x#xs) = (fst x)#(snd x)#(flatten xs)"
+
+consts ds_list :: "(string \<times> string)list \<Rightarrow> (string \<times> string)list \<Rightarrow> string list"
+defs ds_list_def: "ds_list pi1 pi2 \<equiv> remdups ([x:(flatten (pi1@pi2)). x\<in>ds pi1 pi2])"
+
+lemma set_flatten_eq_atms:
+ "set (flatten pi) = atms pi"
+apply(induct_tac pi)
+apply(auto)
+done
+
+lemma ds_list_equ_ds:
+ "set (ds_list pi1 pi2) = ds pi1 pi2"
+apply(auto)
+apply(simp add: ds_list_def)
+apply(simp add: ds_list_def)
+apply(simp add: set_flatten_eq_atms)
+apply(simp add: ds_def)
+done
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/Equ.thy Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,156 @@
+
+
+theory Equ = Main + Terms + Fresh + PreEqu:
+
+lemma equ_refl:
+ "nabla\<turnstile>t\<approx>t"
+apply(induct_tac t)
+apply(auto simp add: ds_def)
+done
+
+
+lemma
+ equ_sym: "nabla\<turnstile>t1\<approx>t2 \<Longrightarrow> nabla\<turnstile>t2\<approx>t1" and
+ equ_trans: "\<lbrakk>nabla\<turnstile>t1\<approx>t2;nabla\<turnstile>t2\<approx>t3\<rbrakk> \<Longrightarrow> nabla\<turnstile>t1\<approx>t3" and
+ equ_add_pi: "nabla\<turnstile>t1\<approx>t2 \<Longrightarrow> nabla\<turnstile>swap pi t1\<approx>swap pi t2"
+apply(insert big)
+apply(best)+
+done
+
+
+lemma equ_dec_pi:
+ "nabla\<turnstile>swap pi t1\<approx>swap pi t2\<Longrightarrow>nabla\<turnstile>t1\<approx>t2"
+apply(drule_tac pi="(rev pi)" in equ_add_pi)
+apply(subgoal_tac "nabla\<turnstile>swap (rev pi) (swap pi t2)\<approx>t2")
+apply(drule_tac "t1.0"="swap (rev pi) (swap pi t1)" and
+ "t2.0"="swap (rev pi) (swap pi t2)" and
+ "t3.0"="t2" in equ_trans)
+apply(assumption)
+apply(subgoal_tac "nabla\<turnstile>t1\<approx>swap (rev pi) (swap pi t1)")
+apply(drule_tac "t1.0"="t1" and
+ "t2.0"="swap (rev pi) (swap pi t1)" and
+ "t3.0"="t2" in equ_trans)
+apply(assumption)+
+apply(rule equ_sym)
+apply(rule rev_pi_pi_equ)
+apply(rule rev_pi_pi_equ)
+done
+
+lemma equ_involutive_left:
+ "nabla\<turnstile>swap (rev pi) (swap pi t1)\<approx>t2 = nabla\<turnstile>t1\<approx>t2"
+apply(auto)
+apply(subgoal_tac "nabla \<turnstile> t1\<approx> swap (rev pi) (swap pi t1)")
+apply(drule_tac "t2.0"="swap (rev pi) (swap pi t1)" and
+ "t1.0"="t1" and "t3.0"="t2" in equ_trans)
+apply(assumption)+
+apply(rule equ_sym)
+apply(rule rev_pi_pi_equ)
+apply(subgoal_tac "nabla \<turnstile> swap (rev pi) (swap pi t1)\<approx>t1")
+apply(drule_tac "t1.0"="swap (rev pi) (swap pi t1)" and
+ "t2.0"="t1" and "t3.0"="t2" in equ_trans)
+apply(assumption)+
+apply(rule rev_pi_pi_equ)
+done
+
+lemma equ_pi_to_left:
+ "nabla\<turnstile>swap (rev pi) t1\<approx>t2 = nabla\<turnstile>t1\<approx> swap pi t2"
+apply(auto)
+apply(rule_tac pi="rev pi" in equ_dec_pi)
+apply(rule equ_sym)
+apply(simp only: equ_involutive_left)
+apply(rule equ_sym)
+apply(assumption)
+apply(drule_tac pi="rev pi" in equ_add_pi)
+apply(drule equ_sym)
+apply(simp only: equ_involutive_left)
+apply(drule equ_sym)
+apply(assumption)
+done
+
+lemma equ_pi_to_right:
+ "nabla\<turnstile>t1\<approx>swap (rev pi) t2 = nabla\<turnstile>swap pi t1\<approx>t2"
+apply(auto)
+apply(rule_tac pi="rev pi" in equ_dec_pi)
+apply(simp only: equ_involutive_left)
+apply(drule_tac pi="rev pi" in equ_add_pi)
+apply(simp only: equ_involutive_left)
+done
+
+lemma equ_involutive_right:
+ "nabla\<turnstile>t1\<approx>swap (rev pi) (swap pi t2) = nabla\<turnstile>t1\<approx>t2"
+apply(simp only: swap_append[THEN sym])
+apply(simp only: equ_pi_to_left[THEN sym])
+apply(simp)
+apply(simp only: swap_append)
+apply(simp only: equ_involutive_left)
+done
+
+lemma equ_pi1_pi2_add:
+ "(\<forall>a\<in> ds pi1 pi2. nabla\<turnstile>a\<sharp>t)\<longrightarrow>(nabla\<turnstile>swap pi1 t \<approx> swap pi2 t)"
+apply(rule impI)
+apply(simp only: equ_pi_to_right[THEN sym])
+apply(simp only: swap_append[THEN sym])
+apply(rule equ_pi_right[THEN spec, THEN mp])
+apply(auto)
+apply(simp only: ds_rev)
+done
+
+lemma pi_right_equ: "(nabla\<turnstile>t \<approx> swap pi t)\<longrightarrow>(\<forall>a\<in> ds [] pi. nabla\<turnstile>a\<sharp>t)"
+apply(insert pi_right_equ_help)
+apply(best)
+done
+
+lemma equ_pi1_pi2_dec:
+ "(nabla\<turnstile>swap pi1 t \<approx> swap pi2 t)\<longrightarrow>(\<forall>a\<in> ds pi1 pi2. nabla\<turnstile>a\<sharp>t)"
+apply(rule impI)
+apply(simp only: equ_pi_to_right[THEN sym])
+apply(simp only: swap_append[THEN sym])
+apply(drule pi_right_equ[THEN mp])
+apply(simp only: ds_rev)
+done
+
+lemma equ_weak:
+ "nabla1\<turnstile>t1\<approx>t2\<longrightarrow>(nabla1\<union>nabla2)\<turnstile>t1\<approx>t2"
+apply(rule impI)
+apply(erule equ.induct)
+apply(auto simp add: fresh_weak)
+done
+
+(* no term can be equal to one of its proper subterm *)
+lemma psub_trm_not_equ:
+ "\<forall> t2\<in>psub_trms t1. (\<not>(\<exists> pi. (nabla\<turnstile>t1\<approx>swap pi t2)))"
+apply(induct t1)
+apply(auto)
+apply(simp add: equ_pi_to_left[THEN sym])
+apply(ind_cases "nabla \<turnstile> Abst (swapas (rev pi) list) (swap (rev pi) trm) \<approx> t2")
+apply(simp_all)
+apply(drule abst_psub_trms)
+apply(drule_tac x="t2a" in bspec)
+apply(simp)
+apply(simp add: equ_pi_to_left[THEN sym] swap_append[THEN sym])
+apply(drule_tac x="rev (((swapas (rev pi) list, b) # rev pi))" in spec)
+apply(simp)
+apply(drule abst_psub_trms)
+apply(drule_tac x="t2a" in bspec)
+apply(simp)
+apply(simp)
+apply(simp add: equ_pi_to_left[THEN sym])
+apply(ind_cases "nabla \<turnstile> Paar (swap (rev pi) trm1) (swap (rev pi) trm2) \<approx> t2")
+apply(simp)
+apply(drule paar_psub_trms)
+apply(best)
+apply(simp add: equ_pi_to_left[THEN sym])
+apply(ind_cases "nabla \<turnstile> Paar (swap (rev pi) trm1) (swap (rev pi) trm2) \<approx> t2")
+apply(simp)
+apply(drule paar_psub_trms)
+apply(best)
+apply(simp add: equ_pi_to_left[THEN sym])
+apply(ind_cases "nabla \<turnstile> Func list (swap (rev pi) trm) \<approx> t2")
+apply(simp_all)
+apply(drule func_psub_trms)
+apply(best)
+done
+
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/Fresh.thy Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,90 @@
+
+
+theory Fresh = Main + Swap + Terms + Disagreement:
+
+types fresh_envs = "(string \<times> string)set"
+
+consts
+ fresh :: "(fresh_envs \<times> string \<times> trm) set"
+syntax
+ "_fresh_judge" :: "fresh_envs \<Rightarrow> string \<Rightarrow> trm \<Rightarrow> bool" (" _ \<turnstile> _ \<sharp> _" [80,80,80] 80)
+translations
+ "nabla \<turnstile> a \<sharp> t" \<rightleftharpoons> "(nabla,a,t) \<in> fresh"
+
+inductive fresh
+intros
+ fresh_abst_ab[intro!]: "\<lbrakk>(nabla\<turnstile>a\<sharp>t); a\<noteq>b\<rbrakk> \<Longrightarrow> (nabla\<turnstile>a\<sharp>Abst b t)"
+ fresh_abst_aa[intro!]: "(nabla\<turnstile>a\<sharp>Abst a t)"
+ fresh_unit[intro!]: "(nabla\<turnstile>a\<sharp>Unit)"
+ fresh_atom[intro!]: "a\<noteq>b \<Longrightarrow> (nabla\<turnstile>a\<sharp>Atom b)"
+ fresh_susp[intro!]: "(swapas (rev pi) a,X)\<in>nabla \<Longrightarrow> (nabla\<turnstile>a\<sharp>Susp pi X)"
+ fresh_paar[intro!]: "\<lbrakk>(nabla\<turnstile>a\<sharp>t1);(nabla\<turnstile>a\<sharp>t2)\<rbrakk> \<Longrightarrow> (nabla\<turnstile>a\<sharp>Paar t1 t2)"
+ fresh_func[intro!]: "(nabla\<turnstile>a\<sharp>t) \<Longrightarrow> (nabla\<turnstile>a\<sharp>Func F t)"
+
+lemma fresh_atom_elim[elim!]: "(nabla\<turnstile>a\<sharp>Atom b) \<Longrightarrow> a\<noteq>b"
+apply(ind_cases "(nabla \<turnstile>a\<sharp>Atom b)")
+apply(auto)
+done
+lemma fresh_susp_elim[elim!]: "(nabla\<turnstile>a\<sharp>Susp pi X) \<Longrightarrow> (swapas (rev pi) a,X)\<in>nabla"
+apply(ind_cases "(nabla\<turnstile>a\<sharp>Susp pi X)")
+apply(auto)
+done
+lemma fresh_paar_elim[elim!]: "(nabla\<turnstile>a\<sharp>Paar t1 t2) \<Longrightarrow> (nabla\<turnstile>a\<sharp>t1)\<and>(nabla \<turnstile>a\<sharp>t2)"
+apply(ind_cases "(nabla\<turnstile>a\<sharp>Paar t1 t2)")
+apply(auto)
+done
+lemma fresh_func_elim[elim!]: "(nabla\<turnstile>a\<sharp>Func F t) \<Longrightarrow> (nabla\<turnstile>a\<sharp>t)"
+apply(ind_cases "nabla\<turnstile>a\<sharp>Func F t")
+apply(auto)
+done
+lemma fresh_abst_ab_elim[elim!]: "\<lbrakk>nabla\<turnstile>a\<sharp>Abst b t;a\<noteq>b\<rbrakk> \<Longrightarrow> (nabla\<turnstile>a\<sharp>t)"
+apply(ind_cases "nabla\<turnstile>a\<sharp>Abst b t", auto)
+done
+
+lemma fresh_swap_left: "(nabla\<turnstile>a\<sharp>swap pi t) \<longrightarrow> (nabla\<turnstile>swapas (rev pi) a\<sharp>t)"
+apply(induct t)
+apply(simp_all)
+-- Abst
+apply(rule impI)
+apply(case_tac "swapas (rev pi) a = list")
+apply(force)
+apply(force dest!: fresh_abst_ab_elim)
+--Susp
+apply(force dest!: fresh_susp_elim simp add: swapas_append[THEN sym])
+--Unit
+apply(force)
+--Atom
+apply(force dest!: fresh_atom_elim)
+--Paar
+apply(force dest!: fresh_paar_elim)
+-- Func
+apply(force dest!: fresh_func_elim)
+done
+
+lemma fresh_swap_right: "(nabla\<turnstile>swapas (rev pi) a\<sharp>t) \<longrightarrow> (nabla\<turnstile>a\<sharp>swap pi t)"
+apply(induct t)
+apply(simp_all)
+-- Abst
+apply(rule impI)
+apply(case_tac "a = swapas pi list")
+apply(force)
+apply(force dest!: fresh_abst_ab_elim)
+--Susp
+apply(force dest!: fresh_susp_elim simp add: swapas_append[THEN sym])
+--Unit
+apply(force)
+--Atom
+apply(force dest!: fresh_atom_elim)
+--Paar
+apply(force dest!: fresh_paar_elim)
+-- Func
+apply(force dest!: fresh_func_elim)
+done
+
+lemma fresh_weak: "nabla1\<turnstile>a\<sharp>t\<longrightarrow>(nabla1\<union>nabla2)\<turnstile>a\<sharp>t"
+apply(rule impI)
+apply(erule fresh.induct)
+apply(auto)
+done
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/Mgu.thy Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,403 @@
+
+
+theory Mgu = Main + Terms + Fresh + Equ + Substs:
+
+(* unification problems *)
+
+syntax
+ "_equ_prob" :: "trm \<Rightarrow> trm \<Rightarrow> (trm\<times>trm)" ("_ \<approx>? _" [81,81] 81)
+ "_fresh_prob" :: "string \<Rightarrow> trm \<Rightarrow> (string\<times>trm)" ("_ \<sharp>? _" [81,81] 81)
+
+translations
+ "t1 \<approx>? t2" \<rightharpoonup> "(t1,t2)"
+ " a \<sharp>? t" \<rightharpoonup> "(a,t)"
+
+(* all solutions for a unification problem *)
+
+types
+ problem_type = "((trm\<times>trm)list) \<times> ((string\<times>trm)list)"
+ unifier_type = "fresh_envs \<times> substs"
+
+consts
+ U :: "problem_type \<Rightarrow> (unifier_type set)"
+defs all_solutions_def :
+ "U P \<equiv> {(nabla,s).
+ (\<forall> (t1,t2)\<in>set (fst P). nabla \<turnstile> subst s t1 \<approx> subst s t2) \<and>
+ (\<forall> (a,t)\<in>set (snd P). nabla \<turnstile> a \<sharp> subst s t) }"
+
+(* set of variables in unification problems *)
+
+consts
+ vars_fprobs :: "((string\<times>trm) list) \<Rightarrow> (string set)"
+ vars_eprobs :: "((trm\<times>trm)list) \<Rightarrow> (string set)"
+ vars_probs :: "problem_type \<Rightarrow> nat"
+primrec
+ "vars_fprobs [] = {}"
+ "vars_fprobs (x#xs) = (vars_trm (snd x))\<union>(vars_fprobs xs)"
+primrec
+ "vars_eprobs [] = {}"
+ "vars_eprobs (x#xs) = (vars_trm (snd x))\<union>(vars_trm (fst x))\<union>(vars_eprobs xs)"
+defs
+ vars_probs_def: "vars_probs P \<equiv> card((vars_fprobs (snd P))\<union>(vars_eprobs (fst P)))"
+
+
+(* most general unifier *)
+
+consts
+ mgu :: "problem_type \<Rightarrow> unifier_type \<Rightarrow> bool"
+defs mgu_def:
+ "mgu P unif \<equiv>
+ \<forall> (nabla,s1)\<in> U P. (\<exists> s2. (nabla\<Turnstile>(subst s2) (fst unif)) \<and>
+ (nabla\<Turnstile>subst (s2 \<bullet>(snd unif)) \<approx> subst s1))"
+
+(* idempotency of a unifier *)
+
+consts
+ idem :: "unifier_type \<Rightarrow> bool"
+defs
+ idem_def: "idem unif \<equiv> (fst unif)\<Turnstile> subst ((snd unif)\<bullet>(snd unif)) \<approx> subst (snd unif)"
+
+(* application of a substitution to a problem *)
+
+consts
+ apply_subst :: "substs \<Rightarrow> problem_type \<Rightarrow> problem_type"
+defs apply_subst_def:
+ "apply_subst s P \<equiv> (map (\<lambda>(t1,t2). (subst s t1 \<approx>? subst s t2)) (fst P),
+ map (\<lambda>(a,t). (a \<sharp>? (subst s t)) ) (snd P))"
+
+(* equality reductions *)
+
+consts
+ s_red :: "(problem_type \<times> substs \<times> problem_type) set"
+syntax
+ "_s_red" :: "problem_type \<Rightarrow> substs \<Rightarrow> problem_type \<Rightarrow> bool" ("_ \<turnstile> _ \<leadsto> _ " [80,80,80] 80)
+translations "P1 \<turnstile>sigma\<leadsto> P2" \<rightleftharpoons> "(P1,sigma,P2)\<in>s_red"
+inductive s_red
+intros
+ unit_sred[intro!,dest!]: "((Unit\<approx>?Unit)#xs,ys) \<turnstile>[]\<leadsto> (xs,ys)"
+ paar_sred[intro!,dest!]: "((Paar t1 t2\<approx>?Paar s1 s2)#xs,ys) \<turnstile>[]\<leadsto> ((t1\<approx>?s1)#(t2\<approx>?s2)#xs,ys)"
+ func_sred[intro!,dest!]: "((Func F t1\<approx>?Func F t2)#xs,ys) \<turnstile>[]\<leadsto> ((t1\<approx>?t2)#xs,ys)"
+ abst_aa_sred[intro!,dest!]: "((Abst a t1\<approx>?Abst a t2)#xs,ys) \<turnstile>[]\<leadsto> ((t1\<approx>?t2)#xs,ys)"
+ abst_ab_sred[intro!]: "a\<noteq>b\<Longrightarrow>
+ ((Abst a t1\<approx>?Abst b t2)#xs,ys) \<turnstile>[]\<leadsto> ((t1\<approx>?swap [(a,b)] t2)#xs,(a\<sharp>?t2)#ys)"
+ atom_sred[intro!,dest!]: "((Atom a\<approx>?Atom a)#xs,ys) \<turnstile>[]\<leadsto> (xs,ys)"
+ susp_sred[intro!,dest!]: "((Susp pi1 X\<approx>?Susp pi2 X)#xs,ys)
+ \<turnstile>[]\<leadsto> (xs,(map (\<lambda>a. a\<sharp>? Susp [] X) (ds_list pi1 pi2))@ys)"
+ var_1_sred[intro!]: "\<not>(occurs X t)\<Longrightarrow>((Susp pi X\<approx>?t)#xs,ys)
+ \<turnstile>[(X,swap (rev pi) t)]\<leadsto> apply_subst [(X,swap (rev pi) t)] (xs,ys)"
+ var_2_sred[intro!]: "\<not>(occurs X t)\<Longrightarrow>((t\<approx>?Susp pi X)#xs,ys)
+ \<turnstile>[(X,swap (rev pi) t)]\<leadsto> apply_subst [(X,swap (rev pi) t)] (xs,ys)"
+
+(* freshness reductions *)
+
+consts
+ c_red :: "(problem_type \<times> fresh_envs \<times> problem_type) set"
+syntax
+ "_c_red" :: "problem_type \<Rightarrow> fresh_envs \<Rightarrow> problem_type \<Rightarrow> bool" ("_ \<turnstile> _ \<rightarrow> _ " [80,80,80] 80)
+translations "P1 \<turnstile>nabla\<rightarrow> P2" \<rightleftharpoons> "(P1,nabla,P2)\<in>c_red"
+inductive c_red
+intros
+ unit_cred[intro!]: "([],(a \<sharp>? Unit)#xs) \<turnstile>{}\<rightarrow> ([],xs)"
+ paar_cred[intro!]: "([],(a \<sharp>? Paar t1 t2)#xs) \<turnstile>{}\<rightarrow> ([],(a\<sharp>?t1)#(a\<sharp>?t2)#xs)"
+ func_cred[intro!]: "([],(a \<sharp>? Func F t)#xs) \<turnstile>{}\<rightarrow> ([],(a\<sharp>?t)#xs)"
+ abst_aa_cred[intro!]: "([],(a \<sharp>? Abst a t)#xs) \<turnstile>{}\<rightarrow> ([],xs)"
+ abst_ab_cred[intro!]: "a\<noteq>b\<Longrightarrow>([],(a \<sharp>? Abst b t)#xs) \<turnstile>{}\<rightarrow> ([],(a\<sharp>?t)#xs)"
+ atom_cred[intro!]: "a\<noteq>b\<Longrightarrow>([],(a \<sharp>? Atom b)#xs) \<turnstile>{}\<rightarrow> ([],xs)"
+ susp_cred[intro!]: "([],(a \<sharp>? Susp pi X)#xs) \<turnstile>{((swapas (rev pi) a),X)}\<rightarrow> ([],xs)"
+
+(* unification reduction sequence *)
+
+consts
+ red_plus :: "(problem_type \<times> unifier_type \<times> problem_type) set"
+syntax
+ red_plus :: "problem_type \<Rightarrow> unifier_type \<Rightarrow> problem_type \<Rightarrow> bool" ("_ \<Turnstile> _ \<Rightarrow> _ " [80,80,80] 80)
+
+translations "P1 \<Turnstile>(nabla,s)\<Rightarrow> P2" \<rightleftharpoons> "(P1,(nabla,s),P2)\<in>red_plus"
+inductive red_plus
+intros
+ sred_single[intro!]: "\<lbrakk>P1\<turnstile>s1\<leadsto>P2\<rbrakk>\<Longrightarrow>P1\<Turnstile>({},s1)\<Rightarrow>P2"
+ cred_single[intro!]: "\<lbrakk>P1\<turnstile>nabla1\<rightarrow>P2\<rbrakk>\<Longrightarrow>P1\<Turnstile>(nabla1,[])\<Rightarrow>P2"
+ sred_step[intro!]: "\<lbrakk>P1\<turnstile>s1\<leadsto>P2;P2\<Turnstile>(nabla2,s2)\<Rightarrow>P3\<rbrakk>\<Longrightarrow>P1\<Turnstile>(nabla2,(s2\<bullet>s1))\<Rightarrow>P3"
+ cred_step[intro!]: "\<lbrakk>P1\<turnstile>nabla1\<rightarrow>P2;P2\<Turnstile>(nabla2,[])\<Rightarrow>P3\<rbrakk>\<Longrightarrow>P1\<Turnstile>(nabla2\<union>nabla1,[])\<Rightarrow>P3"
+
+lemma mgu_idem:
+ "\<lbrakk>(nabla1,s1)\<in>U P;
+ \<forall>(nabla2,s2)\<in> U P. nabla2\<Turnstile>(subst s2) nabla1 \<and> nabla2\<Turnstile>subst(s2 \<bullet> s1)\<approx>subst s2 \<rbrakk>\<Longrightarrow>
+ mgu P (nabla1,s1)\<and>idem (nabla1,s1)"
+apply(rule conjI)
+apply(simp only: mgu_def)
+apply(rule ballI)
+apply(simp)
+apply(drule_tac x="x" in bspec)
+apply(assumption)
+apply(force)
+apply(drule_tac x="(nabla1,s1)" in bspec)
+apply(assumption)
+apply(simp add: idem_def)
+done
+
+lemma problem_subst_comm: "((nabla,s2)\<in>U (apply_subst s1 P)) = ((nabla,(s2\<bullet>s1))\<in>U P)"
+apply(simp add: all_solutions_def apply_subst_def)
+apply(auto)
+apply(drule_tac x="(a,b)" in bspec, assumption, simp add: subst_comp_expand)+
+done
+
+lemma P1_to_P2_sred:
+ "\<lbrakk>(nabla1,s1)\<in>U P1; P1 \<turnstile>s2\<leadsto> P2 \<rbrakk>\<Longrightarrow>((nabla1,s1)\<in>U P2) \<and> (nabla1\<Turnstile>subst (s1\<bullet>s2)\<approx>subst s1)"
+apply(ind_cases "P1 \<turnstile>s2\<leadsto> P2")
+apply(simp_all)
+--Unit
+apply(force intro!: equ_refl simp add: all_solutions_def ext_subst_def subst_equ_def subst_susp)
+--Paar
+apply(simp add: all_solutions_def ext_subst_def subst_equ_def subst_susp)
+apply(force intro!: equ_refl dest!: equ_paar_elim)
+--Func
+apply(simp add: all_solutions_def ext_subst_def subst_equ_def subst_susp)
+apply(force intro!: equ_refl dest!: equ_func_elim)
+--Abst.aa
+apply(simp add: all_solutions_def ext_subst_def subst_equ_def subst_susp)
+apply(force intro!: equ_refl dest!: equ_abst_aa_elim)
+--Abst.ab
+apply(simp add: all_solutions_def ext_subst_def subst_equ_def subst_susp)
+apply(force intro!: equ_refl dest!: equ_abst_ab_elim simp add: subst_swap_comm)
+--Atom
+apply(simp add: all_solutions_def ext_subst_def subst_equ_def subst_susp)
+apply(force intro!: equ_refl)
+--Susp
+apply(rule conjI)
+apply(auto)
+apply(simp add: all_solutions_def)
+apply(erule conjE)+
+apply(simp add: ds_list_equ_ds)
+apply(simp only: subst_susp)
+apply(drule equ_pi1_pi2_dec[THEN mp])
+apply(auto)
+apply(drule_tac x="aa" in bspec)
+apply(assumption)
+apply(simp add: subst_susp)
+apply(simp add: subst_equ_def subst_susp)
+apply(rule ballI)
+apply(rule equ_refl)
+--Var.one
+apply(drule_tac "t2.1"="swap (rev pi) t" in subst_not_occurs[THEN mp])
+apply(simp only: problem_subst_comm)
+apply(simp add: all_solutions_def ext_subst_def subst_equ_def)
+apply(rule conjI)
+apply(rule ballI)
+apply(erule conjE)+
+apply(drule unif_1)
+apply(clarify)
+apply(drule_tac x="(a,b)" in bspec)
+apply(assumption)
+apply(simp)
+apply(simp add: unif_2a)
+apply(erule conjE)+
+apply(drule unif_1)
+apply(rule ballI)
+apply(clarify)
+apply(drule_tac x="(a,b)" in bspec)
+apply(assumption)
+apply(simp)
+apply(simp add: unif_2b)
+apply(rule unif_1)
+apply(simp add: all_solutions_def)
+--Var.two
+apply(drule_tac "t2.1"="swap (rev pi) t" in subst_not_occurs[THEN mp])
+apply(simp only: problem_subst_comm)
+apply(simp add: all_solutions_def ext_subst_def subst_equ_def)
+apply(auto)
+apply(drule_tac x="(a,b)" in bspec)
+apply(assumption)
+apply(simp)
+apply(drule equ_sym)
+apply(drule unif_1)
+apply(simp add: unif_2a)
+apply(drule_tac x="(a,b)" in bspec)
+apply(assumption)
+apply(simp)
+apply(drule equ_sym)
+apply(drule unif_1)
+apply(simp add: unif_2b)
+apply(rule unif_1)
+apply(rule equ_sym)
+apply(simp add: all_solutions_def)
+done
+
+lemma P1_from_P2_sred: "\<lbrakk>(nabla1,s1)\<in>U P2; P1\<turnstile>s2\<leadsto>P2\<rbrakk>\<Longrightarrow>(nabla1,s1\<bullet>s2)\<in>U P1"
+apply(ind_cases "P1 \<turnstile>s2\<leadsto> P2")
+--Susp.Paar.Func.Abst.aa
+apply(simp add: all_solutions_def, force)
+apply(simp add: all_solutions_def, force)
+apply(simp add: all_solutions_def, force)
+apply(simp add: all_solutions_def, force)
+--Abst.ab
+apply(simp only: all_solutions_def)
+apply(force simp add: subst_swap_comm)
+--Atom
+apply(simp only: all_solutions_def, force)
+--Susp
+apply(simp)
+apply(auto)
+apply(simp add: all_solutions_def)
+apply(simp add: ds_list_equ_ds)
+apply(subgoal_tac "nabla1\<turnstile>(swap pi1 (subst s1 (Susp [] X)))\<approx>(swap pi2 (subst s1 (Susp [] X)))")
+apply(simp add: subst_susp subst_swap_comm)
+apply(simp add: subst_susp subst_swap_comm)
+apply(rule equ_pi1_pi2_add[THEN mp])
+apply(drule conjunct2)
+apply(auto)
+apply(drule_tac x="(a,Susp [] X)" in bspec)
+apply(auto)
+apply(simp add: subst_susp)
+--Var.one
+apply(simp only: problem_subst_comm)
+apply(simp only: all_solutions_def)
+apply(simp)
+apply(simp only: subst_comp_expand)
+apply(subgoal_tac "subst [(X, swap (rev pi) t)] t = t")--A
+apply(simp add: subst_susp)
+apply(simp only: subst_swap_comm)
+apply(simp only: equ_pi_to_right[THEN sym])
+apply(simp only: equ_involutive_right)
+apply(rule equ_refl)
+--A
+apply(rule subst_not_occurs[THEN mp])
+apply(assumption)
+--Var.two
+apply(simp only: problem_subst_comm)
+apply(simp only: all_solutions_def)
+apply(simp)
+apply(simp only: subst_comp_expand)
+apply(subgoal_tac "subst [(X, swap (rev pi) t)] t = t")--B
+apply(simp add: subst_susp)
+apply(simp only: subst_swap_comm)
+apply(simp only: equ_pi_to_left[THEN sym])
+apply(simp only: equ_involutive_left)
+apply(rule equ_refl)
+--B
+apply(rule subst_not_occurs[THEN mp])
+apply(assumption)
+done
+
+lemma P1_to_P2_cred:
+ "\<lbrakk>(nabla1,s1)\<in>U P1; P1 \<turnstile>nabla2\<rightarrow> P2 \<rbrakk>\<Longrightarrow>((nabla1,s1)\<in>U P2) \<and> (nabla1\<Turnstile>(subst s1) nabla2)"
+apply(ind_cases " P1\<turnstile>nabla2\<rightarrow>P2")
+apply(simp_all)
+apply(auto simp add: ext_subst_def all_solutions_def)
+apply(rule fresh_swap_left[THEN mp])
+apply(simp add: subst_swap_comm[THEN sym] subst_susp)
+done
+
+lemma P1_from_P2_cred:
+ "\<lbrakk>(nabla1,s1)\<in>U P2; P1 \<turnstile>nabla2\<rightarrow> P2; nabla3\<Turnstile>(subst s1) nabla2\<rbrakk>\<Longrightarrow>(nabla1\<union>nabla3,s1)\<in>U P1"
+apply(ind_cases "P1 \<turnstile>nabla2\<rightarrow> P2")
+apply(simp_all)
+apply(auto simp add: ext_subst_def all_solutions_def fresh_weak)
+apply(simp add: subst_susp)
+apply(rule fresh_swap_right[THEN mp])
+apply(drule_tac "nabla2.1"="nabla1" in fresh_weak[THEN mp])
+apply(subgoal_tac "nabla3 \<union> nabla1=nabla1 \<union> nabla3")--A
+apply(simp)
+apply(rule Un_commute)
+done
+
+lemma P1_to_P2_red_plus: "\<lbrakk>P1 \<Turnstile>(nabla,s)\<Rightarrow>P2\<rbrakk>\<Longrightarrow> (nabla1,s1)\<in>U P1 \<longrightarrow>
+ ((nabla1,s1)\<in>U P2)\<and>(nabla1\<Turnstile>subst (s1\<bullet>s)\<approx>subst s1)\<and>(nabla1\<Turnstile>(subst s1) nabla)"
+apply(erule red_plus.induct)
+-- sred
+apply(rule impI)
+apply(drule_tac "P2.0"="P2" and "s2.0"="s1a" in P1_to_P2_sred)
+apply(force)
+apply(rule conjI, force)+
+apply(force simp add: ext_subst_def)
+-- cred
+apply(rule impI)
+apply(drule_tac "P2.0"="P2" and "nabla2.0"="nabla1a" in P1_to_P2_cred)
+apply(assumption)
+apply(force intro!: equ_refl simp add: subst_equ_def)
+-- sred
+apply(rule impI)
+apply(drule_tac "P2.0"="P2" and "s2.0"="s1a" in P1_to_P2_sred)
+apply(assumption)
+apply(erule conjE)+
+apply(rule conjI)
+apply(force)
+apply(rule conjI)
+apply(drule mp)
+apply(assumption)
+apply(erule conjE)+
+apply(rule_tac "s2.0"="((s1\<bullet>s2)\<bullet>s1a)" in subst_trans)
+apply(simp only: subst_assoc subst_equ_def)
+apply(rule ballI)
+apply(rule equ_refl)
+apply(rule_tac "s2.0"="(s1\<bullet>s1a)" in subst_trans)
+apply(rule subst_cancel_right)
+apply(assumption)
+apply(assumption)
+apply(force)
+-- cred
+apply(rule impI)
+apply(drule_tac "P2.0"="P2" and "nabla2.0"="nabla1a" in P1_to_P2_cred)
+apply(auto simp add: ext_subst_def)
+done
+
+lemma P1_from_P2_red_plus: "\<lbrakk>P1 \<Turnstile>(nabla,s)\<Rightarrow>P2\<rbrakk>\<Longrightarrow>(nabla1,s1)\<in>U P2\<longrightarrow>
+ nabla3\<Turnstile>(subst s1)(nabla)\<longrightarrow>(nabla1\<union>nabla3,(s1\<bullet>s))\<in>U P1"
+apply(erule red_plus.induct)
+-- sred
+apply(rule impI)+
+apply(drule_tac "P1.0"="P1" and "s2.0"="s1a" in P1_from_P2_sred)
+apply(assumption)
+apply(force simp only: all_solutions_def equ_weak fresh_weak)
+-- cred
+apply(rule impI)+
+apply(drule_tac "P1.0"="P1" and "nabla3.0"="nabla3" and "nabla2.0"="nabla1a" in P1_from_P2_cred)
+apply(assumption)+
+apply(simp add: all_solutions_def)
+-- sred
+apply(rule impI)+
+apply(simp)
+apply(drule_tac "P1.0"="P1" and "P2.0"="P2" and "s2.0"="s1a" in P1_from_P2_sred)
+apply(assumption)
+apply(simp add: all_solutions_def subst_assoc)
+-- cred
+apply(rule impI)+
+apply(subgoal_tac "nabla3 \<Turnstile> (subst s1) nabla2")--A
+apply(simp)
+apply(drule_tac "P1.0"="P1" and "P2.0"="P2" and
+ "nabla2.0"="nabla1a" and "nabla3.0"="nabla3" in P1_from_P2_cred)
+apply(assumption)
+apply(simp)
+apply(simp add: ext_subst_def)
+apply(subgoal_tac "nabla1 \<union> nabla3 \<union> nabla3=nabla1 \<union> nabla3")--A
+apply(simp)
+--A
+apply(force)
+--B
+apply(simp add: ext_subst_def)
+done
+
+lemma mgu: "\<lbrakk>P \<Turnstile>(nabla,s)\<Rightarrow>([],[])\<rbrakk>\<Longrightarrow> mgu P (nabla,s) \<and> idem (nabla,s)"
+apply(frule_tac "nabla3.2"="nabla" and "nabla2"="nabla" and
+ "s1.2"="[]" and "nabla1.2"="{}"
+ in P1_from_P2_red_plus[THEN mp,THEN mp])
+apply(force simp add: all_solutions_def)
+apply(force simp add: ext_subst_def)
+apply(rule mgu_idem)
+apply(simp add: all_solutions_def)
+apply(rule ballI)
+apply(clarify)
+apply(drule_tac "nabla1.0"="a" and "s1.0"="b"in P1_to_P2_red_plus)
+apply(simp)
+done
+
+end
+
+
+
+
+
+
+
+
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/PreEqu.thy Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,579 @@
+
+
+theory PreEqu = Main + Swap + Terms + Disagreement + Fresh:
+
+consts
+ equ :: "(fresh_envs \<times> trm \<times> trm) set"
+
+syntax
+ "_equ_judge" :: "fresh_envs \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> bool" (" _ \<turnstile> _ \<approx> _" [80,80,80] 80)
+translations
+ "nabla \<turnstile> t1 \<approx> t2" \<rightleftharpoons> "(nabla,t1,t2) \<in> equ"
+
+inductive equ
+intros
+equ_abst_ab[intro!]: "\<lbrakk>a\<noteq>b;(nabla \<turnstile> a \<sharp> t2);(nabla \<turnstile> t1 \<approx> (swap [(a,b)] t2))\<rbrakk>
+ \<Longrightarrow> (nabla \<turnstile> Abst a t1 \<approx> Abst b t2)"
+equ_abst_aa[intro!]: "(nabla \<turnstile> t1 \<approx> t2) \<Longrightarrow> (nabla \<turnstile> Abst a t1 \<approx> Abst a t2)"
+equ_unit[intro!]: "(nabla \<turnstile> Unit \<approx> Unit)"
+equ_atom[intro!]: "a=b\<Longrightarrow>nabla \<turnstile> Atom a \<approx> Atom b"
+equ_susp[intro!]: "(\<forall> c \<in> ds pi1 pi2. (c,X) \<in> nabla) \<Longrightarrow> (nabla \<turnstile> Susp pi1 X \<approx> Susp pi2 X)"
+equ_paar[intro!]: "\<lbrakk>(nabla \<turnstile> t1\<approx>t2);(nabla \<turnstile> s1\<approx>s2)\<rbrakk> \<Longrightarrow> (nabla \<turnstile> Paar t1 s1 \<approx> Paar t2 s2)"
+equ_func[intro!]: "(nabla \<turnstile> t1 \<approx> t2) \<Longrightarrow> (nabla \<turnstile> Func F t1 \<approx> Func F t2)"
+
+lemma equ_atom_elim[elim!]: "nabla\<turnstile>Atom a \<approx> Atom b \<Longrightarrow> a=b"
+apply(ind_cases "nabla \<turnstile> Atom a \<approx> Atom b", auto)
+done
+
+lemma equ_susp_elim[elim!]: "(nabla \<turnstile> Susp pi1 X \<approx> Susp pi2 X)
+ \<Longrightarrow> (\<forall> c \<in> ds pi1 pi2. (c,X)\<in> nabla)"
+apply(ind_cases "nabla \<turnstile> Susp pi1 X \<approx> Susp pi2 X", auto)
+done
+lemma equ_paar_elim[elim!]: "(nabla \<turnstile> Paar s1 t1 \<approx> Paar s2 t2) \<Longrightarrow>
+ (nabla \<turnstile> s1 \<approx> s2)\<and>(nabla \<turnstile> t1 \<approx> t2)"
+apply(ind_cases "nabla \<turnstile> Paar s1 t1 \<approx> Paar s2 t2", auto)
+done
+lemma equ_func_elim[elim!]: "(nabla \<turnstile> Func F t1 \<approx> Func F t2) \<Longrightarrow> (nabla \<turnstile> t1 \<approx> t2)"
+apply(ind_cases "nabla \<turnstile> Func F t1 \<approx> Func F t2", auto)
+done
+lemma equ_abst_aa_elim[elim!]: "(nabla \<turnstile> Abst a t1 \<approx> Abst a t2) \<Longrightarrow> (nabla \<turnstile> t1 \<approx> t2)"
+apply(ind_cases "nabla \<turnstile> Abst a t1 \<approx> Abst a t2", auto)
+done
+lemma equ_abst_ab_elim[elim!]: "\<lbrakk>(nabla \<turnstile> Abst a t1 \<approx> Abst b t2);a\<noteq>b\<rbrakk> \<Longrightarrow>
+ (nabla \<turnstile> t1 \<approx> (swap [(a,b)] t2))\<and>(nabla\<turnstile>a\<sharp>t2)"
+apply(ind_cases "(nabla \<turnstile> Abst a t1 \<approx> Abst b t2)", auto)
+done
+
+lemma equ_depth: "nabla \<turnstile> t1 \<approx> t2 \<Longrightarrow> depth t1 = depth t2"
+apply(erule equ.induct)
+apply(simp_all)
+done
+
+lemma rev_pi_pi_equ: "(nabla\<turnstile>swap (rev pi) (swap pi t)\<approx>t)"
+apply(induct_tac t)
+apply(auto)
+-- Susp
+apply(drule_tac ds_cancel_pi_left[of _ "rev pi @ pi" _ "[]", THEN mp, simplified])
+apply(simp only: ds_rev_pi_pi)
+apply(simp only: ds_def)
+apply(force)
+done
+
+lemma equ_pi_right: "\<forall>pi. (\<forall>a\<in>ds [] pi. nabla\<turnstile>a\<sharp>t) \<longrightarrow> (nabla\<turnstile>t\<approx>swap pi t)"
+apply(induct_tac t)
+apply(simp_all)
+-- Abst
+apply(rule allI)
+apply(case_tac "(swapas pi list)=list")
+apply(simp)
+apply(rule impI)
+apply(rule equ_abst_aa)
+apply(drule_tac x="pi" in spec)
+apply(subgoal_tac "(\<forall>a\<in>ds [] pi. nabla \<turnstile> a \<sharp> trm)")--A
+apply(force)
+--A
+apply(rule ballI)
+apply(drule_tac x=a in bspec)
+apply(assumption)
+apply(case_tac "list\<noteq>a")
+apply(force dest!: fresh_abst_ab_elim)
+apply(simp add: ds_def)
+apply(rule impI)
+apply(rule equ_abst_ab)
+apply(force)
+apply(drule_tac x="swapas (rev pi) list" in bspec)
+apply(simp add: ds_def)
+apply(rule conjI)
+apply(subgoal_tac "swapas (rev pi) list \<in> atms (rev pi)") --B
+apply(simp)
+--B
+apply(drule swapas_pi_ineq_a[THEN mp])
+apply(rule swapas_pi_in_atms)
+apply(simp)
+apply(clarify)
+apply(drule swapas_rev_pi_b)
+apply(simp)
+apply(force dest!: fresh_abst_ab_elim swapas_rev_pi_b intro!: fresh_swap_right[THEN mp])
+apply(drule_tac x="(list, swapas pi list)#pi" in spec)
+apply(subgoal_tac "(\<forall>a\<in>ds [] ((list, swapas pi list) # pi). nabla \<turnstile> a \<sharp> trm)")--C
+apply(force simp add: swap_append[THEN sym])
+--C
+apply(rule ballI)
+apply(drule_tac x="a" in bspec)
+apply(rule_tac b="list" in ds_7)
+apply(force)
+apply(assumption)
+apply(case_tac "list=a")
+apply(simp)
+apply(simp only: ds_def mem_Collect_eq)
+apply(erule conjE)
+apply(subgoal_tac "a\<noteq>swapas pi a")
+apply(simp)
+apply(force)
+apply(force dest!: fresh_abst_ab_elim)
+-- Susp
+apply(rule allI)
+apply(rule impI)
+apply(rule equ_susp)
+apply(rule ballI)
+apply(subgoal_tac "swapas list1 c\<in>ds [] pi")--A
+apply(force dest!: fresh_susp_elim)
+--A
+apply(rule ds_cancel_pi_left[THEN mp])
+apply(simp)
+-- Unit
+apply(force)
+-- Atom
+apply(rule allI)
+apply(rule impI)
+apply(case_tac "(swapas pi list) = list")
+apply(force)
+apply(drule ds_elem)
+apply(force dest!: fresh_atom_elim)
+-- Paar
+apply(force dest!: fresh_paar_elim)
+-- Func
+apply(force)
+done
+
+lemma pi_comm: "nabla\<turnstile>(swap (pi@[(a,b)]) t)\<approx>(swap ([(swapas pi a, swapas pi b)]@pi) t)"
+apply(induct_tac t)
+apply(simp_all)
+-- Abst
+apply(force simp add: swapas_comm)
+-- Susp
+apply(rule equ_susp)
+apply(rule ballI)
+apply(simp only: ds_def)
+apply(simp only: mem_Collect_eq)
+apply(erule conjE)
+apply(subgoal_tac "swapas (pi@[(a,b)]) (swapas list1 c) =
+ swapas ([(swapas pi a,swapas pi b)]@pi) (swapas list1 c)")--A
+apply(simp add: swapas_append[THEN sym])
+--A
+apply(simp only: swapas_comm)
+-- Units
+apply(rule equ_unit)
+-- Atom
+apply(force dest!: swapas_rev_pi_b swapas_rev_pi_a simp add: swapas_append)
+--Paar
+apply(force)
+--Func
+apply(force)
+done
+
+
+lemma l3_jud: "(nabla \<turnstile> t1\<approx>t2) \<Longrightarrow> (nabla \<turnstile> a\<sharp>t1) \<longrightarrow> (nabla \<turnstile> a\<sharp>t2)"
+apply(erule equ.induct)
+apply(simp_all)
+--Abst.ab
+apply(rule impI)
+apply(case_tac "aa=a")
+apply(force)
+apply(case_tac "b=a")
+apply(force)
+apply(force dest!: fresh_abst_ab_elim fresh_swap_left[THEN mp])
+-- Abst.aa
+apply(case_tac "a=aa")
+apply(force)
+apply(force dest!: fresh_abst_ab_elim)
+-- Susp
+apply(rule impI)
+apply(drule fresh_susp_elim, rule fresh_susp)
+apply(case_tac "swapas (rev pi1) a = swapas (rev pi2) a")
+apply(simp)
+apply(drule_tac x="swapas (rev pi2) a" in bspec)
+apply(rule ds_cancel_pi_left[THEN mp])
+apply(subgoal_tac "swapas (pi1@(rev pi2)) a \<noteq> a")--A
+apply(drule ds_elem)
+apply(force simp add: ds_def swapas_append)
+--A
+apply(clarify)
+apply(simp only: swapas_append)
+apply(drule swapas_rev_pi_a)
+apply(force)
+apply(assumption)
+-- Paar
+apply(force dest!: fresh_paar_elim)
+-- Func
+apply(force dest!: fresh_func_elim)
+done
+
+lemma big: "\<forall>t1 t2 t3. (n=depth t1) \<longrightarrow>
+ (((nabla\<turnstile>t1\<approx>t2)\<longrightarrow>(nabla\<turnstile>t2\<approx>t1))\<and>
+ (\<forall>pi. (nabla\<turnstile>t1\<approx>t2)\<longrightarrow>(nabla\<turnstile>swap pi t1\<approx>swap pi t2))\<and>
+ ((nabla\<turnstile>t1\<approx>t2)\<and>(nabla\<turnstile>t2\<approx>t3)\<longrightarrow>(nabla\<turnstile>t1\<approx>t3)))"
+apply(induct_tac n rule: nat_less_induct)
+apply(rule allI)+apply(rule impI)
+apply(rule conjI)
+-- SYMMETRY
+apply(rule impI)
+apply(ind_cases "nabla \<turnstile> t1 \<approx> t2")
+apply(simp_all)
+-- Abst.ab
+apply(rule equ_abst_ab)
+apply(force) --abst.ab.first.premise
+apply(rule_tac "t1.1"="swap [(a,b)] t2a" in l3_jud[THEN mp])
+apply(drule_tac x="depth t1a" in spec)
+apply(simp)
+apply(rule fresh_swap_right[THEN mp])
+apply(simp) --abst.ab.second.premise
+apply(subgoal_tac "nabla \<turnstile> swap [(b, a)] t1a \<approx> t2a")
+apply(drule_tac x="depth t1a" in spec)
+apply(simp)
+apply(subgoal_tac "nabla \<turnstile> swap [(b,a)] t1a \<approx> swap ([(b,a)]@[(a,b)]) t2a") --A
+apply(subgoal_tac "nabla \<turnstile> swap ([(b,a)]@[(a,b)]) t2a \<approx> t2a") --B
+apply(drule_tac x="depth t1a" in spec)
+apply(simp (no_asm_use))
+apply(drule_tac x="swap [(b,a)] t1a" in spec)
+apply(simp (no_asm_use))
+apply(drule_tac x="swap [(b,a),(a,b)] t2a" in spec)
+apply(force)
+--B
+apply(subgoal_tac "nabla\<turnstile>t2a \<approx> swap ([(b, a)] @ [(a, b)]) t2a")--C
+apply(drule_tac x="depth t1a" in spec)
+apply(simp)
+apply(drule_tac x="t2a" in spec)
+apply(drule mp)
+apply(drule equ_depth)
+apply(force)
+apply(best)
+--C
+apply(rule equ_pi_right[THEN spec,THEN mp])
+apply(subgoal_tac "ds [] ([(b, a)] @ [(a, b)])={}")
+apply(simp)
+apply(simp add: ds_baab)
+--A
+apply(force simp only: swap_append)
+-- Abst.aa
+apply(force)
+-- Unit
+apply(rule equ_unit)
+-- Atom
+apply(force)
+-- Susp
+apply(force simp only: ds_sym)
+-- Paar
+apply(rule equ_paar)
+apply(drule_tac x="depth t1a" in spec)
+apply(simp add: Suc_max_left)
+apply(drule_tac x="depth s1" in spec)
+apply(simp add: Suc_max_right)
+-- Func
+apply(best)
+-- ADD.PI
+apply(rule conjI)
+apply(rule impI)
+apply(ind_cases "nabla \<turnstile> t1 \<approx> t2")
+apply(simp_all)
+-- Abst.ab
+apply(rule allI)
+apply(rule equ_abst_ab)
+-- abst.ab.first.premise
+apply(clarify)
+apply(drule swapas_rev_pi_a)
+apply(simp)
+-- abst.ab.second.premise
+apply(rule fresh_swap_right[THEN mp])
+apply(simp)
+-- abst.ab.third.premise
+apply(subgoal_tac "nabla \<turnstile> swap pi t1a \<approx> swap (pi@[(a,b)]) t2a") --A
+apply(subgoal_tac "nabla \<turnstile> swap (pi@[(a,b)]) t2a \<approx> swap ([(swapas pi a,swapas pi b)]@pi) t2a") --B
+apply(drule_tac x="depth t1a" in spec)
+apply(simp (no_asm_use))
+apply(drule_tac x="swap pi t1a" in spec)
+apply(simp (no_asm_use))
+apply(drule_tac x="swap (pi@[(a,b)]) t2a" in spec)
+apply(drule conjunct2)+
+apply(drule_tac x="swap ((swapas pi a, swapas pi b) # pi) t2a" in spec)
+apply(simp add: swap_append[THEN sym])
+--B
+apply(rule pi_comm)
+apply(force simp only: swap_append)
+-- A
+apply(force simp only: swap_append)
+-- Unit
+apply(rule equ_unit)
+-- Atom
+apply(force)
+-- Susp
+apply(force simp only: ds_cancel_pi_front)
+-- Paar
+apply(rule allI)
+apply(rule equ_paar)
+apply(drule_tac x="depth t1a" in spec)
+apply(simp only: Suc_max_left)
+apply(drule_tac x="depth s1" in spec)
+apply(simp only: Suc_max_right)
+-- Func
+apply(best)
+-- TRANSITIVITY
+apply(rule impI)
+apply(erule conjE)
+apply(ind_cases "nabla \<turnstile> t1 \<approx> t2")
+apply(simp_all)
+-- Abst.ab
+apply(ind_cases "nabla \<turnstile> Abst b t2a \<approx> t3")
+apply(simp)
+apply(case_tac "ba=a")
+apply(simp)
+apply(rule equ_abst_aa)
+apply(subgoal_tac "nabla\<turnstile>swap [(a,b)] t2a \<approx> t2b") --A
+apply(best)
+--A
+apply(subgoal_tac "nabla\<turnstile>swap [(a,b)] t2a\<approx> swap ([(a,b)]@[(b,a)]) t2b") --B
+apply(subgoal_tac "nabla\<turnstile>swap ([(a,b)]@[(b,a)]) t2b \<approx> t2b") --C
+apply(drule_tac x="depth t1a" in spec)
+apply(simp)
+apply(drule_tac x="swap [(a,b)] t2a" in spec)
+apply(drule equ_depth)
+apply(simp (no_asm_use))
+apply(best)
+--C
+apply(subgoal_tac "nabla\<turnstile>t2b \<approx> swap ([(a,b)]@[(b,a)]) t2b")--D
+apply(drule_tac x="depth t1a" in spec)
+apply(simp)
+apply(drule_tac x="t2b" in spec)
+apply(drule mp)
+apply(force dest!: equ_depth)
+apply(best)
+--D
+apply(rule equ_pi_right[THEN spec, THEN mp])
+apply(simp add: ds_baab)
+--B
+apply(drule_tac x="depth t1a" in spec)
+apply(simp)
+apply(drule_tac x="t2a" in spec)
+apply(drule equ_depth)
+apply(simp)
+apply(drule_tac x="swap [(b, a)] t2b" in spec)
+apply(drule conjunct2)
+apply(drule conjunct1)
+apply(simp)
+apply(drule_tac x="[(a,b)]" in spec)
+apply(simp add: swap_append[THEN sym])
+-- Abst.ab
+apply(rule equ_abst_ab)
+-- abst.ab.first.premise
+apply(force)
+-- abst.ab.second.premise
+apply(rule_tac "t1.1"="swap [(b,ba)] t2a" in l3_jud[THEN mp])
+apply(subgoal_tac "nabla \<turnstile> swap [(b,ba)] t2a \<approx> swap ([(b,ba)]@[(b, ba)]) t2b") --A
+apply(subgoal_tac "nabla\<turnstile>swap ([(b,ba)]@[(b,ba)]) t2b \<approx> t2b") --B
+apply(drule_tac x="depth t1a" in spec)
+apply(simp)
+apply(drule_tac x="swap [(b, ba)] t2a" in spec)
+apply(drule mp)
+apply(force dest!: equ_depth)
+apply(best)
+--B
+apply(subgoal_tac "nabla\<turnstile>t2b \<approx> swap ([(b,ba)] @ [(b,ba)]) t2b")--C
+apply(drule_tac x="depth t1a" in spec)
+apply(simp)
+apply(drule_tac x="t2b" in spec)
+apply(drule mp)
+apply(force dest!: equ_depth)
+apply(best)
+-- C
+apply(rule equ_pi_right[THEN spec, THEN mp])
+apply(simp add: ds_abab)
+--A
+apply(drule_tac x="depth t1a" in spec)
+apply(simp)
+apply(drule_tac x="t2a" in spec)
+apply(drule mp)
+apply(force dest!: equ_depth)
+apply(drule_tac x="swap [(b,ba)] t2b" in spec)
+apply(drule conjunct2)
+apply(drule conjunct1)
+apply(simp)
+apply(drule_tac x="[(b,ba)]" in spec)
+apply(simp add: swap_append[THEN sym])
+-- abst.ab.third.premise
+apply(force intro!: fresh_swap_right[THEN mp])
+-- very.complex
+apply(subgoal_tac "nabla\<turnstile>t1a \<approx> swap ([(a,b)]@[(b,ba)]) t2b") --A
+apply(subgoal_tac "nabla\<turnstile>swap ([(a,b)]@[(b,ba)]) t2b \<approx> swap [(a,ba)] t2b") --B
+apply(drule_tac x="depth t1a" in spec)
+apply(simp (no_asm_use))
+apply(best)
+--B
+apply(subgoal_tac "nabla\<turnstile>swap [(a, ba)] t2b \<approx> swap [(a,b),(b,ba)] t2b")--C
+apply(drule_tac x="depth t1a" in spec)
+apply(simp)
+apply(drule_tac x="swap [(a, ba)] t2b" in spec)
+apply(drule mp)
+apply(force dest!: equ_depth)
+apply(force)
+apply(subgoal_tac "nabla\<turnstile>swap [(a,ba)] t2b\<approx> swap [(a,ba)] (swap [(a,ba),(a,b),(b,ba)] t2b)") --D
+apply(subgoal_tac "nabla\<turnstile>swap (rev [(a,ba)]) (swap [(a,ba)] (swap [(a,b),(b,ba)] t2b))
+ \<approx>swap [(a,b),(b,ba)] t2b") --E
+apply(simp add: swap_append[THEN sym])
+apply(drule_tac x="depth t1a" in spec)
+apply(simp)
+apply(drule_tac x="swap [(a,ba)] t2b" in spec)
+apply(drule mp)
+apply(force dest!: equ_depth)
+apply(drule_tac x="swap [(a, ba), (a, ba), (a, b), (b, ba)] t2b" in spec)
+apply(drule conjunct2)+
+apply(best)
+-- D
+apply(rule rev_pi_pi_equ)
+-- E
+apply(subgoal_tac "nabla\<turnstile>t2b\<approx>swap [(a, ba), (a, b), (b, ba)] t2b") --F
+apply(drule_tac x="depth t1a" in spec)
+apply(simp)
+apply(drule_tac x="t2b" in spec)
+apply(drule mp)
+apply(force dest!: equ_depth)
+apply(best)
+--F
+apply(rule equ_pi_right[THEN spec, THEN mp])
+apply(subgoal_tac "ds [] [(a,ba),(a,b),(b,ba)]={a,b}") -- G
+apply(simp)
+apply(drule_tac "t1.1"="t2a" and "t2.1"="swap [(b, ba)] t2b" and a1="a" in l3_jud[THEN mp])
+apply(assumption)
+apply(subgoal_tac "nabla \<turnstile> swapas (rev [(b,ba)]) a \<sharp> t2b") --H
+apply(simp)
+apply(case_tac "b=a")
+apply(force)
+apply(force)
+--H
+apply(rule fresh_swap_left[THEN mp])
+apply(assumption)
+--G
+apply(rule ds_acabbc)
+apply(assumption)+
+--A
+apply(subgoal_tac "nabla\<turnstile>swap [(a,b)] t2a\<approx>swap [(a,b)] (swap [(b,ba)] t2b)")--I
+apply(drule_tac x="depth t1a" in spec)
+apply(simp (no_asm_use))
+apply(drule_tac x="t1a" in spec)
+apply(simp (no_asm_use))
+apply(drule_tac x="swap [(a,b)] t2a" in spec)
+apply(drule conjunct2)+
+apply(drule_tac x="swap [(a, b)] (swap [(b, ba)] t2b)" in spec)
+apply(force simp add: swap_append[THEN sym])
+--I
+apply(drule_tac x="depth t1a" in spec)
+apply(simp (no_asm_use))
+apply(drule_tac x="t2a" in spec)
+apply(drule mp)
+apply(force dest!: equ_depth)
+apply(drule_tac x="swap [(b,ba)] t2b" in spec)
+apply(best)
+-- Abst.ab
+apply(simp)
+apply(rule equ_abst_ab)
+apply(assumption)
+apply(drule_tac "t1.1"="t2a" and "t2.1"="t2b" and a1="a" in l3_jud[THEN mp])
+apply(assumption)+
+apply(subgoal_tac "nabla\<turnstile>swap [(a, b)] t2a\<approx>swap [(a, b)] t2b") --A
+apply(best)
+--A
+apply(drule_tac x="depth t1a" in spec)
+apply(simp (no_asm_use))
+apply(drule_tac x="t2a" in spec)
+apply(drule mp)
+apply(force dest!: equ_depth)
+apply(drule_tac x="t2b" in spec)
+apply(best)
+-- Abst
+apply(ind_cases "nabla \<turnstile> Abst a t2a \<approx> t3")
+apply(best)
+apply(best)
+-- Susp
+apply(ind_cases "nabla \<turnstile> Susp pi2 X \<approx> t3")
+apply(simp)
+apply(rule equ_susp)
+apply(rule ballI)
+apply(drule_tac "pi2.1"="pi2" in ds_trans[THEN mp])
+apply(force)
+-- Paar
+apply(ind_cases "nabla \<turnstile> Paar t2a s2 \<approx> t3")
+apply(simp)
+apply(rule equ_paar)
+apply(drule_tac x="depth t1a" in spec)
+apply(simp (no_asm_use) add: Suc_max_left)
+apply(best)
+apply(drule_tac x="depth s1" in spec)
+apply(simp (no_asm_use) add: Suc_max_right)
+apply(best)
+-- Func
+apply(ind_cases "nabla \<turnstile> Func F t2a \<approx> t3")
+apply(best)
+done
+
+lemma pi_right_equ_help:
+ "\<forall>t. (n=depth t) \<longrightarrow> (\<forall>pi. nabla\<turnstile>t\<approx>swap pi t\<longrightarrow>(\<forall>a\<in> ds [] pi. nabla\<turnstile>a\<sharp>t))"
+apply(induct_tac n rule: nat_less_induct)
+apply(rule allI)+
+apply(rule impI)
+apply(rule allI)+
+apply(rule impI)
+apply(ind_cases "nabla \<turnstile> t \<approx> swap pi t")
+apply(simp_all)
+--Abst.ab
+apply(rule ballI)
+apply(case_tac "aa=a")
+apply(force)
+apply(rule fresh_abst_ab)
+apply(case_tac "aa=swapas (rev pi) a")
+apply(simp)
+apply(drule fresh_swap_left[THEN mp])
+apply(assumption)
+apply(drule_tac x="depth t1" in spec)
+apply(simp)
+apply(drule_tac x="t1" in spec)
+apply(simp add: swap_append[THEN sym])
+apply(drule_tac x="[(a, swapas pi a)] @ pi" in spec)
+apply(simp)
+apply(case_tac "aa=swapas pi a")
+apply(simp)
+apply(drule_tac x="swapas pi a" in bspec)
+apply(simp (no_asm) only: ds_def)
+apply(simp only: mem_Collect_eq)
+apply(rule conjI)
+apply(simp)
+apply(simp)
+apply(rule impI)
+apply(clarify)
+apply(drule sym)
+apply(drule swapas_rev_pi_a)
+apply(force)
+apply(assumption)
+apply(drule_tac x="aa" in bspec)
+apply(subgoal_tac "\<forall>A. aa\<in>A-{swapas pi a} \<and> aa\<noteq>swapas pi a \<longrightarrow>aa\<in>A")--A
+apply(drule_tac x="ds [] ((a,swapas pi a) # pi)" in spec)
+apply(simp add: ds_equality[THEN sym])
+--A
+apply(force)
+apply(assumption)+
+--Abst.aa
+apply(rule ballI)
+apply(case_tac "aa=a")
+apply(force)
+apply(best)
+--Unit
+apply(force)
+--Atom
+apply(force simp add: ds_def)
+--Susp
+apply(rule ballI)
+apply(rule fresh_susp)
+apply(drule_tac x="swapas (rev pi1) a" in bspec)
+apply(rule ds_cancel_pi_right[of _ _ "[]" _, simplified, THEN mp])
+apply(simp)
+apply(assumption)
+--Paar
+apply(rule ballI)
+apply(rule fresh_paar)
+apply(drule_tac x="depth t1" in spec)
+apply(force simp add: Suc_max_left)
+apply(drule_tac x="depth s1" in spec)
+apply(force simp add: Suc_max_right)
+--Func
+apply(best)
+done
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/Substs.thy Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,363 @@
+
+
+theory Substs = Main + Terms + PreEqu + Equ:
+
+(* substitutions *)
+
+types substs = "(string \<times> trm)list"
+
+consts
+ look_up :: "string \<Rightarrow> substs \<Rightarrow> trm"
+primrec
+ "look_up X [] = Susp [] X"
+ "look_up X (x#xs) = (if (X = fst x) then (snd x) else look_up X xs)"
+
+consts
+ subst :: "substs \<Rightarrow> trm \<Rightarrow> trm"
+primrec
+ subst_unit: "subst s (Unit) = Unit"
+ subst_susp: "subst s (Susp pi X) = swap pi (look_up X s)"
+ subst_atom: "subst s (Atom a) = Atom a"
+ subst_abst: "subst s (Abst a t) = Abst a (subst s t)"
+ subst_paar: "subst s (Paar t1 t2) = Paar (subst s t1) (subst s t2)"
+ subst_func: "subst s (Func F t) = Func F (subst s t)"
+
+declare subst_susp [simp del]
+
+(* composition of substitutions (adapted from Martin Coen) *)
+
+consts
+ alist_rec :: "substs \<Rightarrow> substs \<Rightarrow> (string\<Rightarrow>trm\<Rightarrow>substs\<Rightarrow>substs\<Rightarrow>substs) \<Rightarrow> substs"
+
+primrec
+ "alist_rec [] c d = c"
+ "alist_rec (p#al) c d = d (fst p) (snd p) al (alist_rec al c d)"
+
+consts
+ "\<bullet>" :: "substs \<Rightarrow> substs \<Rightarrow> substs" (infixl 81)
+defs
+ comp_def: "s1 \<bullet> s2 \<equiv> alist_rec s2 s1 (\<lambda> x y xs g. (x,subst s1 y)#g)"
+
+(* domain of substitutions *)
+
+consts
+ domn :: "(trm \<Rightarrow> trm) \<Rightarrow> string set"
+defs
+ domn_def: "domn s \<equiv> {X. (s (Susp [] X)) \<noteq> (Susp [] X)}"
+
+(* substitutions extending freshness environments *)
+
+consts
+ ext_subst :: "fresh_envs \<Rightarrow> (trm \<Rightarrow> trm) \<Rightarrow> fresh_envs \<Rightarrow> bool" (" _ \<Turnstile> _ _ " [80,80,80] 80)
+defs
+ ext_subst_def: "nabla1 \<Turnstile> s (nabla2) \<equiv> (\<forall>(a,X)\<in>nabla2. nabla1\<turnstile>a\<sharp> s (Susp [] X))"
+
+(* alpah-equality for substitutions *)
+
+consts
+ subst_equ :: "fresh_envs \<Rightarrow> (trm\<Rightarrow>trm) \<Rightarrow> (trm\<Rightarrow>trm) \<Rightarrow> bool" (" _ \<Turnstile> _ \<approx> _" [80,80,80] 80)
+
+defs
+ subst_equ_def:
+ "nabla\<Turnstile> s1 \<approx> s2 \<equiv> \<forall>X\<in>(domn s1\<union>domn s2). (nabla \<turnstile> s1 (Susp [] X) \<approx> s2 (Susp [] X))"
+
+lemma not_in_domn: "X\<notin>(domn s)\<Longrightarrow> (s (Susp [] X)) = (Susp [] X)"
+apply(auto simp only: domn_def)
+done
+
+lemma subst_swap_comm: "subst s (swap pi t) = swap pi (subst s t)"
+apply(induct_tac t)
+apply(auto simp add: swap_append subst_susp)
+done
+
+lemma subst_not_occurs: "\<not>(occurs X t) \<longrightarrow> subst [(X,t2)] t = t"
+apply(induct t)
+apply(auto simp add: subst_susp)
+done
+
+lemma [simp]: "subst [] t = t"
+apply(induct_tac t, auto simp add: subst_susp)
+done
+
+lemma [simp]: "subst (s\<bullet>[]) = subst s"
+apply(auto simp add: comp_def)
+done
+
+lemma [simp]: "subst ([]\<bullet>s) = subst s"
+apply(rule ext)
+apply(induct_tac x)
+apply(auto)
+apply(induct_tac s rule: alist_rec.induct)
+apply(auto simp add: comp_def subst_susp)
+done
+
+lemma subst_comp_expand: "subst (s1\<bullet>s2) t = subst s1 (subst s2 t)"
+apply(induct_tac t)
+apply(auto)
+apply(induct_tac s2 rule: alist_rec.induct)
+apply(auto simp add: comp_def subst_susp subst_swap_comm)
+done
+
+lemma subst_assoc: "subst (s1\<bullet>(s2\<bullet>s3))=subst ((s1\<bullet>s2)\<bullet>s3)"
+apply(rule ext)
+apply(induct_tac x)
+apply(auto)
+apply(simp add: subst_comp_expand)
+done
+
+lemma fresh_subst: "nabla1\<turnstile>a\<sharp>t\<Longrightarrow> nabla2\<Turnstile>(subst s) nabla1 \<longrightarrow> nabla2\<turnstile>a\<sharp>subst s t"
+apply(erule fresh.induct)
+apply(auto)
+--Susp
+apply(simp add: ext_subst_def subst_susp)
+apply(drule_tac x="(swapas (rev pi) a, X)" in bspec)
+apply(assumption)
+apply(simp)
+apply(rule fresh_swap_right[THEN mp])
+apply(assumption)
+done
+
+lemma equ_subst: "nabla1\<turnstile>t1\<approx>t2\<Longrightarrow> nabla2\<Turnstile> (subst s) nabla1 \<longrightarrow> nabla2\<turnstile>(subst s t1)\<approx>(subst s t2)"
+apply(erule equ.induct)
+apply(auto)
+apply(rule_tac "nabla1.1"="nabla" in fresh_subst[THEN mp])
+apply(assumption)+
+apply(simp add: subst_swap_comm)
+-- Susp
+apply(simp only: subst_susp)
+apply(rule equ_pi1_pi2_add[THEN mp])
+apply(simp only: ext_subst_def subst_susp)
+apply(force)
+done
+
+lemma unif_1:
+ "\<lbrakk>nabla\<turnstile>subst s (Susp pi X)\<approx>subst s t\<rbrakk>\<Longrightarrow> nabla\<Turnstile> subst (s\<bullet>[(X,swap (rev pi) t)])\<approx>subst s"
+apply(simp only: subst_equ_def)
+apply(case_tac "pi=[]")
+apply(simp add: subst_susp comp_def)
+apply(force intro: equ_sym equ_refl)
+apply(subgoal_tac "domn (subst (s\<bullet>[(X,swap (rev pi) t)]))=domn(subst s)\<union>{X}")--A
+apply(simp)
+apply(rule conjI)
+apply(simp add: subst_comp_expand)
+apply(simp add: subst_susp)
+apply(simp only: subst_swap_comm)
+apply(simp only: equ_pi_to_left)
+apply(rule equ_sym)
+apply(assumption)
+apply(rule ballI)
+apply(simp only: subst_comp_expand)
+apply(simp add: subst_susp)
+apply(force intro: equ_sym equ_refl simp add: subst_swap_comm equ_pi_to_left)
+--A
+apply(simp only: domn_def)
+apply(auto)
+apply(simp add: subst_comp_expand)
+apply(simp add: subst_susp subst_swap_comm)
+apply(simp only: subst_comp_expand)
+apply(simp add: subst_susp subst_swap_comm)
+apply(drule swap_empty[THEN mp])
+apply(simp)
+apply(simp only: subst_comp_expand)
+apply(simp only: subst_susp)
+apply(auto)
+apply(case_tac "x=X")
+apply(simp)
+apply(simp add: subst_swap_comm)
+apply(drule swap_empty[THEN mp])
+apply(simp)
+apply(simp add: subst_susp)
+done
+
+lemma subst_equ_a:
+"\<lbrakk>nabla\<Turnstile>(subst s1) \<approx> (subst s2)\<rbrakk>\<Longrightarrow> \<forall>t2. nabla\<turnstile>(subst s2 t1)\<approx>t2 \<longrightarrow> nabla\<turnstile>(subst s1 t1)\<approx>t2"
+apply(rule allI)
+apply(induct t1)
+--Abst.ab
+apply(rule impI)
+apply(simp)
+apply(ind_cases "nabla \<turnstile> Abst list (subst s1 trm) \<approx> t2")
+apply(best)
+--Abst.aa
+apply(force)
+--Susp
+apply(rule impI)
+apply(simp only: subst_equ_def)
+apply(case_tac "list2\<in>domn (subst s1) \<union> domn (subst s2)")
+apply(drule_tac x="list2" in bspec)
+apply(assumption)
+apply(simp)
+apply(subgoal_tac "nabla \<turnstile> subst s2 (Susp [] list2) \<approx> swap (rev list1) t2")--A
+apply(drule_tac "t1.0"="subst s1 (Susp [] list2)" and
+ "t2.0"="subst s2 (Susp [] list2)" and
+ "t3.0"="swap (rev list1) t2" in equ_trans)
+apply(assumption)
+apply(simp only: equ_pi_to_right)
+apply(simp add: subst_swap_comm[THEN sym])
+--A
+apply(simp only: equ_pi_to_right)
+apply(simp add: subst_swap_comm[THEN sym])
+apply(simp)
+apply(erule conjE)
+apply(drule not_in_domn)+
+apply(subgoal_tac "subst s1 (Susp list1 list2)=swap list1 (subst s1 (Susp [] list2))")--B
+apply(subgoal_tac "subst s2 (Susp list1 list2)=swap list1 (subst s2 (Susp [] list2))")--C
+apply(simp)
+--BC
+apply(simp (no_asm) add: subst_swap_comm[THEN sym])
+apply(simp (no_asm) add: subst_swap_comm[THEN sym])
+--Unit
+apply(force)
+--Atom
+apply(force)
+--Paar
+apply(rule impI)
+apply(simp)
+apply(ind_cases "nabla \<turnstile> Paar (subst sigma1 trm1) (subst sigma1 trm2) \<approx> t2")
+apply(best)
+--Func
+apply(rule impI)
+apply(simp)
+apply(ind_cases "nabla \<turnstile> Func list (subst sigma1 trm) \<approx> t2")
+apply(best)
+done
+
+lemma unif_2a: "\<lbrakk>nabla\<Turnstile>subst s1\<approx>subst s2\<rbrakk>\<Longrightarrow>
+ (nabla\<turnstile>subst s2 t1 \<approx> subst s2 t2)\<longrightarrow>(nabla\<turnstile>subst s1 t1 \<approx> subst s1 t2)"
+apply(rule impI)
+apply(frule_tac "t1.0"="t1" in subst_equ_a)
+apply(drule_tac x="subst s2 t2" in spec)
+apply(simp)
+apply(drule equ_sym)
+apply(drule equ_sym)
+apply(frule_tac "t1.0"="t2" in subst_equ_a)
+apply(drule_tac x="subst s1 t1" in spec)
+apply(rule equ_sym)
+apply(simp)
+done
+
+
+lemma unif_2b: "\<lbrakk>nabla\<Turnstile>subst s1\<approx> subst s2\<rbrakk>\<Longrightarrow>nabla\<turnstile>a\<sharp> subst s2 t\<longrightarrow>nabla\<turnstile>a\<sharp>subst s1 t"
+apply(induct t)
+--Abst
+apply(simp)
+apply(rule impI)
+apply(case_tac "a=list")
+apply(force)
+apply(force dest!: fresh_abst_ab_elim)
+--Susp
+apply(rule impI)
+apply(subgoal_tac "subst s1 (Susp list1 list2)= swap list1 (subst s1 (Susp [] list2))")--A
+apply(subgoal_tac "subst s2 (Susp list1 list2)= swap list1 (subst s2 (Susp [] list2))")--B
+apply(simp add: subst_equ_def)
+apply(case_tac "list2\<in>domn (subst s1) \<union> domn (subst s2)")
+apply(drule_tac x="list2" in bspec)
+apply(assumption)
+apply(rule fresh_swap_right[THEN mp])
+apply(rule_tac "t1.1"=" subst s2 (Susp [] list2)" in l3_jud[THEN mp])
+apply(rule equ_sym)
+apply(simp)
+apply(rule fresh_swap_left[THEN mp])
+apply(simp)
+apply(simp)
+apply(erule conjE)
+apply(drule not_in_domn)+
+apply(simp add: subst_swap_comm)
+--AB
+apply(simp (no_asm) add: subst_swap_comm[THEN sym])
+apply(simp (no_asm) add: subst_swap_comm[THEN sym])
+--Unit
+apply(force)
+--Atom
+apply(force)
+--Paar
+apply(force dest!: fresh_paar_elim)
+--Func
+apply(force dest!: fresh_func_elim)
+done
+
+lemma subst_equ_to_trm: "nabla\<Turnstile>subst s1 \<approx> subst s2\<Longrightarrow> nabla\<turnstile>subst s1 t\<approx>subst s2 t"
+apply(induct_tac t)
+apply(auto simp add: subst_equ_def)
+apply(case_tac "list2\<in>domn (subst s1) \<union> domn (subst s2)")
+apply(simp only: subst_susp)
+apply(simp only: equ_pi_to_left[THEN sym])
+apply(simp only: equ_involutive_left)
+apply(simp)
+apply(erule conjE)
+apply(drule not_in_domn)+
+apply(subgoal_tac "subst s1 (Susp list1 list2)=swap list1 (subst s1 (Susp [] list2))")--A
+apply(subgoal_tac "subst s2 (Susp list1 list2)=swap list1 (subst s2 (Susp [] list2))")--B
+apply(simp)
+apply(rule equ_refl)
+apply(simp (no_asm_use) add: subst_swap_comm[THEN sym])+
+done
+
+lemma subst_cancel_right: "\<lbrakk>nabla\<Turnstile>(subst s1)\<approx>(subst s2)\<rbrakk>\<Longrightarrow>nabla\<Turnstile>(subst (s1\<bullet>s))\<approx>(subst (s2\<bullet>s))"
+apply(simp (no_asm) only: subst_equ_def)
+apply(rule ballI)
+apply(simp only: subst_comp_expand)
+apply(rule subst_equ_to_trm)
+apply(assumption)
+done
+
+lemma subst_trans: "\<lbrakk>nabla\<Turnstile>subst s1\<approx>subst s2; nabla\<Turnstile>subst s2\<approx>subst s3\<rbrakk>\<Longrightarrow>nabla\<Turnstile>subst s1\<approx>subst s3"
+apply(simp add: subst_equ_def)
+apply(auto)
+apply(case_tac "X \<in>domn (subst s2)")
+apply(drule_tac x="X" in bspec)
+apply(force)
+apply(drule_tac x="X" in bspec)
+apply(force)
+apply(rule_tac "t2.0"="subst s2 (Susp [] X)" in equ_trans)
+apply(assumption)+
+apply(case_tac "X \<in>domn (subst s3)")
+apply(drule_tac x="X" in bspec)
+apply(force)
+apply(drule_tac x="X" in bspec)
+apply(force)
+apply(rule_tac "t2.0"="subst s2 (Susp [] X)" in equ_trans)
+apply(assumption)+
+apply(drule not_in_domn)+
+apply(drule_tac x="X" in bspec)
+apply(force)
+apply(simp)
+apply(case_tac "X \<in>domn (subst s1)")
+apply(drule_tac x="X" in bspec)
+apply(force)
+apply(drule_tac x="X" in bspec)
+apply(force)
+apply(rule_tac "t2.0"="subst s2 (Susp [] X)" in equ_trans)
+apply(assumption)+
+apply(case_tac "X \<in>domn (subst s2)")
+apply(drule_tac x="X" in bspec)
+apply(force)
+apply(drule_tac x="X" in bspec)
+apply(force)
+apply(rule_tac "t2.0"="subst s2 (Susp [] X)" in equ_trans)
+apply(assumption)+
+apply(drule not_in_domn)+
+apply(simp)
+apply(rotate_tac 1)
+apply(drule_tac x="X" in bspec)
+apply(force)
+apply(simp)
+done
+
+(* if occurs holds, then one subterm is equal to (subst s (Susp pi X)) *)
+
+lemma occurs_sub_trm_equ:
+ "occurs X t1 \<longrightarrow> (\<exists>t2\<in>sub_trms (subst s t1). (\<exists>pi.(nabla\<turnstile>subst s (Susp pi1 X)\<approx>(swap pi t2))))"
+apply(induct_tac t1)
+apply(auto)
+apply(simp only: subst_susp)
+apply(rule_tac x="swap list1 (look_up X s)" in bexI)
+apply(rule_tac x="pi1@(rev list1)" in exI)
+apply(simp add: swap_append)
+apply(simp add: equ_pi_to_left[THEN sym])
+apply(simp only: equ_involutive_left)
+apply(rule equ_refl)
+apply(rule t_sub_trms_t)
+done
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/Swap.thy Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,55 @@
+
+theory Swap = Main:
+
+consts
+ swapa :: "(string \<times> string) \<Rightarrow> string \<Rightarrow> string"
+ swapas :: "(string \<times> string)list \<Rightarrow> string \<Rightarrow> string"
+primrec
+ "swapa (b,c) a = (if b = a then c else (if c = a then b else a))"
+
+primrec
+ "swapas [] a = a"
+ "swapas (x#pi) a = swapa x (swapas pi a)"
+
+lemma [simp]:
+ "swapas [(a,a)] b=b"
+apply(auto)
+done
+
+lemma swapas_append:
+ "swapas (pi1@pi2) a = swapas pi1 (swapas pi2 a)"
+apply(induct_tac pi1, auto)
+done
+
+lemma [simp]:
+ "swapas (rev pi) (swapas pi a) = a"
+apply(induct_tac pi)
+apply(simp_all add: swapas_append, auto)
+done
+
+lemma swapas_rev_pi_a:
+ "swapas pi a = b \<Longrightarrow> swapas (rev pi) b = a"
+apply(auto)
+done
+
+lemma [simp]:
+ "swapas pi (swapas (rev pi) a) = a"
+apply(rule swapas_rev_pi_a[of "(rev pi)" _ _,simplified])
+apply(simp)
+done
+
+lemma swapas_rev_pi_b:
+ "swapas (rev pi) a = b \<Longrightarrow> swapas pi b = a"
+apply(auto)
+done
+
+lemma swapas_comm:
+ "(swapas (pi@[(a,b)]) c) = (swapas ([(swapas pi a,swapas pi b)]@pi) c)"
+apply(auto)
+apply(simp_all only: swapas_append)
+apply(auto)
+apply(drule swapas_rev_pi_a, simp)
+apply(drule swapas_rev_pi_a, simp)
+done
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/Termination.thy Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,205 @@
+
+
+theory Termination = Main + Terms + Fresh + Equ + Substs + Mgu:
+
+(* set of variables *)
+
+consts
+ vars_trm :: "trm \<Rightarrow> string set"
+ vars_eprobs :: "eprobs \<Rightarrow> (string set)"
+primrec
+ "vars_trm (Unit) = {}"
+ "vars_trm (Atom a) = {}"
+ "vars_trm (Susp pi X) = {X}"
+ "vars_trm (Paar t1 t2) = (vars_trm t1)\<union>(vars_trm t2)"
+ "vars_trm (Abst a t) = vars_trm t"
+ "vars_trm (Func F t) = vars_trm t"
+primrec
+ "vars_eprobs [] = {}"
+ "vars_eprobs (x#xs) = (vars_trm (snd x))\<union>(vars_trm (fst x))\<union>(vars_eprobs xs)"
+
+lemma[simp]: "vars_trm (swap pi t) = vars_trm t"
+apply(induct_tac t)
+apply(auto)
+done
+
+consts
+ size_trm :: "trm \<Rightarrow> nat"
+ size_fprobs :: "fprobs \<Rightarrow> nat"
+ size_eprobs :: "eprobs \<Rightarrow> nat"
+ size_probs :: "probs \<Rightarrow> nat"
+primrec
+ "size_trm (Unit) = 1"
+ "size_trm (Atom a) = 1"
+ "size_trm (Susp pi X) = 1"
+ "size_trm (Abst a t) = 1 + size_trm t"
+ "size_trm (Func F t) = 1 + size_trm t"
+ "size_trm (Paar t t') = 1 + (size_trm t) + (size_trm t')"
+primrec
+ "size_fprobs [] = 0"
+ "size_fprobs (x#xs) = (size_trm (snd x))+(size_fprobs xs)"
+primrec
+ "size_eprobs [] = 0"
+ "size_eprobs (x#xs) = (size_trm (fst x))+(size_trm (snd x))+(size_eprobs xs)"
+
+lemma[simp]: "size_trm (swap pi t) = size_trm t"
+apply(induct_tac t)
+apply(auto)
+done
+
+syntax
+ "_measure_relation" :: "(nat\<times>nat\<times>nat) \<Rightarrow> (nat\<times>nat\<times>nat) \<Rightarrow> bool" ("_ \<lless> _" [80,80] 80)
+translations
+ "n1 \<lless> n2" \<rightleftharpoons> "(n1,n2) \<in> (less_than<*lex*>less_than<*lex*>less_than)"
+
+consts
+ rank :: "probs \<Rightarrow> (nat\<times>nat\<times>nat)"
+primrec
+ "rank (eprobs,fprobs) = (card (vars_eprobs eprobs),size_eprobs eprobs,size_fprobs fprobs)"
+
+lemma[simp]: "finite (vars_trm t)"
+apply(induct t)
+apply(auto)
+done
+
+lemma[simp]: "finite (vars_eprobs P)"
+apply(induct_tac P)
+apply(auto)
+done
+
+lemma union_comm: "A\<union>(B\<union>C)=(A\<union>B)\<union>C"
+apply(auto)
+done
+
+lemma card_union: "\<lbrakk>finite A; finite B\<rbrakk>\<Longrightarrow>(card B < card (A\<union>B)) \<or> (card B = card (A\<union>B))"
+apply(auto)
+apply(rule psubset_card_mono)
+apply(auto)
+done
+
+lemma card_insert: "\<lbrakk>finite B\<rbrakk>\<Longrightarrow>(card B < card (insert X B)) \<or> (card B = card (insert X B))"
+apply(auto)
+apply(rule psubset_card_mono)
+apply(auto)
+done
+
+lemma subseteq_card: "\<lbrakk>A\<subseteq>B; finite B\<rbrakk>\<Longrightarrow>(card A \<le> card B)"
+apply(drule_tac A="A" in card_mono)
+apply(auto simp add: le_eq_less_or_eq)
+done
+
+lemma not_occurs_trm: "\<not>occurs X t \<longrightarrow> X\<notin> vars_trm t"
+apply(induct_tac t)
+apply(auto)
+done
+
+lemma not_occurs_subst: "\<not>occurs X t1\<longrightarrow> X\<notin> vars_trm (subst [(X,swap pi2 t1)] t2)"
+apply(induct_tac t2)
+apply(auto simp add: subst_susp not_occurs_trm)
+done
+
+lemma not_occurs_list: "\<not>occurs X t \<longrightarrow>
+ X\<notin> vars_eprobs (apply_subst_eprobs [(X, swap pi t)] xs)"
+apply(induct_tac xs)
+apply(simp)
+apply(case_tac a)
+apply(auto simp add: not_occurs_subst)
+done
+
+lemma vars_equ: "\<not>occurs X t1 \<and> occurs X t2\<longrightarrow>
+ vars_trm (subst [(X, swap pi t1)] t2)=(vars_trm t1\<union>vars_trm t2)-{X}"
+apply(induct_tac t2)
+apply(force)
+apply(case_tac "X=list2")
+apply(simp add: subst_susp not_occurs_trm)
+apply(simp)
+apply(simp)
+apply(simp)
+apply(simp)
+apply(rule conjI)
+apply(case_tac "occurs X trm2")
+apply(force)
+apply(force dest: not_occurs_trm[THEN mp] simp add: subst_not_occurs)
+apply(force dest: not_occurs_trm[THEN mp] simp add: subst_not_occurs)
+apply(force)
+done
+
+lemma vars_subseteq: "\<not>occurs X t \<longrightarrow>
+ vars_eprobs (apply_subst_eprobs [(X, swap pi t)] xs) \<subseteq> (vars_trm t \<union> vars_eprobs xs)"
+apply(induct_tac xs)
+apply(simp)
+apply(rule impI)
+apply(simp)
+apply(case_tac "occurs X (fst a)")
+apply(case_tac "occurs X (snd a)")
+apply(simp add: vars_equ[THEN mp])
+apply(force)
+apply(simp add: subst_not_occurs)
+apply(force simp add: vars_equ)
+apply(case_tac "occurs X (snd a)")
+apply(simp add: vars_equ[THEN mp])
+apply(force simp add: subst_not_occurs)
+apply(force simp add: subst_not_occurs)
+done
+
+lemma vars_decrease:
+ "\<not>occurs X t\<longrightarrow> card (vars_eprobs (apply_subst_eprobs [(X, swap pi t)] xs))
+ <card (insert X (vars_trm t \<union> vars_eprobs xs))"
+apply(rule impI)
+apply(case_tac "X\<in>(vars_trm t \<union> vars_eprobs xs)")
+(* first case *)
+apply(subgoal_tac "insert X (vars_trm t \<union> vars_eprobs xs) = (vars_trm t \<union> vars_eprobs xs)") (*A*)
+apply(simp)
+apply(frule_tac pi1="pi" and xs1="xs" in vars_subseteq[THEN mp])
+apply(frule_tac pi1="pi" and xs1="xs" in not_occurs_list[THEN mp])
+apply(subgoal_tac "vars_eprobs (apply_subst_eprobs [(X, swap pi t)] xs)
+ \<subset> vars_trm t \<union> vars_eprobs xs") (* B *)
+apply(simp add: psubset_card_mono)
+(* B *)
+apply(force)
+(* A *)
+apply(force)
+(* second case *)
+apply(subgoal_tac "finite (vars_trm t \<union> vars_eprobs xs)")
+apply(drule_tac x="X" in card_insert_disjoint)
+apply(simp)
+apply(simp)
+apply(frule_tac pi1="pi" and xs1="xs" in vars_subseteq[THEN mp])
+apply(auto dest: subseteq_card)
+done
+
+lemma rank_cred: "\<lbrakk>P1\<turnstile>(nabla)\<rightarrow>P2\<rbrakk>\<Longrightarrow>(rank P2) \<lless> (rank P1)"
+apply(ind_cases "P1 \<turnstile> nabla \<rightarrow> P2")
+apply(simp_all add: lex_prod_def)
+done
+
+lemma rank_sred: "\<lbrakk>P1\<turnstile>(s)\<leadsto>P2\<rbrakk>\<Longrightarrow>(rank P2) \<lless> (rank P1)"
+apply(ind_cases "P1 \<turnstile> s \<leadsto> P2")
+apply(simp_all add: lex_prod_def union_comm)
+apply(subgoal_tac
+ "vars_trm s1\<union>vars_trm t1\<union>vars_trm s2\<union>vars_trm t2\<union>vars_eprobs xs =
+ vars_trm s1\<union>vars_trm s2\<union>vars_trm t1\<union>vars_trm t2\<union>vars_eprobs xs") (*A*)
+apply(simp)
+(* A *)
+apply(force)
+(* Susp-Susp case *)
+apply(simp add: card_insert)
+(* variable elimination cases *)
+apply(simp_all add: apply_subst_def vars_decrease)
+done
+
+
+lemma rank_trans: "\<lbrakk>rank P1 \<lless> rank P2; rank P2 \<lless> rank P3\<rbrakk>\<Longrightarrow>rank P1 \<lless> rank P3"
+apply(simp add: lex_prod_def)
+apply(auto)
+done
+
+(* all reduction are well-founded under \<lless> *)
+
+lemma rank_red_plus: "\<lbrakk>P1\<Turnstile>(s,nabla)\<Rightarrow>P2\<rbrakk>\<Longrightarrow>(rank P2) \<lless> (rank P1)"
+apply(erule red_plus.induct)
+apply(auto dest: rank_sred rank_cred rank_trans)
+done
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/Terms.thy Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,142 @@
+
+theory Terms = Main + Swap:
+
+(* terms *)
+
+datatype trm = Abst string trm
+ | Susp "(string \<times> string)list" string
+ | Unit
+ | Atom string
+ | Paar trm trm
+ | Func string trm
+
+(* swapping operation on terms *)
+
+consts swap :: "(string \<times> string)list \<Rightarrow> trm \<Rightarrow> trm"
+primrec
+"swap pi (Unit) = Unit"
+"swap pi (Atom a) = Atom (swapas pi a)"
+"swap pi (Susp pi' X) = Susp (pi@pi') X"
+"swap pi (Abst a t) = Abst (swapas pi a) (swap pi t)"
+"swap pi (Paar t1 t2) = Paar (swap pi t1) (swap pi t2)"
+"swap pi (Func F t) = Func F (swap pi t)"
+
+lemma [simp]: "swap [] t = t"
+apply(induct_tac t)
+apply(auto)
+done
+
+lemma swap_append[rule_format]: "\<forall>pi1 pi2. swap (pi1@pi2) t = swap pi1 (swap pi2 t)"
+apply(induct_tac t, auto)
+apply(induct_tac pi1,auto)
+apply(induct_tac pi1,auto)
+done
+
+lemma swap_empty: "swap pi t=Susp [] X\<longrightarrow> pi=[]"
+apply(induct_tac t)
+apply(auto)
+done
+
+(* depyh of terms *)
+
+consts depth :: "trm \<Rightarrow> nat"
+primrec
+"depth (Unit) = 0"
+"depth (Atom a) = 0"
+"depth (Susp pi X) = 0"
+"depth (Abst a t) = 1 + depth t"
+"depth (Func F t) = 1 + depth t"
+"depth (Paar t t') = 1 + (max (depth t) (depth t'))"
+
+lemma
+ Suc_max_left: "Suc(max x y) = z \<longrightarrow> x < z" and
+ Suc_max_right: "Suc(max x y) = z \<longrightarrow> y < z"
+apply(auto)
+apply(simp_all only: max_Suc_Suc[THEN sym] less_max_iff_disj)
+apply(simp_all)
+done
+
+lemma [simp]: "depth (swap pi t) = depth t"
+apply(induct_tac t,auto)
+done
+
+(* occurs predicate and variables in terms *)
+
+consts occurs :: "string \<Rightarrow> trm \<Rightarrow> bool"
+primrec
+"occurs X (Unit) = False"
+"occurs X (Atom a) = False"
+"occurs X (Susp pi Y) = (if X=Y then True else False)"
+"occurs X (Abst a t) = occurs X t"
+"occurs X (Paar t1 t2) = (if (occurs X t1) then True else (occurs X t2))"
+"occurs X (Func F t) = occurs X t"
+
+consts vars_trm :: "trm \<Rightarrow> string set"
+primrec
+"vars_trm (Unit) = {}"
+"vars_trm (Atom a) = {}"
+"vars_trm (Susp pi X) = {X}"
+"vars_trm (Paar t1 t2) = (vars_trm t1)\<union>(vars_trm t2)"
+"vars_trm (Abst a t) = vars_trm t"
+"vars_trm (Func F t) = vars_trm t"
+
+
+(* subterms and proper subterms *)
+
+consts sub_trms :: "trm \<Rightarrow> trm set"
+primrec
+"sub_trms (Unit) = {Unit}"
+"sub_trms (Atom a) = {Atom a}"
+"sub_trms (Susp pi Y) = {Susp pi Y}"
+"sub_trms (Abst a t) = {Abst a t}\<union>sub_trms t"
+"sub_trms (Paar t1 t2) = {Paar t1 t2}\<union>sub_trms t1 \<union>sub_trms t2"
+"sub_trms (Func F t) = {Func F t}\<union>sub_trms t"
+
+consts psub_trms :: "trm \<Rightarrow> trm set"
+primrec
+"psub_trms (Unit) = {}"
+"psub_trms (Atom a) = {}"
+"psub_trms (Susp pi X) = {}"
+"psub_trms (Paar t1 t2) = sub_trms t1 \<union> sub_trms t2"
+"psub_trms (Abst a t) = sub_trms t"
+"psub_trms (Func F t) = sub_trms t"
+
+lemma psub_sub_trms[rule_format]:
+ "\<forall>t1 . t1 \<in> psub_trms t2 \<longrightarrow> t1 \<in> sub_trms t2"
+apply(induct t2)
+apply(auto)
+done
+
+lemma t_sub_trms_t:
+ "t\<in> sub_trms t"
+apply(induct t)
+apply(auto)
+done
+
+lemma abst_psub_trms[rule_format]:
+ "\<forall> t1. Abst a t1 \<in> sub_trms t2 \<longrightarrow> t1 \<in> psub_trms t2"
+apply(induct_tac t2)
+apply(auto)
+apply(simp only: t_sub_trms_t)
+apply(best dest!: psub_sub_trms)+
+done
+
+lemma func_psub_trms[rule_format]:
+ "\<forall> t1. Func F t1 \<in> sub_trms t2 \<longrightarrow> t1 \<in> psub_trms t2"
+apply(induct_tac t2)
+apply(auto)
+apply(best dest!: psub_sub_trms)+
+apply(simp add: t_sub_trms_t)
+apply(best dest!: psub_sub_trms)
+done
+
+lemma paar_psub_trms[rule_format]:
+ "\<forall> t1. Paar t1 t2 \<in> sub_trms t3\<longrightarrow>(t1 \<in> psub_trms t3 \<and> t2 \<in> psub_trms t3)"
+apply(induct_tac t3)
+apply(auto)
+apply(best dest!: psub_sub_trms)+
+apply(simp add: t_sub_trms_t)+
+apply(best dest!: psub_sub_trms)+
+done
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/Unification.thy Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,625 @@
+
+
+theory Unification = Main + Terms + Fresh + Equ + Substs + Mgu:
+
+(* problems to which no reduction applies *)
+
+consts stuck :: "problem_type set"
+defs
+ stuck_def: "stuck \<equiv> { P1. \<not>(\<exists>P2 nabla s. P1 \<Turnstile>(nabla,s)\<Rightarrow>P2)}"
+
+(* all problems which are stuck and have no unifier *)
+
+consts fail :: "problem_type set"
+inductive fail
+intros
+[intro!]: "\<lbrakk>occurs X t\<rbrakk>\<Longrightarrow>(Susp pi X\<approx>?Abst a t#xs,ys)\<in>fail"
+[intro!]: "\<lbrakk>occurs X t\<rbrakk>\<Longrightarrow>(Susp pi X\<approx>?Func F t#xs,ys)\<in>fail"
+[intro!]: "\<lbrakk>occurs X t1\<or>occurs X t2\<rbrakk>\<Longrightarrow>(Susp pi X\<approx>?Paar t1 t2#xs,ys)\<in>fail"
+[intro!]: "\<lbrakk>occurs X t\<rbrakk>\<Longrightarrow>(Abst a t\<approx>?Susp pi X#xs,ys)\<in>fail"
+[intro!]: "\<lbrakk>occurs X t\<rbrakk>\<Longrightarrow>(Func F t\<approx>?Susp pi X#xs,ys)\<in>fail"
+[intro!]: "\<lbrakk>occurs X t1\<or>occurs X t2\<rbrakk>\<Longrightarrow>(Paar t1 t2\<approx>?Susp pi X#xs,ys)\<in>fail"
+[intro!,dest!]: "([],a\<sharp>? Atom a#ys)\<in>fail"
+[intro!]: "a\<noteq>b\<Longrightarrow>(Atom a\<approx>? Atom b#xs,ys)\<in>fail"
+[intro!,dest!]: "(Abst a t\<approx>?Unit#xs,ys)\<in>fail"
+[intro!,dest!]: "(Unit\<approx>?Abst a t#xs,ys)\<in>fail"
+[intro!,dest!]: "(Abst a t\<approx>?Atom b#xs,ys)\<in>fail"
+[intro!,dest!]: "(Atom b\<approx>?Abst a t#xs,ys)\<in>fail"
+[intro!,dest!]: "(Abst a t\<approx>?Paar t1 t2#xs,ys)\<in>fail"
+[intro!,dest!]: "(Paar t1 t2\<approx>?Abst a t#xs,ys)\<in>fail"
+[intro!,dest!]: "(Abst a t\<approx>?Func F t1#xs,ys)\<in>fail"
+[intro!,dest!]: "(Func F t1\<approx>?Abst a t#xs,ys)\<in>fail"
+[intro!,dest!]: "(Unit\<approx>?Atom b#xs,ys)\<in>fail"
+[intro!,dest!]: "(Atom b\<approx>?Unit#xs,ys)\<in>fail"
+[intro!,dest!]: "(Unit\<approx>?Paar t1 t2#xs,ys)\<in>fail"
+[intro!,dest!]: "(Paar t1 t2\<approx>?Unit#xs,ys)\<in>fail"
+[intro!,dest!]: "(Unit\<approx>?Func F t1#xs,ys)\<in>fail"
+[intro!,dest!]: "(Func F t1\<approx>?Unit#xs,ys)\<in>fail"
+[intro!,dest!]: "(Atom a\<approx>?Paar t1 t2#xs,ys)\<in>fail"
+[intro!,dest!]: "(Paar t1 t2\<approx>?Atom a#xs,ys)\<in>fail"
+[intro!,dest!]: "(Atom a\<approx>?Func F t1#xs,ys)\<in>fail"
+[intro!,dest!]: "(Func F t1\<approx>?Atom a#xs,ys)\<in>fail"
+[intro!,dest!]: "(Func F t\<approx>?Paar t1 t2#xs,ys)\<in>fail"
+[intro!,dest!]: "(Paar t1 t2\<approx>?Func F t#xs,ys)\<in>fail"
+[intro!]: "\<lbrakk>F1\<noteq>F2\<rbrakk>\<Longrightarrow>(Func F1 t1\<approx>?Func F2 t2#xs,ys)\<in>fail"
+
+(* the results that are interesting are the stuck ones *)
+
+consts
+ results :: "problem_type \<Rightarrow> problem_type set"
+defs
+ results_def:
+ "results P1 \<equiv> if P1\<in>stuck then {P1} else {P2. \<exists>nabla s. P1\<Turnstile>(nabla,s)\<Rightarrow>P2 \<and> P2\<in>stuck}"
+
+(* a "failed" problem has no unifier *)
+
+lemma fail_then_empty:
+ "(P1\<in>fail) \<Longrightarrow> (U P1={})"
+apply(erule fail.cases)
+apply(simp add: all_solutions_def)
+apply(rule allI)+
+apply(rule impI)
+apply(drule_tac nabla1="aa" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
+apply(subgoal_tac "\<forall>t1\<in>psub_trms (Abst a (subst b t)).\<not>(\<exists>pi2. aa\<turnstile>Abst a (subst b t)\<approx>swap pi2 t1)")
+apply(simp)
+apply(drule equ_sym)
+apply(clarify)
+apply(drule_tac "t1.0"="Abst a (subst b t)" and
+ "t2.0"="subst b (Susp pi X)" and
+ "t3.0"="swap pia t2" in equ_trans)
+apply(simp)
+apply(best)
+apply(rule psub_trm_not_equ)
+apply(simp add: all_solutions_def)
+apply(rule allI)+
+apply(rule impI)
+apply(drule_tac nabla1="a" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
+apply(subgoal_tac "\<forall>t1\<in>psub_trms (Func F (subst b t)).\<not>(\<exists>pi2. a\<turnstile>Func F (subst b t)\<approx>swap pi2 t1)")
+apply(simp)
+apply(drule equ_sym)
+apply(clarify)
+apply(drule_tac "t1.0"="Func F (subst b t)" and
+ "t2.0"="subst b (Susp pi X)" and
+ "t3.0"="swap pia t2" in equ_trans)
+apply(simp)
+apply(best)
+apply(rule psub_trm_not_equ)
+apply(simp add: all_solutions_def)
+apply(rule allI)+
+apply(rule impI)
+apply(erule disjE)
+apply(drule_tac nabla1="a" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
+apply(subgoal_tac "\<forall>t3\<in>psub_trms (Paar (subst b t1) (subst b t2)).
+ \<not>(\<exists>pi2. a\<turnstile>Paar (subst b t1) (subst b t2)\<approx>swap pi2 t3)")
+apply(simp)
+apply(drule equ_sym)
+apply(clarify)
+apply(drule_tac "t1.0"="Paar (subst b t1) (subst b t2)" and
+ "t2.0"="subst b (Susp pi X)" and
+ "t3.0"="swap pia t2a" in equ_trans)
+apply(simp)
+apply(best)
+apply(rule psub_trm_not_equ)
+apply(drule_tac nabla1="a" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
+apply(subgoal_tac "\<forall>t3\<in>psub_trms (Paar (subst b t1) (subst b t2)).
+ \<not>(\<exists>pi2. a\<turnstile>Paar (subst b t1) (subst b t2)\<approx>swap pi2 t3)")
+apply(simp)
+apply(drule equ_sym)
+apply(clarify)
+apply(drule_tac "t1.0"="Paar (subst b t1) (subst b t2)" and
+ "t2.0"="subst b (Susp pi X)" and
+ "t3.0"="swap pia t2a" in equ_trans)
+apply(simp)
+apply(best)
+apply(rule psub_trm_not_equ)
+apply(simp add: all_solutions_def)
+apply(rule allI)+
+apply(rule impI)
+apply(drule_tac nabla1="aa" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
+apply(subgoal_tac "\<forall>t3\<in>psub_trms (Abst a (subst b t)).
+ \<not>(\<exists>pi2. aa\<turnstile>Abst a (subst b t)\<approx>swap pi2 t3)")
+apply(simp)
+apply(clarify)
+apply(drule_tac "t1.0"="Abst a (subst b t)" and
+ "t2.0"="subst b (Susp pi X)" and
+ "t3.0"="swap pia t2" in equ_trans)
+apply(simp)
+apply(best)
+apply(rule psub_trm_not_equ)
+apply(simp add: all_solutions_def)
+apply(rule allI)+
+apply(rule impI)
+apply(drule_tac nabla1="a" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
+apply(subgoal_tac "\<forall>t1\<in>psub_trms (Func F (subst b t)).\<not>(\<exists>pi2. a\<turnstile>Func F (subst b t)\<approx>swap pi2 t1)")
+apply(simp)
+apply(clarify)
+apply(drule_tac "t1.0"="Func F (subst b t)" and
+ "t2.0"="subst b (Susp pi X)" and
+ "t3.0"="swap pia t2" in equ_trans)
+apply(simp)
+apply(best)
+apply(rule psub_trm_not_equ)
+apply(simp add: all_solutions_def)
+apply(rule allI)+
+apply(rule impI)
+apply(erule disjE)
+apply(drule_tac nabla1="a" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
+apply(subgoal_tac "\<forall>t3\<in>psub_trms (Paar (subst b t1) (subst b t2)).
+ \<not>(\<exists>pi2. a\<turnstile>Paar (subst b t1) (subst b t2)\<approx>swap pi2 t3)")
+apply(simp)
+apply(clarify)
+apply(drule_tac "t1.0"="Paar (subst b t1) (subst b t2)" and
+ "t2.0"="subst b (Susp pi X)" and
+ "t3.0"="swap pia t2a" in equ_trans)
+apply(simp)
+apply(best)
+apply(rule psub_trm_not_equ)
+apply(drule_tac nabla1="a" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
+apply(subgoal_tac "\<forall>t3\<in>psub_trms (Paar (subst b t1) (subst b t2)).
+ \<not>(\<exists>pi2. a\<turnstile>Paar (subst b t1) (subst b t2)\<approx>swap pi2 t3)")
+apply(simp)
+apply(clarify)
+apply(drule_tac "t1.0"="Paar (subst b t1) (subst b t2)" and
+ "t2.0"="subst b (Susp pi X)" and
+ "t3.0"="swap pia t2a" in equ_trans)
+apply(simp)
+apply(best)
+apply(rule psub_trm_not_equ)
+apply(simp add: all_solutions_def, fast dest!: fresh.cases)
+apply(simp add: all_solutions_def, fast dest!: equ.cases)+
+done
+
+(* the only stuck problems are the "failed" problems and the empty problem *)
+
+lemma stuck_equiv: "stuck = {([],[])}\<union>fail"
+apply(subgoal_tac "([],[])\<in>stuck")
+apply(subgoal_tac "\<forall>P\<in>fail. P\<in>stuck")
+apply(subgoal_tac "\<forall>P\<in>stuck. P=([],[]) \<or> P\<in>fail")
+apply(force)
+apply(rule ballI)
+apply(thin_tac "([], []) \<in> stuck")
+apply(thin_tac "\<forall>P\<in>fail. P \<in> stuck")
+apply(simp add: stuck_def)
+apply(clarify)
+apply(case_tac a)
+apply(simp)
+apply(case_tac b)
+apply(simp)
+apply(simp)
+apply(case_tac aa)
+apply(simp)
+apply(case_tac ba)
+apply(simp_all)
+apply(case_tac "ab=lista")
+apply(force)
+apply(force)
+apply(force)
+apply(force)
+apply(force)
+apply(force)
+apply(force)
+apply(case_tac aa)
+apply(simp)
+apply(case_tac ab)
+apply(simp_all)
+apply(case_tac ba)
+apply(simp_all)
+apply(case_tac "lista=listb")
+apply(force)
+apply(force)
+apply(case_tac "occurs list2 (Abst lista trm)")
+apply(force)
+apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Abst lista trm))] (list,b))" in spec)
+apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Abst lista trm))] (list,b))" in spec)
+apply(drule_tac x="{}" in spec)
+apply(drule_tac x="[(list2, swap (rev list1) (Abst lista trm))]" in spec)
+apply(simp only: surjective_pairing[THEN sym])
+apply(force)
+apply(force)
+apply(force)
+apply(force)
+apply(force)
+apply(case_tac ba)
+apply(simp_all)
+apply(case_tac "occurs list2 (Abst lista trm)")
+apply(force)
+apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Abst lista trm))] (list,b))" in spec)
+apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Abst lista trm))] (list,b))" in spec)
+apply(drule_tac x="{}" in spec)
+apply(drule_tac x="[(list2, swap (rev list1) (Abst lista trm))]" in spec)
+apply(simp only: surjective_pairing[THEN sym])
+apply(force)
+apply(case_tac "list2=list2a")
+apply(force)
+apply(case_tac "occurs list2 (Susp list1a list2a)")
+apply(simp)
+apply(drule_tac
+ x="fst (apply_subst [(list2,swap (rev list1) (Susp list1a list2a))] (list,b))" in spec)
+apply(drule_tac
+ x="snd (apply_subst [(list2,swap (rev list1) (Susp list1a list2a))] (list,b))" in spec)
+apply(drule_tac x="{}" in spec)
+apply(drule_tac x="[(list2, swap (rev list1) (Susp list1a list2a))]" in spec)
+apply(simp only: surjective_pairing[THEN sym])
+apply(force)
+apply(case_tac "occurs list2 Unit")
+apply(simp)
+apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) Unit)] (list,b))" in spec)
+apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) Unit)] (list,b))" in spec)
+apply(drule_tac x="{}" in spec)
+apply(drule_tac x="[(list2, swap (rev list1) Unit)]" in spec)
+apply(simp only: surjective_pairing[THEN sym])
+apply(force)
+apply(case_tac "occurs list2 (Atom lista)")
+apply(simp)
+apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Atom lista))] (list,b))" in spec)
+apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Atom lista))] (list,b))" in spec)
+apply(drule_tac x="{}" in spec)
+apply(drule_tac x="[(list2, swap (rev list1) (Atom lista))]" in spec)
+apply(simp only: surjective_pairing[THEN sym])
+apply(force)
+apply(case_tac "occurs list2 (Paar trm1 trm2)")
+apply(force)
+apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Paar trm1 trm2))] (list,b))" in spec)
+apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Paar trm1 trm2))] (list,b))" in spec)
+apply(drule_tac x="{}" in spec)
+apply(drule_tac x="[(list2, swap (rev list1) (Paar trm1 trm2))]" in spec)
+apply(simp only: surjective_pairing[THEN sym])
+apply(force)
+apply(case_tac "occurs list2 (Func lista trm)")
+apply(force)
+apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Func lista trm))] (list,b))" in spec)
+apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Func lista trm))] (list,b))" in spec)
+apply(drule_tac x="{}" in spec)
+apply(drule_tac x="[(list2, swap (rev list1) (Func lista trm))]" in spec)
+apply(simp only: surjective_pairing[THEN sym])
+apply(force)
+apply(case_tac ba)
+apply(simp_all)
+apply(force)
+apply(case_tac "occurs list2 Unit")
+apply(simp)
+apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) Unit)] (list,b))" in spec)
+apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) Unit)] (list,b))" in spec)
+apply(drule_tac x="{}" in spec)
+apply(drule_tac x="[(list2, swap (rev list1) Unit)]" in spec)
+apply(simp only: surjective_pairing[THEN sym])
+apply(force)
+apply(force)
+apply(force)
+apply(force)
+apply(force)
+apply(case_tac ba)
+apply(simp_all)
+apply(force)
+apply(case_tac "occurs list2 (Atom lista)")
+apply(simp)
+apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Atom lista))] (list,b))" in spec)
+apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Atom lista))] (list,b))" in spec)
+apply(drule_tac x="{}" in spec)
+apply(drule_tac x="[(list2, swap (rev list1) (Atom lista))]" in spec)
+apply(simp only: surjective_pairing[THEN sym])
+apply(force)
+apply(force)
+apply(force)
+apply(force)
+apply(force)
+apply(case_tac ba)
+apply(simp_all)
+apply(force)
+apply(case_tac "occurs list2 (Paar trm1 trm2)")
+apply(force)
+apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Paar trm1 trm2))] (list,b))" in spec)
+apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Paar trm1 trm2))] (list,b))" in spec)
+apply(drule_tac x="{}" in spec)
+apply(drule_tac x="[(list2, swap (rev list1) (Paar trm1 trm2))]" in spec)
+apply(simp only: surjective_pairing[THEN sym])
+apply(force)
+apply(force)
+apply(force)
+apply(force)
+apply(force)
+apply(case_tac ba)
+apply(simp_all)
+apply(force)
+apply(case_tac "occurs list2 (Func lista trm)")
+apply(force)
+apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Func lista trm))] (list,b))" in spec)
+apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Func lista trm))] (list,b))" in spec)
+apply(drule_tac x="{}" in spec)
+apply(drule_tac x="[(list2, swap (rev list1) (Func lista trm))]" in spec)
+apply(simp only: surjective_pairing[THEN sym])
+apply(force)
+apply(force)
+apply(force)
+apply(force)
+apply(force)
+apply(rule ballI)
+apply(thin_tac "([], []) \<in> stuck")
+apply(simp add: stuck_def)
+apply(clarify)
+apply(ind_cases "((a, b), (nabla, s), aa, ba) \<in> red_plus")
+apply(ind_cases "((a, b), s, aa, ba) \<in> s_red")
+apply(simp_all)
+apply(ind_cases "((Unit, Unit) # aa, b) \<in> fail")
+apply(ind_cases "((Paar t1 t2, Paar s1 s2) # xs, b) \<in> fail")
+apply(ind_cases "((Func F t1, Func F t2) # xs, b) \<in> fail")
+apply(simp)
+apply(ind_cases "((Abst ab t1, Abst ab t2) # xs, b) \<in> fail")
+apply(ind_cases "((Abst ab t1, Abst bb t2) # xs, b) \<in> fail")
+apply(ind_cases "((Atom ab, Atom ab) # aa, b) \<in> fail")
+apply(simp)
+apply(ind_cases "((Susp pi1 X, Susp pi2 X) # aa, b) \<in> fail")
+apply(ind_cases "((Susp pi X, t) # xs, b) \<in> fail")
+apply(simp add: apply_subst_def)+
+apply(case_tac "occurs X t1")
+apply(simp)
+apply(simp)
+apply(ind_cases "((t, Susp pi X) # xs, b) \<in> fail")
+apply(simp add: apply_subst_def)
+apply(case_tac "occurs X t1")
+apply(simp)
+apply(simp)
+apply(case_tac "occurs X t1")
+apply(simp)
+apply(simp)
+apply(ind_cases "((a, b), nabla, aa, ba) \<in> c_red")
+apply(simp_all)
+apply(ind_cases "([], (ab, Unit) # ba) \<in> fail")
+apply(ind_cases "([], (ab, Paar t1 t2) # xs) \<in> fail")
+apply(ind_cases "([], (ab, Func F t) # xs) \<in> fail")
+apply(ind_cases "([], (ab, Abst ab t) # ba) \<in> fail")
+apply(ind_cases "([], (ab, Abst bb t) # xs) \<in> fail")
+apply(ind_cases "([], (ab, Atom bb) # ba) \<in> fail")
+apply(simp)
+apply(ind_cases "([], (ab, Susp pi X) # ba) \<in> fail")
+apply(ind_cases "(a, b) \<turnstile> s1 \<leadsto> P2")
+apply(simp_all)
+apply(ind_cases "((Unit, Unit) # xs, b) \<in> fail")
+apply(ind_cases "((Paar t1 t2, Paar s1 s2) # xs, b) \<in> fail")
+apply(ind_cases "((Func F t1, Func F t2) # xs, b) \<in> fail")
+apply(simp)
+apply(ind_cases "((Abst ab t1, Abst ab t2) # xs, b) \<in> fail")
+apply(ind_cases "((Abst ab t1, Abst bb t2) # xs, b) \<in> fail")
+apply(ind_cases "((Atom ab, Atom ab) # aa, b) \<in> fail")
+apply(simp)
+apply(ind_cases "((Susp pi1 X, Susp pi2 X) # aa, b) \<in> fail")
+apply(ind_cases "((Susp pi X, t) # xs, b) \<in> fail")
+apply(simp add: apply_subst_def)+
+apply(case_tac "occurs X t1")
+apply(simp)
+apply(simp)
+apply(ind_cases "((t, Susp pi X) # xs, b) \<in> fail")
+apply(simp add: apply_subst_def)
+apply(case_tac "occurs X t1")
+apply(simp)
+apply(simp)
+apply(case_tac "occurs X t1")
+apply(simp)
+apply(simp)
+apply(ind_cases "(a, b) \<turnstile> nabla1 \<rightarrow> P2")
+apply(simp_all)
+apply(ind_cases "([], (ab, Unit) # ba) \<in> fail")
+apply(ind_cases "([], (ab, Paar t1 t2) # xs) \<in> fail")
+apply(ind_cases "([], (ab, Func F t) # xs) \<in> fail")
+apply(ind_cases "([], (ab, Abst ab t) # ba) \<in> fail")
+apply(ind_cases "([], (ab, Abst bb t) # xs) \<in> fail")
+apply(ind_cases "([], (ab, Atom bb) # ba) \<in> fail")
+apply(simp)
+apply(ind_cases "([], (ab, Susp pi X) # ba) \<in> fail")
+apply(simp add: stuck_def)
+apply(rule allI)+
+apply(clarify)
+apply(ind_cases "(([], []), (nabla, s), a, b) \<in> red_plus")
+apply(ind_cases "(([], []), s, a, b) \<in> s_red")
+apply(ind_cases "(([], []), nabla, a, b) \<in> c_red")
+apply(ind_cases "([], []) \<turnstile> s1 \<leadsto> P2")
+apply(ind_cases "([], []) \<turnstile> nabla1 \<rightarrow> P2")
+done
+
+lemma u_empty_sred:
+ "P1\<turnstile>s\<leadsto>P2 \<longrightarrow> U P2 ={} \<longrightarrow> U P1={}"
+apply(rule impI)
+apply(ind_cases "P1 \<turnstile> s \<leadsto> P2")
+apply(rule impI, simp add: all_solutions_def)
+apply(rule impI, simp add: all_solutions_def)
+apply(fast dest!: equ_paar_elim)
+apply(rule impI, simp add: all_solutions_def)
+apply(fast dest!: equ_func_elim)
+apply(rule impI, simp add: all_solutions_def)
+apply(fast dest!: equ_abst_aa_elim)
+apply(rule impI, simp add: all_solutions_def)
+apply(force dest!: equ_abst_ab_elim simp add: subst_swap_comm[THEN sym])
+apply(rule impI, simp add: all_solutions_def)
+apply(rule impI, simp add: all_solutions_def)
+apply(simp add: ds_list_equ_ds)
+apply(rule allI)+
+apply(rule impI)
+apply(drule_tac x="a" in spec)
+apply(drule_tac x="b" in spec)
+apply(erule disjE)
+apply(force)
+apply(simp add: subst_susp)
+apply(drule equ_pi1_pi2_dec[THEN mp])
+apply(force simp add: subst_susp)
+apply(auto)
+apply(simp add: all_solutions_def)
+apply(simp_all add: apply_subst_def)
+apply(auto)
+apply(drule_tac x="a" in spec)
+apply(drule_tac x="b" in spec)
+apply(drule unif_1)
+apply(auto)
+apply(drule_tac x="(aa,ba)" in bspec)
+apply(assumption)
+apply(simp)
+apply(drule_tac "t1.0"="aa" and "t2.0"="ba" in unif_2a)
+apply(simp add: subst_comp_expand)
+apply(drule_tac x="(aa,ba)" in bspec)
+apply(assumption)
+apply(simp)
+apply(drule_tac a="aa" and "t"="ba" in unif_2b)
+apply(simp add: subst_comp_expand)
+apply(simp add: all_solutions_def)
+apply(auto)
+apply(drule_tac x="a" in spec)
+apply(drule_tac x="b" in spec)
+apply(drule equ_sym)
+apply(drule unif_1)
+apply(auto)
+apply(drule_tac x="(aa,ba)" in bspec)
+apply(assumption)
+apply(simp)
+apply(drule_tac "t1.0"="aa" and "t2.0"="ba" in unif_2a)
+apply(simp add: subst_comp_expand)
+apply(drule_tac x="(aa,ba)" in bspec)
+apply(assumption)
+apply(simp)
+apply(drule_tac a="aa" and "t"="ba" in unif_2b)
+apply(simp add: subst_comp_expand)
+done
+
+lemma u_empty_cred:
+ "P1\<turnstile>nabla\<rightarrow>P2 \<longrightarrow> U P2 ={} \<longrightarrow> U P1={}"
+apply(rule impI)
+apply(ind_cases "P1 \<turnstile>nabla\<rightarrow>P2")
+apply(rule impI, simp add: all_solutions_def)
+apply(rule impI, simp add: all_solutions_def)
+apply(fast dest!: fresh_paar_elim)
+apply(rule impI, simp add: all_solutions_def)
+apply(fast dest!: fresh_func_elim)
+apply(rule impI, simp add: all_solutions_def)
+apply(rule impI, simp add: all_solutions_def)
+apply(force dest!: fresh_abst_ab_elim)
+apply(rule impI, simp add: all_solutions_def)
+apply(rule impI, simp add: all_solutions_def)
+done
+
+lemma u_empty_red_plus:
+ "P1\<Turnstile>(nabla,s)\<Rightarrow>P2 \<longrightarrow> U P2 ={} \<longrightarrow> U P1={}"
+apply(rule impI)
+apply(erule red_plus.induct)
+apply(drule u_empty_sred[THEN mp], assumption)
+apply(drule u_empty_cred[THEN mp], assumption)
+apply(drule u_empty_sred[THEN mp], force)
+apply(drule u_empty_cred[THEN mp], force)
+done
+
+(* all problems that cannot be solved produce "failed" problems only *)
+
+lemma empty_then_fail: "U P1={} \<longrightarrow> (\<forall>P\<in>results P1. P\<in>fail)"
+apply(simp add: results_def)
+apply(rule conjI)
+apply(rule impI)
+apply(rule impI)
+apply(simp add: stuck_equiv)
+apply(erule disjE)
+apply(subgoal_tac "({},[])\<in>U ([],[])")
+apply(simp)
+apply(simp add: all_solutions_def)
+apply(assumption)
+apply(rule impI)+
+apply(rule allI)+
+apply(rule impI)
+apply(erule conjE)
+apply(simp add: stuck_equiv)
+apply(auto)
+apply(subgoal_tac "({},[])\<in>U ([],[])")
+apply(drule_tac "nabla3.0"="nabla" and "nabla1.0"="{}" and "s1.0"="[]" in P1_from_P2_red_plus)
+apply(simp add: ext_subst_def)
+apply(auto)
+apply(simp add: all_solutions_def)
+done
+
+(* if a problem can be solved then no "failed" problem is produced *)
+
+lemma not_empty_then_not_fail: "U P1\<noteq>{} \<longrightarrow> \<not>(\<exists>P\<in>results P1. P\<in>fail)"
+apply(rule impI)
+apply(simp)
+apply(rule ballI)
+apply(clarify)
+apply(simp add: results_def)
+apply(case_tac "P1\<in>stuck")
+apply(simp_all)
+apply(drule fail_then_empty)
+apply(simp)
+apply(drule fail_then_empty)
+apply(erule conjE)
+apply(clarify)
+apply(drule u_empty_red_plus[THEN mp])
+apply(simp)
+done
+
+end
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/app-nomu.ps Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,4441 @@
+%!PS-Adobe-2.0
+%%Creator: dvips(k) 5.92b Copyright 2002 Radical Eye Software
+%%Title: nomu-csl.dvi
+%%Pages: 4
+%%PageOrder: Ascend
+%%BoundingBox: 0 0 596 842
+%%DocumentFonts: Times-Bold Times-Roman Times-Italic CMMI9 CMMI10 CMSY10
+%%+ CMTT10 CMR10 CMSY7 CMBSY10 CMR7 CMMI7 CMR6 CMR9 Times-BoldItalic
+%%+ CMSY9 CMTI9 CMTT9 CMTI10 TeX-cmex9 CMSY6 CMSY8 CMR8 CMBXTI10 CMBX10
+%%+ CMMIB10 MSBM10 CMMI6 MSBM7 CMR5 CMMI5 EUSM10 CMMI8 CMTT8 CMBSY7
+%%+ Courier
+%%EndComments
+%DVIPSWebPage: (www.radicaleye.com)
+%DVIPSCommandLine: dvips -o nomu-csl.ps nomu-csl.dvi
+%DVIPSParameters: dpi=600, compressed
+%DVIPSSource: TeX output 2003.04.09:1710
+%%BeginProcSet: texc.pro
+%!
+/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/id 0 N/rw 0 N/rc 0 N/gp 0 N/cp 0 N/G 0 N/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]/id Ci N/rw Cw 7 add 8 idiv string N/rc 0 N/gp 0 N/cp 0 N{
+rc 0 ne{rc 1 sub/rc X rw}{G}ifelse}imagemask restore}B/G{{id gp get/gp
+gp 1 add N A 18 mod S 18 idiv pl S get exec}loop}B/adv{cp add/cp X}B
+/chg{rw cp id gp 4 index getinterval putinterval A gp add/gp X adv}B/nd{
+/cp 0 N rw exit}B/lsh{rw cp 2 copy get A 0 eq{pop 1}{A 255 eq{pop 254}{
+A A add 255 and S 1 and or}ifelse}ifelse put 1 adv}B/rsh{rw cp 2 copy
+get A 0 eq{pop 128}{A 255 eq{pop 127}{A 2 idiv S 128 and or}ifelse}
+ifelse put 1 adv}B/clr{rw cp 2 index string putinterval adv}B/set{rw cp
+fillstr 0 4 index getinterval putinterval adv}B/fillstr 18 string 0 1 17
+{2 copy 255 put pop}for N/pl[{adv 1 chg}{adv 1 chg nd}{1 add chg}{1 add
+chg nd}{adv lsh}{adv lsh nd}{adv rsh}{adv rsh nd}{1 add adv}{/rc X nd}{
+1 add set}{1 add clr}{adv 2 chg}{adv 2 chg nd}{pop nd}]A{bind pop}
+forall N/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: 8r.enc
+% File 8r.enc as of 2002-03-12 for PSNFSS 9
+%
+% 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
+%
+% Authors: S. Rahtz, P. MacKay, Alan Jeffrey, B. Horn, K. Berry, W. Schmidt
+%
+% 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 + extra characters from Lucida + Euro.
+%
+% 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.
+%
+% (4) Remaining positions left undefined are for use in (hopefully)
+% upward-compatible revisions, if someday more characters are generally
+% available.
+%
+% (5) hyphen appears twice for compatibility with both ASCII and Windows.
+%
+% (6) /Euro is assigned to 128, as in Windows ANSI
+%
+/TeXBase1Encoding [
+% 0x00 (encoded characters from Adobe Standard not in Windows 3.1)
+ /.notdef /dotaccent /fi /fl
+ /fraction /hungarumlaut /Lslash /lslash
+ /ogonek /ring /.notdef
+ /breve /minus /.notdef
+% These are the only two remaining unencoded characters, so may as
+% well include them.
+ /Zcaron /zcaron
+% 0x10
+ /caron /dotlessi
+% (unusual TeX characters available in, e.g., Lucida Bright)
+ /dotlessj /ff /ffi /ffl
+ /.notdef /.notdef /.notdef /.notdef
+ /.notdef /.notdef /.notdef /.notdef
+ % very contentious; it's so painful not having quoteleft and quoteright
+ % at 96 and 145 that we move the things normally found there down to here.
+ /grave /quotesingle
+% 0x20 (ASCII begins)
+ /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 % rubout; ASCII ends
+% 0x80
+ /Euro /.notdef /quotesinglbase /florin
+ /quotedblbase /ellipsis /dagger /daggerdbl
+ /circumflex /perthousand /Scaron /guilsinglleft
+ /OE /.notdef /.notdef /.notdef
+% 0x90
+ /.notdef /.notdef /.notdef /quotedblleft
+ /quotedblright /bullet /endash /emdash
+ /tilde /trademark /scaron /guilsinglright
+ /oe /.notdef /.notdef /Ydieresis
+% 0xA0
+ /.notdef % nobreakspace
+ /exclamdown /cent /sterling
+ /currency /yen /brokenbar /section
+ /dieresis /copyright /ordfeminine /guillemotleft
+ /logicalnot
+ /hyphen % Y&Y (also at 45); Windows' softhyphen
+ /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: aae443f0.enc
+% Thomas Esser, Dec 2002. public domain
+%
+% Encoding for:
+% cmmi10 cmmi12 cmmi5 cmmi6 cmmi7 cmmi8 cmmi9 cmmib10
+%
+/TeXaae443f0Encoding [
+/Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi /Omega
+/alpha /beta /gamma /delta /epsilon1 /zeta /eta /theta /iota /kappa
+/lambda /mu /nu /xi /pi /rho /sigma /tau /upsilon /phi /chi /psi
+/omega /epsilon /theta1 /pi1 /rho1 /sigma1 /phi1 /arrowlefttophalf
+/arrowleftbothalf /arrowrighttophalf /arrowrightbothalf /arrowhookleft
+/arrowhookright /triangleright /triangleleft /zerooldstyle /oneoldstyle
+/twooldstyle /threeoldstyle /fouroldstyle /fiveoldstyle /sixoldstyle
+/sevenoldstyle /eightoldstyle /nineoldstyle /period /comma /less /slash
+/greater /star /partialdiff /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 /flat /natural /sharp /slurbelow
+/slurabove /lscript /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 /dotlessi /dotlessj /weierstrass /vector
+/tie /psi /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/space /Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi
+/.notdef /.notdef /Omega /alpha /beta /gamma /delta /epsilon1 /zeta /eta
+/theta /iota /kappa /lambda /mu /nu /xi /pi /rho /sigma /tau /upsilon
+/phi /chi /psi /tie /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef
+] def
+
+%%EndProcSet
+%%BeginProcSet: bbad153f.enc
+% Thomas Esser, Dec 2002. public domain
+%
+% Encoding for:
+% cmsy10 cmsy5 cmsy6 cmsy7 cmsy8 cmsy9
+%
+/TeXbbad153fEncoding [
+/minus /periodcentered /multiply /asteriskmath /divide /diamondmath
+/plusminus /minusplus /circleplus /circleminus /circlemultiply
+/circledivide /circledot /circlecopyrt /openbullet /bullet
+/equivasymptotic /equivalence /reflexsubset /reflexsuperset /lessequal
+/greaterequal /precedesequal /followsequal /similar /approxequal
+/propersubset /propersuperset /lessmuch /greatermuch /precedes /follows
+/arrowleft /arrowright /arrowup /arrowdown /arrowboth /arrownortheast
+/arrowsoutheast /similarequal /arrowdblleft /arrowdblright /arrowdblup
+/arrowdbldown /arrowdblboth /arrownorthwest /arrowsouthwest /proportional
+/prime /infinity /element /owner /triangle /triangleinv /negationslash
+/mapsto /universal /existential /logicalnot /emptyset /Rfractur /Ifractur
+/latticetop /perpendicular /aleph /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 /union /intersection
+/unionmulti /logicaland /logicalor /turnstileleft /turnstileright
+/floorleft /floorright /ceilingleft /ceilingright /braceleft /braceright
+/angbracketleft /angbracketright /bar /bardbl /arrowbothv /arrowdblbothv
+/backslash /wreathproduct /radical /coproduct /nabla /integral
+/unionsq /intersectionsq /subsetsqequal /supersetsqequal /section
+/dagger /daggerdbl /paragraph /club /diamond /heart /spade /arrowleft
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/minus /periodcentered /multiply /asteriskmath /divide /diamondmath
+/plusminus /minusplus /circleplus /circleminus /.notdef /.notdef
+/circlemultiply /circledivide /circledot /circlecopyrt /openbullet
+/bullet /equivasymptotic /equivalence /reflexsubset /reflexsuperset
+/lessequal /greaterequal /precedesequal /followsequal /similar
+/approxequal /propersubset /propersuperset /lessmuch /greatermuch
+/precedes /follows /arrowleft /spade /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+] def
+
+%%EndProcSet
+%%BeginProcSet: 09fbbfac.enc
+% Thomas Esser, Dec 2002. public domain
+%
+% Encoding for:
+% cmsltt10 cmtt10 cmtt12 cmtt8 cmtt9
+/TeX09fbbfacEncoding [
+/Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi
+/Omega /arrowup /arrowdown /quotesingle /exclamdown /questiondown
+/dotlessi /dotlessj /grave /acute /caron /breve /macron /ring /cedilla
+/germandbls /ae /oe /oslash /AE /OE /Oslash /visiblespace /exclam
+/quotedbl /numbersign /dollar /percent /ampersand /quoteright /parenleft
+/parenright /asterisk /plus /comma /hyphen /period /slash /zero /one
+/two /three /four /five /six /seven /eight /nine /colon /semicolon /less
+/equal /greater /question /at /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 /bracketleft /backslash /bracketright
+/asciicircum /underscore /quoteleft /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 /braceleft /bar /braceright
+/asciitilde /dieresis /visiblespace /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /space /Gamma /Delta /Theta /Lambda /Xi /Pi
+/Sigma /Upsilon /Phi /Psi /.notdef /.notdef /Omega /arrowup /arrowdown
+/quotesingle /exclamdown /questiondown /dotlessi /dotlessj /grave /acute
+/caron /breve /macron /ring /cedilla /germandbls /ae /oe /oslash /AE
+/OE /Oslash /visiblespace /dieresis /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+] def
+
+%%EndProcSet
+%%BeginProcSet: f7b6d320.enc
+% Thomas Esser, Dec 2002. public domain
+%
+% Encoding for:
+% cmb10 cmbx10 cmbx12 cmbx5 cmbx6 cmbx7 cmbx8 cmbx9 cmbxsl10
+% cmdunh10 cmr10 cmr12 cmr17cmr6 cmr7 cmr8 cmr9 cmsl10 cmsl12 cmsl8
+% cmsl9 cmss10cmss12 cmss17 cmss8 cmss9 cmssbx10 cmssdc10 cmssi10
+% cmssi12 cmssi17 cmssi8cmssi9 cmssq8 cmssqi8 cmvtt10
+%
+/TeXf7b6d320Encoding [
+/Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi /Omega
+/ff /fi /fl /ffi /ffl /dotlessi /dotlessj /grave /acute /caron /breve
+/macron /ring /cedilla /germandbls /ae /oe /oslash /AE /OE /Oslash
+/suppress /exclam /quotedblright /numbersign /dollar /percent /ampersand
+/quoteright /parenleft /parenright /asterisk /plus /comma /hyphen
+/period /slash /zero /one /two /three /four /five /six /seven /eight
+/nine /colon /semicolon /exclamdown /equal /questiondown /question /at
+/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 /bracketleft /quotedblleft /bracketright /circumflex /dotaccent
+/quoteleft /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 /endash /emdash /hungarumlaut /tilde /dieresis /suppress
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /space
+/Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi /.notdef
+/.notdef /Omega /ff /fi /fl /ffi /ffl /dotlessi /dotlessj /grave /acute
+/caron /breve /macron /ring /cedilla /germandbls /ae /oe /oslash /AE
+/OE /Oslash /suppress /dieresis /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+] def
+
+%%EndProcSet
+%%BeginProcSet: 10037936.enc
+% Thomas Esser, Dec 2002. public domain
+%
+% Encoding for:
+% cmbsy10
+%
+/TeX10037936Encoding [
+/minus /periodcentered /multiply /asteriskmath /divide /diamondmath
+/plusminus /minusplus /circleplus /circleminus /circlemultiply
+/circledivide /circledot /circlecopyrt /openbullet /bullet
+/equivasymptotic /equivalence /reflexsubset /reflexsuperset /lessequal
+/greaterequal /precedesequal /followsequal /similar /approxequal
+/propersubset /propersuperset /lessmuch /greatermuch /precedes /follows
+/arrowleft /arrowright /arrowup /arrowdown /arrowboth /arrownortheast
+/arrowsoutheast /similarequal /arrowdblleft /arrowdblright /arrowdblup
+/arrowdbldown /arrowdblboth /arrownorthwest /arrowsouthwest /proportional
+/prime /infinity /element /owner /triangle /triangleinv /negationslash
+/mapsto /universal /existential /logicalnot /emptyset /Rfractur /Ifractur
+/latticetop /perpendicular /aleph /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 /union /intersection
+/unionmulti /logicaland /logicalor /turnstileleft /turnstileright
+/floorleft /floorright /ceilingleft /ceilingright /braceleft /braceright
+/angbracketleft /angbracketright /bar /bardbl /arrowbothv /arrowdblbothv
+/backslash /wreathproduct /radical /coproduct /nabla /integral
+/unionsq /intersectionsq /subsetsqequal /supersetsqequal /section
+/dagger /daggerdbl /paragraph /club /diamond /heart /spade /arrowleft
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /space
+/minus /periodcentered /multiply /asteriskmath /divide /diamondmath
+/plusminus /minusplus /circleplus /circleminus /.notdef /.notdef
+/circlemultiply /circledivide /circledot /circlecopyrt /openbullet
+/bullet /equivasymptotic /equivalence /reflexsubset /reflexsuperset
+/lessequal /greaterequal /precedesequal /followsequal /similar
+/approxequal /propersubset /propersuperset /lessmuch /greatermuch
+/precedes /follows /arrowleft /spade /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+] def
+
+%%EndProcSet
+%%BeginProcSet: 74afc74c.enc
+% Thomas Esser, Dec 2002. public domain
+%
+% Encoding for:
+% cmbxti10 cmff10 cmfi10 cmfib8 cmti10 cmti12 cmti7 cmti8cmti9 cmu10
+%
+/TeX74afc74cEncoding [
+/Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi /Omega
+/ff /fi /fl /ffi /ffl /dotlessi /dotlessj /grave /acute /caron /breve
+/macron /ring /cedilla /germandbls /ae /oe /oslash /AE /OE /Oslash
+/suppress /exclam /quotedblright /numbersign /sterling /percent
+/ampersand /quoteright /parenleft /parenright /asterisk /plus /comma
+/hyphen /period /slash /zero /one /two /three /four /five /six /seven
+/eight /nine /colon /semicolon /exclamdown /equal /questiondown /question
+/at /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 /bracketleft /quotedblleft /bracketright /circumflex /dotaccent
+/quoteleft /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 /endash /emdash /hungarumlaut /tilde /dieresis /suppress
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /space
+/Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi /.notdef
+/.notdef /Omega /ff /fi /fl /ffi /ffl /dotlessi /dotlessj /grave /acute
+/caron /breve /macron /ring /cedilla /germandbls /ae /oe /oslash /AE
+/OE /Oslash /suppress /dieresis /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+] def
+
+%%EndProcSet
+%%BeginProcSet: 0ef0afca.enc
+% Thomas Esser, Dec 2002. public domain
+%
+% Encoding for:
+% cmr5
+%
+/TeX0ef0afcaEncoding [
+/Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi /Omega
+/arrowup /arrowdown /quotesingle /exclamdown /questiondown /dotlessi
+/dotlessj /grave /acute /caron /breve /macron /ring /cedilla /germandbls
+/ae /oe /oslash /AE /OE /Oslash /suppress /exclam /quotedblright
+/numbersign /dollar /percent /ampersand /quoteright /parenleft
+/parenright /asterisk /plus /comma /hyphen /period /slash /zero /one
+/two /three /four /five /six /seven /eight /nine /colon /semicolon
+/less /equal /greater /question /at /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 /bracketleft /quotedblleft
+/bracketright /circumflex /dotaccent /quoteleft /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 /endash /emdash
+/hungarumlaut /tilde /dieresis /suppress /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /space /Gamma /Delta /Theta /Lambda
+/Xi /Pi /Sigma /Upsilon /Phi /Psi /.notdef /.notdef /Omega /arrowup
+/arrowdown /quotesingle /exclamdown /questiondown /dotlessi /dotlessj
+/grave /acute /caron /breve /macron /ring /cedilla /germandbls /ae /oe
+/oslash /AE /OE /Oslash /suppress /dieresis /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+] def
+
+%%EndProcSet
+%%BeginProcSet: texps.pro
+%!
+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
+%!
+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
+%%BeginFont: MSBM7
+%!PS-AdobeFont-1.1: MSBM7 2.1
+%%CreationDate: 1992 Oct 17 08:30:50
+% 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 (MSBM7) readonly def
+/FamilyName (Euler) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle 0 def
+/isFixedPitch false def
+end readonly def
+/FontName /MSBM7 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 63 /emptyset put
+readonly def
+/FontBBox{0 -504 2615 1004}readonly def
+/UniqueID 5032014 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF5B8CABB9FFC6A66A4000A13D5F68BFF326D
+1D432B0D064B56C598F4338C319309181D78E1629A31ECA5DD8536379B03C383
+D10F04E2CB7C8461B10646CD63AFEB7608468CA0FCFC4D3458FB43D22879B515
+27DD9CCF44C2BFCD95A4DE911E4915FBC02335E9999FD9B546134081D6DA3792
+EC4A76DEBA77635BE52E09986268A919CB48B5EFB1A1301EE0683CB5709BC8CE
+D819D799020CBA673BA39C911075501395B1FD20EAD392C9D5A8C9FD1198C737
+D1A614CF0C0432F29DDEB4BF9DB026908DBE89EE522B7D55DE9BF64AFBE6248E
+2E10466655EB9083E7D23E3F0EE26154F191BEBC9987930CD4B4CABE1275BDF9
+8755EF3D531FDF91D54954FC53F15A38D1E8F8D1E36447484FA2C09D34813615
+838B6330FEAE536D08376E4A0FDDF58CDF5647C9F1FF3A7D1ACAD376DB3CADB6
+9459F7A5D4F1864863B79E9F93A1EDE8B99C3138D26227C01F6FE0AAC800F2E5
+94DD81CF7B1355B642CE45CB532FC5B535D66EDFFEA076C009E87406D9772D71
+848C3C53B7496A5D6B58679EF11E114C5F457C6A0D3CDE50278E4A89D5393B1C
+F877CF4E2142A4D045C4AA9138105D748903BACC28FD43DFEDB341E1FCDBE2EA
+D412498FBB5374D6836CFBEB13D4C2B7B9625C25B037FDA9DCC42F5679C4B3C1
+6340E341F73A9215092C0ACC505A859FA935BE5172F4F6D4A30E73914DBD5297
+7FE0CEB5CD0B92176B8174870F9FAFD22BD2ADDE02B5705B5FAFDEE372F17857
+40C1B4024C9F04375B9CF997E9D0C0F7D82465D678BB9810016E6BCC9C4374EA
+6B2CC834894FDCA891643D9417369458A630FD498794823FFA55705315F0687E
+7592A5DFC8B8D6FE2F3C64B4A4F9D37F5F2200BAA277F2E0BA8E5A84ABC10319
+7BA47E1E8D80405541FC64FF71B84C0010920AAA4DD4BEC0D71E447A356CE717
+6FCD3B21A07E2B32DF50D54BDBD96C634197E502FCA75F122E91314F241D6D4F
+C28E2867440D324E80504245CDF43A7106CFF64E8BA7809A292EF017796AE18E
+F2CD78CCF84F82685BB3A147E4A5B7350541966264D191C1BEC22F6F0FF82DA8
+1D43CD551DF0F9CCC2142D03BE81DFBE02CA779D7E6E51CF2F1D7D48A788F42C
+908AF42E7A50C5BC65AE2FDDBA97B01E600D24882D98FFEA0EC3807B36688149
+C199CF15A74FA8DBB731CA8FD16C13491E39309AC7C8C47FEF5AC2CE02A2652A
+7416CBB2486D61DA08323C6BB1BCE4D3CC421DEB14276C7F1294CC3B5BA2BAFA
+C9053BE2615B7F28FF7418CAC67326884A2CC3402BFD71495D6BB7EFC64450D1
+C94E41977843140A6FD0340C39C88E86CC0C058B87234810E1E900ADFB99C232
+12C88EF58715C1998419546382D0CFAE4297C48C2BD7
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMBSY7
+%!PS-AdobeFont-1.1: CMBSY7 001.000
+%%CreationDate: 1992 Oct 22 12:18:11
+% Computer Modern fonts were designed by Donald E. Knuth.
+% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
+11 dict begin
+/FontInfo 7 dict dup begin
+/version (001.000) readonly def
+/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
+/FullName (CMBSY7) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Bold) readonly def
+/ItalicAngle -14.035 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMBSY7 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 1 /periodcentered put
+readonly def
+/FontBBox{0 -927 1542 750}readonly def
+/UniqueID 5032008 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5F086C44216EBE57F4BA37B479BF1E5A5139D8
+91F3E6DDA157B25D359C5E7FE4CFB264DF1707BB6497D3E074CFF95D9FD269B5
+0B1566D1161680C46D1548BBF384EF8653AF29FD474EA2336C876979AC00AE18
+DC87DC0DDD3164B96FC6C3ED826EBAAED383BB3EC5044BA84A0426B87ED04C9F
+4B3EDFC734C241D9B3D2321619F870FE68BB4BA7C060DBE8FBD12F641E7709F2
+C430491944B78B1E59580798F1B40CAC4D59DE95941217EF1A861DDB0156D5C3
+349AAB13FEFF6C646F6401550F5853BC09B267A6C63639228DF55BE60A99E4A5
+308C616892AA0DC96ADB7CD7AB8AEFA859F69B587B61930596A46A905661E4BA
+DAB5E1CF15C94CF060B7FA600B17162AEC2DEE64A156B3F87248E7A7F88C9154
+8C494273B33483BEDF0BEA4DF6A19941F52AA04717623ACBDE926B4851ED05A0
+28698A1C5AE63A46EF473A4F3DCAF3E73C4FAF0C1077EE7A6504074C0D77947F
+940B16425B3F5834763732F26D3385774A1CACA70C07F58887A0301D1BE530A7
+D7AC00A0664617A0CBA9F6281FA4B9168DBA3C1EDCF915778351E6BD8A9CE7E5
+3E56F2FFE0969E1CFFC83F07C01A3873EE1CE4E124565E8F493BE4FAA0A5D099
+A116CEE4EC6C8CB2E93B42771FFF67680A56501A201E12AFDA8448ABE80BEB9B
+80428F48753C4EFB174B693C69DB81CEF0A0B75C53A9D5B4C5F26FA58059A324
+8D4E9D4E9C54AB8F9C21CD66B9B259F9C797559384A653DD43ED4B9C2110BB5B
+C3A6370921186AEE29EB4E62793ECD96935C3D9D89DC6AAECFB745BABDB570FD
+E7E6248B6ED9C09EDA896ECFCAEECE8C1E8BE20BDE6F3558EF5A32ED390ABF86
+3A585DF34F2B8B9567778BF51A1BF9C1018AEEF42FF1F9AAB1F9F73627F6C7DE
+12747A5031EDFF0C8BBB61D651344F0D188BDFDAFE3CCE9916646C67C437A80B
+F00FCB77E47839B2E76E72333A9E4BB4734434BC32E2263F7B27289D59C17B58
+C262B8D8AE6A95FBC5A3D9F72BB299CA31A6E15E647AC3E4975EBEB7C1B7F562
+A238F35ED7C1787094BE1AC142EA716D2784431BCD49E9FB3C070C5F12D0F1DD
+70AC849C3F18153C05C67EF30CB0AB03B748BF8CBC4B80C9E9F79E960E8C4308
+940655DE1405914E2D32B06EADFDFA08B9B7B2F9B88902FC1A6DC36F968F2D33
+5860D1D1D1768392BB9CFA
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMTT8
+%!PS-AdobeFont-1.1: CMTT8 1.0
+%%CreationDate: 1991 Aug 20 16:46:05
+% 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 (CMTT8) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle 0 def
+/isFixedPitch true def
+end readonly def
+/FontName /CMTT8 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 0 /.notdef put
+readonly def
+/FontBBox{-5 -232 545 699}readonly def
+/UniqueID 5000830 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5F0187316F83DDE3E2D27FCDF6C5CE4F95B6EE
+3317BD91B7921F3039DD35FEA387D5CFB6C6E9DC84C178F3432994FC7FAC6E5A
+ED41A1E2EBA350178FBFEB45944511731BA827167DDAC238FC69A5486B995477
+C469E2E27493B0B711DF8E267D3D5613B450011921685147114106C9472580BD
+F531022F6DF5432B2A4EBC51A8032C7F9689B6FA942D849B29709631613DA68D
+4DF7B6F059A19304F40A3C3580CE3B51D79D42984194D4F178801720892FB6E7
+61FF43C63F9256B5E9F4227B1378222BAAD4D52C77462DF01892220E11129C16
+6C9E45BB9F01ED7C1AD5D8B4D72BE0E12969AFEA90FEF170603CDB91CB243173
+B19A56084D10293B80A35275F41BF78A054DDC98F4A1FFF592463D944960FB31
+6BE5F03960F9B1F213CBCC7FD448657FE388F10104D42B0715FC9571CC60CF23
+C72560CBB8835A0CA208FE06676B3B48B093CB7FB2C0C53AF17EC5B372A9771B
+BFD52FFB7062B4FE0106A01A2A1A1DD4EF5C8C7623EC9324A2CB3B402FCC1FCE
+52BFC8662F8A39D5F1B41C97E7CE34E16AC28A1E94007AEA7D4C519399F1B7A9
+48FA7DDB671067244F09C29F95DD60668223F45BBDA8B1C452E930A9F3F341C5
+351D59EA87462FFB30277D3B24E2104D4AAB873BB2B16DA5B23BEE25BE2C8128
+C4CF2F4F438A4E520CD932BAC455BF8775C27AEA6C73EED3EB2F8DB5E356AE27
+41B35C8AEFE73C4CD6A591AAE4F45762EBD6D3636C03F08C552BBFD0A13D11D5
+491F8369B4BAB8ED9D6F1DE7DB7AFD383986C4338D3AA71C9AF2B8A0955CFD86
+0345F16D9798B25156DDF826A7CB6A0CC4CB43078BEBD3E499DA95562A08EED9
+7CA27B7A0CE3FA7EBDAA87A6042231493B69CFEA7AE1992BCCF33B408CD50C7B
+8665EF678BB3B42D17D703B63DD8B76E3ECA117D3DE545A88A49AD0CB49D1B5A
+C920F89299DEBA6CA0B38D0131DE51CC2652D597386B659849699B7AB3465F93
+256D9CCF0A07891D6F310750D72ACFE1A498583FE2C167B057FB3EA6EC057AF4
+0A96B2399719B8902403E7833B7EB6F117C3F96FB7ADC2490E680A7ABA9BEA84
+49BD1D4307AB6F6401F913F85D9DDF0B414F1FF79BF7C5788A5EFF60CEA9B2F7
+5436EADECABB4C56C2E6BDFB6DB26AF884D590E545ECBDFD7425DB89C6B078AE
+FABE6479DC4F184445D71C052BCCEDD9A85AB9C94431E2BC305564F43C9BD36C
+BEB52E69ED948933A791CA47FC35625B97BF321BA9118DE302F526C7728194D2
+69622B146974D1D5105DC891279E8B0F7BF7CDEED7005AACDF972745E6290F2E
+38EFA0EA88D22A366FDC299883554B30D1A9C0EC5F4BEBBB72BF069E84EE3460
+1FEABC64557779EA9871F92473F25DB12E53EB4E0067D72EF60BF798DD4B96C4
+A051F4877F9040EA9881602F814996038FD60CB721B2BC34DB5115DCCBE09322
+43C75F2148B8CED5D250932CF91235019CF81300388EA235D0B0CD180A8D7ADD
+FE8EEDBAC7828A02B86A150FC415179B8083B3AD2D8C438E19CBC1C110E132CE
+7ABD9C28BD953636EE20FE72F56DE320C2E85690F15B774B89A0046F4EFC2D2D
+160FF39BFF8E71F2E5096F2B9464257703DA24538BB4764A43AD58A1C66CDEBE
+570D84E2EFC62A29274AF650A0843276C6A1EB4257704655EE1D11BBE8B8A0FB
+7AA7E93971EEF39F3EF2DC11B795D10581F89850C8E691C759B266E4DFDA8006
+AFDE411B07E822021E27BB6608F9206F239DA5E3F743D8780428809AE6649CC8
+628853E96C15B56223DC454EAC9750BB27457FBE40C34149848E11320BF06883
+1A1BA06C3A2BCD716E52AC4FA9DF97F82203CEC315A046C215A7A474A7BE6CA2
+D837D6CA08651D247FE2132CF89A530EC5CE24309CD42A63B38BD42B4B9DF104
+64A32984AC2DC3B4285D54815EC8D0702A1CAE2D0B1B42BE55E56EABC131D8C4
+282C11C95997CA005109CF71D8D277A468E23D101DB285250AEC26A778A2F3B4
+007A88C1B7FE3B4CB4DCE318824700203E4555816ADEFFD13A1C4D2A3485D1C2
+681482F0606BA2936543A623E121D140B70AAA459B0060FEABC5DF65C2470511
+07277A4E5B30B31A1DF870B571EC276A8EC80A11794DC5F1D444218690D89DCA
+4501F78C16A3118BF6A557F5A1A30E38A52145A545BC101DF51A51DE563001FC
+DB2773064DE8EB005CEF2D1F822E85B889104DB7DFB72D7E93068495531315EA
+C4983BFA1186ACE72A45E709B846D2D5861AF4745D360B8E39F14E37575041C4
+F72697AF00EC75F46C4021E4B03A6E1FE26326667B08F6FB584F6B7E0133D5FC
+DC9CB6C5B67BB5C7C52D115D189B796C
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%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 0 /.notdef put
+readonly def
+/FontBBox{-24 -250 1110 750}readonly def
+/UniqueID 5087383 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D03D6A8F05B47AF95EF28A9C561DBDC98C47CF5
+5250011D19E9366EB6FD153D3A100CAA6212E3D5D93990737F8D326D347B7EDC
+4391C9DF440285B8FC159D0E98D4258FC57892DDF753642CD526A96ACEDA4120
+788F22B1D09F149794E66DD1AC2C2B3BC6FEC59D626F427CD5AE9C54C7F78F62
+C36F49B3C2E5E62AFB56DCEE87445A12A942C14AE618D1FE1B11A9CF9FAA1F32
+617B598CE5058715EF3051E228F72F651040AD99A741F247C68007E68C84E9D1
+D0BF99AA5D777D88A7D3CED2EA67F4AE61E8BC0495E7DA382E82DDB2B009DD63
+532C74E3BE5EC555A014BCBB6AB31B8286D7712E0E926F8696830672B8214E9B
+5D0740C16ADF0AFD47C4938F373575C6CA91E46D88DE24E682DEC44B57EA8AF8
+4E57D45646073250D82C4B50CBBB0B369932618301F3D4186277103B53B3C9E6
+DB42D6B30115F67B9D078220D5752644930643BDF9FACF684EBE13E39B65055E
+B1BD054C324962025EC79E1D155936FE32D9F2224353F2A46C3558EF216F6BB2
+A304BAF752BEEC36C4440B556AEFECF454BA7CBBA7537BCB10EBC21047333A89
+8936419D857CD9F59EBA20B0A3D9BA4A0D3395336B4CDA4BA6451B6E4D1370FA
+D9BDABB7F271BC1C6C48D9DF1E5A6FAE788F5609DE3C48D47A67097C547D9817
+AD3A7CCE2B771843D69F860DA4059A71494281C0AD8D4BAB3F67BB6739723C04
+AE05F9E35B2B2CB9C7874C114F57A185C8563C0DCCA93F8096384D71A2994748
+A3C7C8B8AF54961A8838AD279441D9A5EB6C1FE26C98BD025F353124DA68A827
+AE2AF8D25CA48031C242AA433EEEBB8ABA4B96821786C38BACB5F58C3D5DA011
+85B385124615C1B216CC43CEF394B2DC098149B707283AE24ED2EDA26BBC4DEB
+500B32CD72FC969DF53D9265D208A96F86123F96A29C9B119606A65F968C206A
+14304460C26AD6A835094168ED1DB681C827A47715E9D0B2BD7CC130A086B5A4
+810C9A85A450B2D4E33FDE76FEF5E8A9B1F17C1D97A7F7D02A44CBF4CD2B38D3
+61A6A5E5076BC04D7BBA6B908DB8B44E725AF39D2BF6716FEADCD01A0BB0DA5D
+6ADB5443E333C877FAEF437C9C20AC0C0376F2CEF26C20C3DB0CF9D2FF5AEE54
+9C424B3E9F7D7C799F3373A513A73CC7E3E12C6A69C2A7B77B66C43C916AC102
+08BAC14188F6976CA3496876BECBB33404D17B6267FD57E49091E0C34A6B1993
+37A80608F0EE39B9B9D697879ED565F5BB5AE78352AD941C631289AFF8B1FA78
+06262A4AB0520EA586E582BD03AFCCDA606CC6D1037FA7AD2893D2136F5F6275
+B31F8848B8D0FB0EF5B50389DB77C89010E1FAEB7DD6B0383DD2E05A2633B8A0
+AA88382182299E12CD43D048C310EE75BD63A684CB79F8663C6A9EE87A33F1FE
+81731211FB99BEF57604F0360B5BD72EA53BD52E15EAA18A3044BD6354682D6D
+BA1EB4A0FF5182D0465E4C9EC658EB329922A44379BAB59EBEC22230BB6D7AA3
+6C0AA156F01FC97CB7D5F139F68AAD63FE3A2C812DA71744DB92F6CA8C803D4C
+EDF0F95906A45F968CF6DC5E03495A326656FC5A5031E6F2BC9CAA1215714EA9
+433D751B2D13CC1C7FEB8A0E05BB416E2C9189F1EF1281883FAC5D4B9331BE02
+D7F586C7402C197B4CCF41556E1181E2DE99D8C4ACEF13250E15F066B1C814F6
+607A8AC269917EF48DCF722FAEB3681D8E936E63F4F1BF4BAF1124BCA7F9917E
+5E103092E4F8DE83FB89F70BFD65F41C85B32AB071F094CCEEC01ABC7A80BCF8
+B8FE028BEE1BE87A546B7659415CC3FDBD9F4966DB45905496B2211DA0B9DA2F
+C018ECB798B7B0855AE9C2EAF3FE9D616B70E47E1CE58BADD52E0A7F395598FD
+35B3D28C185FAF23B9F5F637B563E88652D17B5C7C97135BA3FC2BFFCF166FEA
+4E1AD6677FBCAA6955DCFA22BD16A5A5F6A695D6CF9FD9FB840DD5E847449C86
+E93FF01531490C39462CA855083F3E26F87CDD6852075A6193038015325C10ED
+B70ECC11E5945F3607D4D5ED0A8F4B21F70A5ED0D4284F6FFB084317855DB040
+ADB763E28259D7168C7E81D36CF650B0A8870887A69596D8710EF17EDADCD702
+F8DCD43B9F4CADF091CA0730A537EEAAAB1291CD0715A5772AEA204B4B8E600E
+3705FD6CF889F34D6C19B3592FC4B959E25B5C830F9674E636A189C866FAC222
+A847843A712C50E4AA0764F9047BF0EF9C3C976CE31BC8EE43F7A721E4EC5092
+DF6644A6FAC4E6AD4D35A054A729A97F5B5D760A8E7A2F387CF294D75D04CA7F
+F657E348BADC737807E31C6639B0169E557612C185369E8C0B1DB33C8119A4C3
+E0242C7A08054010E34680CAB60F9F0DBD3DD4B22035206A84F2DBBCFF995060
+3E28A4E91C95FDDA07A28A745F363D60B9D4E13B84CFD02D58BA138547BC55F4
+504FEE6C6E483759C21BE46D283FE78D7F4C60E49461520E5894F8FD4EB94993
+DA0BD0F619BC8555D1925EB31E7F00F0CCC78C3A37FB6FEC198489AD9E0BC573
+E268E35688E8D93B070D13AA1727E801153EACB130308F555D14CD499312D83D
+A61CC3291EBB1E776A75801DE27156156F4FEEA183640CB92793EF9905F73320
+8D65A4EE322A787CA5879B5F51D147A3F12622E3E3AFB9BCBA2A54C40F629961
+B1FC21C2630032E238CCA525E3B92B3E9F372C8884271E3A4B1FF467459F3CDD
+275DC4172F1B5537142CAE4DF9D82789C5E58B13B2CD1F3173E95F258ABB2833
+C7B24A65D2FC81A1F3A992C4235A47EEFF071830B35D99230DCE906457102A3A
+36CA9F76D1FA5D0082F2496A37868162B2DA9F4CC5CD4D8DFF4F1F8A41756C69
+
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: EUSM10
+%!PS-AdobeFont-1.1: EUSM10 2.1
+%%CreationDate: 1992 Nov 20 17:36:44
+% Euler fonts were designed by Hermann Zapf.
+% 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 (EUSM10) readonly def
+/FamilyName (Euler) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle 0 def
+/isFixedPitch false def
+end readonly def
+/FontName /EUSM10 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 85 /U put
+readonly def
+/FontBBox{-8 -194 963 738}readonly def
+/UniqueID 5031988 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF4E9D2405B169DE483FC3C5DBA587E58D683
+9E11948C1A3E8B5360ADE57410E9F910E561F679EBFBEB0730125606CD072EB1
+0A975FC5186A70BA3CAA77642D606B57469AD72289DA911758623D8C66DCB000
+A804D330FA2310AB619846B77495D97E7065EFF7E9863C2B64CF640C06301DE3
+EB4EE29B52BEA3204AF4CCEA9C444F88AD777F296CE544EE4A1E88E093DB16EE
+35D8C6473B422A716796536C1347DEF0115B8E31BBDD72D574FF3800E34CFA96
+9322738BF68A0EC2D78FB0CAB62512B656837C6184CA2E03BB30EC2396428702
+0EFE3DD99016DFF0BB433068C23FF079C1EA255E1510957CA6FA2B6307274252
+0E255FEBD892EC4434A09579FBE9BF52000882448823D29A95205EBAC8E4CDAB
+252C9C21441B4E3616AE348DE997113556FD7D4A14256BE54BA2EFF1F19A01A6
+4824F691EBBFA99BF593940B4A376C6FF81E89B0E672EEF95D12E8E962873673
+A80D036D1687EA75A40ECCD264BB3228450B425867FE2CE54BDA69C86A2210D8
+3B62AF190E473AE7F045B7DCC33B892F8009272D22B951B29B74690CC2FB8740
+0C1A7B0327C30AD0CF027A06E8CBE8CBE7C2224BA8E0FE349B5BF30033D027C2
+6EF80F00AC06A2AC1CE7C237ECB5F3616EE0C1CDD06A47F952F8F455172C4191
+59CC7C7D9378CB4CF4F7FA17D9FDBC94024D0323F115206DDBDEE9C97DB4A29F
+174BBD203ED0FDA1BFF0AFF74781F045E3D3F23AC926F472734F1E3C858FA60C
+1C5CC4DECF246DFCBBD6E7FDA25AA240952DA6118ECF0EA4D3F589A81E541018
+43E874DBE23F815D1728DE8699A74BE80B2F63553EA85A982FFAA0679AC682ED
+A87E0015ABA520C142F0F21214AA560C249278AD062087755F734029D3A9342C
+A45F4894A7AB991A660F4597ECEDF4D95E9866BD9ED2BA1BFCB7B4F60BBBA440
+C91605E3B358E1F81250DEF10D877374E9F61F161231B704F67DB722D50172F6
+A1B5C0E10718A5F8A8F7A0FF8C63CB49460AB402064F1C141E2BD409534AA315
+9801843C2F833633B222AE74F93A1A74B27935E9E133ED7E1C7FD93187C3E4E7
+B26DC68E12F835BCF5D24EAE7FFB4BE275AB5940B39D829C85F21B68C8B8757E
+889BB1AA153FDFDA189EE0840B591F2A9C9765641B560C177200EF108B7FCF5A
+6B3A42E3A61E22FE4771CD6BEB64E15640A52905DA14F4FAA6BD086244DA9593
+757DB7A4770E04BAE20A030E328A3DE2F6C9726F857373A0AA578F83FDFC0D41
+FD20887CA66B59C0F001AE122EAE21C92F5D54844EC4B33E0D09C248FD0A3E58
+9C865B9F37211292FF2BDC19E178469FDA43175560D02E69EF7C0F2E9B51E77A
+B90DD9B214A2B640B6F24AAEDB415F21BD2B81CF9996DE91042230BC20639EF6
+F2E0ED12EFEB6991C484992C5AFF42FC2F4CB67406DB8277EE0D743CAEEDECB6
+C1F662FFDFB8147D00131C20709598FAF2F16D656BAE8321B1F10B852738F436
+00FA43B92FADF4C987484716026CA895FF550062E81CB64C1EC9F97E7626D454
+CFD1F657E4802AD5834851315AF4C97A5C
+0000000000000000000000000000000000000000000000000000000000000000
+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 63 /emptyset put
+readonly def
+/FontBBox{-55 -420 2343 920}readonly def
+/UniqueID 5031982 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF5B8CABB9FFC6A66A4000A13D5F68BFF326D
+1D432B0D064B56C598F4338C319309181D78E1629A31ECA5DD8536379B03C383
+D10F04E2C2822D3E73F25B81C424627D3D9A158EAB554233A25D3C6849ABA86F
+1F25C1667CB57D2E79B7803083CB7CC0616467F68450D9A3FEAB534EB9721003
+DBFEEFD050F3AC3492F5C74162A9A531ECEC0F47610B4940E946D21CAA771D30
+A6C27ECBA11708CC46C62396BF9D1990D579D0C394899D24FE7A4382EA18E7E1
+160E7283AF5BE17254790628E79FCC206F28B5566075B3A5697D5209062544FF
+D85FD89D6F43D6588B242AB2666B5D2861CD38A8CE676503EDFAE84D12A71E77
+8405E468FE391F4F3F50D2C57ED55512036B0DB8E76A7EF413ED08673E56DE2C
+16A3B65CD478433C0D2F9FEC4E662D54DAA43CFA6957D2A9AF8979BE06F70B68
+ED4C8C493D6DAC4971A3F1D010A7726D084EC1074FECD7D12D72AE16C26194AF
+21AF5774D9B860EEE8608D34F150092F09C19959BAA670022B9A9F263CD391E3
+74DD1D1B4CD4D75273CAA4E37F68C631723E08FA35AD34C0AFB4621AE6689861
+854D16CE1C375FD159A337E221A6FF1CFFB5693A0623E7EBB58C2969F590D081
+AD92DD9E5322E26D6A15023664AC73A355998BCC48ADD0E7A4BC79790519606F
+A1FEF6075033BCD422EE8233B83D1E7C20043280D531223D5AD4D5B41669F884
+95CE4D6DDE819B588742B930C579EDF743F2C74C95F717FAA6154FADC3FE2975
+F59CFB1C1A29059487E75C48505BAEAD7145667D4E18E46E610C868A257173ED
+0D30EAA4C090854DD8378E92D0A376226EA7DA63798F247BAC770FE26D70E72F
+90CCFAADF118304646955B0310C65F6CA51BEEEF87AFFE294D08C443636DF75B
+DE5A40A671EB63B8FBA940417A22BE3706D2341E62C59351946035638A25BE94
+60C5AECC478C553CF2E3B4CB088D9484D697E831038F8A531F3EEFD41C31449D
+6DDF83FD724F4C81B0D713B73E56FCD2329063033F6A4DAEBF007FD440CF3509
+5A90257034F5827C28890A0994B32EA705AC88782F3AF039F94AF3C6EEA0857D
+3EAA1D1B61597E6AAB0C0CE342E7332A947102F48F28D169AE3575F8EC625964
+520ACB3FFF574B9D39D7BA6CEAD0ACE429F4F147B848FD26C3C2E187CF76E7A6
+FE0CDD9A53C3E4586B1C252E056BF41347B2B8A28510B8E53183D24F5B758BF9
+2A4B744E2DFE23911B0D40A6AA8559D89128A29EE51A806F533B599CA94C44EB
+3EEBB6304B96F2FCEBCE45929D139C3107DC10ED807E1D8B5229216787C20B3E
+E4413D68D0A7BDCF429BDCD1E9AF9B9590BC9D6211B1960CD0FCB32633EB7B23
+67DE12C994F006556E99AB540962208FAB61
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMMI5
+%!PS-AdobeFont-1.1: CMMI5 1.100
+%%CreationDate: 1996 Aug 02 08:21:10
+% 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 (CMMI5) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle -14.04 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMMI5 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 0 /.notdef put
+readonly def
+/FontBBox{37 -250 1349 750}readonly def
+/UniqueID 5087380 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA06DA87FC7163A5A2A756A598FAB07633
+89DE8BB201D5DB4627484A80A431B6AFDBBBF23D4157D4AFE17E6B1C853DD417
+25F84CD55402AB88AB7EEFDEDBF2C2C731BD25567C53B474CCF739188A930039
+098A197F9C4BE7594D79442B2C8A67447DE44698321145D7689B91EF235EA80E
+B600AA8E238064F154284096C4C2554EFE8DDF13AFF8D3CE30E0999375C0FEE6
+F992DEA5FC3897E2CC8B7A90238E61E41622DE80F438DD994C73275CC52249D9
+F6686F87F394FB7BB668138B210BEC9E46415A1B58C990B81E7D7DD301143517
+4C2A259D2A0A1E200F8101469C10D7D537B0D4D39296A9AB3F132DA9A3B459B0
+F850E2B3A03BDCB35AEF82285D19C38F474FB414F8EC971B994D1C7DD753B271
+2B71549DF497C665DF0F266988209D9EB616E4D9BA229FF984E7A886DB01FD21
+48ED2E4859FD6416C2CE52537464EA884C8C9C2D1083E2B83BE4B766474C23B6
+6E8EC5003200AB10514BB44D14CA700416AB6B2683E80862E7D5B49A05526A32
+554BB23AB8B0824BBA198E3825CE82380CC0FECF46651E3E5D77F09465E73164
+20342822F29572BC7F73F2C3BF95ED3BB6FDEADC20C6AC866C4F2C679594D7E8
+8D944704A3C5D771DC39503BECAB89F34D8CDB8FDB91AFE21F3F0260D05E90C5
+73E2C13DFA022C4522E5918EE25038A0498FBB530DA33B0AE238B1C6ED03FC04
+2BFED8236E07820C5BAB411EAE1B31D93A2FA7C374B1725FEC359ABCB88E2C89
+214529A263D795AACB0B95A3AB2F4E08EF350C282CE521716DBB06E5B8291B3F
+5D4ACA230FA192F64BC902A4C8842C0F916F92FBD002ADD408BF0401D0284FBB
+F05D4C6DB631420747CC902C5E1617E6573612FB26C8378DF41FFB5048D3CF06
+4893DBA48EF4B043D760F60C75712169D16C83EE020C45369E443E853E1809DD
+F395B812067D6FDBD26111B34F42C21036AF952D0D767FD17F6959D9FDD46005
+D64FFF54772B50BB9B173AE79702981F58F9F235C591F476A31852174DF0619C
+A470359153DC32610E782B204E7945515464DACE9099B81EEECC7EBD4B5126AF
+C3FD9DDFB329AF1C95C41FA4A5F6958869509A23BD7210386329771FA46FF926
+0E54AC35106253EE140449425A8670E1F92B178A02A58EB57540F4BD8110E548
+BB584EA6D625C5F5FE0124A98E49915F1A1B95D2125874360EED1C4379FEF3C6
+90E5780C20309F11F2F23FAD635C44BA030B39EFF083A3ECCDD2641DCD35B24D
+59A1A8D05205EE919E493B61A71A619F810ADC0C8CCAC8D0E68C82D93CB9EB5E
+5DEF632142B89D317F5BA5363998115AFACB146B840C7DB39B63CE13CA81C194
+75180F8F4A1E059CE9728E64E7EE98523F2CB9108160666FC56DD0ACC8DFED41
+E5F1D0CEDC1F897BFBF406C62F2F8A7D743550B941193F45E787D43C7B29E6B9
+900E72B445315EDC6D94B430E707AF74C2424B8CD7197AC53CE8936CBBF2285C
+4B94F5DE080086D7F96DE48D145B46BD58C1343F1B5DFD928984E88C488BB2AE
+DDFB62A57DD6D0FA0171FF9F63617181395F2170CF1715047501250455E18748
+6CE1B0A8512B4E61FAA76CBF2295FCED2C6F7902010838CD74848580E8C982A1
+98F74525CD02C031B917B0F3B713F779C6F688A6DACCBA0A7CB924C1F4B6E52C
+D1A52A2216B92AAB440499E8EA04033624E9DA018543FD401BD38496FC59ABBA
+4C19614E0072FCDE2FD770EA52807278F203F0703AFFA233833AA63FF44812AB
+CAA3A659434A08E17123CFC2CEC39C706212FF8EB2FA4A1BD87D6FFB8EE2CA2E
+FA9192DAE432B288FAC7E3FB423B8D1BD5B2D81FAE1047EA843971F47BBE9E4D
+182FF0D3959C6328D632FE1C1DEE744A50D54C25AC7129FA1EFCC4A2A51BB055
+4BA5DC6C5558A0DED3A9EB3B70A71EE14C762FC564A6E05613CF37CC4DE74B17
+727DBB90B3EE742F91D5D9A62898970398DC953B5352DD8ECAFE2C7721D1FB1A
+8800A3E0D0025890BBFF0E36B04134226BC5698D146232A1247E128A8AB77E19
+F4E35082B473A91F0B28336BFE89EF6DFCAEEB4ED013BDBFBD20763A448C71F6
+7F8791B394432D797FD77058483F4AFB37BF8FED18CDABE950320B89557200C8
+0504241C52F215F487F5A917DC641312170746806D4555EA5EFC5E6B1C3E48B9
+0E93C9400B1334FB62E82F0C956EF81E6702A933193D66B49471181242
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMR5
+%!PS-AdobeFont-1.1: CMR5 1.00B
+%%CreationDate: 1992 Feb 19 19:55:02
+% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
+11 dict begin
+/FontInfo 7 dict dup begin
+/version (1.00B) readonly def
+/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
+/FullName (CMR5) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle 0 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMR5 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 0 /.notdef put
+readonly def
+/FontBBox{-341 -250 1304 965}readonly def
+/UniqueID 5000788 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF7158F1163BC1F3352E22A1452E73FECA8A4
+87100FB1FFC4C8AF409B2067537220E605DA0852CA49839E1386AF9D7A1A455F
+D1F017CE45884D76EF2CB9BC5821FD25365DDEA1F9B0FF4CFF25B8E64D0747A3
+7CAD14E0DBA3E3CA95F10F24B7D5D75451845F1FB7221D7794A860756CFBB3E7
+704A52A22448C34812C3DBEDD41892577AABA7D555E9298C1A0F7DA638078167
+F56E29672683C51CF1C003764A8E7AD9D8ADE77B4983F56FE2D12723AAD8BF36
+682CFBB71B1D12210144D39DD841A971F71DB82AC6CD815987CDCF29ABC3CC96
+5EEBD5D661F452C6E0C74F9ED8D0C5B3755551A172E0FE31EA02344176E32666
+14B6853A1C303A5E818C2E455A6CF8FC9A66DC6E279101D61C523BD9DB8EB82F
+EAF4D7FDF6372383C0794C4568D079648689A199D4B65BA646CF95B7647E4BEC
+83856C27A8EF177B3A686EDA6354FE9573E123C12EC4BA56A7E8BFB8F9B75147
+9DD79A743968F36F7D0D479FA610F0816E6267E5CE327686A5485AB72201525C
+FB3B7CA10E1BF26E44C24E1696CB089CB0055BD692C89B237CF269F77A31DC81
+0F4B75C8400ABCFDCEC6443CD0E81871CD71AA3064ABDE882C4C52322C27FA8B
+41C689F827FB0F8AAF8022CF3C1F41C0B45601190C1328831857CBF9B1E7D1AA
+246117E56D6B7938488055F4E63E2A1C8D57C17D213729C68349FEC2C3466F41
+171E00413D39DF1F67BC15912F30775AFDF7FB3312587E20A68CF77AD3906040
+842D63C45E19278622DD228C18ABDD024DD9613CDC0B109095DB0ADC3A3C0CB5
+AB597D490189EA81239E39202CBC7A829EB9B313A8F962F7879D374ADF529BD0
+5533EF977142F647AD2F5975BA7E340419116099B19ACCCC37C5512599441893
+4BB8166C90763910DBD81A481642621D3A0F7F9D4F3D785A3180604F676DB0A1
+A6E146DD376E5882EC89A81B94B432FEB4DA8618CAF0EB6D714AA8C5925A3A69
+1924D3746847C01AEF0BE048B30E8A510BFEE69A8621977D2B44B90947503BF3
+7B9278CD826D697E51A9793F27E53F96BAC74FAD80A0B0D9AAB302D73B2F732F
+300CD65DB8D3378B36252DCC132E1E5F36621EC849C7D18CCB40AC2333BF1C4A
+175955246013F6E44E1499AD97553298052A697B3542718C524F67C4FF26B09A
+6379CA97DCBBD646D800506A5812CEBB73C21FB8019C52DC725C82E5315CB160
+6C197979898E0BFADB0D4D09B22C2D1B8B751FEE2E9A5ED298AE360298BD36F6
+D7BBF112A93ACA4834794457FA02F6A5D420FDD955D27D37602F85A8FA3481D3
+E5F5E2FC734668D1D0F41FFE913B0CF48EB41FA8E30F1193B3B6B9CB189CA920
+732070A0F63B2A091FADB797F9940120B77223F7E625282F1FB96523054E8386
+4249F8F02C47B3FB43E64FBB47FAA521BA6809DEB56A33B45479D203A70981CB
+4F26E61272B329FBD509B81CE13384
+0000000000000000000000000000000000000000000000000000000000000000
+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 0 /.notdef put
+readonly def
+/FontBBox{11 -250 1241 750}readonly def
+/UniqueID 5087381 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D03D6A8F05B47AF95EF28A9C561DBDC98C47CF5
+5250011D19E9366EB6FD153D3A100CAA6212E3D5D93990737F8D326D347B7EDC
+4391C9DF440285B8FC159D0E98D4258FC57892DDF0342CA1080743A076089583
+6AD6FB2DC4C13F077F17789476E48402796E685107AF60A63FB0DE0266D55CF1
+8D0AD65B9342CB686E564758C96164FFA711B11C1CE8C726F3C7BB1044BBD283
+9AA4675747DF61E130A55E297CA5F0182A3F12F9085AF2F503481071724077A9
+387E27879A9649AD5F186F33500FAC8F7FA26634BDCE1221EC0ED0E359E5EA5E
+6166526FEB90C30D30099FBDC1BC2F9B62EFEEC48345160804AA98F8D0AA54B7
+A480E715426651865C8E444EDB798C7E11040AF6E5A7ED1888653C6DBF5E6169
+70BCD9C063B63B561EF165BF3AF11F8E519F37C6FDA2827685739DE2C48B5ADE
+EE84F067D704D4511DBFA49E166D543CFD9ECD7417055D8A827F51E087CD2927
+BAFC7E6CFBD70B0FE969F890A11149D3D44D422C3370495DA9951AEE7253A49F
+3A9444C8CD9158D84117299F7F2332FEB0F94E6ED8BC7AA789A3219BC2F227D3
+3B5BC75FB53B55D72AF4A6A7BB613FA235B11BB37D059FD87127CEF73D5B3FBF
+9F91ABAD78BD9240BD9525EBA78095EA0BDB25D1A19E876F292882EAD5619D46
+D20317A345D931F4FF4EAE6216C27044CBA525E3B917CEA25A04C120466C4B93
+FC720E6BA832A06CCA0A3916CEF0968D49085AEBD243C41A448289A6F05CE3F5
+79148DC112A3CC7E8FF810B8C1A09E05F496C0F1EBA334E42E05C376C98F5F69
+C06C71BFC0A2F3AC9951CFBB143C66FB84F9C4ED27DF70869352D61BD5E11508
+0797B87C709E3C151EB44E478CA576D257DF226C00BEF712425856AF0377038F
+7A17D416F208B7718C27CE7D4CB4DA2755D80DD0B9BA847D5EFC18A0985F26B0
+D383A2BB84CA8D0ACABE15ADE0A4E16A21C67B63701C22EAB3D4965FAD4B4013
+1DB24106B490D120406A76942E3D22E6239729C22475CB74A740F737E532440E
+ED93D7DC7E66828533B1E89492378286A329DD59517B9AA2AA50BA846F47514F
+31499C288FED7E87012280F97AD60D2DD3A7D88EA4326FA258DC8FB0FBEAD3E4
+0593A58F2B958009E288C7EBDF997E48FBA5304879386FE51CEB9C33DA3935FB
+F5E66146FEA53F33D7F98EFCAFA2B74E40BEA16A9956004AB554A50FED460D4D
+34806CB786909E74C75514489104E16B9DB406398D43855D9D868FFAF17236E9
+E18065C5FBEC0A5A22C6AD7764BAB978E5190881A388DF7BFBC664129B099686
+A17652E7EB5DDBDF557CFC0A93F8B703EA7962D8F1343158359FB63C893F8E7C
+D32BFA93B63A9FF4CA03CFAA0CAD0F1A113A4626B5DA09148349C90F68DD0D3D
+749B258484F6E383A21C2DBA7278685C29026D43E9B76FBA07A1B448627A489D
+3500A4217A45F355845308382219E5C66F39518405354F4A32B3237878BB73B6
+614CE0A2276709B17605BE8DADF01E261141E18C807F5426DA9857EBAD9F36C4
+077C32E5495F7B004917FDC05135E1C066F1618621DD643B190CC07EFCE250DA
+6C48
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMMIB10
+%!PS-AdobeFont-1.1: CMMIB10 1.100
+%%CreationDate: 1996 Jul 23 07:54:00
+% 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 (CMMIB10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Bold) readonly def
+/ItalicAngle -14.04 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMMIB10 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 0 /.notdef put
+readonly def
+/FontBBox{-15 -250 1216 750}readonly def
+/UniqueID 5087392 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D004B836D34E88C25F6CE738846C8E2E59A2BCF
+4ACF80A26D78872E9343A0537BC3BD7715F32ACD958D5AAED865BFE129278935
+063A31C2634DE2F9077E0AAAAEB224466B779096D8E3FF0A12AD5157F6603DED
+1A82F3511359143311179080C510740B401C930C96270FD1AB3ECBCFEF5DE53F
+E846BAAE95828D5790922640EF8AB9D7CEBE7669FEA02B72F86872D3D8754A18
+A1629C40A7C00C956F140BC63362478279C36EE353638CD3E249897207A94504
+4400668C8E702058EBF7284C9BDF830A3FC79C7EE900CC4C3664F9767A237275
+CEE3671644A75F1E696DA906B4C66870DBE87F5B4A176920C078ADBE24F55C09
+3D18CDE21B5FBC1C6A8AB18E05EDBEF9D1C1C18B3E6377BA2A688579D4F708F9
+A5CF4F56C5E39E2726106E9713E638775E606464CD674E5DC25CE9A696A65806
+C8E9D206B421E246F18013ACC6C7B2985BA93B1B7D7745CCB25B09957F50128C
+B523A55ACA6A7A2A0193A536E590291ED9D577B527CAD0372E05BFCA1829FED1
+662D06144A5FFA628C587A4FA05B179F1A7E3B23B47765FDC054271A0DBF9C2B
+B4F4771F80D1F7AAD9024868C30DAD5CF728DB2A71D86D53B0E674996E8C01F7
+EF97B225A28872F8AD4752A466E078C2B020EB832F237CB9B5631EB2D2EDDB00
+709D3864CA3A6C3EF18085EAEABC011E9F35C9BE4B5D0B608361F329B5784DAC
+5557A602E9E3C204909D84DB988F0BAB914E87CD685C7DA55C5E0B9F0176184F
+FC39B570873BBF346A0D1DE3942DA05434949A65CE64D8BAB0A091C40F7FF47A
+4FC57CB4420221C7B3EB8B891044B5FB0227009F0F6028D3F28545E63C975F11
+6BCCB67C58A544AE706BA5309FE3919CAAC4963FE7BF37D138B055BD8E4C0DDC
+BBA1C35908837F9AFDE32D80E21BD13D2215F54F9BE68FA2838AD858EF22B4E2
+F4C54C5CC38F832B385198149A52397A0B0C68099B357BF5B169E8592E55919F
+1E71171FC891885BA189E296245D2688E688AA754E424492D1FF161AA097BACD
+B519FB0C11388FB59D2C4F302F5AC067610780A4EEEE49A6A7222A8CD8386D6D
+A6B962C033B250CBD0EF97A76B68059A3C1D19F928D8070DB6CAA55AD41ACDF3
+9F88BD4F52A2ED36AE70B3FFFECB80DEDEF495715C90B5FD00F830FC867917DE
+67902D015FB148A1665C7E306BEEB39589FC817BFD1D82224F14D9B3EF5A320D
+D96CED34F7FA3B429AA2770E1F3040C8E8779441A4A31A80369C2C3FCD3ED3DF
+62
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMBX10
+%!PS-AdobeFont-1.1: CMBX10 1.00B
+%%CreationDate: 1992 Feb 19 19:54:06
+% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
+11 dict begin
+/FontInfo 7 dict dup begin
+/version (1.00B) readonly def
+/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
+/FullName (CMBX10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Bold) readonly def
+/ItalicAngle 0 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMBX10 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 0 /.notdef put
+readonly def
+/FontBBox{-301 -250 1164 946}readonly def
+/UniqueID 5000768 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5F00F963068B8B731A88D7740B0DDAED1B3F82
+7DB9DFB4372D3935C286E39EE7AC9FB6A9B5CE4D2FAE1BC0E55AE02BFC464378
+77B9F65C23E3BAB41EFAE344DDC9AB1B3CCBC0618290D83DC756F9D5BEFECB18
+2DB0E39997F264D408BD076F65A50E7E94C9C88D849AB2E92005CFA316ACCD91
+FF524AAD7262B10351C50EBAD08FB4CD55D2E369F6E836C82C591606E1E5C73F
+DE3FA3CAD272C67C6CBF43B66FE4B8677DAFEEA19288428D07FEB1F4001BAA68
+7AAD6DDBE432714E799CFA49D8A1A128F32E8B280524BC8041F1E64ECE4053C4
+9F0AEC699A75B827002E9F95826DB3F643338F858011008E338A899020962176
+CF66A62E3AEF046D91C88C87DEB03CE6CCDF4FB651990F0E86D17409F121773D
+6877DF0085DFB269A3C07AA6660419BD0F0EF3C53DA2318BA1860AB34E28BAC6
+E82DDB1C43E5203AC9DF9277098F2E42C0F7BD03C6D90B629DE97730245B8E8E
+8903B9225098079C55A37E4E59AE2A9E36B6349FA2C09BB1F5F4433E4EEFC75E
+3F9830EB085E7E6FBE2666AC5A398C2DF228062ACF9FCA5656390A15837C4A99
+EC3740D873CFEF2E248B44CA134693A782594DD0692B4DBF1F16C4CDECA692C4
+0E44FDBEF704101118BC53575BF22731E7F7717934AD715AC33B5D3679B784C9
+4046E6CD3C0AD80ED1F65626B14E33CFDA6EB2825DC444FA620E40822E15C2AB
+8BF88EF4FF68D9C81564F0AC5495C426E3ED091CF4AEAD7696C3B00A0D6FB6CE
+60547AD516EC9CC7B904D6617F81D2AC5086C30E40057B3C8854C4893D3B1D82
+2D61C4414046FFBFDCD7AD22EE19016B50905C92287F9946F2DBA4CB1B18E118
+FD01CB5EB39E77C77A4BE276258F7EF35269272729704F2A26BDE6159D0570E8
+665CC6C2582807B75C4D1EB2BA83CDA1A75141865EDF30C292A5891068DDB8D6
+EFE1AC19BBCD228DCEBF1EABB35DEC8D01CE99F3D836F008BE8B7B46802784DB
+A23580F64D164423C813FB2A92EB6C92B7C7CF906A8F51BB897AA6FA652ADF35
+75F468373F5395AC796F2106C04F7309CBDC13CBF28F70F2178E6DBEADFB3D97
+9879350745DC7050393D7ADD4214B2CB1E2C297BB89D61A3FEBFAD67AF5D434B
+4E875E3382481BD89F3FA82446388893D6B2D5A298D934A2C0024FE3B21A1C9E
+D86F020B142FF68B0604A21F88AE36C0C887533136C540E77C5DE78C2B2F1BBE
+7A63C554C0A46714C26011B7C74D46F43A7CE72857A0EBB961A145D3441D87D0
+811F69DC70BE61636E41E29D33643BF327556F724FBBA79BD1BB7DD49155EDDA
+540949C03B64C991E8E6A0324A41A2A6831C241E1CDD7CDFF4219649B5D36E11
+4CBB4F2E25AEFCDAF56C4F0074103CFA58117204614502CBB7291566181C7158
+FA3DF47EB064FA057627214C6FF0468728E66ECC6741732DB4CA9971F8A1BB67
+E11F08912BA27F745080CF2E3346B8E7BC2BD93579AA0C28CF7038C60A819D4A
+53950828
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMBXTI10
+%!PS-AdobeFont-1.1: CMBXTI10 1.0
+%%CreationDate: 1991 Aug 18 17:46:30
+% 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 (CMBXTI10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Bold) readonly def
+/ItalicAngle -14.04 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMBXTI10 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 0 /.notdef put
+readonly def
+/FontBBox{-29 -250 1274 754}readonly def
+/UniqueID 5000771 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D004B836D34E88C20EEB527CE1124209388A2DF
+E27A8DF298A2693A9D529916AA0B2176E6ED237F69D84A8FEEB36861D1847207
+BE2BD61C6A412FFFEDFF13AFEC32AC7735BCCE5965F5966418A62ECB99112AB3
+3BC938EC590FF6922659125EB67E260BF02885E49BA6019E696D33F0B53606A2
+F515E0C45F323311613A94B838491BAB9FE230C5CC79D22925E3D882799F2707
+C32975A494F0F9513E4D8332E7E54470D9721FBD345CDBB48286F2F19CC6D66E
+BB631DD6476A509167A49CA525A72CA50E82C1D08C2B372DB54C5949C753B632
+2009B761EB90492ACD3CBE6A35CE1B66F3BC4D8DC36827CE4261A703328451D1
+879438479917C1647772999171DCCF1491A1C9086E0C6393506768F8757BD81D
+141C46EB9BF507EEC29962A0072B6C5D8C8588F3D68886CD2606DD3BD2FECCEF
+63245494E93EEA12AAFB06110E54ADC444C7E7619627A48A464394E5DE06EB46
+4C76A2FF010318BBE48B3776C826A265C66515717F7F2E943C60EBAB23D96B5B
+FD514A1C4E79BB3D3D2DEB936F90CD3FABF7B09FF7F564AB5CF4AF6A40E869FD
+395885A88F4A138B3CA6943A2D430BBE43D91F7F17621CAF52FB7161DA3B2003
+82244FB6EE792DCA1722C03392C296C029A2DCC5BAAB3EA03F8DEB039DC83AE1
+763AAB84776A2CCFFAE9EAF0BFDAE417E8BE682D237FFEDAF224AC09C9665019
+165CE32F5349E857177D94AD6396570932E1657ADE4D3FF57A3419946CCD210E
+57E5A1D91CF708395942527D127606350924D71BC21C6F969288B1C8CA3404ED
+E6219985F7301A20621368F74747EAD38990A4C9F2B62913B8FDB93657409FF5
+178DAA7A3A1177EFE49258513602C658739CB3F9AEEE681EE53141DC07F66A60
+4CF9EC7DE660D031BCAEC4145E31E6F55505663A1F3B3CFED3B831E61B26C210
+1946B7EE0405F1009D0B7D0BEB5E76F145813F81E67F72C9C78D545545EE1E3F
+47C96DE9A92BFAA73FDE5EFDF48DB33F03C1E8DB8DE83FDF2EB4CF9DBC77A88C
+9DEADE0C16FB2CD360EAE622282185221400970AA54D8CF1DCF153C9AD91241F
+86B14A0D3B7915A06D979046731386D12EBD547C3C1EFC088C107A26E0A9D07B
+FF0B6CC63D3F7437C530B88AD798947F8ED78C1BA0A1C0CF36A1D5B7A06798CC
+9E968F1849A1F689ED73B445C3AAB49B79DFDF7A306BAF90CE1BEAC3D7A7AE45
+21495CF72E06B29CDAF3FE5744F22F9265094F073EFA10D471F8606C3962C140
+F41F49DE85FADE764206291DE6C7972E90BD2D8A0C5EF6EEE5818C7B6829AC2E
+118CE5B4569B1008079C7F655DCAA850BB9D178471FA1D52D6F0839FDA4F025A
+A4CE0CE57C374F00778D55397FC14D120A106210F13756BE0FC12C139FE2F26A
+D100817D9F6120A2BA168AB199FBFA79FB3F852253604F64AC90D5F51DAFDBCE
+7EC90ADE4ADD0476E486D4E4FFF60B268DA2F966103A653AD55A2386BE3535F3
+23ED41B3A22AAF761FC3579E1F9D8404C5F09989FD1FF35AD09E82AA2FBA1467
+7CDD9BA843B56713A4F2162D0FA5804DE4A552CD76B375C690195D453C78094C
+4560D8F15EF076181D072BB3D154BFB6827B995059B571FB1E43A69C6F60F828
+6194F7269C352D0DFB0A8CB4F5D0EC267E95BB12D6B696B70E21E6484FA9B4A9
+78B7B666C2694EEA77698DD33E9B82961A0B745884F6891A489E5944B9C08937
+0A4132E0002739E2C525123CA5BD0FC108AD5F9D5A47CFCAB015C3E031CD7B43
+93EDC8320DAFC1DFE88D47D448E5B1FAC1718F5205BE90665EFF2F13D2145384
+999B2B04B31BAC35BDD640C0C215C12977EE07E937A32BF97A4F9739E9B06FDC
+0D577A34F6CAFAF1043561D2A66A3EB620226396F9B1F56B81D530FCC050B511
+003C00FA4FFEC509C018E16B295F2D03DD0246353440EA885699
+0000000000000000000000000000000000000000000000000000000000000000
+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 0 /.notdef put
+readonly def
+/FontBBox{-36 -250 1070 750}readonly def
+/UniqueID 5000791 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF4E9D2405B169CD5365D6ECED5D768D66D6C
+68618B8C482B341F8CA38E9BB9BAFCFAAD9C2F3FD033B62690986ED43D9C9361
+3645B82392D5CAE11A7CB49D7E2E82DCD485CBA1772CE422BB1D7283AD675B65
+48A7EA0069A883EC1DAA3E1F9ECE7586D6CF0A128CD557C7E5D7AA3EA97EBAD3
+9619D1BFCF4A6D64768741EDEA0A5B0EFBBF347CDCBE2E03D756967A16B613DB
+0FC45FA2A3312E0C46A5FD0466AB097C58FFEEC40601B8395E52775D0AFCD7DB
+8AB317333110531E5C44A4CB4B5ACD571A1A60960B15E450948A5EEA14DD330F
+EA209265DB8E1A1FC80DCD3860323FD26C113B041A88C88A21655878680A4466
+FA10403D24BB97152A49B842C180E4D258C9D48F21D057782D90623116830BA3
+9902B3C5F2F2DD01433B0D7099C07DBDE268D0FFED5169BCD03D48B2F058AD62
+D8678C626DC7A3F352152C99BA963EF95F8AD11DB8B0D351210A17E4C2C55AD8
+9EB64172935D3C20A398F3EEEEC31551966A7438EF3FEE422C6D4E05337620D5
+ACC7B52BED984BFAAD36EF9D20748B05D07BE4414A63975125D272FAD83F76E6
+10FFF8363014BE526D580873C5A42B70FA911EC7B86905F13AFE55EB0273F582
+83158793B8CC296B8DE1DCCF1250FD57CB0E035C7EDA3B0092ED940D37A05493
+2EC54E09B984FCA4AB7D2EA182BCF1263AA244B07EC0EA901C077A059F709F30
+4384CB5FA748F2054FAD9A7A43D4EA427918BD414F766531136B60C3477C6632
+BEFE3897B58C19276A301926C2AEF2756B367319772C9B201C49B4D935A8267B
+041D6F1783B6AEA4DAC4F5B3507D7032AA640AAB12E343A4E9BDCF419C04A721
+3888B25AF4E293AACED9A6BDC78E61DA1C424C6503CC1885F762BADF2C8DF883
+72667B753F6ED2E112D5B4C52D7D8F12B6F6A035E4059396F735BDF380A19F55
+DBD76BB28F403FD5E1A6F9598E37877EF7FED390AB50697BD4DB3ED460B4BF6D
+0A48B487E717B65DBAD50769458BB38C64BB61B18FDB52793622BED41D2D82CD
+381401E20D1961927C146111D3269973204B3961B0E883CB38A31C7E761D84DF
+FBE36F0E69E30A5F76F936832311615BC41833CBD64D57904685DA4AF6D6B586
+2E3DEF012DC23CC757FDE1F2873E9C85BAE00BF2D57ABD41BF8D4BA86C55B37D
+DB1A2812BB5B6FE702B1B7C65ED238E63374EED003D4BCBC76CCF1493FCC6EAB
+122D60A6B2C471AA8DA8F8F77D96D1F9C30516FA656DE86A5621CA3D1C8422DD
+8C2B27F6A4F82B97B772A37658DB570A792E47852120845407FDA783926EA3F3
+4C9DAEE69D115692907B4E02AB3323D4A7F376C238662245AF012FD9774E5528
+46428164A066C6650FAFED817514C72144D9E22B4AFD79E3046616808A34C424
+2E3939CE2A7329E7FF3904118B4161D2D54FF3965577E93BC3E1B6DDD1F36653
+BD616627D05AEA14AA89D84F88274783F054BCE4839E7C9FB661157E08B5BFDE
+11B3C6BB81CD061125D0425C3B171C846986BC878259F9E9A6385237819F4945
+031C0182AF6FC299B1B4180FC19ABE56EE278A95D861EDDF235A286E0CC9995A
+AFC29DB40D4D582D126446E0394D39F7C0D776490563045662F45E15CCD800A7
+D70CCEB689A234A0189703FBF3A4269E391C779EDE2AA2F3BF31E50C2D114F0C
+DAD93F2D09231C37319B122987BF8C752C3C6718BC272F26EF7DFEB0C03A8184
+A9BE2DA313BFA4BA39432DF20B2D86CEB3B48F4F8A29DF1C6E5B248B39031854
+D7F7FDD2EE10DE9E6F01A0F9C3BB7BC645334BB19C472CBF7901EFEE3400F56C
+301C8435A6232492256466999AF932DABD83F00F06A15E2AC6F6F1987F80329D
+A912207B7E0F2B449407E6495C43484300369C0C6F50D1A002BEF47E0044C7D1
+1B247FD8F7769729F0E32B6D4974F1D765CA9103480F08403F80C03B7AB0856F
+53774EA5579B4B824F7749DBB633425CAF5A390B7E22365BD59A5D9102B19307
+2D280F12B20321572D811AB934648CA916DD2BA99555E0B6BBCB90A03A0D8734
+FD74B8D6AC4E169678A758BB1CEB9598015F2C2D5BE1573B296D79232193B438
+905B86E51B7451B577986E15CFA0D801F9C71F9F94012DACEF31493E40034804
+45A07869D9BB1DDEFEB1C72A241906B8162C732BF424EC1D8A1EC90D9171A223
+B7AEEA8DA678CE3AB570108B33F3C7D08B00CFC50F23892F912CE1B19730816C
+8FDDD6FD1A38D43D65C2E8D915428B8B15949D3E89843BDDF96BDA615E61D8C2
+58124DB2968F22B92B45CF2CBB3970BB16E331B9CF433DE928DBBA186B34468B
+2F0E496D2770B5C80F079E2D3952A831A172EABAC6E0E4B87A8456F26165C2B4
+416F3046DDD25717800B0E06A613D5F1A700B881AA849BAFB037A563B0C63846
+6F5F32D5C2542BD820DD6A808BAF767E62BA9FC14316F5098E6F459C2850DB8F
+DF8A12EE156645BBF48DF68BA6A9529CE5D9BE1DD6029823B073642717EDDD42
+E1B0CD0F7855DAFEB1BCA638B78CC4CAA34461C5982B503CF9542EDFDC05A208
+ADA27C66C75698D80377B2F272131AF7F8FDAFB2D1A571C7994202BECB5DC25A
+1EBD560972AFF379030742D639C8142A162BDF6D65AC83B1B84C4D756C6BB87F
+8648A5210B9A832DD1109DE955F1CC0DF547D37214F9378C8039F0A677A79CB2
+5FDBC57DA6D09BDDCDF38097E20425D436F589C05AFE327ACD8E385AFA0A9A17
+11928F48C6B74AD648D86CEEAD0E8E4D6FF1833E2A0CA9B2AA54BDE166C1BDAB
+F1BB8C0C8314FE95E453FC2D8B99A18D435C210B1D786365DFE302E263EAA10B
+9C4CA6F8C6FFF5337B56A7BDD183144BD4648A8EFB5F132CA36B5205DF89DB3F
+8B65BF3C2838E308C9C9E8E3CF5FB052E38AD8096A3F0A79EC8A5E552E1871AB
+304CB5D340F1B334855F5CD48EAFC0DFC745DB56252704BCF89CF41D459275F7
+4A068AF7032BEC79F4FAEDB3278A8F9DBB9BE8B5A06D5E685E13CA29ABDEC8AF
+079FB8BF771A53B4ECA27DA7A2578A51FBADA0206A58D755E7410ABC1E08A88F
+B1D8E446ECA1B1067EA28865AE01C7FDD717A2A5151086E18847A8F22C32A3E3
+C871CB6631D3E718BD89BEF95F566BBC6BAA2B795DD4C84B2C6EA96FAC0E65DD
+75DB2C0DF9D9C73CC50A03539D29413875DBDC4C58DC47E6C2C7C09A7AF7F893
+376DC4DEDB527546F059F029A32ED4A1F5090B323F22C2BAADAA6C7EA6191587
+61911BAD66E23AE24BAC77EA1779414D9E49EF6548A0DE1AAAC658032B6C8D6F
+F77B1DF27549263CC3928AC6FCDFC14C2513E1190DB4CB7A72B49756D43BBCA1
+8B7DC2C2A7C9C7839E87577EBAAF07CBC7DC4585F21E59A7A87C85C5DB8B46
+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 0 /.notdef put
+readonly def
+/FontBBox{-30 -955 1185 779}readonly def
+/UniqueID 5000818 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052F09F9C8ADE9D907C058B87E9B6964
+7D53359E51216774A4EAA1E2B58EC3176BD1184A633B951372B4198D4E8C5EF4
+A213ACB58AA0A658908035BF2ED8531779838A960DFE2B27EA49C37156989C85
+E21B3ABF72E39A89232CD9F4237FC80C9E64E8425AA3BEF7DED60B122A52922A
+221A37D9A807DD01161779DDE7D5FC1B2109839E5B52DFBB2A7C1B5D8E7E8AA0
+5B10EA43D6A8ED61AF5B23D49920D8F79DAB6A59062134D84AC0100187A6CD1F
+80F5DDD9D222ACB1C23326A7656A635C4A241CCD32CBFDF8363206B8AA36E107
+1477F5496111E055C7491002AFF272E46ECC46422F0380D093284870022523FB
+DA1716CC4F2E2CCAD5F173FCBE6EDDB874AD255CD5E5C0F86214393FCB5F5C20
+9C3C2BB5886E36FC3CCC21483C3AC193485A46E9D22BD7201894E4D45ADD9BF1
+CC5CF6A5010B5654AC0BE0DA903DB563B13840BA3015F72E51E3BC80156388BA
+F83C7D393392BCBC227771CDCB976E9330253ACACB1811D98821C8EAA6108C84
+166B3851634883E1DFD97D1819EEC112FEE6AB74138AA57C32C40B53238142B0
+FFD1C3CEF6429A695FA14A644342233A08FF35F7AAF790722003E849BEEC4488
+3479B4E1149C8E3A67F8A9B81A6E11D007E9D721B1B0C2B21234C7A4A4BFDED4
+07DCA44310553284009E5A6B8283D275CA9504B1295AA2CFDA1E4AD96C96ECC8
+863BFF85AEC30DB31819103878F12028B2BCC4DF7107F0F2BE4BA911E97E8ABF
+201E0D9CB730EA4F4C2419B9EEC4A5EDF821F2CF7AEE08EACF1417C946D93404
+765CEBC79E85742E9A8BB21A01032D4E4A7DB663F7872B1AEE63E66FE0A9EE2D
+DAB71E938EB532F8DA4BE950F8D54A778B2ADA262279A64E6BC6BF2DEE05A03D
+0EDB2759743BD5BBD8A66A5A714F046B8B946C0A2C3A6AC855CF820E0DEC4594
+BB0E8AB3ABABA56BB5D5D400551A86DA941E823C86A5D1BBA9110A78652189C8
+7035B78DAD18ACB5713C30B0CF1662DC8D4BC67F108783717CE2DF4BF6043F2B
+7E5147CAE9E13BE2DE6FA346337D925B10D11ACD206C1ED1A608BF197FDF471C
+4FA1621B03853209B0B17236F88F73BB2DE7F282A9846E24C4BAD4CFF9630BFA
+877B8E2199094C3C3410950025CDD49BFA1CCCE20A14B9F111B776673AE9D60E
+4189385F91501C15F52124A66C6BD79681CFDD97361827AC7FB23E374A934BEC
+2466EFC46A640D4646ACE3C0AF3D1438A135673895BBFDCDD9D452C6C35F0CD0
+72019DBDEFAEA873DBCB1402CA1F6739D49BC3F951B8CAEDC2FE1C5E13DF7DA2
+2411C041B782436AC96FA309EEA8E3E7724272EE5D277C76099045986B835E0D
+5101ECA2CF42E56FC1EB7326E81FC2BC24449AD2F554C74FB8E7EC83D33A2976
+7AE44A57C617628810EF8DED0DCCF628D5DC20C0584946A8A3E1634D0AEE5063
+D2118C760AC05B138ECF6D862EC73EE558EF80E9E25C12EACD5C0B07E06DE0A6
+97B32514A170D7D7A253D09FF665C536C9B628F2428305FD4AF5F38395CEB664
+4484A320C6C779FC1EF8D9D6D5F25E78A59B520AF7B49705072BBA07B3FC94C5
+57C306C697F3258BF17D17E794C5A023F6B4BA203F8F303ACFDA31A4BC2B8626
+9AEC1F2DDC90BF45588425DED6589E59E41CCE53E3F3C92854BC8B72E62603D0
+931D9FD867A7556E09CF5E61FC6AE4E1E589B2CBD912A71EF10A6AB0F554D058
+00F73501F1A95CEBAA75269CF87A6BCC32162BFB44B4EB0D4B058EAAFA38E6B2
+EC9D0707264BEF9044AFF0C7A65CDE973759E1E53C5BBB2272C3CB63731A6D7B
+C0FB5370136A2AAD9E1C7A7D521F683ECB4560A85B559167688812B6D1E4CD2E
+0E3DE6D110E0CCAD3A1A716CA6B49909A32607F9978A7FA913A6E8F30930E7FF
+C127E28B36098B6347DAD6B3F82908BA26356451F9A49B82739F80DCD2792383
+B9593B64B1C42B840DE962AF3452BECF040D8031761ABCD1B943603E3E285102
+779906BAD4613BD0FB80079A098E42F682824FCBFFDAB3D4DBC9CB7A599CA704
+57F6EC1603B9105DA837C7EF52EC545027D613A0F67D08708FB018852AABDB1A
+37D245521036881E8FB39A88E38E57FB01D179924DE9339910AFA525EE495679
+268D80CD9B0D5F6F37890636D31B43220E9129B9C46607F9D01273580D7E3717
+C8C5D2974B4D0947F12FA31CEBB2BD7250718CF75D94F4D498436B79C73495AD
+F9F30BE7481C9BE301EB539B0E2371A20C9EF6B573F6B8A235657EED50D47369
+AC0DEBBB87A35187C750AE04BAFE3B85810800003FC6540AF5648C2FE73BFF27
+679A091269F71592511FF9702D60550F970A9965F9B0CFA0ED486565C7A46D03
+46244FD6828C65013BB22945BD0712B97B5804092F86FEEE1E5C476C2177A898
+EE3137E6790F5189A35C2314D445AF35DCF6F8543D266B44BB8FC728E7BB0326
+9A47B3B78B769CC9AE80E81EBAD96AC8182E1CEB0E
+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 0 /.notdef put
+readonly def
+/FontBBox{-4 -948 1329 786}readonly def
+/UniqueID 5000816 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052F09F9C8ADE9D907C058B87E9B6964
+7D53359E51216774A4EAA1E2B58EC3176BD1184A633B951372B4198D4E8C5EF4
+A213ACB58AA0A658908035BF2ED8531779838A960DFE2B27EA49C37156989C85
+E21B3ABF72E39A89232CD9F4237FC80C9E64E8425AA3BEF7DED60B122A52922A
+221A37D9A807DD01161779DDE7D5FC1B2109839E5B52DFB7605D7BA557CC35D6
+49F6EB651B83771034BA0C39DB8D426A24543EF4529E2D939125B5157482688E
+9045C2242F4AFA4C489D975C029177CD6497EACD181FF151A45F521A4C4043C2
+1F3E76EF5B3291A941583E27DFC68B9211105827590393ABFB8AA4D1623D1761
+6AC0DF1D3154B0277BE821712BE7B33385E7A4105E8F3370F981B8FE9E3CF3E0
+007B8C9F2D934F24D591C330487DDF179CECEC5258C47E4B32538F948AB00673
+F9D549C971B0822056B339600FC1E3A5E51844CC8A75B857F15E7276260ED115
+C5FD550F53CE5583743B50B0F9B7C4F836DEF0E4DAFB5591D919AA50F0078928
+03C748D6E926C1F493B1C7F9573016B615ECF2B62C88534FCAD46221D7A20BC5
+82BE6756538242A42EE1D9B0622CADABCA834E12346D0E0751A8A67B8A222235
+062F00F8D0E42E76FABD3BF8693D5D7D2C03F8119FF0BB7ECD6B879093CCC39B
+2679CD7732C3AA7DEED9549853B53500ED2E67477540E986D30D30668CBDAD49
+2BC50C25F5CAB5CD68887D3B97C8F3FE73E8E83829E3F3748CC26C8A577489BA
+B80CCE5B07DA4845A391A40D7B3ADD509960DD2DD9E880B3F19CC65C4345FC0E
+259E0921EF29391F592989A3B9D5894C6B3821D51E7F9449C6D974ED6C5B55E6
+35D396FBC8F605DFE9A64361FEC5087F42C3C0CE6D5ED9E3A9244403B590D73F
+49A4D1A5C84A23F4B41FB8DEC2586BDD8FABC54FFBD35EEC40964B50CFFB80CF
+307D124F4EA5CAD3061C4A3218D36E199E209D991C24885B006BA6591B714001
+1AE090EE40075BB269086D1CC8D46461D273E53FE575B3A6A6EAC60418D053FB
+0785725CC5E0ED8DEBCD50E7FCB38A16502D52EC454B603AA95624B77645D7C4
+146BDE38D73AAC4C12759C98D2531EF87F9B25A41623E54A84E2F6F340CD0415
+D44A9A2BC4A9E011E7937AD601F8432DC5E096DD37B5B99ED19E0ED513EC262A
+50D3D896E249D230E688F8555A8A14C3EC5D3EC689B51D0DE3981EB4AAD38D5B
+B2928E8BCA4E9A10EFDB3C70B597FDCA9B9F64627FEDDB0F0E64A65FE749B5B1
+D33AB3B756C0FE2C3309AFE0984A74167BDEA0
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: TeX-cmex9
+%!PS-AdobeFont-1.0: TeX-cmex9 001.001
+% Filtered by type1fix.pl 0.05
+%%EndComments
+13 dict dup begin
+/FontInfo 16 dict dup begin
+/Copyright (see\040copyright\040of\040original\040TeX\040font) def
+/FamilyName (TeX\040cmex9) def
+/FullName (TeX\040cmex9\040Regular) def
+/ItalicAngle 0 def
+/Notice (converted\040after\040April\0402001) def
+/UnderlinePosition -100 def
+/UnderlineThickness 50 def
+/Weight (Regular) def
+/isFixedPitch false def
+/version (001.001) def
+end readonly def
+/FontName /TeX-cmex9 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 27 /ff put
+dup 56 /eight put
+dup 58 /colon put
+dup 60 /less put
+dup 62 /greater put
+readonly def
+/FontBBox {-26 -2961 1503 773} readonly def
+/UniqueID 4314405 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA06DA87FC7163A5A2A756A598FAB07633
+89DE8BAE4F093966CD2192CE95EB0F323A6BABFDACCFCF27D91F7869A0E46CA5
+9AAF6905783E8AC1F3F9875A76F97187738432F8D14E61574CB292FFB9740871
+66839799D8CAF6E0DFE00012EE6D46A2B3655F29705BE37FD5EDA1C765AA2CF5
+C5AD37207ED1EE9DB82FF31A33307FFA16911406557336AF92F50B603C7BD336
+73EC060F68318378A6F599DDADA5A21504CADBA1E1F4B1A22962BA1BB39ADC7B
+E8CC92F196549457877C9636A8A7EFAC1C3745644C0FD151C70B9FAD69B02C1F
+FE5ED071CA1CF3D4A70909B6A3986687D8FB87391E0C665A4EBD2993161FA81B
+2B8F7221294CCB11AE65A31E8903DCC3AC1A47765E880ACDDC88509418B04F0F
+2D84FD3007D411EEE2DFF237D99A10430524F07480302DFD698B571A023B08A1
+36F84F09BCECBE34ECC1CFD251EBBF338CFF7C197D9B33CBD9ACA7171370527C
+CFA0F7FD5DE8B62790D7DF23F004AEDA717F35B5E2B260A8AAFDE31164AFA47C
+686735EC47454F42AC5037D97B410DC373EB0CEFE3C41E243EDA86FF582AB53E
+7B56D103AD816F2D7DA35DE239DE30CA5645A377E77A1980B984C195E605841B
+36C82FF23D95B5FF770C3AB37C2D657FD1731E91FA0446C316C68F62626B9447
+743EF1214D5DAE1E55DAC83C3F4233BECD8851018F83825EF41597174C8AE0AB
+3479AAD95424879FD9202B4262BF1570588288CBF492192A72D9F9319DDC6EC9
+920EC6AB4C366A816148992F21C4EE89ACBE56D6B25BAF1DF7254836AFB1B2F1
+A9B0741E5527D45E3DCDE5AC0694FB91A11E9414C108E8F22E7AED7BAE513822
+314577A941F3EAFF49425FA0C579A83C7B0CE2DD19198D31B2B57A73B867BCB9
+9700B1CEC07880C36FCDF287DD1D3AFD5297DA570C6CDED54CB6A020A8D38552
+BFAA69FC34334087D26ED9E289BF3C68CB50CCB909C0D2992E7A5489E3134B0D
+2691EDC113CB03E923E7A944FA48D06E45F88FE0A69FCC925C5D254887D91DE2
+0DF4D3FD2A329A9453704FE4EE360FFD3510903CA86C59EDA303B015CD715FCE
+CD34C6B71E87706297E2AA87A25C1E7EC2447EF2CCDC68D23FA2C6895EBCA5C3
+F3D41BE42830555EA7BBD9C61BA42F4C880CF1D36EE28D86AC905E2DF3617539
+20DC5C9A09DE5919E7CAC4E68F2EAE3B30E138950E823E9384B233013F010DF6
+596291E4B10466D8D3A43C8F3020F9319557315A07F7C64155258D6FC49022DF
+A29A319D14ED56EDB0A1F8357DF8B5D9C78B18496767F028AE4D987CA80DC032
+944F9D31C0F3FA47B1C5E2EA84AAC7A18E351B69046481766A2BA8C657820103
+688A7689D1661792E0E3DC6EF8F0F42B5EDB2BCF3BB991E590AE035D4B9B5C06
+D10D1E6885419EB41734E77206BF05F22DB1A60DFC9CD85FB6B7671D150E34D4
+125240728DAF679AC15AC8686B56FE7CB0EDF9C5128D7863B3598947EB8FD8B1
+DE83CAB6042E84C35F4E7C7B258A5A28CDCC44AE90DD0F1DA91F01A5C0441F23
+5D1DB63EFA06B37417FAFE20E6519D72DE124657EBA395298127A0FCB3386F30
+3605BF883A9DEA6DE588197841691112A02B3A20E0853DA51A0488C7B0B949DE
+73C9DC9C91018F324F6255E72138142ACDC9961A08591461DDBD633BADA66A51
+6EE8ED2A6A66F00A874432BCDE4A0DEC3E959A21428A188EE491108F09E33104
+B6DC0EB76F989644836C4754FFC6408D234A6CCD7492F8808C2AC7B338592630
+DD89DAC767AE555D44132349D09651689AFDD90C8CEF60A532DB4D21DE1B8AC7
+D5825EAD56993F19AF31F2A34DB11C72FC7A64BB722B9CA5C9938AB056792062
+DA167DBE8807E604F26C38BB54AD2C9C20CEE9B1D4C4941F91D9FF8A157A54C0
+35073124D6B331C67642B7D7181820CFB612CB006801AB4341920F60FE8DA498
+537127B7BB7479A1EB2793BB224B266BA2F2AB63DAAB7338A6FCF5694A91EE7D
+ECBDEB494CDDCC127F4C1FB27BF4924D61BA494C710DBCBAAD3D73972FB4DF88
+1F9A77E4A2D694F86575EA98835582FC59F563D2300796AEA2966179ABACA3EE
+E55F7CFB408E84DB87E1E5C129CFECBB3F9E599081E3433F3BFBDF0AB4B90503
+51EBAB4DC1BD9198692768AC0BB36E16F6E00EF90F6B1F424FD5094F995F44AA
+C744100D77A9362DE5F3EF0A436B8830B938AC6077BCE3D85849706A5CD052DF
+AD7C827E39224E63DB43788D6B6AB67C63EBE63F6B1F0C1A3969E2C24AB33F32
+70C2B1248F4C96790DFBDE0240A9A3D291F25776CA49FFF060AF6AC9183B78E1
+A2EBAD33C92B106BC6237A9AF71AA7326AF4B50434ADDB2056F4EFDEABD5F60A
+5E926391B6C8DDAB53489F4A77679229C0F605479B0CA13E5CDA1DEB7EBE032B
+0991C68832526C5EA68A94233625E02D8C539347B4D8C8D5902796841F1737FF
+148622C733F3DE
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMBSY10
+%!PS-AdobeFont-1.1: CMBSY10 1.00
+%%CreationDate: 1992 Jul 23 21:21:18
+% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
+11 dict begin
+/FontInfo 7 dict dup begin
+/version (1.00) readonly def
+/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
+/FullName (CMBSY10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Bold) readonly def
+/ItalicAngle -14.035 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMBSY10 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 1 /periodcentered put
+readonly def
+/FontBBox{-27 -940 1332 825}readonly def
+/UniqueID 5000762 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF17EFB0FFE6C69FEBA8389DCC1923D30683D
+A8CD93F7195D5A07BA2F18CB3FD5FFEDA4D83BF758062134D84AC0100187A6CD
+1F80F5DC15B47D73F69655445AD218A8AD78C16EF96F385C9E2D46F8A330C7B5
+A859EB0610C78FC5CE39715A1C5458D30498C0A339504A74C7E8F84B3DEC1516
+B3ABAA0A06DEDCD5F9FEAA5AC4AE8D5A5BA5EC0B64784454F58049E13467D705
+8F13A22BDED5F93EDDCAB7A1886A5168D25B120F8BBCC23546BC7398D4E3EC17
+138921404C390EB84C3CC243C0FF3DEC9EBFFF3DEA73365F1E4BC2F3AB911B2F
+780946F4F6F49935A54EF955D9894FEB37239C896CF98240162F6A6E9677EA24
+06BEE1F04463C033047F7F972C560213C7A02BFEE5AE5AE5BF72377CED942A6D
+8059E59CF03CD6782BD34BC02AA4FD1BA25A5CBE32569D7FED28EFB4C0F5F7C8
+6DADC1A047CB514E19B36A84D4DB390FFE5B841C390666FE27C712E23E22FC84
+A8670626E8B72700B9EE9F06F2121264C1CF69FEEC3E20897D0D9057032830FE
+A18A4BA2AD5CE10EE4FED4BB9E2A9C06965779827D7CBA93926793A7161454E3
+C5AC6A3AAEB75EC64556142508DE6E37B71058F8B97C1A9B4CEBF74FBD2D6D84
+F5DAA2B04AD30B313070B33789935E83DB470FAB8EC65165679F247964BD0C20
+78291B6E13C29E8B86429C1B90C396729D6BDE4CCF24BE000390D798DA73BBEC
+AC5C9B1AC19B2C660CF1CDEC05289F6CAEF0E43465E3627DE26670BAA825429B
+4B8FE57928267D5EBE38C5BF93F90304EB89DE120F81362FB5A3D374AB25B33C
+D03A8E9E176E41C964625E58A65EA958EF2B089933C06B71E29249A96D5A2395
+DE687A0C60B837B5657B90F8642A27B037E4FFFA82343351B7C36566DB55E543
+704DF628D0D6C4A672B6BF5C32E797279E72EEFD88551A3B581C615C3D9A11C8
+270ECE7BBDE9ED6DAAE1E81635A51F046840086FC9FFE90840982501EACE70FB
+3495CA202A5F29CA2A4F56C99CE83F882A551087BC666D0A90C14A4AC08F5158
+A2903B69BA116FEF3715532F5E441037A44D2648D62E14A3569E9D57ED99D92A
+85DA381440E32FFF9546B9BFD2B14508D42F198C975076E2269C8B2BBF1AE20E
+C463B0EBE68BEF1F29F27E86E7600E0A7ECF879B5350A8B74101625D3DDDAC09
+083BCA5E10DACF
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMTI10
+%!PS-AdobeFont-1.1: CMTI10 1.00B
+%%CreationDate: 1992 Feb 19 19:56:16
+% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
+11 dict begin
+/FontInfo 7 dict dup begin
+/version (1.00B) readonly def
+/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
+/FullName (CMTI10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle -14.04 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMTI10 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 0 /.notdef put
+readonly def
+/FontBBox{-163 -250 1146 969}readonly def
+/UniqueID 5000828 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D03D5993DFC0930297866E1CD0A319B6B1FD958
+9E3948FFB0B4E70F212EC976D65099D84E0D37A7A771C3101D6AD26A0513378F
+21EC3643079EECE0C9AB54B4772E5DCA82D0D4ACC7F42FB493AA04A3BF4A1BD6
+06ECE186315DBE9CFDCB1A0303E8D3E83027CD3AFA8F0BD466A8E8CA0E7164CF
+55B332FAD43482748DD4A1CB3F40CB1F5E67192B8216A0D8FE30F9F05BF016F5
+B5CC130A4B0796EE065495422FBA55BEE9BFD99D04464D987AC4D237C208FA86
+0B112E55CE7B3782A34BC22E3DE31755D9AFF19E490C8E43B85E17ECE87FA8B9
+1485831624D24F37C39BF9972D74E6EC4784727AC00B9C4A3AD3DA1C22BD6961
+7E0ADAF55422F22ACA5E4DCD4DF9FCD187A566B7FB661D0530454D0DD6C6C50A
+7A3875C6CBF8EC7769F32A1F3F7FC1C072BADEC97794D4E90E0035282A170402
+356E5A9CD9ABD80AC4342A5283E458A7269252F4541CBB6452B39ED54D336D0B
+19928E9CD1AB26AD83EB209E2EC75011A2643813053B5DBB0246097C4821B5F2
+C92554E9140BE35B2DBFCD98809A8EC9FC910FDE9E0D86457C70ACB056EBF90F
+244DC0A5BBD455E15D6E3180311D52CF50B0BF7D0A7F64F3A1821E0AEDBC2E7B
+AEB549FE1D51088C153799C6E089B5D5D65E1C4E2D2B430CDF1FFA23CCB25D95
+5C43C8942435D0AAA3D9055FF808F2C3C887A3C469BBD98F026D0A59E26BA9F9
+C2144CFE49A9AD892D4D31764F0AE3A10644AE3966B0A790684B14D11FA49785
+EC5565D2B2E584CBFD85125F3FAC133338DE35361943DCE9AF05FCF2840CE512
+998D42CBEC52B57B79DD63F00985881E8463396ADA47189A94DDF951A78866F0
+B8A3D9197E39335277EF2294308DA70065D910943A34F7D5F2090FB4AA42ED70
+CBA469A9F64B95A6FBA4BC89DBC93765E3AE4723162DF3F9D6BDE77DD5870ADE
+C8900D6346957B84C3CE88A8F9A12D46B8FCA50DF4433B0B8AED6A63B3DA102B
+6DF94E62408E24154BAAC66B2B249C695BC0FA37A28699D9C0F3EE94AA32E3C5
+8F8D7F803B5D25014D43A353D719B14B247A87898A960DF68C0C0BAF70C83917
+6E9F7B3ACC64DBAEF3FDCD3A80C0AB907EE342E543D607556CBE5A9089B86D1D
+E768F27D74A613F3ABF883222A8596B542EBF54E9DCE327B5682AEE5F1A7A225
+BF26E2AEA0F352B9C950B47ADF650E1B2AE31E883EDD884EC90F94761C470EB1
+72F27B74049C2A13EC522271032939B656020D617F4E58DCA88D138F4C84979D
+5EE89221BFD28AE8117B69C318A5EF626F95048D62CEA5EC1E8EFD8955D3ABD8
+0A98F103318B58ACC4D520AB607550EB6E430CF787310F2A00E75E6399514B04
+CAA55903B024D8EB6889518CCD4BB93FFC2073A7E8A29EEE2295AA80B9880A77
+95E12BCFB23A029E0245411CA7389B00CAC4E4EB98B35C5EABFDDFC75C04CCDF
+522B50526860CE41E607EA50F7A17396A67407302E624FEC49C59A30EA1F4843
+E6E1CE968A94F76CFC54D4DA0B7725CE574F247AE90E16523D51D4ADCDC8EF7D
+282903CA6390DBA8DDE0025055CAA90D4A1170C7FD891FE699D55AA42FF93D2C
+E6713D48560660B6C94FCE873094498A33662621C9556C59028E01C2DEA1EFFF
+A67C0B5AF6D103A67BD25EA16E4A5EA8C73FFB7188E67C7373AE2B6F045CD2AE
+26ADB068BF9899C14A00F83BA6DD10BDC3E62434331C34642A882210C522A44C
+1AFB96B8DB802F6BD1DD431C07F8664325548CC419BACBE9A5D2E9AB16B3A5FA
+87169D2824F8112DEFB5000FB6C5BE298156D057E2F75F123543C098B1E865A0
+340B298CC6A373549E762EFE3495FFA4103130A20FBAEECBCB5A59B7E2D3799C
+3D6D0B692F14578E0E74CAD6BDD04297EF2E2BEDD821A81A60F1CC03F0985943
+B520F940E71CBA0944800046BE533A0C55DF4D7FA9FF28A2404BDF50F5EFEEF8
+B64E4EFBA2F41983B1AE564D9E192D61C5E2BDDF439DA6AF7FE4A2944D7A2633
+C373D22C1776F84D9A74FA21A6CA78CE70C29BE2F3EF83C08FAFF81EBDB9BE24
+2FE44526DE800CEC0E80A78021FA8D6578FAEAB3A588EAE6568E484E4BE7347D
+9AAE4EB945F51558A63ADFF47BD790DED4827C31741C6B7EF977EA02D78FC18F
+412F50F36F5DA790BF36E9AED37AEB98729178EA18D061E1B7A594D1923479F4
+4AACEF69DE62DF88FEBF5BC8DDB83AB9793A83CDB0D2DCA7845C8CE752578F56
+E435FE7C2145BA82724B36389E628C64B6A9C6DCEEE75B73D64804B8131971C8
+5E7E2804E80486B5DD5E558CBFF0D3BDE9BED97BE314EA726927FCF2105441D2
+14C003D20B4E2D0CC4F4B19D8C48BD6B59CC410661610A104755B252D31BEEA3
+28AB99519B5DC7F45654154897F80067096F4C113B9A2AF53C713B953EBA59BE
+2070121B39E3468F9BF6A9366CFAA80A3AF9C85B001FF3F8B8126D65C1DA505E
+FA9E978CBBAB83347792AF5D5DF8B320BC5E40CC57CD00211EE576457832D1BB
+C5A52DF4EFB0F0D3180043B1B3D3882518E368F5A748B92BFF665EE2871166DD
+778EE3323297411B03538E831819A435F60AE79889F5006CEEF90CDCA18173A6
+BF6FBB9965A39C062524B24EC429A20EE7FA9E56369FD2E95FC3C407C1688EEA
+B24A6851F094BF7D88BAAB8912D12AAD081DF01CD191153C9943DE73F2326956
+460E84C140E571021815FE96EEE97F5CC3D19256760C63EDBFFF392F77D0C5F3
+A7786512A9FFD935736A6074C611BC1DB24E0B6095B7408CE2F25A3599949F1B
+D2BC59AFFD1F954D948612AE61AB7752201B836691015A4F39775CB3F2255E83
+7487B248EB3DD5CE803D4321E13DCF20B2D9EA143AF18F6A3FECE97733E4EBBD
+4E5F491606B743230D389ECB285B500A1AFB9B643FA392077976E882284229FD
+508534123AD595A3D8176D24FDF58926CC640D1A1F30E61A4E265C270F621CB6
+0551EF983DF2AAAC0E54CBA33A712AF1E9B31866B700B57A8FF6EA756A8A216E
+ED769E560DF54008487355BB6778CE9D67543CF99639F1A1F7C671A9339AACB3
+4A3F3F45B9DB072D8C2F78E707E930004C9D0B5F714BFA164088244AAAF9FC09
+15F3F6F138EE837C23B1A5C13FB7916122A042450BAEBCAE93AA2DE4BCF9A4DD
+573B43A2472D49AD7575D50438D2E136CBE29D464F3411906DDE47F58223225F
+D2D980522CDDD6C91C24DF561EBCBF1B9A0183315ABA4823CF0AD0F259A7477F
+04CBC553A05066E640A2A58E9FE32344BB272ABB5BB3E9C00D6013EC2C7B3F42
+C6A9ABC8D6AB3CC518774896FEAC9DCC32A159F1F06BFB7B79F596C7E7C60AB8
+1ED435E7BA996B54BE3EBE83A732E794B20F658A4F07E6A4D99984A27CE60E33
+B9DAD536934870A853842C6A840CE51E90302CB9518972BDEAA708D28A545DA1
+66DBC2816B8B0243E078915C50592191BD6308170E826BCF196AD59E2BEAB2BB
+E73E0369BF31709AC8605FA1BC71B054A54C629A42DB2979B180B8A57A682833
+71D29EA53903A88E799828DC65E4616B0F239583E05C62FCBA963DAC189005C6
+A04B3F68BBF900CFF5A26A5CB053859CB94B0802D1C7F052CAC988EB82C36135
+136932688DB2034ACDF4163A8428844B4B0CB46405A158DFBAA05AA92F71FD44
+A5A13F58C46DDC7F35283117AACA926E89ACA9C2EAB5235AA306CB9973F5502B
+8F63053617F8ACC3DC563A32C1C000A40EF12741FD113428D05FE4E6980AF342
+086958F9C46DC6532ABAA7002D8DF4FBF5F0A89B3431076FEDF79A9AE7D34510
+C65FB7AA564E7302769479235FE05220CB92008F625E9C5FB90B10234D58EDC2
+0A62628462E902F3573F88CE3E049A98D5B72D25B54D7682F9444FB21508FCAE
+62B916AFBB96615108975BEB799065A350FA5EED2B42C22DA1AA206D840E6DA3
+117F86298F29B4993417F6AA7EE4764677F8FDD43D15E33CABC8DA1045DECDCC
+E5CD09EA90D9B9397804D89181C1E313DD578AC7B232AB4F7C3917E6756AA7C9
+CFFAFE997DF21620B1C0C74979E9836204BF112D3DB351233A2561B73FD3E0EE
+0D0669A0EA9DDDC9A4911971DE10F8E45BF0B73980FB6E0D7A809CAD7090F3B0
+82B53CD1760790E5B3B57E79F9F49101EA8D107BC47C9C3D8E2EADAA79709552
+D59A56BD2701BB331FB2859855896E2D1B0AE6A360726F4020420EC27FD62093
+1F7CB4D71ACA2185991061F4BAA6A14AEE25FB44C5233F872FFE69F7AA772499
+8C727FD48C9027AD4FC6659CCBF40BD7609A01AD6D42522B5314F6FEEF46907F
+0E031F76A06B789FB749AE05A8D07FED69B6EA57AB27B45F826861E94A3B4F3E
+98B81E228840DF98D616EFB555EFE9F36D8980F9EF6889A5D071E56787473400
+8089489236BDC5607179BEDF7E9F2AEC9DE47BE311E1296ABC88AD81C4A77327
+67E6FD86E83DD06D40BD32E8543EC2B1FE8B9564EE55A29F4C1A114380232B4A
+130123FBC32B5F1871EAC20ADB2178E3768E5C877134E0748DE882E3E1DC46BF
+E479F0D6052234A8479BCA9862A99EEEA3EDFE26560A03B277531581AB7AF294
+39E39B25B2381999FE5438D171F8B98996DF6BC925A72F1653E29F00DB91AD4A
+695E16EA4A0831AA806E3CBC06F555BD17911FF6C8F34FEA7F5E27FD12A8A4FF
+4415D9EC041707B1452C648DEF4045D9B960DF2331BD37E537C10618158E109B
+
+0000000000000000000000000000000000000000000000000000000000000000
+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 0 /.notdef put
+readonly def
+/FontBBox{-6 -233 542 698}readonly def
+/UniqueID 5000831 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5F00F963068B8232429ED8B7CF6A3D879A2D1E
+2931CE5F5D18C658602059F07BE66E6EFC9239D7AB2FB8A4CBD41675B8ECF279
+650C29E53B14AC0E392A664848C1844B1CECBB2D5CFB72D0916B675C9A9A1E35
+F12696A6F628473C604A95376468E06E295AD6F76CEB939D94113532050B9D5A
+D2F41A9EFB9424D986612313B89EFE9C8A71313340B248F6853B1EDBF02B7F9E
+F447220FE131D7D54CFB8AA1281DBAEA73E665BACB1F164552CC0CEDB63BD4B1
+4A9AE8AC6FA02242DBE8DA46B64B6BFC11762F0784F216FC8B9120D688D1705A
+438B14F5E5DEAF2A98408B3B64620DE3732A4DAE6D08D5D97E34C75DAE19EABD
+BA0796165C1151BCBFB1DF8D29A63A8300DBDB9E3323CB82D0337598B83F4F2B
+A97CF5196D4D1CEC1EDB8966E548C0D9C194C932319610FB43EA1B86322FE641
+AB48770FF13BD475A7267E142388563D1A400419C585B22A9886074687BEDF74
+D905BE8EE440BA2ABF28EAB673399B7F129B9729DD5564C681954621903B84BB
+CAF89AC5ADB2932472DF29ADA2BDBDB4D05F65F28F5F4C529613D61858E0074A
+082A852710A62A147C966F2B85B51B0BE85F11D2057C66FDD61F6C5755367980
+9F4DE680601D4DA41B46F8D2148450000413C27AA39B586B74B977B25F0FD3C0
+4BA1EBFAFDBEC531EA13DFBD6700E53818CE04D23886B8AE75DCC36BCD3189B1
+0D55FAE27D0D126E82AEF31D7B5DF27E58C30BB0867D6D7AC1DA9EFB8A2DF095
+B5B934A68EE122DA0A83B36C952431586B957990206194E89339048AA6EE4C53
+703763505ED57C494DD907D0EEA04F6B1D4C8F3BA778F4E7AA832AAB4D75F024
+61E91C6D25FD6823CB24FC863B545568B81CF5BDD00BB653944D15A513A204CE
+7DC12C7F291C41F2B4CCA1F15D365AF34EDBEEE7E887ACE56A4ED3B50BA92E87
+06713FA90462A310776683ABDCEF85615F471AD11056D8FC1799D214A3353D57
+0E61109749C437016BB22F44C415D963790D2E9C3C482227BF1F37EEA61D87A6
+4B30891C8306E4C56B506C3ADBCDE557C7CA44817EF7F91A66AFB6BB81ECEA46
+D8DB4216D7346273EF7ABB54A6ED452137ACB00BF452D24FCCC63EA7AA24CB92
+4C82BD978803CBF98A945C860000EE294A37E78087F0F30D2CFC52790A76E975
+F94FDCFCB4F174E58D33D9C7EC2DE77D279E6D28FB27806B6B74F4B310920367
+38D9E0EA2964BFE5BDCC7AC8209B1C8986F91683A9F6979DA478609A714DA85C
+75AA37CD724FEA2BDA74F46E5C8F3C1B296E12CA5B80A4A3AB0411F81D0D76B7
+5C006ECF2DE2552FE18CDC733A8983174669AB2121CCBEED918D5C5D5BDBF0A6
+0E16127240BDE1CA6575B016AEC13DC0D6DE3F8D23E2A90DD660921087778FBD
+7E06DAF89366DB825E62019565EB8027FA93DA554F6CF29DA2B1D8E1F1336FF2
+0AC79B13B2BD7ED75051DA1B78143AD9484C9D023E013F389553D5EC314DE9AB
+B3740F07AF837E46A56F4000020A25D4FABE545BC79CC03828D7470608F493DB
+F1BC3E3982F42935AC3320B94D03F9CAE88AF372993339AD356DD82A38C9E238
+8C1A8F860DF2246E98929D18222265CAC17CB595D419CC959B256520B886F686
+7BEEB9FB11B35C7A64ECDB6E555C636AB4856C05DBE93B88B1EF176402923439
+77998EED459905324C390EDCF28DF8461BE1F7627250BF95FF308FC9C5387D14
+2B6BB4354570BB21D6BB954131B7848457704D4523CB0BDF9440E5CE20AB12D8
+5A40E077B91323D494F0EF748AB8F030996FF48EE9B181B6D8347F86A33329D6
+2E9C09CFF31A77E66339BA01BC4AF4D7C6AE40C1E1EDAE88BAFEEB5E6B529137
+D7924EFB370D9C1E84AE17B54B2DE335FE65065537DEB5435E743592FF8CADA8
+E6884EB354C2B0A20296E4587607252CBEB79EE7A4301DD130224AF5223B5FDB
+D0ABA9326A98E37C70ADF94CD9DAAD554BF2B21DFF3B02F79B3D3133D7221206
+0225E7F4ED79B616A55EB747FB897B8A97D9F219717A2C127BE023D9D42AE162
+13BC4FB38A756DFFB4DCB2C917E99DDD79CD2D8EE76053864B5946051AEAB7C2
+30CEDF0D56A11103263B7245957B3202594C8ED6E4498901FC89D0F8E5525504
+772497699B8639E66318060CACFEF3F407D5C7820186B907238B58A5E33D93ED
+51BF6555041610955F76427105861DF0770C450C836220314578CBFFED1DA2E0
+F428DA402D4AD571DE646C7DA3DB1510A0224313BE13FE2BEC1E2ACD61C8C3EE
+612FC823243F2BA35D311F6B2F36A7D195B513EEE0C4577F852942CA5512E852
+2009164C06A29D0E139440DAEBD5EA16CE1B51C78C597EFD10E55F88E4F9DB12
+FC8586074727D6B2B3C1E1C6F71CB7347E4AC9F9CA9133FDF861213E547B22F0
+EF22F24CE59B6886A9F72CE4D4D6EB5D799F107F153E5B96780A1C9953262593
+2A8C9DC9836571232C7E98213A51E0B39E8C95B83B2D97AB5069CD59EB5AD770
+061B865C972E7D9C4FB0E09814A3DB09999D62CEA7E8F0B18BCFF3E141A0D021
+CFF37838C7121C25A30E294E4039975979BD2A6C67842B32AAE72F97A3E7306F
+5DF519B15A8B659F259D445ED8A0420DC6BDCDCFEAF69204959BF42AEB53FE07
+E9D9320A1E5A2A34B4E61FC9AF7E80D2E908F41EE7FB7C70A0EAB65A13607F6C
+DD0DADD8C14A6D25DA3B88FB023EFE58C3523B4976DC248AEE322165C7F69C63
+D721A255CEFBCE483816ECB7D0C6419011866A295CD1B181E1FD911777BDFB4F
+4DBFE9B96FF5B75652280D6D9CE066F3038FD655CA518080E6356B5B5D35A225
+5B37E495538222F6B9CC5B6BA4305554E280152989B205BAC2CCF2D66C5B2654
+AF6696B4B93EE6CD89ED0711D1276EDC893CDE441239F6D407F7CDC53A1253CE
+47745F4F17BECC113871D3C5201FBBE4C2151B66520CA0984AFCF98B996DA64E
+FAF8A39F37993A6F435045E356BA398133C8ADC82022B8C8F2468E7BE0779FDC
+753F16CA960C96983BF9930A738BA14972D6B954DB8847987F31F87493325808
+53E97FA8AE8806B64E07BE8AFC189B5C4AA71B811B8BC369D9C0DF44683BFAEA
+EAD516898E2E1831879595C6B22950C53C422A423F240F1557265B0752DFFE53
+8BF1B161E1A2F3466B5390495DB250345536952D7C69550672B91B4A7B272158
+A82D173B31931B5BD34C9F5DB432104E31D55CA3EBE12EB3DCA8A17199DDC170
+9DDC9F1195FFED02E2531580387FCD95BDD6C9DDAB71D058348CC3454FD96495
+7573E4C520D809DE37E0B2E62D9AAD98B150B69450346703BD9842FBA046F4F3
+56D1EA9ABEB8E917DA94F70755ADA0C59FB8057E1E920BF360940DAAC3D17D9B
+A34C8E1F21A890F38E9B9CD4C5667D01812D56B6759B69D80B1AD0E5C0552CF1
+C6A3D41CA2E1F94CBFE974AEA2B56E37FA199D2F9BEAEB9A8E7BC9F447088375
+9E2A397E0CBD44FDA702FD61F1A3322DA6D1639C2F78D896747C158604852748
+011BD5E9D3BDADD2F41D6A45A9C9A5BCE696448E3312E63B70A978547DD93997
+76850CB9C756C928ACD32D5B1AD998864DC67DE4D21537CECB44808D19EEEE85
+C38707EFD83D0A64B2C680A501187D3B10D4C0EF056FCE3A94DB42B8278A6C6C
+8892B9A46078B0A29133339FF126C2E913D215F752D632958CED9C424D693C04
+523BBDD1D223CC3F87A02D0D17215040EFA3537AD7CE1295E6F622707C32A84B
+D457D30B49161351C22F2FA893E7C6CB4454009716083EF08F50CF7F1B6A6C55
+4A024064C903FDBB36162150949EFC17C8C8B78F746EE025D7368C2AEBE20875
+8621B014D1A41D6EC7EFD8EB77F09B7E1D4BE2F7D501F4F1B08ABDA217F1BE3E
+EEA294184F5BD274B1CED275B00F94C8BDE4D07225F823729F5298A4003F2A9D
+06922904823F3DA497C5024912CE2A5E725D9126E4E50D0BADA13FBE65A99A3D
+8F1EC4E01FEA2C8AAEA7C9A0F7A9AAFA257A0EC17DB25F3B1B5FCC260490AC2D
+807F240C25D08353FE4CB07FB251B9085416EE52C2FCB32A0FA233FAB4CF1E2A
+B84251C6416470E4FC63FBE35D5610FFE1EBC27E3EB7D6C40C0D5E6EC3F30952
+FAB0F3D1001C9EB5E2C6BAEBCBE63A97367B19DDE736F71660214200988208B2
+B6C8F825C57AE368167DF167A1C51572A1DE92DFF2E8125D7BBEA55B9D0396FC
+384B94A89EDCDA2ADC213EA054008C294D702943004FFB089CA163D300E74EEC
+088445C2E34FB819A6D6859AB0A5A01EFD96F1B090292805235C18F6C0CDDF6F
+AB2FB7BC99919C62859A692F4A88E4B19DDEF7A3CA531ED932306CBCA2336BC3
+23A4123052B47596E2708013CBCA438EAD95E9C850C7F2467C66D8891E28B89F
+17088EBB3A53A369D1A3A60E1226FC63757C3E3F76BE90605019E22F4110B345
+66B2461EF308A6ED330F1798079F2923C19DDA8B0614BE6B783844753B8DD1CA
+5964694E718D1F00C8F9BC06F2AF9EB167FF3B85E185D11EA81C614354AD2A0B
+AF94D7ED87F4BEDF6632167D920DBDF396569D17AAC1C567BA6ACE2815E161F2
+94CE44B15EF6EA87DE3ED2D7A0D8F6A37F667C1FD64B9DE37015D64FBB191237
+21EF9BA69BC9CC2C4B1F126438370977DEDD7EC72A7405AFCC082FC4B0D18F8C
+E9E4F26D301FC78C07BE705636D5B8E6051B49853E8AD4EBF9EDD35B3C672406
+17CD22FA770E0F8942CECF099B4712B837793F84DDC9F3295015F614691B408B
+43A38F73739DA913BE320900450F896668958092A3BC657311C485EBA8FC0DD9
+2A4D
+0000000000000000000000000000000000000000000000000000000000000000
+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 0 /.notdef put
+readonly def
+/FontBBox{-35 -250 1148 750}readonly def
+/UniqueID 5000827 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D03D5993DFC0930297866E1CD0A319B6B1FD958
+9E3948FFB3DF7BFF10C9BDA4EFE5F68A8CB1526990D1357AE6D2F7C2D2EF8496
+4E47B39E6712EB8908A3265E5FAB40567E866C244814449F1E993AAB422C3F1D
+DFA8C7118584F2E5197FD4BFA3A8AE9E953C6CD4672C0FF51E41C3A919749C1A
+F06650DF4C5E17492164BDBCDF22609A74BFA7F69960A64B9F949FFC2A807458
+8579366C4F41BDE1FDFBCC4845FA19BBB6963D65EE8532549274BAEBDFF24FA6
+03235D1BE37C06B1938AF369DA75BF38DDBC87A1FF445EAA16E1895ABE9506B9
+211955753E447865D33CEF007391D2666A046277A30A49804FFCED3FEA5EB2C3
+E52EE14A9F75241EA10C91974CDA6236EB840FD44D6DDE4D9B3266C3B99BD38B
+D835BCA8CB819C073480FB972CC028D218F6A1D344CE1B63F4FBF2C826F412E1
+6E0B05A26125865A14FD7B7030B478BB8BC6BC395335C3BA940E1C348267F4F9
+0AF97BBEE253511940F1048E175D3569F7D05A28851B6F50765FEB6C9654FEDC
+1BF52F535DB5BB90C1BD5D2EBF75E0AEBE82B20507F3C28A03746781018D4EB2
+298E4F2C27ACF73FA73EBE43F014BB575AAD516C0407B29E1653375135ECB74D
+C91372F06FA8EF37C31AF3FA48AE65318EAA6C34830A5377ABB2DFA5DA53A574
+433484BA1466709A4B186761655C8E482833B697673E847C691079E7F1DCB8D6
+1AD91101D757B83E2090337D525AEECB028FB3C9F6A6E6AD2F322CFDC5A833E6
+1CE4EDBF41FD34FD61630581D222F854A76C2EA9FD72796A7C9CC1F6C2FCCD16
+E95CA05826A4ECFADA6A5FB83C41A7131E52BA6585DD6DD78515D8F7327DFC6F
+9404F898A96A0916FBC8CB072755CDF678C51A4C873DAE4C4EF4289EFF0E3427
+F2B06F8F9534A80E2346F35C37946028FC69A6C587BC81B55A8DA8D1860B4C59
+44A9416140A8197147DF90200F21E665C98BF5CEC5335E0390D6E649185DF81F
+01C0C37C30AF1A05DAC2E8A25194AD0A30540793802F05B3A4F2ACC18E8E5150
+111C64105E3790A1DFDD8B209AC1EF21E11AE6B1FB60373BA243E92ADBC432DF
+3C91C3A053F3BC594A2342A860159AFE149D73F3D10118E35442F14FF50D4D4D
+365C1EAB513C17FB6E1503187F1D339D54A36CA9E6F4F1F7D78C384F2F51CEAA
+AC0355E4B074A47B579436D954ED72E3132AE640BAE5E044E9C9EC4EDDD605F9
+49C02C5ABAE0C049598B48E8C94E63C922CB4A33ECB34DF0AF18483947814ADD
+B5B4772B1E36F6F9D3B6C1868E5A9FF74DACE292F432DA2240999AA83A21A285
+DD3DA44E52DA479059143EF1960DDA149DDFCBE834B72F2D0BEAA80CEB8AAFF4
+702B1EC9D799D4E1A1AF2881811CA9EC5DA9756458950348EA259A87F0288FAF
+96E65CD9AFDAD306A8ABCA84E02B4DE3DE614F8F567E09BC8E3C0EFA0DC7B838
+80B4F927BB7C11D74CA5C4EAB87A807BCBB569C598EADDE3C42182298D2500CF
+97957C6CE7A4FBE5B307C27578E5D410B359BB4DD711E2222EA323FCCD30C3A2
+2B08764EBCBB234EB34714739498F1A7E31EEF99661EF528D54A950906BBC0D6
+C6B96DE517D97675CF5498B011D2A36F35F81A7C23A4E02F3BC76D8DED90F8AE
+77D0F9F0EAAC2B7F329B8F8A06F85777800DFFC2ABAA0C89FD0BBF93E458686B
+205BA3611DA3105A920E68666C59429460DB25B931B99F13E8849D6DDD594E58
+6CECF980CCB9188E23B5D795E3B9E3EED93F4A928FE907DFF1CC28501E576BC9
+68215917DA8A548514C8C7A7976F42B0CE558ACFC57E0AEEFF4B253C1928E99D
+0FE365225045D42BDFF2C6A7F607E206BF7E21E85499B4C4E7B781EE2BC09411
+B54C81DAD548AED1195329061027EAF21798899304E5A74A5CB2BC02381EE039
+81541CC1ACDCDD3B48F14BB7BAC8FE9A80D70B04719CF775D85799E88FF4C824
+586B2C2BD88BAD68EDA66B0688A054494F05083306525B27754945421D998C20
+BBB2361DE17CBA8E9FCD76C82647676566B5DD658FCA1C6AAF0109B0285B7333
+DDA2D8D89562F260733DE9D1E5CF8A65E1F115F783002A4CCF871CD0C56C6A60
+E50F0E66CB275004D30BAD866B9C3E22EDA8C6D869D7DA21BB8E09CFAD1F04A8
+76C1BEEBDB25F695F940B368B57A257E4AE9374624A0C56CF9DF01F5859222A1
+20AD4B0EE0D291D9042F13BDAA3032E9D260C57BA7B4163FDA78B6A15B71C1AE
+1F0CE65DD462A250A03E057EB8D8E6737D0D37DCFA8506BE9346192C8B5222F6
+C4B82C3DFA4E4264D58C8FEA8C9FB4926D49BC7A6A177F5C583E3903440BA667
+C29839448CB7E91D3DCE2D972634570B2498AE1321E842F5B37086039C7B7852
+3A8C0570E31D1B2174F7CDEAC33718A6DF12A95CFF42BADC9FA85633C3174BCB
+7DCDB9CB6099A11170277295AFE43F975C3F971F52FC6B68E31DBD228E7DC2C0
+D6231BFB4E8C71FEDC0F7977FBAC88F79DBBA064FFAD87DCBB747A02DC6DE115
+61D851414EF87F32C1493E36546DD8BCC331008153E94DB4F180201DC7CD59BC
+FA46FD284ECF509B940B2991F2C218DB36E89B3B7F6C88C9CE77CCAA484F504B
+5657FBE479E5B1764B8076FAA69BC6B3F444F0609C6642B355C79928FB2A2260
+516416208AAD8A7C149B7E30A75E946646A4C818F96BAA73DA0A6BD418788B02
+4A83221758B7241FF214AB2E398B9BC7592D1C65331C7CE5169EF239E6C3B662
+611C4694425709864EE9C9AC40D5301E91E05207643E5BC23B32EA74A8DDCAE2
+8281C844074B64F51EFAD6B8BA2183406EB4D6E7CF181B9DA67D981AB5292A35
+5B6FCAF5495DC264A0F56974F7C6272BDD3B9F73BD8EF38E11F20EE82858FF4D
+053D180D1C208243F14CAE24F8B63DC07F7E451317A2F374C7CE4A7D75D9A655
+03A9968122F934C35112E08008F1F9D6B7AE5A2FEEEF2C15705170F209F6E650
+356BB92FCDDA6DC3EA5501A500B45E38FAA888E8C27D6BE396069C0845349633
+38E1F4F7FAF1B9E7E7154A66746B9D7F3A4E21CC955437352B756FD3CDAB8E5E
+ED9A3CDF
+0000000000000000000000000000000000000000000000000000000000000000
+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 0 /.notdef put
+readonly def
+/FontBBox{-30 -958 1146 777}readonly def
+/UniqueID 5000819 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052F09F9C8ADE9D907C058B87E9B6964
+7D53359E51216774A4EAA1E2B58EC3176BD1184A633B951372B4198D4E8C5EF4
+A213ACB58AA0A658908035BF2ED8531779838A960DFE2B27EA49C37156989C85
+E21B3ABF72E39A89232CD9F4237FC80C9E64E8425AA3BEF7DED60B122A52922A
+221A37D9A807DD01161779DDE7D31FF2B87F97C73D63EECDDA4C49501773468A
+27D1663E0B62F461F6E40A5D6676D0037D33F24E2FAC2B0009AD3C8350CDF8CC
+65BCA87979C36D14CB552E9A985E48BE4E88ECA16DF418749AF04FDD2B0E1380
+D281BB2476BB45FF30946B247DFD7F57305FA87E50CA338121C71CDFDF927A9C
+77FF14CB4A1D6D80356FB1171ED38C37702350497B44E42CE31DB2F493807DAA
+15B887C671199A54C4C1294BC520F5538C15556BC43C9F62342B121C6DCD6C5F
+491DA47FF360201EE21C08A781ED0589A6DF91B99FE118B9B29E4F068672E52F
+1A06C514D91C4C937D4E642503392B1CD1BDE732563086F6AB251AD87A685592
+BD0B2E5F918D919249293F893BADDBCA8563E807BC5BE28FC0C41A634EEA0106
+079196A0E6BA21FDFB809CCBF82CD0EDBB3F8410FA0A109F786799D5C03F089D
+4C1F51D2C9393F93948B8CCAA12209A6A9CE91FA3A8ECCA4360DFEAD4579E594
+FDE490897F3322CBC8DDF65FA8B8EC7C27622A9F081F52279D99D918A375CEAD
+A3CB0245D52924996E9C107331E234DBE97A07504294321FCA0F375B30A26862
+B3548148DA2AE7D7FB1562DE48AA1B512461EDEA30173E9FB1651FD9140ADAC8
+359568999D296B3DC8E0EF65855EC7CB642E359BA4D9811E0A8850BE4CCCF222
+027700C510C58A911F36DE4E12172CC44738B1E4E3395099BA971DFEB150DCF5
+28EDAAB524A931AE1DFA1EFDE27E16A331278A2DA40126341D8174E7D8494B53
+B1DDDDFD086E8FFF984ADEE8E48989113228E5FB46ABA6FBC90E9F05D8372664
+5A03824903FE6A086C57B9DD2145880F207DB42756F6F73156F49DC53E780EC2
+EA4779128E8E9C433135BE99E10DE7D3A9B18F88BC1BCCEF667820D211CB48A9
+B9BFFE7E1F527C809FFE5C8A895B3558B00229E4C483C212BB7709E604D31397
+A4064EDDBE13E6FA61339E5DD3DF72A33FA84BBDB4AB82FF096D2D44CBD36E78
+AE78365935B8E78374194324AFC5D65DEB197B694B3BFBDE80FA00AA66ABB3B0
+7D7375D3FDB9A82D16F08382717A7E703228027C3C9FBE41F401388DE67EF571
+C8A0BB9B7C9045B40B6ABAB8581FFEEE2C652758BB3944434EDC2BC8E33C143F
+551D868A84B3AB71C2DCE99F0FDFC34890FF403C164F09052F820DD18FA6BF15
+86031D4B0FA966DD7FF0A9604EB422645B0AE46A3BBF4E39E47C2208649CD670
+94CC385A11A34BFA7C0333F37CC69782DFB00A6C575BD13802A43C2966C1A7CD
+4E80312765E3236625D040B089218A0234492A239210C4F70EA409110425470E
+9D1D085CC886CF0899EBADB96A1C26B143261621ECC751F98BE47354D33FECB6
+6EE91E3EF0D5BAFAE1E49FF45F2D2743E6A6B23364FA4F11FF4C5545920F4299
+633B28FCC82C69C06897FED47ECBD6D3EFB751AA5F39694DB066AF5F4B345A4D
+AEFC771B6FF2826232210F2E16EC04626E7A78F971945D3665F715CFC5AE7447
+25248BBDFD6D9B48D42A57DAFA57F39C79225C852F4DEC90A55F65CDB826627B
+B6355BB7F9FFB113BF156FA01C3A2C35CCA1A27B3BA58722316B21F5CFB09802
+04295B292784EEE294EB7002D07F0C862F126775C84C8029125F88FDC60208E9
+37992711E8CCAB7B072787003D178C11CC76078E55FEE77C38EBB6795A22E656
+76EE5DD4109F2DD4F506679A839261462C3E31FAC1EA911267946FD7C955BAC6
+02120302D9902900C68609ADA33BA6B2E08FA48F9410E3E695CB42BF948BAB12
+BB35BE3B7AE5C733C3F6946A92C39137D79E1E5CF76ECA4978DCC3048C8A0186
+7999210812D2694F1EA59A06861762B906E4D6832E5B8F9913A76E20C975AD38
+3F0D83D9C09D6D371DAA0291392F2D20B40C410694415B74CE147F1AAAF31EDF
+4D3DC7EA4595CF0EF760695F1C92FB6DF6212EFFB72E99E43F5D340B73EEB869
+3CC4AA451EA55EFB2CE2B41FF993BAE2351A7C1CA903CB17F5D9FEDE37B05717
+62A30534B4B7E6A2D4F1EAE794E89A25876F6D42A472ADE09F4E4A2302FA4DD1
+5CC2A95E086AC1861292759B595184FF658B9860498E7E0603202F0228C84692
+6DB66D50CBE929FABB8C17E888D0D2ACF370909936FBCF44E078FECED0739507
+B665E0CFE1982BEFFE33B25492594ABF4173DBE8914D21C3FF53ACA4076AE374
+A5F7368A445802A205E042830A66D9A13CD78A956B79C98181656D26B5A34BC0
+7A78AE3D6C4E48621BFBD22D7CA1748A461E256DBA622794BCB80F2C33211B03
+CDE0E6B7FFFCC82A17E7FF5768F6075EC48E6D97DD3420F25B46A3D784726778
+5C7BD42371C6D59DC2EDF267AE7F3165E1C91A3796B0C4CB0EC9EA647FA36F7F
+9FDF6A3A8540C915097CDD713E9667E72F8B564967D58CFBCF30814738F713BA
+773FC9C743F4580237B40BAB0AA29A0C3A782F2DB60D74E1DBAE029FF9CE65BF
+98030E880D9F1172128A70332AA63EA69F6F3C5121E544F1F62646F4E6CE89BC
+958AAAD51F81952B0BC1E4C47105DA8B16CC96C44A6A9742D242BC824E9F043A
+8A31CB6C4E792EF4C8CB02C4220F0D8CCB250B1C69B5211F026C2D2DA9DB4294
+DF0F6FFBC5DEB1FD9500973A455ADF2D88B9B320C38641D3F7E2761C59D06679
+992811D99203F7B79A83C3A4512E6C849C108A4561EDEB62604EF36784E6606F
+AB503A4CDB6C20611A88F5BE8BCE061B7DD73D51C11547B605FE5FD15260B2DC
+916E73531E33345E56162446A37E92C5058FBA808FDDBD52A3ACBB782E7BE17E
+7EC9B9F8053C242EA85CEC946354D7913E6F830E2B1342772F3287EAD0E50B6A
+06626824D87A62B9A5EEE2F3B765A742F413CC5CF05716360C9D2538E59E46C4
+E83B0C2B2CBA8EECDC377AD6B7FA1A981838C3CE3738EE93680E7A53D962F468
+B9DA735D574D6E4E4A80672F5EE45D9D8B4AB1B095CEDF6BE0F46ABF70537807
+1F69785D58247237D3DEA0F4C9D5D48CF09E6EE9F3F0E7BCE2A95EF8D0DF9979
+EE1592D193F62B18B537A0AAC9C478E1313C0DDE2098F16DDCF9DAF130FAEC43
+DA3611584224DCCB2ACC88D7F9DC371C30F1C6A8A23CBC5054EF40E784A9F546
+7117C6349FE9D9B02B9088584C8C74A831CA8CB09271BFB2AF460C3F1BDCE356
+75A3B43B36308640468094ED612BD04D6FFC4786ED2038498696804DB5C438E3
+971182792D10AD98AA40140DC03EECF4648033053FD2FAB99037AAA7AFE41F64
+4BBFA1BD36B05436467B00570518B93A9BE6BC7969D7B15CFE011AA92DC14278
+DDCE3A115512B347215A360E59AD7A67696DADBC3FCEE1C971FA2154F574C4A7
+5F501C0F99E39450F2086F66F613355315F530DB22FD2CF6D33E9CFFE7C686D1
+E776FD83EBF7FEA4250FF50B85D141A9F65BB963BF57C43D46750A41DDBB99A6
+BD03B2ACF338FCB81B0442F3CFC2F8FDEFEF354E7FF4F490BD91BBF9BCC76EE1
+59E4AA3595B802AD61AD91C64B66048E4F8FCD7420C57BE11E6C3AE7BFE15C7F
+8961F913E62EF3CE9945AA4FAF486239D55F922F868DF797B53C9DFE524FBB3E
+9729C599B7576A737A0DEC597442C476A818E5B98EA70BA9FFAE75A73AAA46D5
+73F6D7573F416C4E8A9C8AF8DC33E3E095ADCB22DE661F70CD594440C3642E0E
+10BCEC492756EBE189C01F09BC3A25EB247DF8FE0F1389598BF502852A446235
+F2912E5CCF01ED62182103A709282E99BF0FC0F2CC0D5347EF175B3FBF9B5CD2
+E994037E35D47162F7DF0C62A59C5A34B862FAB81220D2CD36D695ACFC8F5D9C
+C7584F00628C46D3AC4D16F6AB8753B4346FA7394D0DCBF02C4051F80DC8E32D
+F638207DB78D7C6718225358C34169F67B261A490C3F0C4BA4215807A2F6B449
+F71D593145EE744F2FF1D66C7FD0334C56B854F09BDFE15B3135C2BC12E9E9DC
+EAD53743FC7FF287BEFA92250332A075270CF557B27DC23192F76BF3795B676A
+30DE313198E06A3ABFB738C3B5F25494FDE37DCA093F90D5BAD75D803E226549
+46B674C093A1893D9975CFCE9183E4061CA406A1444570B62B92F2193769BF2B
+BADC8285874CBAEB79E8782EBB9EA5C2844B0D6DDF462936B4DB38C58EE300
+0000000000000000000000000000000000000000000000000000000000000000
+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 0 /.notdef put
+readonly def
+/FontBBox{-39 -250 1036 750}readonly def
+/UniqueID 5000792 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF7158F1163BC1F3352E22A1452E73FECA8A4
+87100FB1FFC4C8AF409B2067537220E605DA0852CA49839E1386AF9D7A1A455F
+D1F017CE45884D76EF2CB9BC5821FD25365DDEA6E45F332B5F68A44AD8A530F0
+92A36FADB679CF58BAFDD3E51DFDD314B91A605515D729EE20C42505FD4E0835
+3C9D365B14C003BC6DD352F0228A8C161F172D2551CD1C67CD0B1B21DED53203
+046FAFF9B1129167921DD82C5964F9DDDFE0D2686875BD075FC81831A941F20E
+C5CD90040A092E559F6D1D3B0E9BB71733595AE0EA6093F986377A96060BF12A
+A1B525CD9FA741FE051DD54A32BECD55A868DD63119A4370F8322CCBEC889BC2
+A723CB4015FC4AA90AE873EA14DE13382CA9CF0D8DFB65F0ABEDFD9A64BB3F4D
+731E2E1C9A1789228FF44116230A70C339C9819676022AB31B5C9C589AE9094B
+09882051AD4637C1710D93E8DD117B4E7B478493B91EA6306FDB3FA6D738AAB1
+49FBB21A00AC2A999C21445DE3177F21D8B6AAB33869C882613EA6B5EC56476B
+5634181ECBF03BFEDB57F079EACE3B334F6F384BDF9D70AEBD592C8ECF21378B
+54A8B5DBF7CB9282E16AA517E14843909339B5E7C55B038BF3BB493F3B884A1C
+C25F9E8FB912CBE23199AD9D2C3E573727701BA301526C66C3617B9514D6F11F
+11930B1D97C17816C85B1BFD9B973A191B33CC3B391815AD14F1CBE935942AEC
+D4004E6BEF379066FD72209DC88D2E634E79BCC2B98C766CBD92C561F2703F8A
+109E6C6CEC7B866F2FC7ADF646BF492E520319F3B949AB5D84AE990B33344A40
+3971F58DFDF8D8D67FA0B8F2A0D884F8C09A5A721319B911DBA0A35903877343
+C37BC36C5EB32353272D1E6ED5FCA611BE319A7E1E842CB7576E7F57C4F43525
+58BBDDBF843D36F01726CADD8F2639300D80714CA947B9E2D3D2DBD97E6C1294
+ABC4D44796DE99164ECBBB1D7A2D29A2BA789089D467FC15E396795F8A471417
+89E194C2FD914C6B635A164A8E93A2097EDBCF612E6036A531B8710B354D40F4
+ECCC45A40F0FFD8119C94C92AFE513147CA7707C95CA4DBE3C9856B186CD67BB
+6A82E78C3F3355FFEBE3DE70EABACE0037C9805DE2BC9422126119D659FD4A33
+8A40784180492341526E33C7427C3E741A8C62F4F7257086374C7ED77267B1EC
+D38037D7D970217E24E5322BE26E6AA7C9CB37151B1AA59F56C6A03B951ACAFA
+0C73F2BF0C07DA2DEC38C5F6A0B40BC74BBB259C8B0632CB0486EBD116A33B2A
+C5739B7B70C3926E7DFEF5024CAB5CA10EAE804AA7C963C2B3B1659B7CD6A331
+9BD8F7EE51C8FBCB4F9BE1E52A086EAFE8FC9CB4D654C31E3F09B7FF65DA37A1
+345C0286DD6957C323EEA672D185E9CE4A5BF8EF17B785128AB68CBE32B321F5
+3B7AF349940C3FE46F6C0C6B944FD7A214EF7AE3CAE5CA2A1A74560D5405F743
+489E8F7A0F8CBEC76238559EFC87B79E9A657D567620937F69A18955F7424C4A
+81EE01311DDC5D052DC87A523FCAB19224410D1A37C01635AF5F3D73D3696B8F
+214E798A94566405198860E6FC14840770F4AFF1E243141729DF0F2780F83F83
+89073C09E24EB7EA477234F8069DF28FF5F5C556E71A271F9F541271E69949C1
+8E129CA7D6A464ACC75FBE59F5340F39C9117A7659DA7001CEA903C83A271376
+2A488B018BE9FE97000476EDC8EB79AB467DC1F07ECF55472BF1AD8BC06E5F64
+9D8769D868C8059EBFF8099A62E20C36827CA5D35B590C6D8922166FABB4E61F
+E281DCAF7B1DE49D9FED46EAC3A1AA1BAC648147C285491F68DC041E9441C523
+4F78D90E73DC2C8371F71B67D9463383DAE9C8821DE0358979E6F38C16BDD30B
+7374D73326EAAB1B6CE6841EA2B9DB040625552C962D66B138ACA45C5440D075
+CECCE5877E48D7CD5A4AF0D41A7F1DCE69A9D9142BCF20F081821B50FF06B3E7
+E10CA846519D1D4C7485728D7BA827857B733BB62546FBD3CAFA007761111F8E
+A5B07A0EB8D10D17A3828737218E9E98D4584077AE161B1B193A6DDF1AA2D01A
+B3B64AD33B4F6B8C34A16C626BD68BE14D920D5A6FBC49FE5C92C42A469EF518
+0A5484F0AEB4C1F7A0C74D72193640892D6585A7B41BE0B92592AFE260F98ACC
+44FF8949D4345463C6C9B156495C8748BAEDD9A5A2E6D31FFB9E68B14EAB9223
+40B38D5729CE637C9245B96B46F55AA4E40E7AFBEC362C833D46F92C2E8972D7
+798C0046D2146212E43E83F31CA35E2659E7986FE3C2C69309E6E873CBAF2852
+87CA22AC1F3FDDFC144440BEBBBD1A83FAF63FDDD3C2C904C796CECEE545C512
+5BE4BE63C65EF2AA010308631E95C3AC12B1A198B82E013820CD8BC2AD9443F0
+8E39674FECDA00F6895B5F041F7286EC37F395EB42B3030875DC8CD0BB0B57B6
+966B993555E6DCEA0DCEE4B41B62102C4AA0A347AF6BA74DCDD0D624C188E075
+28C44E30F18134AAFC2F4C84E636BE433A9455B61C41638386644C4381EA1A20
+736E5EEEF4494EDA641782FD28855384F41388091973553A1109D298D9FBD959
+1F70EF81D73C5512D7B5B20BDCD2B8BDADA1425C050397C007AD5C4BFB5B49F6
+E279EE24D0E14728B67E9B43E5F930413CC880AD477B17CF2339173CF6BB64A5
+3A8A4BBEC445B6332BC4095DA1004BEB593CE87E3FEA0452ED9941E292EA564C
+50BA31371AC92659DC86BF214A1454EB26F0C69F98043ED653ADB5B83EFB09D1
+F39008F4F75381A10D9C0750A49CC2264540590BBA88E2939AEC8ABE75751F41
+9598B10AD4C92B89704B72218E4B5629F741D32B5B619D6CCBB8289D8C43E8E3
+D618BA501B9DF76EDB10C749AF1F615D71C3505B15034D96176802AB89625CAB
+70AC70E09EFCE8C91BCCDBDFEDC457C03906A45F1180ACF979698762A3553B66
+5B8EA00CF56B0C3F09A333EE6CA06867A1327576F62D2E5A4367C31EE48F9FC2
+AC78D601FBD8991B998C400F9470B2CC9FDB9F52A8A50327A099517FC6F9CBC4
+E3DF965AC621A07E225C5B780B11E9ED9EBBC208787DE06907F3E5A984D390EF
+F528DD806EB3829542584A578F48B5DB7DD42748E35C5A31A57FD93D1FA4C763
+B56189938F546F15D8584FD62EC1FADE410425DEB0
+0000000000000000000000000000000000000000000000000000000000000000
+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 0 /.notdef put
+readonly def
+/FontBBox{-20 -250 1193 750}readonly def
+/UniqueID 5000789 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF4E9D2405B169CD5365D6ECED5D768D66D6C
+68618B8C482B341F8CA38E9BB9BAFCFAAD9C2F3FD033B62690986ED43D9C9361
+3645B82392D5CAE11A7CB49D7E2E82DCD485CBA17D1AFFF95F4224CF7ECEE45C
+BFB7C8C77C22A01C345078D28D3ECBF804CDC2FE5025FA0D05CCC5EFC0C4F87E
+CBED13DDDF8F34E404F471C6DD2E43331D73E89BBC71E7BF889F6293793FEF5A
+C9DD3792F032E37A364C70914843F7AA314413D022AE3238730B420A7E9D0CF5
+D0E24F501451F9CDECE10AF7E14FF15C4F12F3FCA47DD9CD3C7AEA8D1551017D
+23131C09ED104C052054520268A4FA3C6338BA6CF14C3DE3BAF2EA35296EE3D8
+D6496277E11DFF6076FE64C8A8C3419FA774473D63223FFA41CBAE609C3D976B
+93DFB4079ADC7C4EF07303F93808DDA9F651F61BCCF79555059A44CBAF84A711
+6D98083CEF58230D54AD486C74C4A257FC703ACF918219D0A597A5F680B606E4
+EF94ADF8BF91A5096A806DB64EC96636A98397D22A74932EB7346A9C4B5EE953
+CB3C80AA634BFC28AA938C704BDA8DC4D13551CCFE2B2784BE8BF54502EBA9AF
+D49B79237B9C56310550BC30E9108BB06EAC755D6AA4E688EFE2A0AAB17F20FE
+00CD0BFF1B9CB6BDA0FA3A29A3117388B6686657A150CE6421FD5D420F4F7FB5
+B0DAA1BA19D638676E9CF159AC7325EF17B9F74E082BEF75E10A31C7011C0FFA
+99B797CE549B5C45238DD0FADD6B99D233AC69282DF0D91EA2DBD08CE0083904
+A6D968D5AE3BD159D01BDFF42D16111BC0A517C66B43972080D9DD4F3B9AE7FB
+11B035CE715C1218B2D779761D8D7E9DEBE277531BD58F313EBD27E33BEF9DC5
+50C7821A8BBC3B9FDF899D7EAA0B94493B97AFEAC503EB5ED7A7AB6D9A929A8D
+D186043A1A15F89EDE60FC17A13CAD7F1FCE081BEFC2864447FAE30B0AEFD59E
+3F1A3890297A203003CE8D309197234A7F5D2D82C831907B3E5C1D27208F094C
+F5A41E0BF1FA2DACBE0AB52EE28A0D87E7BE07A2EFCD52016DCE87182C6C3FF6
+D0246B1A71FDFFD46444DDC4B560663C5FB49826D54F03764C975DCF1D49BAD4
+388965C52B7E1A13581695D7598DCCF499743D52B5E0297A0A0DC0034ED2082D
+18822809266C51AAD54308AE120190F9CA896BAAF35538558DF06A3C2316E61A
+9AD6303C7280C8F734CEDC4BA3CDDFA89183B5EFE420FD4D6F838A43FA50460A
+6E16C83F85FAE72E745CFB88FC13D6F837870BD55E3F2DE62C67953237538105
+979CE2859874FAC2F90593C24C009A126B066B8298C9708F03F9605906648B6F
+BBCCFEF6975D9DBA0F54ABC3F679ED9B84ED209F83EAD5E0AA4D309F24F21FA2
+C2ABD1A14E9D3DBD43DC9360B21E8B3A8F3E95678368638E3C6F824AA6061E0E
+26DBE54CBD233EADA3DEF5E22F88050770D83272E39A4E2768A68F74B2FDD39E
+A2980F56690C25FDE55BE282E2842D7B0CC6125C7EDC137124125112313F0601
+F587BBB2627F547567C968E34B23AADCD7AB84C2D7235FA21FAC77CF4AED6E6B
+CF47A588F2738B127178763908BEF394D7999B971BCDB420D0C0CC1122140ED5
+A60B0E440C314D32B4D3FD1CAABA2B01901E55B812AF042442889D7D850B2035
+445ADA4F52A4A89A5279A0B84F612DE970BEB4DCF466DE06CBB89060139CADDE
+3528FDA94C281ED71D3D961361A0DEB949AF48A8E4015769274AD8145822CD9F
+D9C571D2C860BD1EDA9D39C0CEC1F5835FDD750F9AEE6D658D05AE0475920E24
+53E443EF23D74127FAE6EE2DE134F721821191C7BEA795DFA5562226F2CC6CDC
+7BFD7ECAA1455DD046C41967F3A1ADF24F8A2EBAD19A197C19301D51C5E3FD0D
+4CF3101D607EA3D1CE223A674934E3C1030D2FB46B02A1B85D1DC489B2A79D07
+FB021EAA22F8376A6F47EBFC78D38ACB8827B6408E7CB8DAFB809E30FCD99CA1
+002012749AD80858537F798B9762083E8969867AC2B4CFD5AF2028D0CBEB4630
+5FC08AE7BA8F68C2003C5956D644D568D732FF488B5295BD8FCAFC9D978EE6CA
+5F8924AE11B875C2717D053C1F0A5A0EC051CD6B542A45A75E3A2103CBB34243
+656BB24B014023172A31639E630A8E0E0C594B7602D007E287978EE88AF643EF
+EBBED2A4380E756878495E851D8CC25800F82131A6FD8C616DC3FE68D7006BD4
+200A29376CC460CFC3FD3A611F3978A656348FBF04A0AD5EC880B4E5D13C5E38
+8B26266CD59A84507B9CF3B6101E04F3A18FC24BABDCC71B3DC82198B5F5D7AA
+9ECBE68821985ECE0E0F1427BA2E68E9325EEAB3E3E90DA568AFE9162BC60CF1
+FD412469838985D23544FC1B7665CA88E000E1A1F24C4B8C6F12F02C65B20189
+58251C06E92C262D08878D0C789AD17A2CB1A4C36F7E6DC16977767285421CB7
+3EE915B73EC552BD2BFBC9BB7B5D6FC9D4DDD5F018F7B530038165A93E476E2B
+3A21AA4FD9D1CB2A09B25B24029580F581F62E61D785977C37D02B0108AF392C
+EF1EA5027562773353B32743C9E06225A9441C34A7126A5F66B4CB3017B834C5
+129CEFB47624B7C722527E40AE09E224568EAC35C5C0ABD474FCCFC77D209B3E
+CB5FEBC186689394C900489A19C87DEDBC51522D506EDBB33B236D594739FCF9
+4461ACA1FABCBAB699540B4B341C86D52F36DD66D42573C4B757A8D92817648C
+7E459778CC7FA3E829925D3FBE21266546EA54F3F254C94AC8359494E1E35A49
+5CE8B576EBB16D767A3CFAF555D34F6CF9245CF40D3537EE315B93FEEAAC2C86
+8EFE4D27E85F52E48520476E87B9C9A35D3DA06023877E6B7597F728C0107C7C
+C722B6CA30EFC9B5532436F61C5C5DFA9DC45E1C4AB092409D4903C0AA1FE833
+3D73A0FAB63AF2D5BB83647E4914735534DEF6F9B47B22B7DDF4B1A90124AD32
+6225264265913FEC81839FE8E6EE2CF3A585E3733208963F
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMMI7
+%!PS-AdobeFont-1.1: CMMI7 1.100
+%%CreationDate: 1996 Jul 23 07:53:53
+% 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 (CMMI7) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle -14.04 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMMI7 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 0 /.notdef put
+readonly def
+/FontBBox{0 -250 1171 750}readonly def
+/UniqueID 5087382 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D03D77639DF1232A4D6233A9CAF69B151DFD33F
+C0962EAC6E3EBFB8AD256A3C654EAAF9A50C51BC6FA90B61B60401C235AFAB7B
+B078D20B4B8A6D7F0300CF694E6956FF9C29C84FCC5C9E8890AA56B1BC60E868
+DA8488AC4435E6B5CE34EA88E904D5C978514D7E476BF8971D419363125D4811
+4D886EDDDCDDA8A6B0FDA5CF0603EA9FA5D4393BEBB26E1AB11C2D74FFA6FEE3
+FAFBC6F05B801C1C3276B11080F5023902B56593F3F6B1F37997038F36B9E3AB
+76C2E97E1F492D27A8E99F3E947A47166D0D0D063E4E6A9B535DC9F1BED129C5
+123775D5D68787A58C93009FD5DA55B19511B95168C83429BD2D878207C39770
+012318EA7AA39900C97B9D3859E3D0B04750B8390BF1F1BC29DC22BCAD50ECC6
+A3C633D0937A59E859E5185AF9F56704708D5F1C50F78F43DFAC43C4E7DC9413
+44CEFE43279AFD3C167C942889A352F2FF806C2FF8B3EB4908D50778AA58CFFC
+4D1B14597A06A994ED8414BBE8B26E74D49F6CF54176B7297CDA112A69518050
+01337CBA5478EB984CDD22020DAED9CA8311C33FBCC84177F5CE870E709FC608
+D28B3A7208EFF72988C136142CE79B4E9C7B3FE588E9824ABC6F04D141E589B3
+914A73A42801305439862414F893D5B6C327A7EE2730DEDE6A1597B09C258F05
+261BC634F64C9F8477CD51634BA648FC70F659C90DC042C0D6B68CD1DF36D615
+24F362B85A58D65A8E6DFD583EF9A79A428F2390A0B5398EEB78F4B5A89D9AD2
+A517E0361749554ABD6547072398FFDD863E40501C316F28FDDF8B550FF8D663
+9843D0BEA42289F85BD844891DB42EC7C51229D33EE7E83B1290404C799B8E8C
+889787CDC2B4DCCBA250138361A088A4CEB00DE6235F7FCF11245C5731EBA766
+009B35FC28B420918AE888B1A7EBBCB0887AEE1D370A7CEE900987A1AC611D19
+15E1310E985894FCD43778DB50FAE725D2C01583EDCBE41AD025738ADA2E2144
+BB6FFE85ED43F414C191F544C5AF14AC5104E3E56C64271EA21D4B0EEB76FB67
+32F0F0273D4E95CD356E9A69BC2899A4F226A2D084DB4E31557824AE19379DDC
+921EDD6FECBABB378DFFEB52D5B39B2619020B7585E8D328C66E2B59F02A469D
+8E5CC95210B8222A5809D6296CC4F81411BFAEC0A12C42DF00097D9D2CFF58F0
+58CEBD8C213261A4837F6D0B0330B46588467B9783836D225CD292A1A6B2831D
+2B513394A5AE4F5CBE484E4008B9C5A9A7D34E74EB3471A32558801CBE04F633
+C6F0185C3D50A266189D89D7BB9209A831CCEC162005F5D0EAC2B9C28E2096AA
+0CC08D3DC1E26AE4948184A3D7984FEF5E76F2A9C096D193F5DC85A9869D68D6
+0964D091BC63386638B73F55C14BCD739A1FDA77F334211238E22C400A1AB805
+E3DFD7C5CD4A0816DE5ECA37C985492754066732E64635EC290A390F0FBA5CA4
+BEFBF18881321281C9F1BA5DEE48782C30720765CDA224B31C130AA240B41CBB
+48EDFAC7503C445977078605B1F754A0188D444036F79F216276ACB95C51B7EB
+82B534C26348B50807307AB8D589ECB7D987E7F6E65AE288A876D556177C8F77
+9B4456C1509FA699C20634841263031AC9911BD549E06451C4CFD40F1889B36F
+12846050D1052999F93A8817591E45C0505F7C95EC4F7523A4E31AA990330407
+552358C63478662C7BDD5520521F6CBD5682FB835BBE90F460037A8F9E95850B
+4267C8E193279D44BA4A56BBA90C3FEBFDAB75B79C3C461058AB47FC6C250953
+BBAE8ACA33655275A23040A68D476305EE3130E029B11DD094206176829BB6FA
+1EEC885933965CA3C9B725359585D4AECBB1C9966F7816C47D866F3953D49D06
+8B111EEA917A37DF7E33D45F8FC2C8083A7995DE0A3B4DEDBD03589C47BF1A85
+98FC5B85E1550874B7EE6DE7747A09F2D2B96C9D435CB806F42F5AD7CB3FE7F4
+48C827616048B11061933BEAE5948465CD6C7A38CFB6A8710E8125EA2348098F
+D044A5C08AD2DF5742044AB0A849A09214F7444AF5C59167397EAB96C662D257
+C5D96428722BA2243D4D0BED9C255E6D5D8612E028CEC3C48D3115E2EEFC0A4E
+63819BFBAEBC1FBFE5B9E26645F0D069CB902158C6BC2663464FAC2CEB43373E
+7A81BDFC8A03DF6ABD960227881E7DAFDC7BE49E36FCB2EDB11043CF66503ADE
+DFEB3EA33C968AF09A5A7B4706B6A7C1806FC9C1CBDD3397D42117E04D4677B8
+C13633B85CE011BE4C046850E023DFC6EB9E82C1498674E11B0295195C68948C
+F46849D746BF2648182BAE0CD4430484FD1852AFAF3EA0C6FC448A7F89B9A6F7
+94B345EB8CE765C9D9629B537654D2DE91AD4F77A86440B3B3B8919F9273B7AE
+9D0CACDEBF200C17A0961ED8381A23FFCEFD0558E4982C77BAD368F2DE009C88
+39BF1C5DBCB7514F2C6092DD0707AA28AD82A35F79E0934092DDCBD181A82223
+F205284E209307B33FBAE45789F732E24316F8706E57520FF08AAE73A5D348E9
+3E7B79BA7A26313C7BEC1D924A1EDDA000C2B91C76B77F4E201358EAD4EA07F6
+FA0349E387246F9121F961FB021EB824DE1E6006BFD47EA776F4CDE7B85A766A
+C5CDDA65FC28D57C0C26C2D554073B003701B2203939BD66A32FB41B4E428E1C
+0F0A69867C88F0352A08CCB5B44250E7D292AAFB3580884AF6756B8BE17CE037
+322188772050502267CFA5F41B387AC7433C3C637F13D14E1F2AE685C79A4306
+AC418D6831A6B583C744B504920D6E9578BF3408413098CBBB576878F27A88B2
+735C4E0AC3E4598E0BF8113526267B359DAB5A6C80455CC8B2245C18531963EA
+EB42F71277176F0695F26E40853E400281C4C253F35AD21F8F35F661447F22F1
+A99B310C609961C2E282908721EE08F39DEB56F8B210969E5CA1E39E18ACF2B4
+7473B4A951B1F401582207D00D1DEDC7C5EB24124DB5524EC24659FDEFBB64DC
+581067314D3F80D776C5AE8D40FCD6BCBA647FC33592F22713AD7A15F2867861
+0EBAE8DEDC1082FF5AFD4DC0912B472AB94314D549D7C537716CA34B95ECD27E
+931F242CAB7B86C28D3BE165819A816DC44E9C663832CBC00C9E4B024CAB2D70
+A5E28556B189E6768CB1E2508A431F98668F03016D88ABA2552674F679473875
+8DC0B8271701BB8ED2670FEA09F3E566EE705A60B9DF720B338BEAF2FB349650
+BD7EFECA92B34B8171B2E37332C28CF55AD27D84D3EB63B8E3E95AA6C122469C
+2A6D4DED4E17745F727DF3660945696834D3CDCA99C71D9827DE568A59E1BFCA
+2FECF291BDE69237865AA7D407E321025556E9E37EBE23E7FADA007350CBCE2D
+E8BBB7AEC25A95C67319D6A019E2F048548605433D49FAAFC4CB807E311C309C
+8105A0EE68640E7FA16F54196CF3F40BA40E8D36DDFDC2E9F0E5D6774DD2B1A7
+34BA01E2A2223678331F55DE013D1475AB7C267B27E840D9239218FA9AF21066
+49E356787BCA5AA835D5102EC7C05E1FCE50E76CB17C63FE5529B95A2F33E479
+228819813814057B8E4F6DE3581CA31700EE8F4C2F43A197A7B3852D65FBB370
+69142E15ECF2D182559938AF99EF629BF947CEB09688E5DDD9C3244784DB0C59
+0016BABAEA12DACF3AADE9740D87C965EB63E1500C9697B37A89A9B8D15E302A
+5E4ED4F944A3776A7F7A23B4F3783CCEA6A6D7C269828049DFCB983DEF2AF6B0
+65C36F1DE137BF15C3BEDE16D321D4666F982115EEDB10BFF481AE58A6482534
+3F2DCE4282D81B961921E6FFDB8F8CA7AD72CF553E459C7DEE22E2BF7DC2B480
+CC34172BA9E2C9EA1BC5DDAC4F586515E61B57CDB27B5C4D494AF735B2D8A486
+AB45C459565625F0770A481E630FF99DF85F43E025215B2F3DDA3BBE6A26CAB5
+97FFBB00198F1F0AE425EAA0EB1CA1A12E916D19CA287FC1FB0C7CF7F8DFA0C4
+06F747ADA4448DE0DA98A984779DC8BC19166E5AD9D2B58D4547D341AEF71988
+EDB9536416364D0C4707BA66036711A969C179FCA59B621DC0BEDC07D2C4811F
+8017BA9E9BB6C7F4458A0EBB1DC7BEF6BFCF053A5EEC52010C3B423C65EDFE84
+146C7048B62B831E2DDF633695D6E1E8EFB34C6DB3BF75FAD7E774C2734D6160
+1C503050E700BB4226463E60F38396BC33078F1865812BC826F75A4C3A9A12C2
+845BA3380D4DD9172D2FDCB0CFB1557EF0CE2035C334AB647868F7946278C336
+AAABC01B0F9162E7D4934EE5CF69842186A052593297309E165696BD46981321
+37B1D45E7C895A73088BDA8DE7B23ECE43ADCFF621D1805FB5E2AEC152BE4550
+DCDA2F5AECE4D95202F9807EADA0CD8D14C7C14D3FFBA5124700CBC3247DB915
+52B874C428D60E48FE84CAA048C16FCADC13414E12F21685549CF5F2936D72DC
+5ED5EAC59115D13228115517A2CD809B9C63E32E72ED924C80C9DC36E11439E1
+161ACF9D3BFA9DE90530B2C62A8619B5CA9535F6B4B066665B74EA3B2FF87DB8
+2E835B3E1DE22DDB969B579E096EBAFFA48FFCB50D2AD883A49D1E8643F9CFD4
+84749D930494BAE7D8845602EEFDDD152FE4ED4C73EF7C4CC927E326D1BA14E4
+DA542B687EC2DDC23BB4A2DCBD6F9C1ABBC2ED663F1CDDF6872B6197775F5784
+1E9B6B7944F542ABB60C130774148E1900419DD9C471CCE6F300DD24A3B7B914
+B3F6586DAF7DB3FBC052A3BD83E4D3A9CF6096AA7460817F50DB8A1315AEAE81
+06F3A265F4B416FA3808429197DB72AF5AE1F3C9BDA047AE8893073265ABE109
+E175DE6226E8006CED49A63E6958E6FB43D842F520B34889A09666B884D3C4A9
+671ADB8EAAAA5F1D5FF28D2F72427D63FAAA43D99E8F05D07C93FDF02FA01421
+D1911055B1368C71B076E69F1CF039F7682F4C0049CDDF2B018DCFC9AF260814
+6186F93533A8E69171A3118E920F1ECCCE5A5B5D3129941D6CD3CDF9CCAC1ADA
+14FA85D8BD719DED7082B72A181E7BC96B02A7C4E6BAE426FEAE6FA275E8B42A
+D19343B67E67B176C541C3218B206442391A07C6718854F0402FD4CF04FFF5F2
+0BAFE7A8E3548210E33938E8F3B204201573B4FB321088A5DEE51A2C32697E46
+A46CF131F9B21A38FE3BA3F40186DD0C7DB3800C492CEE52196F3E8E14347512
+1B81B92A8BF4C1326032EE9214F7F942E614E7F58176A617C6585235E0257B33
+1EDEAABA1ADBFA216E2E6F094E7CB862063CB7CE8EA4B6BBEAA51617180D4BB2
+FD5ED5042C0A980D7E3D6AE86ADB4EC1CD75E60A4BB9D8F9EC0BD3091C72581A
+2693976B83FF4F4FB2143F34A4BD0019C2DEFAE797F0C23F11AC2E57
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMR7
+%!PS-AdobeFont-1.1: CMR7 1.0
+%%CreationDate: 1991 Aug 20 16:39:21
+% 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 (CMR7) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle 0 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMR7 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 0 /.notdef put
+readonly def
+/FontBBox{-27 -250 1122 750}readonly def
+/UniqueID 5000790 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF5B8CABB9FFC6CC3F1E9AE32F234EB60FE7D
+E34995B1ACFF52428EA20C8ED4FD73E3935CEBD40E0EAD70C0887A451E1B1AC8
+47AEDE4191CCDB8B61345FD070FD30C4F375D8418DDD454729A251B3F61DAE7C
+8882384282FDD6102AE8EEFEDE6447576AFA181F27A48216A9CAD730561469E4
+78B286F22328F2AE84EF183DE4119C402771A249AAC1FA5435690A28D1B47486
+1060C8000D3FE1BF45133CF847A24B4F8464A63CEA01EC84AA22FD005E74847E
+01426B6890951A7DD1F50A5F3285E1F958F11FC7F00EE26FEE7C63998EA1328B
+C9841C57C80946D2C2FC81346249A664ECFB08A2CE075036CEA7359FCA1E90C0
+F686C3BB27EEFA45D548F7BD074CE60E626A4F83C69FE93A5324133A78362F30
+8E8DCC80DD0C49E137CDC9AC08BAE39282E26A7A4D8C159B95F227BDA2A281AF
+A9DAEBF31F504380B20812A211CF9FEB112EC29A3FB3BD3E81809FC6293487A7
+455EB3B879D2B4BD46942BB1243896264722CB59146C3F65BD59B96A74B12BB2
+9A1354AF174932210C6E19FE584B1B14C00E746089CBB17E68845D7B3EA05105
+EEE461E3697FCF835CBE6D46C75523478E766832751CF6D96EC338BDAD57D53B
+52F5340FAC9FE0456AD13101824234B262AC0CABA43B62EBDA39795BAE6CFE97
+563A50AAE1F195888739F2676086A9811E5C9A4A7E0BF34F3E25568930ADF80F
+0BDDAC3B634AD4BA6A59720EA4749236CF0F79ABA4716C340F98517F6F06D9AB
+7ED8F46FC1868B5F3D3678DF71AA772CF1F7DD222C6BF19D8EF0CFB7A76FC6D1
+0AD323C176134907AB375F20CFCD667AB094E2C7CB2179C4283329C9E435E7A4
+1E042AD0BAA059B3F862236180B34D3FCED833472577BACD472A463E26637C66
+FAD3F83D6DF42D2272506260C16234A2DD01EF4B7C2CE47E1D05581329B003AB
+66386D9CD94A1C9631580A34D0B4FAF3FEBE63D81BDC5FB75DCD092CF3931640
+E738910524BF1DF55A57CF8DC8E966B7C0708CD41A37CCAFBFCF3FCE3142B95C
+9768BB674C1FCADEF8066087AF1CDD3F1D30E06CD5F60642C3AB0716881FBC53
+0186861A1C7BE68BFB6B172E66E31F4003FF836E7A3A794C40CE877BD8121147
+87D5732E808314B09562E8280608B52F430E93BE7D3591CC86064F6656F9F197
+227E2F42CC06EADFDFCDDFCC20A88782312D628CB09492A4230ACD39C7D0D767
+B4DA640D966F838DC3757870867D0755ECBF58DD1AA7A96E76B3480E1462C22D
+7317CD112DA2107E11EA821D5BACFAE97B68CEEB573F525061B2CCEBACE22D4E
+4BF7C254D6C2F56D38CF5E53EA59D0575FF57703BBE6EB9EBF8207B4A5C5517A
+FC3987EA3A3BEE8547398518F652FA4FB527015437A740E12721F747899290B8
+9C2E8215A3A845E6BEDE2319FB631B2AD0E9D8E6F16E518225876A2587882297
+9C8444DA1B9D9A1988AB36371FD56E61FF0EA7B6DF3FC51A7BAB65ABF08E8E45
+DA6DF06787D291854EB43258331F5D52353E6EDF6D44266307D4E5AB9237130A
+E903D5438B519A8AA100C916107A8CEFE67E86F78A357E016093B00F531DEB76
+72EEA9102B60697377C5CFEFC9DB5362EA5DC9A194A7EA680C9E7F3814D258F7
+7A0EBD57EC9EDE06D35B641BBE4C17A758426DD80C5F7FEEF2F791967C9A034D
+62E4B2168AF1CDF5CA8BA1667A835E387C78B27F9DD0170936304A6EFA70C316
+8EF07DFE83B6D3596B8DA21CC92E505BC58D8E3A662353BD3A5518038123AD4F
+2D4832CA87AD9093497246B6837C70B56A9F23520BE34D55F71BAEE93E41FE90
+0C4762DD5EFC95D69D376C9660DB0A6D5FA8EFC1027CAA4C5A10E0A1022B989C
+1B6C519ACCF6701973542C390B7A54BEE74D40BFB7A400012CD770AE7DAB5A61
+93C4D76C7FBFBDFD2E068D418FCAD289199DEF8E242EDE44B040764D13E867A8
+ACC218D0F65076009735F1B6DAD43645FEE77374D501ED26CF4ACB0E13495C94
+BF8140E0C078EEF46FB11099DC2B9A691CF718909604AEE70428D43BC46E03BD
+88DF7E7B95F00231BCF83C935193A7029EE0BD8A81BBE51045FF0FE5878BE9A7
+0BD49F32B2C47911880F9EA1F41F4A1272501D6CA9EC95DDA2536AD78333BBA8
+1A909E99189CED1707B8DE4E4E21DE3A89BF2D31D3BEE08A8BA966436CF89C78
+2B0034F5B84EED28D31BFC24EC39F3B2EBAF55ABBCC8CB95BB4160FD45BD8A7E
+C10E6B1460453F760C4BBC01666C412C1FA7FA662C98E7FFE745737FCA96AD44
+AFEF828F2111D808E276AB87E1A572A2096B9649F0CAD5AC96AE04299B456FA1
+454DC7B981C6B1338DA7D773FDF86941E016A528C79C963DE11525E23C3AFC23
+991B1B202731B7D31BEDA0FE267F909D282BB16CB0893F9446B465C7AEB942D8
+6F14D738ED87E73406D7BC46D3104BCBA64400D1092B2414B14222D2A20B80F1
+FF9309184BF5EA6EE13539FE8AC9AC4FC7E743ED7CA48C608C2556E5DA5ABD79
+04D3DA5CE8737B3B4EC2A5C16CB91FDCD2F6B4AD6BBDDF904DCE389328546564
+EBA18EC09E0264C27144CF43147EA1142BBD17A840FD9762C32AF999B66BE9F7
+1E779A55834CDCB79626ACA171FAEAF8BC3D05CF477DE2AF82D4DD290B117AFA
+7D59EA61CE8E5DB62C1FC8EEF581077A80D03AA377F71BFF886D7CAC24EDA7E2
+0E56678403B3D00850C6F1FEF4A9B1622A65054E8D629B68060B71B20534BCE5
+A8D7A6498FE644460D0E08F3CA382B1C5644A580EF8CA55A1ECFB9AA7BF02CC6
+BF229177FC721A4B757386369739618E67AE0938F9DDC970C81101465CB691EA
+A8380C8C41D263852C95F5F70835C1DEF6103C4994E923BBFA2DD5A24E043230
+222A166F59A2794DAA151672F10E8CEEA9816E4BE722A2A45B8C9063B0C6934E
+169DF729011E1E59B21AFC4B0A725E2193FF
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMSY7
+%!PS-AdobeFont-1.1: CMSY7 1.0
+%%CreationDate: 1991 Aug 15 07:21:52
+% 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 (CMSY7) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle -14.035 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMSY7 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 0 /.notdef put
+readonly def
+/FontBBox{-15 -951 1252 782}readonly def
+/UniqueID 5000817 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052F09F9C8ADE9D907C058B87E9B6964
+7D53359E51216774A4EAA1E2B58EC3176BD1184A633B951372B4198D4E8C5EF4
+A213ACB58AA0A658908035BF2ED8531779838A960DFE2B27EA49C37156989C85
+E21B3ABF72E39A89232CD9F4237FC80C9E64E8425AA3BEF7DED60B122A52922A
+221A37D9A807DD01161779DDE7D251491EBF65A98C9FE2B1CF8D725A70281949
+8F4AFFE638BBA6B12386C7F32BA350D62EA218D5B24EE612C2C20F43CD3BFD0D
+F02B185B692D7B27BEC7290EEFDCF92F95DDEB507068DE0B0B0351E3ECB8E443
+E611BE0A41A1F8C89C3BC16B352C3443AB6F665EAC5E0CC4229DECFC58E15765
+424C919C273E7FA240BE7B2E951AB789D127625BBCB7033E005050EB2E12B1C8
+E5F3AD1F44A71957AD2CC53D917BFD09235601155886EE36D0C3DD6E7AA2EF9C
+C402C77FF1549E609A711FC3C211E64E8F263D60A57E9F2B47E3480B978AAF63
+868AEA25DA3D5413467B76D2F02F8097D2841BAFE21CA9B940712F2E5C0EE674
+F89B3247B69A86E24CD2A3B7FD864727497CE1EC298601477FD97FC32D7866FD
+8ED1F7CCF92B61B0E16F60CD717C9C63049916F6BDDCD5D34DC2528A08A74E1E
+4DA5B7730109B739A72D3C25A5FF0CCE1DEEE2BBA598B2645DF5B81EF0C9480A
+1C6A24DC27E9A01223B6B3D268BC97B21B27C063A06788C60E09E02EE2AAC6FF
+F378ACB658097260AC1EEA8B5407E6E7E8A5CA2068B9943353E816C58191CD13
+9BEC68E7C147B3096B0D5585D38E38372EFA5046F5131B2CD7384CC27F7DE0BE
+FD7E058199C6634638166E021F4AD709008819B705CA4D2DE21926ADB073DA5D
+57F4D7038EDA03AF325017193205CB9D2C33D75CF0A0135B96319DAED586EFC1
+C29EABA78C54E83DB8320D49DE2C4B32E87EE459E61DEEE68BF573AFFF61B053
+B8BA1F3ABE1DD007AC7D31B51855216B448EDAA69D0E6199F40ED759ED32356C
+967EDE53B4F9E144DE735D64E7CC0FB17EC42906E2197FA14F7527978A325737
+906D284E04C7058834815FC92BC933D5683F3064DBC9C310644F6771B6D5BDDF
+5609DA98A3C9052B0C16525EC32F01A88816705401293AED0D53B8B9ABCB9CC6
+4FEDE96A9E63EC6E1F3E39A3C0981117785B59A5DDAC3635AFEFCDB93412ECDB
+D4C966886689283139890C1E195FCA76440ED08922C2D1E078E735339AD3DDF6
+6E49069FDE022121FE61A836548156673C7A6D9BE216642CB9734381A6EF4CDF
+68C417A0AA13C0F79157FFC2769A174C70F8B325975E676A10439A3B700BB9DD
+20BAE3A695CD4A9D864AF5A5AF01F290EBF00606AAF2D27C8C21E60A4BF037D6
+12DAD9304D4D860EFD898C3A71E4A9D014E2175520709C907D7241F982D6FE98
+F157E90640046E53365F7992C477ED04A7C2572142EC7ADACF6DD562F96C252A
+0AF90FA4DF6132F1B63690010043950C04CD6D35924DE2CDC0327B3DE603B88D
+B7755236BBA8284AC2F2F7439B048A045C
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMR10
+%!PS-AdobeFont-1.1: CMR10 1.00B
+%%CreationDate: 1992 Feb 19 19:54:52
+% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
+11 dict begin
+/FontInfo 7 dict dup begin
+/version (1.00B) readonly def
+/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
+/FullName (CMR10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle 0 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMR10 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 0 /.notdef put
+readonly def
+/FontBBox{-251 -250 1009 969}readonly def
+/UniqueID 5000793 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF7158F1163BC1F3352E22A1452E73FECA8A4
+87100FB1FFC4C8AF409B2067537220E605DA0852CA49839E1386AF9D7A1A455F
+D1F017CE45884D76EF2CB9BC5821FD25365DDEA6E45F332B5F68A44AD8A530F0
+92A36FAC8D27F9087AFEEA2096F839A2BC4B937F24E080EF7C0F9374A18D565C
+295A05210DB96A23175AC59A9BD0147A310EF49C551A417E0A22703F94FF7B75
+409A5D417DA6730A69E310FA6A4229FC7E4F620B0FC4C63C50E99E179EB51E4C
+4BC45217722F1E8E40F1E1428E792EAFE05C5A50D38C52114DFCD24D54027CBF
+2512DD116F0463DE4052A7AD53B641A27E81E481947884CE35661B49153FA19E
+0A2A860C7B61558671303DE6AE06A80E4E450E17067676E6BBB42A9A24ACBC3E
+B0CA7B7A3BFEA84FED39CCFB6D545BB2BCC49E5E16976407AB9D94556CD4F008
+24EF579B6800B6DC3AAF840B3FC6822872368E3B4274DD06CA36AF8F6346C11B
+43C772CC242F3B212C4BD7018D71A1A74C9A94ED0093A5FB6557F4E0751047AF
+D72098ECA301B8AE68110F983796E581F106144951DF5B750432A230FDA3B575
+5A38B5E7972AABC12306A01A99FCF8189D71B8DBF49550BAEA9CF1B97CBFC7CC
+96498ECC938B1A1710B670657DE923A659DB8757147B140A48067328E7E3F9C3
+7D1888B284904301450CE0BC15EEEA00E48CCD6388F3FC3A3A4AAE8B33E71E4D
+02DEC8C94BCEA6B88FCC6940D4700E371B76295C26C6B423098A3A57AD41F0DC
+A7EF476FAFAF30B36B2268B9055F55100F35A8925085317CCF8666F132417AAB
+A5FB3F070DF1CC8CAB13CC24B4FC62AA06FDCC588D847D0A4C232E3AD5677DE8
+593A4E99A03A35942CE0F211067D815E95F80F722072DE3D6F652FB0C9C9843D
+A6DD2482804FE10153467FC0AFC951B645E30ECCABC655E444F87B8B46C868D6
+D2FAEF826C12ED0AC7340059522DE65BA324640B70E9B902E9FB32F894C4CF49
+421240850599926BFB60AEC2B30D2A6340553EF7512A32BEBC5C892A004DFE16
+DD9CA9ADB963D51374A806711890C1AC5997F21A14E7523CEDB295474BF8C3A0
+5017A993D91DABC265A39EE4A494BCDEE0775982D2EB9D2112ECA3C2E597CBD7
+A9BBAAF9A808FBE68DB5E710CD843327B1F360133AB8050418585ED2A27ACD28
+09CCE1587C2197F6D65CC586FC6A559ED67B461D0C5B373C069AB10F484D685C
+09B6A74B99B5AC793F7177BD0C2BA51F9C2EE1202FC0201083C42B70DD3A9FD6
+89455D4C12693F89AD5DDDB68447DA380B07D4EB2DC6C52C626E395DFF3C3E8C
+45D649B7709164B8EE423D319798AD8E8838055E2BF8AF666913C95B13962F03
+9F2CBC28DA9623870123BB1F74D0B6783057D77E416BEBF35C72C4107123759F
+3FDEF05D530F2504362EE1BEA662B5BCED70A87CD0D4798B97D61B2BB33D7298
+86568307954003001CC236BD3016C37C9A52426D73FA201302382872B3C31E86
+FACB377A74258F2A8F05FE1B2953EF2921C66197DBCDBEC06EC6AA5C7D14FEC1
+58D3B98B0E716184023E82B3C9FC62DB703DEF0DF436967EE67EE00782F70B7F
+DAEC2B9A287735823661790052A2C72622EB41EB81360FF0B4321340021AEE91
+51BA238FF85B87F71A1FC7D807E25624AE06120BF711462AD9A5A895E80FBE08
+815B8EFD6C61732CCB103BE9BE33F0D0F43B65C016EEB85D0D765B5067DF543F
+2C6E8267DC8F585A2796AD110AF157310EC95A4A953616A78C24FB401A1D7155
+EA2A0AC1585F41F489AF13BD7744731C00FC9FD7C374F9C23E4065635F671F52
+811F16D333A30CB8628B59F522D43B5D73C1DFF78A4046611112F940D54BA533
+D02E4D17F440A2E42C0D34CD1242A01B82AB0AA0D511F2E35697A53012D6FC9A
+958D668D5EB74C89659986B53FFDC40FBBF8EFA91E9D73A6037C5506CB248259
+A00457C3D16BA47195F3E2FDE36F5E41DF571863AB622F6B4CAAACEA67C87281
+2F04AD08B10D8D123EADD2587C3EAD9C35F9D81BE13CC04B6A60E5F530E928F2
+6D3FF1879DF768D0CAC7AAB9457A947859C5B88F6607539C05DA2FA2980EFA37
+68D7DE76B0769742EFEACBF54A929A9F47FB85557AFB60A5488F615B52AF26DE
+FE6F9E7502CA1FAE41062985438726B098B509B59F5196C007222A98A17E1728
+BC7A8D9067C09C0617559F2EA60D5A0DD4663DEB0C49B0CD7105220662CB1BA9
+346B65C2D4404DA1238E527EA0AABB2B79809627E06ABD2F0B8512AA269FADEE
+73BBCE65C95A5CEA44EC76376622EAEB695C8565729363D193BB9A8B1FDA9E3E
+1C85BFABC19B32EB9C04D6F3F3E0990887F2FEED617F9FA15C9D00A56D516A07
+D078760D88C6983B8E48601A3A358EBD3CF30F52E2DDFFF3022CA7C9A8795147
+4E5F08603C2691DA719995A219D393FD71C467373B6FDBCF7D6945D5FBB524BE
+17F4D0AAD8D30B202BF36218327DEEB5008AE1B7F7BE5F2D3B418D18D7A7BC9F
+E32FC93AA3EC6A5DC470A7956C69C774AEA8B798411DEB03803C4B327BC43588
+8D9E4DA626FE630F497B3F4D9C145BBC57597972DC2C55B9902B87296AA47CF4
+AFBCD10BB51AF09FD39C75ECA974723B5FD321CCC8CB7E79D6598260A4025BE8
+D51BDB1319F92D15EECE53E656263C187571D6A8A04F2CF0D25C158567D1C5C1
+DBCED064851B0B652579A0CCAE1BD85E73AD13FE2778A7B3BC3B1576188A2378
+2B366060A72E2CF6C06DB63F2A4191D08282211106B524A3A462AFF4FA624CEC
+71FCD1254FFC3B91FE9548E162888F4143C80EB5A1443231CF1FD5FA12BAE538
+7A0843E44916C5880021EF9AF97FA3FF5F6E316E0793C5DF4551853ED71BDA62
+391768EE1645DDCDD33172921906F76912D1E6F25DE0B5E89AEA3D248B0D1CDB
+C4E2CE5BCADBAAA49C95F85F69584A26D6BEA952721EA71DAFB052DB47FB09E3
+5303AA0D2B205F0753C6526B5FB038664484A3C72F72126088E48EE5E7E4FEF2
+C817E214D08DC007158182FA487AB21D630D1D01AFA564FB6BB517DA264B964E
+F911FA1EC3ADB1B22F4CADD2976D985EE3AB277C3A00449B758A6B5847020414
+27FDE6F5D850AD68ECE112DC032334C231E30EDB7528E1A93C45A365BA1654F2
+8329E8D202A18DA16BA37BBDF2BEC71F78E286A33DB6B1BAD825A0170209C097
+EFF63EA6384528E2FFD2F465B1DB45AB1307A057C9E6425D284B68307B38948B
+129011CAFBDB268043C10F8E1D608580CB090DF627E0C3ABD123267954A7423D
+108F4E9C91829CC7C89D142E5734F8DFCD5D474526FE57A2CD615AEE46562E97
+4FF11097BCA90658FFC10EA7750ACBAEBF747A7EFB0F7AE7995FA2B5794B926A
+5E9BA091835A3D8CAD31340E8BFE912D61FE8256362360520E89F41A42FD6A42
+1E491BBFF1783DB1089DD7C898DACADCE08BBFC73CB4DB6344B7FD4A2D247E35
+42C2F600274EF2CD86B4C5B7EE0DE2C151787B9E086C6FC539D86340A68C5114
+8587626C40C105377C2A0E21F9C86582AD59F0A2B7769B56BCFE01CECFB8EA4F
+A9413DFB7A8D880ECD4850D3EE7181522E9E57BBCFCABEF2142FC880082BA893
+E1BD7B0C86DD8275FB2BAE542C7885E3E4053F0D349A8D11755AA86FB900F731
+F629387896523607C5D025AC495D46F4EAC2DBD04867D942EFFB82EA3D9C3772
+8F5052D5C322833A3F1BF61DD574E724D269BF2DAB475E0570FB300000C7E01B
+BB6A231B44272AF94FBB419209B36F5428E735A886E6239CCF5AE7E193DD5A2A
+488FE7AD235A9D3FFC0658B28746E17801C9E88F89CB0B6825DDE18914936032
+6C09708C0A40BD07C094F7F43589E4306D60C925365AB1F949576D62E3820CAF
+C110319885A639B59FC0A14D65D7EFF7952FB6BDF6D8C36A78ABE2BD5496CA66
+2DF5CB1BB78AB1BE87AA40DAA99BA08BD873337112068F230271F1028198C33E
+6FF348262725B7E010F4D6A342032781D08DC9A40B3127F9870D975FF928C6BF
+5F070AE2DFF517815E51D7BC83D147BB30B716936AA091BE1CB441D95CA42FF4
+A438F09FD688ECD663C5F2165A59679066BCD36EFF5581637559FC8DD1C477C0
+5589B2A62C3E79D4138B11E4E7927AF86CEEE1343C143471F1F90DAFA841AB3F
+197085E14A03148B0C1A15D2232EC67980AA0DCAE9F4AC293BEDF9A28A582D2B
+C2E960622E8EF8C7C10C99F304BF75725A844D67C468AB95BEAF9F8A555544DB
+23A4A376101944DB367E514A39AED1D446C95E591D77ACA0BBD5908128B22D37
+577CCF44BC85A038C84108D1BE6D5B06749C47E91581A6FE59E064DF9568F556
+27661D3A831D3A143AD1A48400BEFF549E88321A086340688C0847D3040859DC
+F8BFD83C6278BA033FBB0065AAC68AEC23746BFCF253C93D4E2073075DA5C18D
+16AE9A06F9EEB2272CC7A928CC2B4EE898CB7C03360A0D5BADFCBE6BBB8D7452
+B95658E83B2B98D2BA7C5B6185C3A7472C778C5E416C9DE8F5CAA34E6CEB118F
+6C2B2E9478AE775B4AB698880AB7CA2A5719A173098BD5440428AE5B3AED9A4D
+170810599327746C627702915658E9BED0CFEEC0E15E4A183F2556DD264B239C
+9C01E30DBD6795C4C28AB66FB31C852A61721AADF01E3241F9A88F5C289540DE
+00179B9EF26B7660BF4D7FB8E95E14346DB719D52C3365A72E542CDDCCA2A20E
+6D6F3EC329A51895B0C0962DEE7E22A7488922854991565E2D32B5A81093934A
+3D6F8DB44FD08B19099BABBD2086E0E191E0D360C430D122507E32D3E2E846F7
+1FECEB39512C621EB29167B0AD910EAB3B127B919D8E6402042E990370310D9A
+7927AA4FEE605DB46EA1A9279A9D735C91FA76E63DD48598E07499EF5BE46300
+D45C2D2759611E3B44388CB4E9F0484470B66AE7D693227E0102DE692CD9C820
+27D146F5BB0C6EE06292755BF3015826637E5DC5EC5BF8AB95D614A5729A69EE
+6698D84A2CEA46FFD20D883D20A15D95FDF6F270D8145C7DEE41A790E0B3D0B7
+EE4E49CA39040390BDA69192D9CD28447C507E41C2E4E64625300E23E9EA239E
+C88EAF4C4367F8C7
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMTT10
+%!PS-AdobeFont-1.1: CMTT10 1.00B
+%%CreationDate: 1992 Apr 26 10:42:42
+% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
+11 dict begin
+/FontInfo 7 dict dup begin
+/version (1.00B) readonly def
+/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
+/FullName (CMTT10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle 0 def
+/isFixedPitch true def
+end readonly def
+/FontName /CMTT10 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 0 /.notdef put
+readonly def
+/FontBBox{-4 -235 731 800}readonly def
+/UniqueID 5000832 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5F00F963068B8232429ED8B7CF6A3D879A2D19
+38DD5C4467F9DD8C5D1A2000B3A6BF2F25629BAEC199AE8BD4BA6ED9BBF7DABF
+D0E153BAB1C17900D4FCE209622ACD19E7C74C2807D0397357ED07AB460D5204
+EB3A45B7AC4D106B7303AD8348853032A745F417943F9B4FED652B835AA49727
+A8B4117AFF1D4BCE831EB510B6851796D0BE6982B76620CB3CE0C22CACDD4593
+F244C14EEC0E5A7C4AC42392F81C01BC4257FE12AF33F4BFEA9108FF11CF9714
+4DD6EC70A2C4C1E4F328A1EB25E43525FB1E16C07E28CC359DF61F426B7D41EA
+6A0C84DD63275395A503AAE908E1C82D389FD12A21E86999799E7F24A994472E
+A10EAE77096709BE0D11AAD24A30D96E15A51D720AFB3B10D2E0AC8DC1A1204B
+E8725E00D7E3A96F9978BC19377034D93D080C4391E579C34FF9FC2379CB119F
+1E5BBEA91AE20F343C6420BE1E2BD0636B04FCCC0BEE0DC2D56D66F06DB22438
+452822CBEAF03EE9EAA8398F276EC0D92A7FB978C17805DB2F4A7DFBA56FD6AF
+8670EB364F01DE8FCAFBAF657D68C3A03112915736CEABAA8BA5C0AC25288369
+5D49BD891FABEFE8699A0AE3ED85B48ACB22229E15623399C93DE7D935734ADA
+DA7A1462C111D44AD53EA35B57E5D0B5FC0B481820E43222DB8EFCD5D30E15F9
+BA304FA879392EE0BCC0E1A61E74B3A1FC3A3D170218D7244580C7AA0DC65D19
+741FA5FE6F8CBF60250ACC27454BBF0897CA4B909C83A56672958752ED4B5E79
+E18660764F155E86F09EFA9F7685F2F5027EC85A775287B30E2069DE4E4D5712
+E7D033481A53A2702BA7542C71062173039030CF28D8B9C63B5596A9B42B33E7
+D922944A38713383D3648A4AF160A3B0C8F3379BA4372BE2E7EA49AABA75AEEE
+C5DDE1D8BF68483C3D21271280ABB91D54CC819680322EAB72E1250A760BC8DC
+FF798F2ABFC4F3539392985C4CB324B0007229586D1E0321559F67C057FD7902
+194490A4C133DA790FF3BF23A13C2B1B69EEB75950F9106F2BA1E3CA65C90FF5
+931DADF03DA48AFB8561FC2E710087251BFC42B80B297A3DB0DA138A7622A931
+DA293B0C740987ACE9F2A8EC2DB98F85783C01623FD3612C7E4A84FD93446770
+C3DD7431F955A5F3734F6931BD790F0A45B8D17CB74BDAA4BFF6DAB5380CBF61
+72F37CB67A909E2842E0AC5D9D07D01A4BABBDE2AC70FE5753460D7E1A708B7D
+0EFB2B5FF55F9E4571C466AF1F91E545585845B09D855C3A01F713C1BF081EB2
+7E2A0E598708737D475BEDAF60BC100FD0A0628C6001A203348CF6A3AFEE6DEA
+A2EB57E35599FAD0B8A52BE1B77757E92EA2F51BF07A285E26A452F417D2751B
+3D53F9D671EEB920B5F0325D4D4C2634C07508492279701623E5224F9DA08DA5
+E0A6923FF60DE967B397E5FF727EDA6331CD8A7AFD99B233FD2F29DFF4DAEE2F
+19077C666DF3D5ACB7C7A5D8321D0F0A3BCA7A7AB68700D672AF3BBEF1809303
+753E5DF7A93276F4DAA3D09B72A1248945922F39E7AB9B8A6C79FE2124E63F81
+0E6A3000E8960E71B73E9003276E22FA43B41CCD82192393C6F68918D74284A8
+1C6DCACF6E4A04FAB353367691E77B3E5DB662E87168ACE89E6A568B132E7BF1
+FA92FAE466A025FC48F8D95A20314AD2D0324653A4D9D6A7E2A9CF11A56A6C2D
+F790867A086625553513AE95C1A9335EF6EE11DF29540A0AC0897B7894B0BC07
+3FD16A89D54F7187DA852A526D351F7DF03B26D6C7E4C28BA37619812F9C6729
+6CB735C4159BD3E3497C822AE436CB6822542C6F20D695B2209B499D4F15B882
+072F3F3E6FC8ACD527AB0B309CD62328EF03D48CA3BA524CF25A591FBBAB1B03
+85707D76685EB2A44A1E9671D986927AAED4E779FD23ACCF40156DD7F4F20649
+01ED2E11A94DBABC0736D8D8C4B2947DF17AE0298F5331E314D82FDA476F387A
+411CF772F6679D540B9379281DFDB0C273FE930D5FF978CF9DB0EB3855F84214
+5F37A191A0BB5C32505BE1BBBE32C5C7ED10524771E33BF697195130BE135AD9
+2670F2361EB61D268A9BBAA13CC1746993366BCE051FDA255FB9DF108F59405B
+6DBDAFB5B404299DDCCDAF7E3886C57AD0806D98B27856AD22E13A109E10F6B4
+78DD600DC32363E8FEEEC5007AAD9E6AD54C6E7A7F4544E2C7DBE0BF37438EC9
+81074DE85E6530DFB8132290B4F2B74F6AE2851EBFACCD90D2845ACEC1594E0F
+4BCA60B392A8FA0EB0140B3CEE5BE7ECDF2B2F948A9A9BAE4A7FB131D2E5F1BF
+5589B5732F27150F49B151B04D645767EA0C6D7D1AB9F93AAA16AD20199BD772
+8230BD475D9AF0E1914D2DF09BF82CD56E3A5650BB03635C286697DFE72101B4
+E48D5A11A38820BEC0566EF716F98AD9975D7AC26D99EC250BE8EF16B3DF5E66
+8F68D4DF0D6ED285E3739ACEF64256CA9941D25673C189149B5C553382E7B623
+A570600EF4683E3D4AAE282804F30DB28E2C028B9FBF4465FADA2BB3E27B7DE3
+9124DEE2F85A7A5705FACF1FA8F6C72AF140BD9000676AD17F8D9D743EB15A70
+BCFEA6CFE8F85CA0FA53CBEF37F683386ABDBA237768AA1F45B95364E0D6189A
+F5A8220E38AD1620E05ACCE428A71039FB71E48F306E3188B696529B03795D68
+F1B9C73DE6CEDE8741F080069EE850AC352891DC75A58D6AC869E731074FDF17
+6AF15BA998E6DA5179214D17A9603267E374D118A14065B9012E2036ECF26910
+3AC979102EE37859E800F1E7A52E3F571BD1045B0B739D7EC6C67900CFDE0AEB
+194A7AC269B723CBA99B17504F42E05C8E472B05E95039A95EAD0E3E8180485C
+B6F166EE27D9F2B97A1CC1D0EA09EAD81A21E579C458466FBDE740691124A7E7
+915021DBFF343FC7CCFDE1B9C82DC28612412A07BD931E0E2C99DBF80699F9AA
+CC0E9FF52C3CA7E354FB69545EBE57315E2F3769400BF2C48D9ED5ACAC9381B6
+083DEDDC463898148B40AC45DFADE6079A65E30F7B1E42895A9E86302C400A72
+C8E049AA66EA516058110132590E425A47A1D774DFC40E54C3DAD5485AEA119B
+DA095843223F2A3493D79AA4FB9A7712537CB55F64E549E634A1A15015A160A6
+B16BB76A4307FD663F3E486A7F7A0A406D04BF8A3C89A9240047C0699513B681
+1983A900B008436A2B9217377F78F2CE94736D7DA8AAC244EAF0F130B4F774B0
+AB9959FD6B9ABD0B702D9A53ABDBCFA72B02A84A889EED2A59DBB6F57BB8CB32
+87E4A073F9BED4EADCF6182AE72FE3B197152C8F983E8DA33D21CD5AB578273D
+59425CF45C6E4AD6DD6EEDDBD59821CADC190750A685BAC77F1ECB809E01CB54
+EB18B483645ED2910FB3FC2A316C14F2630685DEADCD51886FA78E961E2B0ECB
+A98ACEBE5CB5C549E5BA41B3D2D407E1BA08633DE5F574B6611279E6CFAD9921
+E0BA191E0B49B220593ED6B6AC68827346920BBF412A4AF36410F9F6C20945AE
+D052D07B883FEB56CFCE56EC59E4CF5A2B9BA9F79C558E12BE7503EE398503AA
+D8D1563695A90B6F1BEED60762994EF96DB83ACF3D14B9E5C174CF15709AB3A1
+D1E3BACCDA09423C9405A0C897204C8940E594FD8D4E8657A4405A5D009FD532
+AF50AE933B31A50072867C67DAF1B74EAF1C45F776BD773998821B6EA0D567A7
+C8A62204FF65A8B2374991BDB93DFEE20826FFDD820615B2DAC5E67E01379688
+341F07A47EFD5E69C5C3B37C1F07B75A05B334153C7EADC9E8B34A6A09528BD1
+C79F85BF256F587D099ED8527847583E349556EA14C071AB5C6DA57A5206F00E
+108EE3CFF337C71DF4A560029F686886619249C0983FA6E81C9FDDE67AE046FD
+6E9C963171C651D93774047DE471BA14F1C429BE4031CD75A028C6EC21195BB9
+1C34FBFB8D6B04E25F100784366144220A2418B289947B28947901CB249C2A2D
+79EE7C7DEAA2DDF6162D87DDE783E96452557D9DACD88CEAE010644A0C5F422E
+5FE82F46CBC55664CF2692893B38DE2FAE71F26EE9D4ADD7EBAAEF7FEF33A953
+271910D9A0C71D4466FE15498F7A821347C0212C697DB1DB0FAE2E5E06DA5BA3
+67475B957A52A98506633795D1E97E90BA106EA5927406228805E263F9D5ED15
+02F6FFC39A22ABC7C9FEB61D62DF6BDA1788C83B097BD949C7C557CD6E1513C1
+D538546DF914A83A17D9AC0B6C842DC9C982B7B4D22F3267595C3826EC7DA8F0
+FB9857ACBAD46820059007CB043FC05B4E7E3F51C0842FC877AF79A502FFEA70
+3CB8CF747BAF9936AE10859C2FF796AF4616F9089A6342066549D299A45105C7
+53C2F88B036FA22824A204871A97552023A7D617F5FCF37677654FA8C69B59B6
+5D3D8D92B608911B12A4F58B34361C4F9007BDEC0A9114C41C48DD88973839AD
+2E11F0DA3FEF8C7E2B8E907C3219785A11785C6D02C0AC4848286A6D8B774B61
+1A3EEAE4C77E71B805C4647B9AB57AB32A4EE4E60725F827D879AC30DAFAA68B
+5B4AA9B0CFFFA90C4DAFE5F6BC845C0B3F5829E1A5091C7AB25999E3DAE23FF5
+6CF9A9D3859436A121A56CCC42B82B4D171F7FE8773219D09F13B97081AC19A0
+68CCC59EE9E9D76C8EC06E88B54865F5FAC92A8AF3CEC29ED25C136E94944573
+CFFDECBAF737F95860460B53833E94D1A975BB27819C2204082D3C3D266A1F02
+25244E82004F396596087AA2920D56013BAB99A70DC8EA792E7B287D532B7858
+65BC7CFDEAF2B1FB5286F820139A17CB50B228F6E05E1A31B993267FFD12BDA6
+35E137FEF2DC5A69656011FD0B4FDC10A0B74D6D5F501A2F6D9B8A7D4939333B
+9AA3D61518B365B7B2E6180E085DD49DAC465F2CD0F42267FEC44B34768E7E75
+FF6307897A5FCCE673DFF5D9C044829403C37B24D754580DC34534C610AB2826
+A435AF5698500ED0E057392B4109B7E892B3EDF264897800B3814311E70EC178
+DDCC5A3B4540157CBFA4BC22DDA3546102080FD5C77E94C09B97ACA1FF132940
+CB85DA88C9F7F08A69AA8A7692DD51B1AF6A6C53D0768AE23A2B3725683077D1
+BC09088A111DECE32CADA70D7199B82010D0B9F675C76D054A60F0CDD4DDBB14
+A66BB4C922B14CAC9F468B8AB84BFA14DC808174C41DD360CCCF3AF615058F29
+2E91C3F476200FEFC52E31CC748FC31EB26B36302271E7F900AEE6CE7E9F290F
+C88A8D63BE2F51412961851CFCB5413BC581FAE04F45FF6DA4A763F8FB53D4A9
+25D1C4B125C8DC1D55D2BCB284E3206E84B993B9E6D145F0847FA27AA045858B
+7B1594FA9B2996946C7CBEC88425D2FBB6A65A89592953C37BCC6933055A679A
+F5DC546F578280760EC954F9F42AC3640996A56075789F4480F141CAD086C097
+B4295FBD689E73B7CFA54582728AFACB7A3E46402398800BDC8ED50F3D7019D1
+03D34D7C9C953F07E00B45A32DBB62A051EAB095AAC19711B6D2038ACFB8FF41
+FA7951054322BC8438458A536237D38DFA92A1792E01D299260E77E3E73CAAC0
+BAB7EBABF2FFF1EBA4DD9694B2D769AF61CCC2769EBC626A43A292765D5A45D0
+CBBF4133E213BF4B566755BA393412B4326C7C58AB5D29B4469FC1D17A5D41BC
+1252C50645D8D84C82A317D8E1A9927FD28EB42432C7D4CE7FD35D4714AAE069
+1CCB6FAF98493F52554B39783D6E14F0DFDBB8E9785AA46C7C827ACA7D025B4C
+4F5E2226DAC0FD8C307C527E88EDF2E26AFD2DCC57FCE42FF6EC8DA982ACD57C
+BD90A7272269364D1FC10094AF55F598AB07B55E6FC5BCE1A989CF666ED2BFA6
+27B57DFBEFFE6C1A05859A8F1C0A78ECAE45AF09266CE82C09A11BF9A635A796
+539AC21C997550F7D5432E12930E1CD5FEB83A77257C0CBA3C64DD717752AE05
+5BF719E6F202A598D91A2E158E624FDEDA5C49DA80BFD13B9CAC43D15A753EE6
+7887AAC3A1D756ACC12AE2F65D8CF9DF1E2EF860F7DAEF18885A1DF8E5B012E9
+17C6D8EBC5CE636C9C38A580ADAD4C3981FE8CD1B9A8F8DD3FCFF76728AADD3B
+E7F93E0DB9DFFF694D5FA6A621CF3FD0E736648C7696F36D977B0CF52389894C
+C62B0257D049F85E79C5598472101508FCC41AE008B0ED2AF8454CDF40B7B0B1
+0F7146D2A353D0CF42D26CBC1EE2AB9C17FDEDCC3C350350E7015ADF8A372B51
+27684CFBA200347766BA5BAC1DD423778B2EE7CEB6F4B0365D9E5EF56DCE95F8
+6EB08F92E4A0019CED357290A6FF59214B55192B7E57FB78A1806A75520687C2
+E8EB90351DE419A7C98EAFD5A0CD6D4984DCE1E40E6E8BBB45DC89051BB3ECA9
+997E3CA7B691FC553D11D583E77EFD5E6D0E768FF6D603F5B447389E04B3442D
+A26A0F2AFA70E5449E25B28A04B8F2CD8AB585021891DCE11FDEA393BA1D0A4A
+0B96AB84EBA33EBDABB2FF77C989061F639F43820568BB4F215D5EDF66A4CCCE
+269BD9BC959D0FB4D0082C8AC999A81EF2ED9C22A9FD2261465B0F29766592B1
+3B77D70A9E7B5E99B765553F4F6498D4FC0BB7BDD237F0035DCE0F7E990425E1
+C1AA2E1355DF6A5B1F8FBB29B2456C11993A894BE2A4A823B5080C9F507298D9
+E1C2FBB91A12EE56BEDE9693284A8F414E686645993A6ECBA9B44F18B1EA
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMSY10
+%!PS-AdobeFont-1.1: CMSY10 1.0
+%%CreationDate: 1991 Aug 15 07:20:57
+% 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 (CMSY10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle -14.035 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMSY10 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 0 /.notdef put
+readonly def
+/FontBBox{-29 -960 1116 775}readonly def
+/UniqueID 5000820 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052F09F9C8ADE9D907C058B87E9B6964
+7D53359E51216774A4EAA1E2B58EC3176BD1184A633B951372B4198D4E8C5EF4
+A213ACB58AA0A658908035BF2ED8531779838A960DFE2B27EA49C37156989C85
+E21B3ABF72E39A89232CD9F4237FC80C9E64E8425AA3BEF7DED60B122A52922A
+221A37D9A807DD01161779DDE7D31FF2B87F97C73D63EECDDA4C49501773468A
+27D1663E0B62F461F6E40A5D6676D1D12B51E641C1D4E8E2771864FC104F8CBF
+5B78EC1D88228725F1C453A678F58A7E1B7BD7CA700717D288EB8DA1F57C4F09
+0ABF1D42C5DDD0C384C7E22F8F8047BE1D4C1CC8E33368FB1AC82B4E96146730
+DE3302B2E6B819CB6AE455B1AF3187FFE8071AA57EF8A6616B9CB7941D44EC7A
+71A7BB3DF755178D7D2E4BB69859EFA4BBC30BD6BB1531133FD4D9438FF99F09
+4ECC068A324D75B5F696B8688EEB2F17E5ED34CCD6D047A4E3806D000C199D7C
+515DB70A8D4F6146FE068DC1E5DE8BC570370285ADA3D8C8E7D5C08077B0148A
+A1859140AD05C54B92D3B813079A40716ABF2EE475814FA63DC9DE004A484038
+86F4951E72A7F52F2E5619404520BB9F6161F630D831D342E79BDA2348C0EAA7
+2C68C0D2FD315FA30A508997DFF68D01535933389751B364BC640D4A48B137E0
+D372D16872FB66DD248E97E4B8CD2DAAF600C762105E7C43FDD396F83B79B2FE
+069830E688A49C4D2F2647305C53FEF69E5A48282514780C73E2A1ABC8D1A979
+4333CE01758BBB08392B3E8E1CAC6190146E974D3785974BA86094DAD6DE04D8
+6E098B9014CFE9845A8E851D3233957FA3AF8A98012B88D7EBA3700E3B83FB2E
+400A8B813F329147060BCC4C7FD62701F2932F55BD9B6909435FE592551A547A
+F3AA2B70AE001BC5E68B224D0A73651A38934FD187B9A77CBA8BE486077A1A6D
+6D942EA405405EF6812F6EAFDF78D5B696F6BD5911576E56521EDFCF277A5314
+CECAA28C46AE9BC2E0BEE03B8E82510B3FF288129B2CDEE1B7CB7A632CF1B873
+690D4F884D41C8A50295902F98C4F8E2A724AC4A68BD5042F58941597D13EC98
+BF3EA97B03CC86501E0E4E7EA007044DCB931D7B5182BC03357E138D011261A1
+E8E19EEC2C2FF1ECA3F96AF6902A2E37AA72F39A9457EB9D60C90F8D223B6001
+79243A530E6162A22E0FB44182FA896AA948FE4477EEF3AD3C5597BDE3029F23
+943BF7A244AC6FBADC73261B9B6F917FA14B910AE543785177BCB99B6ABA42E4
+5C523B29272243B42550E6C7601A32E7F782931E3ED32CFA8149B711F60B1680
+13E3AF3B7CF994F02223E4F1D44E45375BECE29A975A681060E1BC698E7666DE
+6958D0EFEB9D6A78B4B490BB256F30939D246165AB12B9C0859F7558ECF84705
+138404B7E16DF0975D0EAA92CA90DFF402D3EEAAFBA98E9A26D3951E66164B5C
+A49F4D82295D77BA2C3B80BC00CA7A40E7842E33A4D4CC4BEB06F4FDFA98A9AA
+2347E385591BB2A505F591015C0603CCFE987D4FD6BC7E6CEB2552EB340ECEA9
+752E7085A2AB513BF916BF5E72BF398A1BDE3A23BE303D62F77FFA5F6DDF404C
+281121D9E927C223200D25C24D9843F0BFF9E5ED39C3BF692E794C56AA1A8798
+358A66D0FC4BE45D8F71822F785AF06D6DE934C49C210B74800CBADD4044E6EA
+327D7274BB30B64CA509D797FA2CE4C4EB0D03E08C1FE5C502219637B9C38A2A
+052442336B00898420562151486A335632C3C635B241E2708A6F0FF4D3706C54
+900467A130FCDB56E42ABD747ACE5D2839AB370CBB86B4C3A8E0A3091AD5434F
+9FD9A774FCD2865988DCA75D14D067405CD7857D6BD9961E59453D410A566967
+56F6A03F7CB232267A6BAFB3037CDDFDCF265C76F17798AD4CD38404E5454CFB
+FC91DC4E6722C1EB6AA724CB756EE877CBD1F1672F313289093B60DE37323457
+0E3A0497A48C37C67FE5E906CAF1BEA2F4954ABADA4224BDB1B05521BB79A30E
+477FEDD74EE4917B90D87D92A959170DDCCB62ACF71895F35DC9B577A33D21FE
+E506ECDD45176B8CB3F63BD659732AE8A67EACC80629314EC87F203E50AB4EC9
+77931D54163463654ED87DAA331DDB2507C311F623680B03E1EA52A08125819C
+3BFE11A89ED9BCCB8504F8A87F2BF6815BF5F09BB566058FB7D5964FAE6203DB
+B32F6CB9C2066F16028D1BDCB3E32766EDAD3C0A03EDE559D3F277DA977E8654
+A04371AE9C75DC5C20C48B1D70EA4360E1C1ACB05F10018F677E9151247595FF
+A084ED281189456E8A543DAD0D4C78C34E6367BCA3914CE2C6AA2AE0C4FB313D
+EFF383EE7A44EB7700D3FFBBCCA3F3090CCBC2170700D9F1BA341B14B8CBC355
+C0174255BED3B41DE1EE036897344F8ADEA045726FD15F563BE349A2FBD7E093
+0FE4BE1D4C38DF8127F198ECE3F35E1EE02A8E24496673965DE93E8C35D48E59
+CC20AD41A8DBC6567CE9279AF980AAC6E926D2955BFED409232944BA5BAEE6DD
+74486BB5FD9A436386BC1865E0F68419496F096C08F6D419E00BFFDA2344DACB
+10062EA3EFAE1F5751BA90EE9CD4CF1FF437B1EACB2C135C24AD4F321F50084F
+3AE11C17C8AE54C0F8C06B294E8CE27B95FFB40D12CB2388514FD5F72F92260F
+A7DE516B788F9CB83F4004C184BB057EB96C904FBC036DE46A0D2A3A59F09381
+77C260A3539C94D19F3D0400913CBCE527C23FB794A652ABB23F4A076ED2FFF5
+6DDA7B03BAAC89EFF8C53F6B5AD497D51D0749191BFBD19B671BA61F1C6779EB
+837ACE4A0FF1D580756F8C5DAEC2FDC00B0B3460ADA6D16BF46AAFC1D5F793BE
+622518F6915C2602CE63C432DE3399E1315149CC3A2E9DAFB1C2E8E4CF24EC08
+C760D7E0CA4F2CA8B4CDF545C3AB946CA7B612276E64C7A0FF46EB5040020541
+C68CC1F0A74C109BD492CC516D1DE03BE5B18E02B284790ED1A45C45BA13DA11
+5E51987684D44E3FF55653F3E610F4A9272F3625073F569E03147BBC65B5155C
+902851CDB15C68FBE95B3700EE967B4CE772F20358976D97F83DFDC8D5E19A5D
+F83BFE6843AB65A401620146A5C7C0ADBD648FF5734110DDDD9FEDEB7EBA5A9C
+B7E0947CC576B09A8D1F865AC93C03F4CD8454B99FA8F46877D8C978ADE74C7A
+4D5E0643A0F0B9F96D2C6F0FF19884B0318CC3C4408E13C805BEC3411B9DA464
+C184CC2AD28041C6BDCC53C71F7749C3C65918EB491C1DAF91C9ADE140FC74FA
+0BC21A67B026C8819C60AA4E79937838FCD627FB61B30AE8C751FA029D8A24FD
+8226C98C7012E93FC438D6CD4B4B10749995848F6175CB893DEA9785E6A01131
+58ED9F51D5168925BEB51B2308E06357040B073520C8606A15392453B3E8A894
+31E3DDFB3D38FFD7F3046480F2A926896943798A7B738E348EAD0728F2EBA5B3
+5CBD4EF085C21DB265884ECB4141C282B70C4ABD2E220259880A1B5683212031
+3A9CD73FFC5218A2D6C934C3F0363DA701454EB693AD236D7D15829301807DB9
+C475FF82CC67531285CD986D1BAEEF3A7A44723FDEDF6FEE2D79635E4A73C729
+3465E2EA7E58F765DA587C4A2B686413FC21071C8845F20DBC07CE243BB0A3DC
+FBE8A4068F7FE00C2F8DEA26AA88BA8ECCF9890403BD1DA89ECE210AEC1D897B
+7529734B1B82A3981D7B8D3C9A3C7F8F336CD81815D5F8A7353A31BEFBB40A48
+A74A1D71B5FBC1A2C5C4C9F09CCD825FA08271BDCBF1CB66C7AB84C96ADDF704
+0AA2F682E9110A0410DA6C503D4AA9AA97FAE7C3719A6C32C5C293A7B012FB88
+5A6DEEECF9B5A9A34C9253EDA59336256C94DBA8829BF380085674294189DE62
+FC68FB41732199D40ADE6777DFBE1D799642F2AEBEB161A939E387234A70EFA7
+E3497254AB2EBC3D3EB0B3402C038BCD3507AD7E244DBA04D73886A1BAA0B9C4
+CB7AB4FB78F86F6238E67EA057E3D0406E4D6471DDE5D5A1F93E21376C707E97
+DC5C517F0E77ADB8E4A87540B4D944456FA29B0B4251C05A356CCF441DBB00B2
+39D382F7307EC0A4421EFC6FF31801EFF0FD843381C8990892409A5B2757C2F2
+36D2E3AC0D32C2FC4E2D3A3E9F99AA3F31AC25F873671AD7DB9EE223C922990C
+77533E454A3181CAEC199626A5A5470AE9C65D547D59BDDCA539D175274509F1
+FA6F292882D87346AB78F4BD8BF4D831D0BD7AE584B2D44C7AA5676C28E8A814
+7AB35817FD5841EC96F8B17C5CD89A32F63D953226C61DC7CF63BE0F3B583068
+B2DB5D1886E0F2ED7EAC156B8759A60AE670B2CEAB4C56D10476AD2AA1E4D44D
+F405283DE20E74DF15874058CA501C63AB2A5497611AA6BD6299E24E133954A5
+BD7AC46D639069345F36EEDB07D78CAB3C4B78B97A99B7C57743A90CD5695FDA
+61D35827CAF95C31E702C9B31A96E8F286C86A5A79125110C86B204955A795F4
+2F0A897CEB45FE1065ED44FE957778F8D87342EB065D30CD027D874E9F533839
+19F26D9C4FE56FA7F338EC513A6E5B583FFF4C7E8C6160D910238F242474A908
+8F617BC50E924D6B5EF2EA26FF481E9117B85B74784BC5242428531AB827FE37
+59F07E5C5B0BD87BCFE181BDF4F936A485AE3F6B3D97C2D0892E1A7B154D06A3
+FFE21DA58CE7FAD95E7E1FF3DAC89D4CEE83682350AFEEFC8D7D2073C98CE6EF
+00F4521287094D10A1326F0ADABB01A0214E9F17ED07259A757452BDFB36E602
+771BDE69E91D695ABD28AD56C9D01A5131F3C8229A1E3A377A701280D9574603
+4E7DA052341B1AAAFD8B88700D20D92C70
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMMI10
+%!PS-AdobeFont-1.1: CMMI10 1.100
+%%CreationDate: 1996 Jul 23 07:53:57
+% 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 (CMMI10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle -14.04 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMMI10 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 0 /.notdef put
+readonly def
+/FontBBox{-32 -250 1048 750}readonly def
+/UniqueID 5087385 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D03D5993DFC0930297866E1CD0A319B6B1FD958
+9E394A533A081C36D456A09920001A3D2199583EB9B84B4DEE08E3D12939E321
+990CD249827D9648574955F61BAAA11263A91B6C3D47A5190165B0C25ABF6D3E
+6EC187E4B05182126BB0D0323D943170B795255260F9FD25F2248D04F45DFBFB
+DEF7FF8B19BFEF637B210018AE02572B389B3F76282BEB29CC301905D388C721
+59616893E774413F48DE0B408BC66DCE3FE17CB9F84D205839D58014D6A88823
+D9320AE93AF96D97A02C4D5A2BB2B8C7925C4578003959C46E3CE1A2F0EAC4BF
+8B9B325E46435BDE60BC54D72BC8ACB5C0A34413AC87045DC7B84646A324B808
+6FD8E34217213E131C3B1510415CE45420688ED9C1D27890EC68BD7C1235FAF9
+1DAB3A369DD2FC3BE5CF9655C7B7EDA7361D7E05E5831B6B8E2EEC542A7B38EE
+03BE4BAC6079D038ACB3C7C916279764547C2D51976BABA94BA9866D79F13909
+95AA39B0F03103A07CBDF441B8C5669F729020AF284B7FF52A29C6255FCAACF1
+74109050FBA2602E72593FBCBFC26E726EE4AEF97B7632BC4F5F353B5C67FED2
+3EA752A4A57B8F7FEFF1D7341D895F0A3A0BE1D8E3391970457A967EFF84F6D8
+47750B1145B8CC5BD96EE7AA99DDC9E06939E383BDA41175233D58AD263EBF19
+AFC0E2F840512D321166547B306C592B8A01E1FA2564B9A26DAC14256414E4C8
+42616728D918C74D13C349F4186EC7B9708B86467425A6FDB3A396562F7EE4D8
+40B43621744CF8A23A6E532649B66C2A0002DD04F8F39618E4F572819DD34837
+B5A08E643FDCA1505AF6A1FA3DDFD1FA758013CAED8ACDDBBB334D664DFF5B53
+95601766776E1DD2FA640BFAE151DCDC3D0A1E883D3A6589AAE5336A39DE9E8C
+9E83EF1C643315D88AA29D838C8A02278F1D3EC7A04F19C8926221C11C6BCB09
+C6E03DF722D91D75BCBD2B978593680FDB2C3662C39BBF12AFB2C2F7A8409D38
+57AD23B8D96F0790978A427F934749BF6022C501BD1B14D96326CA236A0152AB
+337C28F69ED4A9ABBEBF6EE8D49C30756099A476FC44A4181BBDF64D06F4F5B7
+C572DE95EE0F4339F32B6DBB3D89C00A22034B78ED5D3A7400B4B9F2EE35519C
+910D1F8656C8C0DA3EEFAD0E05A302A10AB7E105D7213BB0D9D5545069E917E7
+7DB377010C1E48FD42D16783057BDC73C4871E012416E42E55D50005B8B17D1E
+C9A470487B10AF65F05B1D254DFC5161262FECAE9D90D18425448C11F8B0BE78
+BFBB1F813EBD53A56BBD79C4BC2E916E50A0BE909CF51E99C659FB69C9A23D10
+E1A14731CD1B068781BF2C689B9BC02463F87629EB1C3B4B096B4774538E5339
+0F4F9462EE4AD87B3267493BF37A78EA1DD1AF69EA7297C6973A5AF06221AC18
+4C128F7570DAF3B403D4DA70A3DF581345E3B7FC47E8E5872099107B8D39870F
+1A5D3D08CBCAB7086F91469AC38509E7B47A5B917EC31E21E134D23088FCDB03
+5805E0D732865E9F139CE28604A72E989F7E7748DDBD5F4D981F45F302A25BA3
+728A7FAB727C3776F11DA776663098741586D8E7BEB70385D6175ADDCAE71FF4
+BC27A207C5EC68D500E26DE0F51E583939A296D24CF33B7A272DE629D0748A3D
+97689A90B309BC74B2DD7D48EDEA4940E07B7E12825500A2A58E1F02B01EE112
+7A321FE2C94EC46465F1EB6EFD78811C6C59A458AE125064425B9FA8259BDC69
+DC2AB11F5FFA286780043F74932F55F01B0B19E9D951F83A1A0A96A8455BC1C0
+2F8BD634144872E58C6AE811B9E3032D0034415BF588FBC6931456575B0F18BF
+C42F701D25C1A0C680DBDCEEB48B2614E6664E56F9029E311B4548CC4449ABAB
+0000034331BBE91DF242FAD0623CA0B2741B0E2061D2FD71FC513BA51EDA75F5
+F315F4800452E44A6D9E6613A6A973E18F217EFD84F504B974D6C0F497D6CF53
+74006ADFF0C9CE5A8B5AFF439F66CC97361FB2CEE3BF2F39254A91BEEADB21A0
+FA3B51E5F74CC8CAE8CBCBD06EFFDD5ACE22A10C84DF5EF365E2953334D8B5EA
+C9B7CB3E44486D2C0A695DB5ABDAA5F3C50080848C983F382FA5C721EBC739BA
+FC522B3C75B63258F68616C2D606515EEC43D5D270D54508AD10B46FD18CA4E7
+75288FFD32F497C7D31D8BFAE0226E2E22C764A341696BAD2DFCD383A92D1C56
+8069FFD99AAA4F2C3B16D3DAD1D6600529872A399CC36759B83C9EE77C1BEF81
+AF3881EE60A0307B84A8D9D8CAE2BD3A84E7798D99CD0501542E5ACDC146E710
+3DE9D14A84519D4C907D9F28BEEFB72EB07A8B3695DC141139DF3CB6F36AC623
+BFF9B05F2E94A5D3A34385FEFC0EE54BEB381FF419F34A17B919A9057F072CD9
+A13039FC734AAC3FC2E66E8B42810ACDAA8EC4FC198940C87F98727FB228F29F
+87DD9A13951F0E9CEA1B3ADF82269FD6CA6D7015279C02DE17599EF0F728E79B
+8602FC23562217AC60D42FA750AFA89065FA47192A755C0B8C0D80097F5B5126
+B5E4A5E92FAF07E3D8B40EEE78FD4D49D39F501CF10AFE8EB911C22850A67029
+0B0548D782B4F7D2213475B86AD3CF5D13C74BE04A5435F00AD01CE19F4A3A77
+7412F1616E329E4C30DFDA6195E28EAA49781CD43450E15371FF913C09435CA9
+B4A955CB11F7022ABE8FAADB58887CC43B7127C8EE7B4F61460B487846711DE6
+BA698D6746CD92A147E76E65B53B87464CE45008C3842A36331691F7A575D361
+69016968E8AE57DC68CA99D96C47EBDD61C9B1668382208DD273017318DC2643
+6E92CD119D2D1A882C4E95D1796825C917D45073D0072D9154A725162FE260BF
+EEFD3001F1BC6D285C8C2065883691B9C8506EAB1B8AFDCB6C03D1D7C0D5217A
+EA58581D838A8AF7AF24433E9BEA4109E1A7395C09422668BCC2634479195A37
+9754FDDCE4F2BD02AE41C71BAC2B247626506AB6388CFE116EF6E262FBC8C48A
+EA0AF6A6835AD2F7DBB408CEF9E995049B29959F03B3CD144A7CBFCCF2E8837A
+76AA05F3DD4131AC52802BDF6DB35308F8E73A1E898D0857720B1E6ABC0A85CB
+290877D0D4B21BB643CB84588224B7D47BDC40CC620D16CC494F3ED8886B0D1C
+737F654F93BF3CC1C790780E93B58A438CE0B6FAECA8A9787BDB56B6F64C1902
+8D9F6ABD6DEBF75F4B6DB5EA1FF27BF800916699D5EDBE04ED1C74BAED4D05AB
+E5DF22AF1725B2C3B26F58AACBBDDA2DF05CEB83AD34C6F9E009E4D87658A483
+02C9CA8014B9548E268AA7E51962988E2E3FBC6F52324BB897FD3AEA792EBE38
+DE327832A2C49E626EB56DBC71897461FC483C0E2DDC3873A9136DBC0D3FD258
+149298C6171C13002B8926C8321528EDBC8A9557025D3219DBEB6BC6224F4048
+66DFECB9E8FC95AA9C11911857B813864A3757A93BBCBA468ECFDE5DA9588C61
+8CA8CE7BC8C13E718C1EB51F453815996DAA0B767011DDE290EADB832354910F
+1FCC9F7E01F589C1E545F061DCC99771D050156E6D2019F3C131D41F4D700E5C
+12D3893CD89703361404B9F2BB7778AE8697E7A68700A09343788D13B0002D31
+64717526C636DE27138B5C888DA602CD41920ED3D1F9806A8FF7291B82D45865
+1702094BB1CB3BBCD8A970B6BA30BB1AB49985523D51D375F98F4F7CF3AE0265
+4BA3ADA1CFD663CC6250C93947356209573A8B1419520D19CD7E8FF696E4B30E
+2B380E1ABEB812908029A384FBA5AD31CED5021E3C83661B083FDCADA5526FDF
+1A0738C1EF5453C1D49316FBF6D8B53E2DB0D98B7056A03FD789903FD3D837BC
+433A3BF4B39A233F3E0376FFDA2A4CFCA679EA0F69EB8E5C5F49651DB78D2D06
+53569991376B996ABCA2FC710B64A40FB78EACABCE93A03F37DE712F5DE7D362
+61391401EAE2DE225CF3500E81A32888F47898398CA353418B869A7FDDAAF86B
+889987DFDB0A1F480D81357CFCEDCF7BED3FD36CF3979F4D2C9CA6C42A05969A
+F8243317CA732BFEFCAD4C153B01E71D56C28E88407EAB6470E1082EDF5A3A8F
+58B804D1241FD157693C0BFD285CB95D43EC8BD7B54FACBCE3BED963E269E23E
+D97C9AB60B9F8C4A8F526B89296C6465E5A2D34283C245A2E55672DCACE56840
+E1E911EB53998D212441EF69DE3B53BC770C700CA9D22D4B93AA764AE1BD5B69
+831A9738927A58862B16820C93C3ED0493A00546470AC1FB38ABB1AAAF2CB6A6
+D8D961E61E2D18B1F4592A2219EAB2300F85D57E06004FB1FFDDB2EDFECA394C
+25340649C5104DF6288C71FC3B69571D8F97DA33C6696A5391E305E338878AF5
+BBCA2C176AFF12129A19A52184249AD23B43A3CE483CD4BC74AE417509664745
+D265E1BF3336BDF9EFE40104D8F184FC8911FB44B31EBA5FE7B17AA3C5828CBA
+9AE7B79204D461B88118914757F82459ACD2D080FDB67ACDBE92641633E46881
+508C56E082653D70C0BC6F46A0F9CB384A3B0A4105D26457F524E755568A2780
+9BEF9949E9C89867C0D948304768886A601F470A449C696273541E8CD10BAB29
+4CFC0BFD695EB8192985DB2A1F60FE49FCF95660292CF06B90C3EAA4695A8FB6
+5075A78D44F30F6A33862395367977C8B8DE22FBB29361586AB712053FB02BBA
+3D2FF6F7368BAC3079108C7EA0044A4A320FA653624399C11AE316AA6F09711E
+6B930C81568AF47BF59B73020B9F8AFFB30047FF99E16106D59A7D987B117418
+62716741466C40CA94B5EBE20E6457C92CDCAEE940EAEB3E16DE01B9D0122484
+52B1F0EADF148A0E1EC5F31FC60719F4F8D5F80C833DD82884D89BB05577FDEE
+E496773897321803991523FBAA5116CD1F85CD1FD4D3A62E477E211267BEED02
+A88A291C8FF1CCCD8F77AE028D75EA6CA242CF88F169E440F2971CF5ECE833E5
+35A4B6EE09110A6025CAF708C033FC75AED223E356F1F31969BB54C98B79C923
+3879E19A1F788933D1AB510E84777BE42E5B9708BF26F4B28EE81C145E16CC0A
+2B8C2222AA976F0F9C3B049D7243395A1BA2DD8B039516CB6DE0EDB5BE640F22
+D728BBE072EF0DAA68696FC484482C2603BF4E45A020B4DC6EE714F41D62EB20
+6361E2874A90FFFEA2A0452B0C7936C6A1E4C1510E2E7E2705C26D56704B7DC0
+DFAC619123A3DD728AA15329E5522D380C3D392C7AA86A2CF74126D747552442
+0B9C82BFFA9B99D92AF7813E04FD670441055BB3E559220D4479A1D646BBFB0E
+EE955EC56367538A807B579BF486FB62A7A39CD6EB935EC3B2CC2A6EFFE36C24
+8FADE48E949C5D9CA69E10A583498F70E6214096C999E63E82DD1A67478AF7C6
+C881CFD302A66B87745F279F5832AA134C2F98D83F51C2ACA954B1274462B041
+A96F12C01130485FF1850E810728233746E7C59A420890006FC721B605404C7C
+E9853A8F012FE4371C8F661A774AA445BDE48A4C9C6B82B631EB6832625E084C
+79866D45655020E2086CE2E2B41CBEAB6BB7E169184488135299B035289A51DA
+494648CE5B1A0DBEE5103E943FCFB82E692293FCD42436216E958F4603F8735A
+4A5D94F0BA8269615FC1BD230839F7032CCA7A0FA3F85D1F024018D078E8A377
+E3C0570ABD921E9B1FAB05CECECD988970A4AE3827A2927C55E4F9026E45F509
+59F42DD0043134686460038414955D4E200F140B2A369C5139926C30786A4F38
+7CD1E10D9377D1BBD5A3369020B24F0D5C0C7709DD805191334E161C29608580
+5478841A4C556E2E1594210253C2018C12076CA72CA6B1B4CF669D10C7E01E11
+70162DF7B41C65EC5C89E3110D960F21387E78C7D7A88970E7FF1001125E4EBB
+BA6ACAFFED6290D99138AEE365B72A7BFF5818AE5312AC94826560FC9BF6A40C
+B913A52EA0BD22EE6F46BE44D6729A8B7EE7C3CDD753C5E13717E5E0FF26AE31
+D987EB56390545D9F6BA63B1309FFF7E7D7D6C3B1A86473D76459FEDFAF7EC61
+AD953AF2AAB8C4B785C2242639BB750E354E0D759C1E57AFF0163C9242983E87
+710E5F0FBFAEA222AED4911A41DBEC84E16239350E41D311FEE062EE41CBC3CE
+10C40A29160FFCD445E79519584AFD817B9AC5492078BA0776B47CD3535C9B4B
+E2436F2BD34A0DABBFDC37FA67267240C7F69758B052A1205468CFEF55300BB9
+E94B6BCC0B1892A9AA50A11136CF5C29D7BA3F869C3372C2A16B2BECF9BBE69E
+6DB523CE1007DFE3DFF2180653C0E57480EC134EB203A4C3AF91713A3B14FD6B
+104B015CED7D2F593882581465580E41B64A7755D9D7236B5B891E42976EDC5F
+419B83EEF692C5899B30AC2106B139EAA8C1F1BE24107F558E01493300DBB491
+0129B2F497813ACE2A784750E0B36D87DB92FE488A6CF424279187375A2B5164
+B473C8D698BF2D6D3F0088988DF347A8C28003E43E5CA707E56C43C932F77CDE
+152693D57CAA1E42C3999339373E5CDD3C2F3E06221D2ED9470FE72C3BDFD5CE
+420CDF556A1123720307B43F739D9FDDEA8670FDEFD2A9818F617E932210585A
+C7FA3E383B44AB014F6503BF3ED46F4EC90A9F8792E19A4C13EDAE7818641D6F
+8CECAEF71DF7FD30822B56053659B26644AC9338A2089C6A43E8E0781A5DBF9C
+F9BB90A999E8252EAF672CAEAED23B8FCD7C023068DF734B0A7848818B9F862D
+03B84CACBE24C4BB23843C7AAD113D7DE21833E35585178049A36ADB8BE29188
+1AC8D9DF4CFE1BE92652C59679926043146D72A708897FE35321F3BC328D5FA7
+D3C3C28844C387367971E282E950DB9615F3CE2AB7304162AFB8CDD0E817013F
+06B0F7BB26825BA9812CE8DF87B7A69E4A3BB3B088CCD5FD4FDA38C5C643A63D
+74C65967C0F606E2F2470CDE27E8CF5BF5B118255E29EE63B5CF8C6AEEC5358F
+6A53A220024850AD2773903003460E45D342CDB4A644A59128F65D30552461E1
+61996406DECDD87EF22FF15166E0EF14C9599FDC744074AB5DAB39E34B202E37
+E3A7BED6B23D2494EFACBBEB1FEF4D891CA659DEAC9D866E4FFB9BC384C1B9C1
+C72AC7ECCB22CF400EF73AC793D805B7478F8B609C1E5DF959D7CF146BB85D68
+3A5CEFB22CA35724508F2156BF330A3875F1262575E667125B8C0E5A43219655
+5F80241140E6E2505458FEFC3683D0F73AA9AAEB87297F5E47A9D73633A4EA80
+2E4E92FAC826D3A0DBE7869D0FD6749B0197DA1EC44CB55C920264BFF37C972A
+1EF7584A3948D08A940C8F434536366BA5C7A49DA695ADE2269111988105AC5B
+A7712D6E795FC482D91955396859BBE140385B3BC4C0F18C883E065A5E18DE11
+5BA8DE2921CDE46A38A4AC51F215B054E9F1BD7B02C21B36C92F53F07FF60930
+634F29C080D5DF21075E089090E0C6443AE57DA192A69D1E7B2518190BC50341
+5CA603F5EB6E9A55EA6963F9E9693763D3C19798B163CAC772F44507C623CB55
+F085FCA9DDB4034DB557EC390316B5FF29E47C06B19C199873FA190B9B29BD3C
+8388474E414913D06ADEA8954934574A34FF01B4EBAF5DB950D002488F135324
+278DB86805FD6C53D9A4D0712FA372C70180D97607171CA030C5EA05B54D3C12
+1C5F2B69201FEC7EB94DC508BB6A49ADF55B8D1F524AA136484DCA3ECDB7336E
+49B2F3029973CD903135BD8120DA8F6B799F42C5297D032E453257B459E612A7
+2974A41F38C2233BF9A181B8535645BC9C8BEBF24BB74A52A5CEF23FA1AD65D0
+1FED1E1E204511204BE728CA4704230039B80033685975A73DE17EAB1CD57E64
+D1170D653E1C022AD4E377DA75E371D531CAAF49A54E6EB6D01E631EE2CADDB9
+EA11A4F6C9EA5E73FC91E140CBDCF01569AE869403656B3A8AAE79C844152271
+B7587D518B4F46DFFB57EDD7239D4A96544600D5265484E7B9DC21286699C90F
+456E0DDDB6240EB64BD6C13E0E64C596F2AF4609FCE9B8BD22179F597B8AA2B2
+5F418463FA9162161C95B15C1206317A3EDF7B201FA20FE9744DF930EE6AB4C2
+73313470BA9DA3F9CFF01B0EE67FD4582B6B74FA30E346E89CC5042F4141C624
+6A653FD1EAFBD89F782D964A2A42BCE6AD6EB8537E4F435F09E8FC28989214B6
+C9C052E218AB84C940F7B5994081304CEF0D0554AE1DBE86E75CD1FC6F7A8970
+15FFEFBC55B61008A039BD52C6EC388C19911B492A14861CFF5A82D1935568A8
+E8FC68B90083D4CC6AB2CED2D564C383988EF7D6DEBB17D3DE6C53168EC4C06F
+EF1763513ADEEEA6794C29F58DBD30C0588764A2B8548B4D950D6505E0D46323
+E69167792E61562E655ADAD672BFFF5D5C734F663E05028C27FAAFBE697D341D
+B38DA0BE7EE3CEFF6ABD729D7C00FAB94946BC478F3AE9D8212A666BB8714FC8
+A2C816A139844EC55B73DB6F2AEA581BB0DA7620B2331418AFD4ECB70B803C7A
+219AB2B7BE43CE92ED72C294726FDCB66E50FAB48879E6CCF1419977B920FB08
+22E1DCEB478786484469500E252F8EF832878C8C0D6D44F97D53DA5D6E786AD4
+F247112463CC61001955C641AB6DDD93F2524ED67A555F6DE87424A83A8419A4
+6B306FCAB5E2466C66DE1793C6BE745D85893310C706BD6DEBF75F4B3D23E6FA
+79F75ED25386934D8A588E83ADBCFE56623BD18E24172DB5E8605589F0170057
+AC5C26C1EAC6A11FDE0DA875B0FD927ED477B2B7D110046B54C10B42EA417E27
+AC9D4C53C18AAA76086EFB251F12AF5A5AA42DAD2CCFB5F1F1C1212CCA4ADB27
+2F14DC0CC44FD88009CB8BB2A4A08CC3C6437F5827AE7B39F532A9DC59D06CB6
+40CC8031FE35B357B6E64331070CA74D3DF9742EB1CA21008100ADBAC532A24E
+D8BBC05C4729BB86D530FEB9E409223990970F970BCCA606BBE8DF660117F7F8
+44A8452F7D841B791FA2C40716A8BE5EC5E1F29B68D53915C86C1CC58B835E25
+1BDBA8CA83254B0D4C5A829DDD64F260292B31297EF4844CB83EE9E9C1F030FE
+E9846B6FE6BF56C6591F6289417DD5B83D53E194AB73416C08F545E204240A31
+8F607A50FEC20C343C78B3822EE3082C42D8F370DC59CDF396ED571FD902CD44
+A3A6F93EB2BE25E63DE6E1B834E6454D1D35E0C5ACF8003C908BEE51B3A5581E
+8630B69ECEA510AB5B5E4B3D896A89FC859D92D8400FED47AFD9C4BF734E791F
+B959E641A1652208A6B20BDA8BCC119076DB68BFD5E9EB26466AD7CA58C907FD
+72444C5CDB94AB6AC385E16662BD8FBAC79A1ECD9C2508D389E1A5DFC7E5EA68
+238B4949711CC9F1E6A79D0BF72224B35A7133D323794C11ECB0226342B54CF9
+56798A5A7D9BD3552128EC10445E0D5F2C72C9A7343F41BFE25985BF31C1561D
+269ACA132561B2B46D2FB0CFA99963F96D0E30F48B5F655261BC9002C14C9672
+9A67040656C5512AFBC6D020A1FFFEE336DF49E7F2A1611A82B31894055CD278
+5D0960CD7E62DAA551E231EE1FF4E911B9F6ED7F9D6A9D6D7A3BF0AD71786EE2
+4A66A20F09736C4A4FBFC5A46CAA02ACC60EC3F76EBD6A7886DB411D75CDDFE2
+36BB49E1FE95C722E985446B2E9F046B9D9345E9FC8E17C7B6C5E12A5FC0022A
+DB99F020B3BB3B76BE86BF70158CCE5D3D21E605C693A09054CCD516C24F1839
+B98DF1F474E38C915B0BA3371EF0CBB878A3D970DF533905BE85752C5EDA214C
+30C76C296B18927B9C5116959A07F136B736596F8DC02E5AF50CBAC85DC550C4
+30D2E3A937A6975564DA240BA7F933F43AB56AAA717D701FADA7DE1A5166D678
+AF8E16CB8954B21DCD3EC0570528B54E76678AC0EFDDDAB8BCC93F3832EA39E0
+C344E47DE579830B4EA994E307CF930D522945AAA95ED40EC28E167250D78543
+1D7F1DCE3F5C3CB7024FB5E8D112F15B99CE5AE08B34745D591E933511F1664C
+FB4C2E3B191587BE93A75E6E00469F75352636595FA88078E0E96643D27C03A6
+F6260ED2A850D2AB54298587C2A8E5329CEDCCD00C31FDFB9677B59236477ACA
+8587A60B6669545564FACAED19EF06C83A6D5E8660D55722882EA38F7067B7E0
+170F3522CBB2B4CF89AFB1F747F9E92731B57524F79073E494426CC17B222BAF
+EE363422DB5E8420CD933D4E7EB262CAC604122F1630619EBF1A111082F6DF22
+2B31224009B8A93A71A41BCB5B39A0D346CCACD15FD6CA540F1B2AE8D64EBC35
+B2F08470B5C9A7BBC38AFB4B5F1B541CD53DDD69DF2786EAC3AB11ACAB3FE782
+1D6F895A66EEB676D1D4523694A4F6A1DF7BD5389BBBEF51A3DF33C30A55C1AA
+B12171648F3004E569C662FF929E2A4C99F387505D1EA0B37847A4E3FE16AD03
+41E28355A7003511A4E61987EF158DEC64AFCFBD104A74929AA51E7D5D278BED
+000EA9F922B739C6191FDEC098002E8B3F7A8F4B73176BAD3231540E7A867895
+955BB73B928781859FB846BF491B0232A0BDC07E70EC1122D5EB05ABB1D2DDA6
+BEA30FDD152F1A75F13D4B149AE4229460944C086CAE12E904EFCB69233B7D1F
+15A163DF683481FA8ED5980E807DFBCBD5D28BB8E7632DC2019A8DD258A7561E
+90FB5DC9494BF9111238112E137C7D3D28B5CA93A86D86FE3EB4A3D1EA46C743
+7FD1CE7E3606C5EB1120E794871D7E3D3A3ACD8907E8670FEDCF6A168A04716D
+33F31C23B02B5E45B6AFCCFF2249FED98849BC7F0942CA42EF0FE9D07FC30E96
+01A7838AFE3967BD95EEF974C16703977578BED3E1A6179E316C2D9B98104E24
+5A34A94118AF772FF63AC5D79C2E7797CB222352A635C3ED847463B469D4A5B8
+7856DC50D8B90AC19311F31EAB833B7A9D2468C04635E4
+0000000000000000000000000000000000000000000000000000000000000000
+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 0 /.notdef put
+readonly def
+/FontBBox{-29 -250 1075 750}readonly def
+/UniqueID 5087384 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D03D5993DFC0930297866E1CD0A319B6B1FD958
+9E394A533A081C36D6F5CA5FED4F9AC9ADE41E04F9FC52E758C9F45A92BED935
+86F9CFDB57732045913A6422AD4206418610C81D882EE493DE9523CC1BFE1505
+DD1390B19BC1947A01B93BC668BE9B2A0E69A968554239B88C00AF9FBDF09CCD
+67D3B2094C11A04762FE8CC1E91D020A28B3C122D24BEAACF82313F4604F2FEF
+6E176D730A879BE45DD0D4996EF0247AEB1CA0AB08FF374D99F06D47B36F9554
+FAD9A2D3CE451B7791C3709D8A1DDDEFBD840C1B42AB824D5A0DFF0E0F15B0B7
+22AEEB877FF489581DA6FA8DA64944555101EB16F7AB0B717E148B7B98D8DBFD
+730C52937E226545CF8DC3E07C5BA30739BAFCD0F2B44275A6D503F582C0FB4F
+449963D0AD2FAFDE33BA3D77BCA9D1DF878DDAFCA2E22CC4BACD542B282164C7
+97C2BDE318AF9D501CA21F6E662E7AAB75A5F24D2C182E598D175D44E88AB19A
+E7CD59584F95B389183EE21B525BF52A3F23C0FE5383A5565A19361D716F508C
+AAB78411CA5A4D27552CC1C435760D5A89D535B71C593E755C616661363308DA
+A683F54ED0C23FB2C225A008392B0B719F66F11A946A090B7C00B662A3C69599
+B4ECB0CC70C85C4BBBF207E0026F6C7A19F2ACFB7A60804FC98A4BFFD7BFFF2B
+9529E6D9D4238002BBC255BC62959D6F3381FE06E0621B879D5FE5B541D45A1E
+759A6E7DC32B1D1632368D09A97039DF255B6492B1B2B7E2C1434E8306ECA7D3
+5A79B6D614B4979F10988BC76ED53A5F45315CD7DA216221F842FD0F3E050DD2
+BAC23C984D506D8F7D614BCB6B244F5F41321549BB0BD041FBF3053307168680
+3435E9C9444EB0DC6AE4AF2557BF6F0A0B0F39DC626DAC1551BB0026CFC147A3
+8336603036DD711E177A006452B62CA294F72E32662216E8BC64AED865221333
+F960D01774FE3147783D6CFA0C983BA5CE843410BE4B597C9B8216D104106963
+9AA786A8EBB3A607BDA1CEC5E13B4F4859A8EE9F34122EEB6AC8C553ADB892BB
+9497115A0D995EB43F9F73366D9D85C3D4C7644E613DD554670E4BDEA4CA869B
+2DCE5F4FAB0A3623F9D0AA89B41218538C7F37114C0EEB790672D177845041AD
+92D6D240BD395E97B626F190DCFBA949CA9CE0C40EE7FE4454D05B227906AB0A
+E27663E28190347EF558B9D33812D3F6DF9630047D7FF58191B65E2AEF42D393
+B5132DF2C57498D1C686FB041DC8FD0099C5C02AB95F37EB1711410043D5D3EC
+88027AA113A73EDA88F38A3D827E671061428E0B38480534A60008635B424077
+598917FEF70FC19F8490674DB895BFFEB4F48D610E5ADBE0D4CC28D24D3FA45A
+133311CB21A0AC31AA849053E09AAFDD8B21D5FF41479E5663202A6153F0622B
+5C987A8FB9EEB58E53447D97AC450B05276A95EA095F7931BFC64F595814E0A7
+DC3DC78112ADAD4E6DAF41F0D9567263BE9C91398CE9D3DD7D3129E8E7357B59
+525556F0639C43B00FBAEA9EC3500485D7A4CF72F8EE8603B9710E4CABC5AC00
+FC52C850350083C58DFE0B722E43E7A2CF008483E591EFAB115F2BDDCFD6C8CC
+4F479E7D753A33EDD2FC4176E029A5D787A159CB6C8A107CD117636D119B52CD
+E6FB1748917FB9A9941426DBC9C38B3DD6D9BB816E2830D8B0AC393C50D26EDE
+F7AF56F5814A31D0DCE05C275F1468BA607A5958FA872C0B460E08B69B74A55F
+431ACE0B45B102472384F38FCB497E3506DF92679B67BC074CE8C3BE7287EB5D
+5FD6110263ABDA1292FA3FD28185CBB75F4BF8216134E4CE772BE723BBFD401B
+4F1AEF382EB8FA72B67C51899BAE9676D4D635A21A8CED51BAF78302F9D48A46
+C335A98F94386A2BDBBF17F7D9F00709BDD409F6336FBB68687F351511032869
+C5E8C6FF1E37E7BD12A7A28D4F263C53BD44DAB3993E65762DFDEE4BE8C5D904
+DDBBC529698CED42D02DBA0A17134FE9AF270DA9689CA32E222CE8F9CF8CE001
+B4D8A70E3D4CC7DFC0D96D1EC4630D7B8591FAF9C08E989373AD4E7E1610C43A
+8BD1F9E05CD20F89B4A08DC7F99C91442B62BDE5C3260ACFF2DDC5E58DB9B0C3
+4F0D484D0A6335B6B461DF556C072FD95C1D27A6D5F49FF555784366B5725584
+AB2979F38F66A548E35A351B0BDAD6AF88493F0D38809568A0A167A545F85F11
+B460625A45F8E8ADE47C4193F398EFD15BB37479F40F0627F81F45F760E3CA8E
+EB99F8ECA08934112E82CA723107E77585100725D03C2C286261CA5639AFA02C
+C686C56DE6C2B7A1E641055174A994677658DB9113D3573712C3723DDA33FE5E
+1D67863C7933FDFB3C344E2382A319BEE4F8961B973A8DBEDBA1031E84489913
+48C440F93FC476E36CA8660E2DFFC20F7F0524A50C11EB1A1CA876C75D0E51F7
+4F2A450AEA07A51FB2CF298B6F47B7A29E22840C566DEBD402DAAA07DE7DE2A1
+ED268702E39BDB19590E33608A8AF0EDACF5723D0BF7D78B8A97BA6108A11283
+2D46D261441AF24B447F9831EAE4BACCF758EABCAE1A3ADA924F3C795552DFCA
+CD3A4FA73A1AAD24B78C5F6631DB227ADA4FDA79825CB7D549FA6F2B95E85D4A
+8FFD9E22F8EBAE086A90BF0562696509DA6846F798622E703CAB3ED73A9592E1
+AF1EC128C4121160FFBACE86279921376FD1B866A154F3B5EC6A4453A4F74047
+1308639C24815F45B415308E598E604518038ACB94BA5C1F817F47A9E3CC0785
+6869988FCBF623EC00FC38E11826489D2A614C1DDC52BA414BE078EC774AFFBF
+F91F4686F967D57762780171DEFB6B308282E3DD31CAD53669FC7946BFEC1D90
+CA41BF444BD6DB8ECE6BCD690A7B4FE21436D64CBFE42C82977CED4149755D9F
+A50BBBBF9D3BAE96D34D397547827EA9D54127F9EBEFECEBAC435D58DD4A0CDF
+58D1B95FDB126E5C1145983C116095FE978334F6F0D0EA1B94E739996067A481
+8677409938362FC2DF109386E543DA5202DC2E1541A1519E3ADBF67157043250
+CD7B71A26AE5B7A4DB4B26EF2A491B1E99348A0AE9A11ED95E900C4B41EB022F
+4ED7AFD0797E8E75F37EE3DD34DD7DC93F617EC927C9875333E425D8A835369E
+6DBFA9705B0E39F6343841716ACEA4801FC8257129D5AD7A241AAFAF01AC0DC6
+1E789EE2401B6343E81051C91D6AD44F6130542FE6FD801C5DFDCB0740A62196
+192D5093480F3BC8CDE5B6536165C60299CF4B078A3635C4A9D5AE2615D29B18
+E9DDDADDA793B6100791C02977B61D265AAC15F5D372A7405B43E1CD3C0BF9EF
+EBED0E14B4C33A7DDA76E244509B31BC05420AA0FC495E6647440C7BA8567116
+6B08BD452AFA111141368BA97760C39E8E308B5D188CD46F150C0A752807895D
+6633A0512E345D48A9040165C93E8818D844BF4F71930F6EF89A3B8554A295A9
+C7ADC3F479EE3A36DCCBFF8B8443CF1833843A96650AEBB3EEB6CB739365D8BE
+7813D55070CB18B0B768E28C6A7A5BE6E8C74C1046D3CB5A641C0EB670A5B1B9
+70D9FC839A30DEA5461B0466385A5E53AD14911481BC7384E9FF996432A595CA
+DE8F0C56F0F362B561237DC71E2DC25914C3C14244ABC26A171308FDF2F684B0
+EB0AFBA346F38E0DF12B5182B7A0D6AC649CA770E7CA48EC9293D23FDCA49FEB
+1661940497E973040D2265D01B36633787E71CA915E8CFDAACBB6A2D10B0B5E1
+8AD09F6849B6FFE0AD9D6E2312F433D1ADD1E063103C3F4CEB67EDC12F3F49F5
+4B12BCBAF3068D0896FAA6A9DD4EA016D9F9B4D7A6D86306ADC6894A4E966BFD
+B1FAC27B556FE59C36384E843261E84655961AD3D9F4DF039A2B2B9B52AA30C7
+9ABD41B483EC6F3BC23E9B2A8073920AFE31CEA57C2E45BC79E9D8ECAAAEFBBD
+8D4E2D2D0E9F119A7D3D74C15843747530666137F423B14EC21A6912ABBD27E9
+67C1BCB4FE16BBBA63D063826B2A20FD7BD1CD0D2A1115EC39D8B4B6F421A00B
+AC9A932EC0D6F76C62E232C22287F5567E77D848D7AE50A44DC85BA6D47C150C
+E9C2C17DBF20D5BA5B7423E58F2B19D15B5E5799C82C74FFF66D9130264D2AD2
+3AE58505ACB3CD50AEA332A144D4C6E3A8F70BD98A7840B49BB0D0E52CF21E48
+72C73F5F1F62EA33B6DBCD654643F5C34682968F0CCC07D978AB53628AD753C1
+188F1BDC145094097B9AFF381A448C20782C4A2663EAA85CB072EB5F3ACBAB65
+CEBCF025255DF726F2E17959D9834FEF9CED1470A39B28B1935505DBBACDD608
+AEC7CD0A1F62132C973F3F568A84F8B9AFE6DBEF88C3B49DDE719337625B1FC8
+E0D10B4D54A56E881A2CDB578271688D6BECA3C455DAFDF2E3EE21B88BA6BA31
+62A33B2C656F3786E8DBD61D408CF8586FC257DC17541ADB159131E5ABDA1E5B
+BC3992EAC4DB60B134750D781AE1C32A1A8C77A29C93A9F618A837709BEF664D
+789FEADD3F4CE62810EB29E392D0BC22FB2899B8B430FC3FF0276CEAEFB040FB
+B778A57F418633B1A75752D1A97DC5D5ADCA82E86EB9CEE2C194B617F5696076
+BA035986DCC36DC26D54B911E9CA8A5355EF3AE33BCF773F66D09D7C073DD99E
+DAA1089DF166552174536F5F71995C0359161535CC844AF4F8A3649D5E0D7CA7
+86A2CB000B5E078B9DD237B2897BF9309D831DA08CD29F73E0C166D1303ABC0C
+76D7ADA8D58F54DB7A2DE34C568D28D4A92CF6943BCCD569D0AE512FF97B30DA
+26DE7311003ED974D06892FA382AB7DB9F6B571F33699143C3BB1A5B69F58AF6
+BE6E0009C3389DFE7F59A4F53D0E289382F673A2E13D2206B940E35499723213
+8AB7DCB5966862DBC9C7F06278ED081213160C3D4FDC218ECB69FC15AA86ACB9
+B0DA9B88AF33804A0CDEF335A1924EDC4744B7C674A9FA1FF85713C9EB0ACD00
+C28F98B3AA48110CD7FB981BBCA7884387D06A9CA17D72DB37C0AB47A2F4A1A5
+5C16F73E8D76FF8E1205A550519CB8168CAAE4A775CF5F23993E2D65F5CB9A41
+7E6DA4CA4E023C12A7C894F76FFF6991D474FAFB8DB38986F961CFA79C1A3EC7
+0D956377F2BE924D00F08EAB8EBC22A4141E40A8A8C8515DDBF85D1A0CE14454
+326F31D98DB56A3C32AE46D3F096A9BB84490E33C86A24BA602D925A9880ECAA
+11B53E969FA26F9EDDA0A2C8C4767939762FB9A50685F66556AE931A15DFAA06
+E119C240504E4F547F75575EF317C198FA63BBC7B2477ED56F412CD249C835D5
+593B0F6E8E0E326B27E51BD8DB4583D477FDFBEEDE07D34601EE699A7257DBF5
+C865F000BF0C94A9114CA0A66116D18F6F699D3E0A16CF4545DBF44E684FB4C6
+3FDFC5F00A43018D8AF0583712EA46A3297E0F1F113F2A0E533D16DAFEACC8D6
+7B92B3516C22F2A9B0F2C3D22C2A967C87D71C8CC29C71DE34D4D456533FA8AA
+DDC7C333DEB24B101011ED0FF486C66607FF5AECE81074D73AB20C5067050F50
+7E6CD544BCF1F53F73AD94D1A7F182C15FA0DF49B8BFBE9BB973B0AB3FD8A0D6
+E7421B2D1EFA961D5CA527591C657E9E235601775800015B169D1E9E90D4F38D
+35CAE80BB4CB7C4955F3EB213A7BE560A5816647E6470DF38ED3452190A04930
+7CA57D172FE23FC693EF662B45AB4E039611A6AC584D62DC8853FCBC203B7E02
+3CE36918289D877749CB80E2C104F6F6B3DA4B38D1C6C9D78B621FD3FDC6109B
+0742FE0CF2A7D8B4D7B656D60DFF5562789409A95B53986B4264A2C23428007A
+9302B133D219B4988422C05262A4155B2965CE87940513453445C0AADA68A32D
+44E60C6E3170D565DF5C61A7636271583FAA64F043DB193030E6D70F637AC82D
+D9962D663E6025BB9E235AAC9E06AE1E2C9B9BE13524CD710956962B162A6CD0
+A3D6E59F5A217A980F3E4E15C3F920DE39E41925D48D3068760D21B7427CB747
+D354BDA1427CC6CBB71CB28AE2BFAE20078AB262D7C932897A21698E92307947
+7FCEE97E2110A14E74D8431962D12DF952206118E40B3A2D3F025689FA69DF0E
+1DACDD55964E945682A53D2695CEE3AFA99895FC3B36A5BE1A0319D08FB41504
+81B5917AAFB5BB6F0A167522A73315EB4BF8011CF0CE85F3DFAD25543C7C4947
+965F5E7278DE0756B0116A4520A7C071DB8278CA273E7EA3B1A973ADB9952112
+9CD5E60A7712399935D140E8C90025C75976E9EBF1A1466E77D4B620B1845063
+9EDBC20CD5D4B9E159025682F51144E9D7CBC7F50CBD6ED1253505F47AEFFD4D
+D657F0E3D45B2634606685EAD6937837D744C07F0EE0080F1E6DAAE28E76865E
+1965C492DED77213B70784F3E159E3ECC9E97D312DD7F84988563F247E0D77E8
+DECAD01DDE5D7343280C2ECE131765A111AB5CDEB7CAC479D19A3B1EC3D9BF9E
+47191D050640F4D9474AD1F1F4E0CF439C51FC1100520C3FD9E461EF45CD946D
+9EC6FD9A6C3BCE7AA0AE89B9747C425C471843F33F192C2F1561E041CD3FF04E
+7879796B04F32DB3972726369D552AABDA6C2312C26838682076B4A2E9FCC96B
+9CE97A2B3F45694088E1119B74940E18648DF535A5EB04F96B4D3C4A7B1D659E
+EB6370856F545E4809E68FE55AAF62E3D93E5555741FD4AA883AE6019CAD9FFD
+E2943D1D0811DE80C952B6A9D9C72E2A4C4DAB7FA87C81E5CEA5AB8F66FB93AA
+5895A3ABDD0579DC3712312093DE09AF368A5E569590F6300C9648E0910C02E6
+6687BC46577F80FFE810CAA7320379D0ED60BB6BCD34310911D1A145ECE67216
+45A0A197A61D2CEA0B422AE6BCE8D109FC4EBC7BD80AF761AA2399F66F3991FD
+631317FB9591B1E7C553019CA671B914265C517A5088CBBFAE9EE3D880DCE012
+6156DEB52A45368A37B29D2A5AFFA38E2B12EAA028A92C0B0632334D20E5E564
+7FB61F661529A268B018E7D91DE5F780C64066FAF5C7FF5C7C7B8C76B87A17A6
+32890CA0FBE282CA2E7D5389D130EACC264C097077706DEFB71F55AE154912C8
+28A4E22AE6DFD590068B9B5B11842303C337088DB0353B2A75C78D575B4B4BAD
+F6FCCE33C08B83E0893DBBB28EBCE02843E2DF1B6551E3A9D6E2DF54243B2D42
+1B116AED9B92B785D8F244CEC10A70F237CDB4826C8D38B6A45A786D163E98F5
+BF093D7B1BCB6A51C190C0B38C2AF368399427F12C3AE13C3A3DA1D0520A15A9
+1D5C72B00587B9A2690104D81E84DD12F1C872D0FA6982F221EE397AC07C37FC
+2D52CC4484DDB59E3B7447ED1DDD2ACC9D
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+TeXDict begin 39158280 55380996 1000 600 600 (nomu-csl.dvi)
+@start /Fa 134[50 50 3[50 1[50 1[50 2[50 50 2[50 3[50
+3[50 19[50 5[50 4[50 50 18[50 4[50 50 4[50 35[{
+ TeXBase1Encoding ReEncodeFont }18 83.022 /Courier rf
+/Fb 134[40 40 40 40 1[40 1[40 1[40 40 1[40 40 2[40 3[40
+3[40 8[40 3[40 5[40 40 5[40 4[40 40 6[40 10[40 40 40
+40 2[40 40 4[40 35[{ TeXBase1Encoding ReEncodeFont }28
+66.4176 /Courier rf /Fc 192[55 63[{}1 66.4176 /MSBM7
+rf /Fd 254[23 1[{}1 66.4176 /CMBSY7 rf /Fe 137[35 3[35
+1[35 1[35 7[35 4[35 97[{ TeX09fbbfacEncoding ReEncodeFont }6
+66.4176 /CMTT8 rf /Ff 157[30 37 8[58 7[45 20[20 20 30[40
+27[{ TeXaae443f0Encoding ReEncodeFont }7 66.4176 /CMMI8
+rf /Fg 170[56 85[{}1 74.7198 /EUSM10 rf /Fh 170[59 85[{}1
+83.022 /EUSM10 rf /Fi 192[60 63[{}1 74.7198 /MSBM10 rf
+/Fj 145[37 51 109[{ TeXaae443f0Encoding ReEncodeFont }2
+41.511 /CMMI5 rf /Fk 206[28 28 48[{ TeX0ef0afcaEncoding ReEncodeFont }2
+41.511 /CMR5 rf /Fl 192[48 63[{}1 49.8132 /MSBM7 rf /Fm
+221[29 6[36 27[{ TeXaae443f0Encoding ReEncodeFont }2
+49.8132 /CMMI6 rf /Fn 192[65 63[{}1 83.022 /MSBM10 rf
+/Fo 228[57 27[{ TeXaae443f0Encoding ReEncodeFont }1 83.022
+/CMMIB10 rf /Fp 214[37 37 40[{ TeXf7b6d320Encoding ReEncodeFont }2
+83.022 /CMBX10 rf /Fq 144[49 1[78 8[49 100[{
+ TeX74afc74cEncoding ReEncodeFont }3 83.022 /CMBXTI10
+rf /Fr 162[20 1[20 14[44 2[25 2[46 4[53 1[33 1[55 2[20
+16[27 27 4[59 35[{ TeXf7b6d320Encoding ReEncodeFont }12
+66.4176 /CMR8 rf /Fs 141[59 8[27 27 35 35 60[71 15[55
+23[20 1[{ TeXbbad153fEncoding ReEncodeFont }8 66.4176
+/CMSY8 rf /Ft 134[33 1[48 33 33 18 26 22 1[33 33 33 52
+18 2[18 33 1[22 29 1[29 33 29 46[33 33 3[22 3[22 22 40[{
+ TeXBase1Encoding ReEncodeFont }24 66.4176 /Times-Roman
+rf /Fu 204[25 25 25 49[{ TeXBase1Encoding ReEncodeFont }3
+49.8132 /Times-Roman rf /Fv 153[16 26 29 100[{
+ TeXBase1Encoding ReEncodeFont }3 58.1154 /Times-Italic
+rf /Fw 141[51 65[18 23[48 23[48{ TeXbbad153fEncoding ReEncodeFont }4
+49.8132 /CMSY6 rf /Fx 193[68 1[68 1[68 1[68 28[58 27[{}5
+74.7198 /TeX-cmex9 rf /Fy 153[19 26 29 48[29 29 29 49[{
+ TeXBase1Encoding ReEncodeFont }6 58.1154 /Times-Roman
+rf /Fz 254[25 1[{}1 74.7198 /CMBSY10 rf /FA 135[39 1[38
+2[34 2[42 42 1[68 3[25 3[38 42 29[54 4[62 65[{
+ TeX74afc74cEncoding ReEncodeFont }11 83.022 /CMTI10
+rf /FB 136[39 39 39 39 1[39 1[39 39 39 39 39 39 1[39
+39 1[39 2[39 1[39 11[39 34[39 1[39 48[{
+ TeX09fbbfacEncoding ReEncodeFont }19 74.7198 /CMTT9
+rf /FC 135[35 1[35 2[31 2[39 6[24 3[35 39 100[{
+ TeX74afc74cEncoding ReEncodeFont }7 74.7198 /CMTI9 rf
+/FD 141[64 7[21 30 30 38 38 5[47 2[51 1[51 36[0 3[51
+8[77 7[77 5[60 1[60 10[38 11[60 21 1[{
+ TeXbbad153fEncoding ReEncodeFont }18 74.7198 /CMSY9
+rf /FE 134[37 2[37 46 23 32 32 42 42 42 46 65 23 2[23
+46 42 28 37 42 37 42 42 94[46 2[{ TeXBase1Encoding ReEncodeFont }22
+83.022 /Times-BoldItalic rf /FF 162[21 1[21 26[60 36
+1[60 2[21 8[38 7[30 30 4[64 35[{ TeXf7b6d320Encoding ReEncodeFont }10
+74.7198 /CMR9 rf /FG 200[30 30 30 30 30 30 30 30 48[{
+ TeXf7b6d320Encoding ReEncodeFont }8 49.8132 /CMR6 rf
+/FH 87[28 46[42 2[42 46 28 32 37 46 46 42 46 69 23 2[23
+46 42 28 37 46 37 46 42 11[60 55 46 2[51 1[60 1[55 7[60
+2[60 6[28 3[42 42 42 42 42 42 2[21 28 3[28 28 22[23 14[46
+2[{ TeXBase1Encoding ReEncodeFont }43 83.022 /Times-Bold
+rf /FI 145[41 59 3[23 2[32 31 35 30 29 36 38[20 30[38
+5[39 3[34 4[38 43 11[{ TeXaae443f0Encoding ReEncodeFont }15
+58.1154 /CMMI7 rf /FJ 200[33 33 33 33 33 33 33 33 48[{
+ TeXf7b6d320Encoding ReEncodeFont }8 58.1154 /CMR7 rf
+/FK 254[27 1[{ TeX10037936Encoding ReEncodeFont }1 83.022
+/CMBSY10 rf /FL 141[55 58[0 6[19 14[66 17[34 14[52{
+ TeXbbad153fEncoding ReEncodeFont }6 58.1154 /CMSY7 rf
+/FM 162[23 1[23 26[65 39 1[65 2[23 3[42 42 42 42 42 42
+7[32 32 4[69 35[{ TeXf7b6d320Encoding ReEncodeFont }15
+83.022 /CMR10 rf /FN 135[44 44 44 44 44 44 44 1[44 44
+44 44 44 2[44 44 44 44 44 44 2[44 46[44 44 49[{
+ TeX09fbbfacEncoding ReEncodeFont }21 83.022 /CMTT10
+rf /FO 141[69 7[23 32 32 42 42 5[51 2[55 1[55 35[0 0
+3[55 6[51 1[83 6[42 83 7[65 9[42 42 11[65 23 65{
+ TeXbbad153fEncoding ReEncodeFont }22 83.022 /CMSY10
+rf /FP 134[37 37 55 37 42 23 32 32 42 42 42 42 60 23
+37 23 23 42 42 23 37 42 37 42 42 28 2[32 1[32 3[69 2[46
+42 51 1[51 2[69 3[28 60 60 51 51 60 1[51 51 5[28 28 4[42
+2[42 42 2[21 28 21 2[28 28 28 36[42 2[{ TeXBase1Encoding ReEncodeFont }
+54 83.022 /Times-Italic rf /FQ 134[41 47 3[30 5[50 4[29
+2[41 39 43 36 36 44 6[57 48 69 1[48 5[53 2[81 57 5[53
+2[59 63 62 2[65 42 1[23 23 23[39 5[36 47 1[47 1[41 1[48
+3[41 2[37 1[47 53 11[{ TeXaae443f0Encoding ReEncodeFont }36
+83.022 /CMMI10 rf /FR 136[72 1[55 33 39 44 55 1[50 55
+83 28 55 1[28 1[50 33 44 55 44 1[50 8[72 2[72 2[72 1[61
+1[72 4[39 3[66 72 72 1[72 11[50 50 50 50 50 46[55 2[{
+ TeXBase1Encoding ReEncodeFont }34 99.6264 /Times-Bold
+rf /FS 135[44 3[28 13[37 3[33 41 8[63 7[49 2[74 10[58
+58 3[38 1[21 21 11[38 17[34 44 1[44 1[38 1[45 6[34 1[43
+49 10[47{ TeXaae443f0Encoding ReEncodeFont }23 74.7198
+/CMMI9 rf /FT 134[33 33 50 33 37 21 29 29 37 37 37 37
+54 21 33 1[21 37 37 21 33 37 33 37 37 9[62 46 54 42 37
+46 1[46 54 50 62 42 1[33 25 54 54 46 1[54 50 1[46 6[25
+3[37 1[37 1[37 37 37 1[19 25 19 2[25 25 37[37 2[{
+ TeXBase1Encoding ReEncodeFont }55 74.7198 /Times-Italic
+rf /FU 133[33 2[54 1[42 25 29 33 2[37 42 62 21 42 1[21
+1[37 1[33 42 33 42 37 15[58 10[46 4[54 6[25 5[37 37 37
+37 2[19 46[{ TeXBase1Encoding ReEncodeFont }27 74.7198
+/Times-Bold rf /FV 105[37 28[37 37 54 37 37 21 29 25
+37 37 37 37 58 21 37 21 21 37 37 25 33 37 33 37 33 25
+2[25 1[25 1[54 1[71 54 54 46 42 50 1[42 54 54 66 46 54
+29 25 54 54 42 46 54 50 50 54 1[33 3[21 21 37 37 37 37
+37 37 37 37 37 37 21 19 25 19 2[25 25 25 36[42 2[{
+ TeXBase1Encoding ReEncodeFont }73 74.7198 /Times-Roman
+rf /FW 71[28 32[83 28[37 42 42 60 42 42 23 32 28 42 42
+42 42 65 23 42 23 23 42 42 28 37 42 37 42 37 28 2[28
+1[28 3[78 1[60 51 46 55 60 46 60 60 74 51 2[28 60 60
+46 51 60 55 55 60 1[37 3[23 23 42 42 42 42 42 42 42 42
+42 42 23 21 28 21 2[28 28 28 5[28 29[46 46 2[{
+ TeXBase1Encoding ReEncodeFont }74 83.022 /Times-Roman
+rf /FX 139[40 4[60 66 100 33 2[33 5[53 1[60 11[86 6[86
+75[66 2[{ TeXBase1Encoding ReEncodeFont }11 119.552 /Times-Bold
+rf end
+%%EndProlog
+%%BeginSetup
+%%Feature: *Resolution 600dpi
+TeXDict begin
+%%PaperSize: A4
+ end
+%%EndSetup
+%%Page: 15 1
+TeXDict begin 15 14 bop 523 448 a FR(A)100 b(APPENDIX)523
+641 y FW(In)21 b(this)h(appendix)d(we)j(wish)g(to)f(guard)f(the)h
+(reader)g(against)f(misconceptions)g(about)g(the)i(relationship)523
+740 y(between)c(nominal)g(uni\002cation)g(and)h(higher)n(-order)d
+(pattern)i(uni\002cation,)f(such)i(as)h(the)f(follo)n(wing)f(one)523
+840 y(that)i(w)o(as)h(e)o(xpressed)e(by)h(an)g(anon)o(ymous)d(referee:)
+664 1012 y FP(.)12 b(.)g(.)g(the)29 b([nominal)f(uni\002cation])f
+(algorithm)h(is)h(a)g(minor)g(variant)f(of)g(a)h(well-known)f(algo-)664
+1112 y(rithm)f(for)f(uni\002cation)e(of)i(higher)n(-or)m(der)e
+(patterns)i(in)g(the)g(sense)g(of)g(Miller)-9 b(.)27
+b(Ther)m(e)o(,)f(fr)m(esh-)664 1212 y(ness)20 b(constr)o(aints)e(ar)m
+(e)i(encoded)d(dir)m(ectly)i(in)h(the)f(term)h(langua)o(g)o(e)d
+(instead)i(of)g(being)f(carried)664 1311 y(ar)l(ound)d(separ)o(ately)-5
+b(.)15 b(The)h(algorithm)f(her)m(e)g(can)h(most)g(easily)g(be)g(seen)f
+(to)h(be)g(a)g(minor)g(variant)664 1411 y(of)21 b(prior)f(ones)g(by)g
+(comparison)f(with)i([Dowek)f(et)h(al.])523 1585 y FW(Miller')-5
+b(s)24 b(higher)n(-order)d(pattern)i(uni\002cation)g(algorithms)g(solv)
+o(es)g(equations)g(between)g FQ(\025)p FW(-terms)h(that)523
+1685 y(satisfy)16 b(the)f(pattern)f(restriction)h(\(see)g([20]\).)e
+(This)j(algorithm)e(solv)o(es)h(such)g(equations)f(modulo)f
+FM(=)3384 1697 y FI(\013\014)3465 1705 y Fk(0)3497 1697
+y FI(\021)3537 1685 y FW(-)523 1784 y(equiv)n(alence)k(\(where)h
+FQ(\014)1234 1796 y FJ(0)1290 1784 y FW(is)i(a)f(restricted)g(form)e
+(of)i FQ(\014)t FW(\227namely)f FM(\()p FQ(\025x:M)9
+b FM(\))p FQ(y)27 b FM(=)22 b FQ(M)9 b FM([)p FQ(y)s(=x)p
+FM(])19 b FW(with)g FQ(y)j FW(being)523 1884 y(a)d(v)n(ariable)e(in)i
+(the)g FQ(\025)p FW(-calculus\).)e(If)h(an)h(equation)e(is)i(solv)n
+(able)f(with)h(respect)f(to)h(this)g(equi)n(v)n(alence,)d(then)523
+1984 y(Miller')-5 b(s)16 b(algorithm)e(produces)g(a)h
+FP(captur)m(e-avoiding)d FW(substitution.)j(In)g(contrast,)f(nominal)g
+(uni\002cation)523 2083 y(solv)o(es)21 b(equations)e(between)h(terms)h
+(that)g(are,)f(roughly)f(speaking,)g(\002rst-order)h(terms)g(plus)h
+(`nominal')523 2183 y(abstractions,)26 b(and)h(looks)g(for)g(solutions)
+g(of)g(such)h(equations)e(with)i(respect)f(to)h(the)f
+FO(\031)p FW(-equi)n(v)n(alence)523 2283 y(de\002ned)g(in)i(Figure)f
+(2.)g(Moreo)o(v)o(er)m(,)e(a)j(solution)e(produced)f(by)i(nominal)f
+(uni\002cation)h(is)h(a)g FP(possibly-)523 2382 y(capturing)21
+b FW(substitution)h FP(and)i FW(a)g(set)f(of)g(freshness)f
+(constraints.)g(Belo)n(w)h(we)g(shed)g(more)f(light)h(on)f(the)523
+2482 y(dif)n(ferences)c(between)i(these)g(notions.)581
+2654 y FO(\017)41 b FW(The)19 b(dif)n(ference)f(between)g(capture-a)n
+(v)n(oiding)e(substitutions)j(and)g(possibly-capturing)d(substitu-)664
+2754 y(tions)k(can)f(be)g(seen)h(in)g(the)f(follo)n(wing)f(e)o(xample:)
+g(consider)h(the)g FQ(\025)p FW(-term)g FQ(\025a:X)27
+b FW(and)19 b(the)h(substitu-)664 2853 y(tion)g FM([)p
+FQ(X)30 b FM(:=)22 b FQ(a)p FM(])p FW(;)f(applying)d(this)j
+(substitution)e(in)i(a)f(possibly-capturing)d(manner)i(yields)2014
+3028 y FQ(\025a:a)1235 b FW(\(15\))664 3202 y(whereas)20
+b(applying)e(it)j(in)f(a)h(capture-a)n(v)n(oiding)c(manner)i(yields)
+2014 3376 y FQ(\025a)2106 3342 y FL(0)2130 3376 y FQ(:a)1211
+b FW(\(16\))664 3550 y(in)18 b(which)f(the)g(binder)g
+FQ(a)h FW(has)f(been)g(renamed)f(to)i FQ(a)2114 3520
+y FL(0)2137 3550 y FW(.)g(This)g(seemingly)e(minor)h(renaming)e(has)j
+(ho)n(w-)664 3650 y(e)n(v)o(er)28 b(drastic)g(consequences:)e(the)i
+(terms)h(in)f(\(15\))f(and)h(\(16\))f(are)h(not)g FQ(\013)p
+FW(-equi)n(v)n(alent!)e(W)-7 b(e)30 b(ha)n(v)o(e)664
+3749 y(illustrated)g(in)g(the)h(Introduction)c(that)j
+(possibly-capturing)d(substitutions)i(occur)g(often)h(when)664
+3849 y(dealing)24 b(with)h(schematic)g(inference)e(rules,)i(which)f
+(are)h(uni)n(v)o(ersal)e(in)j(the)e(theory)g(of)h(program-)664
+3949 y(ming)20 b(languages.)581 4047 y FO(\017)41 b FW(Despite)25
+b(ha)n(ving)f(de\002ned)g(our)g FO(\031)p FW(-equi)n(v)n(alence)e(to)j
+(capture)f(the)h(traditional)f(notion)f(of)i FQ(\013)p
+FW(-equi-)664 4146 y(v)n(alence)18 b(\(see)i(Theorem)d(1)j(and)e
+(Remark)h(1\),)f FO(\031)i FW(and)e FM(=)2290 4158 y
+FI(\013)2357 4146 y FW(do)h FP(not)h FW(coincide.)e(The)o(y)g(coincide)
+g(only)664 4246 y(on)i(ground)e(terms.)i(T)-7 b(o)20
+b(gi)n(v)o(e)g(again)f(the)h(e)o(xample)f(from)g(the)h(paper)m(,)f(the)
+h(equation)1840 4420 y FQ(\025a:X)30 b FM(=)2119 4432
+y FI(\013)2189 4420 y FQ(\025b:X)664 4594 y FW(holds,)20
+b(while)1723 4694 y Fn(?)j FO(`)g FN(fn)d FQ(a:X)29 b
+FO(\031)23 b FN(fn)d FQ(b:X)664 4838 y FW(does)25 b FP(not)h
+FW(hold.)e(This)h(is)h(a)g(crucial)e(point:)g(if)i(latter)f(equation)e
+(were)i(to)g(hold,)f(then)h(a)g(possibly-)664 4938 y(capturing)31
+b(substitution)h(could)g(replace)g FQ(X)39 b FW(by)33
+b(a)g(term)g(containing)d(free)j(occurrences)e(of)h FQ(a)p
+FW(,)664 5037 y(equating)e(tw)o(o)h(terms)g(which)f(are)h(not)g
+FQ(\013)p FW(-equi)n(v)n(alent.)e(So)i(e)n(v)o(en)f(if)h(Miller')-5
+b(s)31 b(algorithm)f(is)i(re-)664 5137 y(stricted)g(to)f(solv)o(e)g
+(uni\002cation)f(problems)g(modulo)g(only)g FM(=)2501
+5149 y FI(\013)2580 5137 y FW(\(instead)h(of)g FM(=)3041
+5149 y FI(\013\014)3122 5157 y Fk(0)3154 5149 y FI(\021)3194
+5137 y FW(\),)h(nominal)p eop end
+%%Page: 16 2
+TeXDict begin 16 15 bop 664 448 a FW(uni\002cation)32
+b(and)g(higher)n(-order)e(pattern)i(uni\002cation)g(disagree)g(on)g
+(the)h(equi)n(v)n(alence)e(relation)664 548 y(that)d(underlies)e(the)h
+(uni\002cation)g(process.)f(Our)h FO(\031)p FW(-equi)n(v)n(alence)e(is)
+j(a)g(non-conserv)n(ati)n(v)o(e)23 b(mod-)664 648 y(i\002cation)c(of)g
+FM(=)1110 660 y FI(\013)1176 648 y FW(on)g(non-ground)c(terms;)k(its)h
+(major)e(feature)g(is)i(that,)f(unlik)o(e)f FM(=)3005
+660 y FI(\013)3052 648 y FW(,)h(it)h(is)g(respected)664
+747 y(by)g(possibly-capturing)d(substitutions)i(\(see)i(Lemma)e(2\).)
+581 849 y FO(\017)41 b FW(Nominal)30 b(uni\002cation)g(solv)o(es)h
+(freshness)g(problems)e(of)i(the)g(form)f FQ(a)43 b FM(#?)h
+FQ(t)p FW(.)31 b(Such)g(freshness)664 948 y(problems)c(ha)n(v)o(e)h(no)
+g FP(dir)m(ect)i FW(equi)n(v)n(alent)d(in)h(Miller')-5
+b(s)29 b(algorithm.)e(Also)h(note)g(that)h(the)o(y)f(do)g(not)664
+1048 y(only)17 b(arise)h(when)f(abstractions)g(are)h(uni\002ed,)f(b)n
+(ut)h(can)f(be)h(arbitrarily)e(chosen.)h(The)o(y)g(are)h(a)g(useful)664
+1148 y(programming)e(construct)j(as)h(illustrated)f(in)h(the)f
+(Conclusion)f(by)h(the)h Fa(type)p FW(-program)c(in)k(which)664
+1247 y(the)k(freshness)g(constraint)f Fa(x)p FM(#)p Fa(Gamma)g
+FW(is)i(added)e(to)h(the)g(third)g(clause)g(ensuring)e(that)i(the)g
+(side-)664 1347 y(condition)16 b(of)h(the)g(corresponding)d(inference)i
+(rule)h(holds.)g(Note)g(that)h Fa(x)p FM(#)p Fa(Gamma)f
+FW(does)g(not)g(arise)664 1446 y(from)i(unifying)e(a)j(goal)f(with)g
+(the)h(head)f(of)g(this)h(clause\227it)f(is)i(an)e(independent)e
+(constraint)i(gi)n(v)o(en)664 1546 y(e)o(xplicitly)g(by)h(the)g
+(programmer)-5 b(.)523 1736 y(The)16 b(conclusion)e(we)i(dra)o(w)g
+(from)f(this)h(comparison)e(is)j(that)f(it)h(does)f(not)f(mak)o(e)h
+(sense)g(to)h(vie)n(w)f(nominal)523 1836 y(uni\002cation)26
+b(as)h(a)h(v)n(ariant)e(of)h(the)g(higher)n(-order)d(pattern)i
+(uni\002cation:)f(the)j(algorithms)d(solv)o(e)i(prob-)523
+1935 y(lems)j(from)f(dif)n(ferent)f(domains)h(\(equational)e(problems)i
+(of)g FQ(\025)p FW(-terms)h(modulo)e FQ(\013\014)3045
+1947 y FJ(0)3082 1935 y FQ(\021)s FW(-equi)n(v)n(alence)523
+2035 y(vs.)22 b(equational)e(problems)g(of)i(\002rst-order)e(terms)i
+(with)g(binders)f(modulo)f FO(\031)p FW(-equi)n(v)n(alence)f(and)i(in)h
+(ad-)523 2135 y(dition)g(freshness)g(problems\))e(and)i(also)h(produce)
+d(dif)n(ferent)h(solutions)h(\(capture-a)n(v)n(oiding)d(substitu-)523
+2234 y(tions)26 b(vs.)g(possibly-capturing)d(substitutions)j(and)f
+(freshness)h(constraints\).)f(Neither)g(the)i(domains,)523
+2334 y(nor)19 b(the)i(solutions,)e(can)h(be)g(seen)g(as)h(v)n(ariants)f
+(of)g(the)g(other)-5 b(.)648 2435 y(This)21 b(comparison,)d(of)j
+(course,)f(lea)n(v)o(es)g(open)g(the)h(question)f(whether)g(there)g(is)
+i(a)f(suitable)g FP(encod-)523 2535 y(ing)28 b FW(of)g(nominal)f
+(uni\002cation)g(problems)g(into)h(higher)n(-order)e(pattern)h
+(uni\002cation)g(problems)g(and)h(a)523 2635 y(suitable)20
+b FP(decoding)d FW(that)j(translates)g(solutions)f(produced)f(by)h
+(Miller')-5 b(s)20 b(algorithm)f(back)g(to)h(solutions)523
+2734 y(of)25 b(nominal)f(uni\002cation)g(problems.)g(Miller')-5
+b(s)26 b(algorithm)e(can)h(then)g(be)g(emplo)o(yed)f(to)h(solv)o(e)g
+(nomi-)523 2834 y(nal)g(uni\002cation)e(problems.)g(At)i(\002rst)g
+(sight,)f(there)g(seems)h(to)g(be)f(a)h(simple)g(encoding.)d(Consider)i
+(for)523 2934 y(e)o(xample)19 b(the)h(nominal)f(problem)1842
+3039 y FQ(a:X)29 b FO(\031)o FM(?)24 b FQ(b:b)523 3194
+y FW(for)d(which)g(nominal)f(uni\002cation)g(produces)f(the)j(solution)
+e FM(\()p Fn(?)p FQ(;)14 b FM([)p FQ(X)32 b FM(:=)25
+b FQ(a)p FM(]\))p FW(.)d(The)f(literal)g(encoding)f(as)523
+3294 y(the)g(problem)e FQ(\025a:X)29 b FM(=)1220 3306
+y FI(\013\014)1301 3314 y Fk(0)1333 3306 y FI(\021)1374
+3294 y FM(?)p FQ(\025b:b)20 b FW(is)g(unsuitable,)e(because)h(there)g
+(is)i(no)e(capture-a)n(v)n(oiding)e(substitu-)523 3393
+y(tion)j(that)g(solv)o(es)g(the)h(problem.)d(Ho)n(we)n(v)o(er)m(,)g
+FQ(X)27 b FW(can)20 b(be)g(made)g(dependent)e(on)h FQ(a)i
+FW(gi)n(ving)e(the)h(problem)1662 3580 y FQ(\025a:)p
+FM(\()p FQ(X)7 b(a)p FM(\))23 b(=)2049 3592 y FI(\013\014)2130
+3600 y Fk(0)2162 3592 y FI(\021)2203 3580 y FM(?)g FQ(\025b:b)523
+3766 y FW(which)17 b(can)g(be)h(solv)o(ed)e(by)h(the)h(capture-a)n(v)n
+(oiding)c(substitution)i FM([)p FQ(X)30 b FM(:=)22 b
+FQ(\025c:c)p FM(])p FW(.)d(If)e(we)h(further)e(apply)g(to)523
+3866 y FQ(\025c:c)k FW(the)f(atom)g FQ(a)p FW(,)g(which)g(w)o(as)h
+(introduced)d(by)h(the)h(encoding,)e(then)i(one)f(can)h(read)g(back)f
+(the)h(original)523 3965 y(solution)g FM([)p FQ(X)30
+b FM(:=)22 b FQ(a)p FM(])f FW(by)f(applying)e(some)i
+FQ(\014)t FW(-reductions.)648 4067 y(There)j(are)g(ho)n(we)n(v)o(er)f
+(se)n(v)o(eral)i(problems)e(with)i(this)h(encoding.)c(First,)k(it)f(is)
+h(not)f(ob)o(vious)e(ho)n(w)h(to)523 4166 y(encode)c(uniformly)f
+(suspensions)h(combined)f(with)j(abstractions,)e(such)h(as)1819
+4353 y FQ(b:a:)p FM(\()p FQ(a)14 b(b)p FM(\))p FK(\001)o
+FQ(X)30 b(:)1157 b FW(\(17\))523 4539 y(where)24 b(either)g
+FQ(a)h FW(or)f FQ(b)h FW(need)f(to)h(be)f(applied)g(to)h
+FQ(X)31 b FW(depending)22 b(on)i(the)h(problem)e(in)h(which)g(this)h
+(term)523 4639 y(occurs.)20 b(This)i(hints)g(at)f(an)h(encoding)d(that)
+j(mak)o(es)f(v)n(ariables)g(dependent)e(on)i FP(all)h
+FW(atoms)f(occurring)e(in)523 4738 y(a)27 b(nominal)e(uni\002cation)g
+(problem,)g(re)o(gardless)g(of)h(whether)f(the)o(y)h(actually)g(occur)f
+(in)i(an)f(indi)n(vidual)523 4838 y(equational)18 b(problem.)f(F)o(or)h
+(e)o(xample,)g(if)h FQ(X)26 b FW(occurs)18 b(else)n(where)h(within)g
+(the)g(scope)f(of)h(abstractions)g(of)523 4938 y FQ(c)p
+FW(,)g FQ(d)p FW(,)g FQ(e)g FW(and)f FQ(f)28 b FW(in)19
+b(a)g(nominal)e(uni\002cation)h(problem)f(in)i(which)f(\(17\))f(is)j
+(one)e(term,)g(then)h FQ(X)25 b FW(needs)18 b(to)h(be)523
+5037 y(encoded)g(as)i FM(\()p FQ(X)g(a)14 b(b)g(c)g(d)g(e)g(f)9
+b FM(\))p FW(,)19 b(although)f(\(17\))i(contains)g(only)g
+FQ(a)h FW(and)f FQ(b)p FW(.)g(Secondly)-5 b(,)19 b(and)h(more)f(impor)n
+(-)523 5137 y(tantly)-5 b(,)17 b(we)i(cannot)e(see)i(ho)n(w)f(to)g
+(encode)f(our)h(freshness)g(constraints)f(into)h(higher)n(-order)d
+(patterns:)j(the)p eop end
+%%Page: 17 3
+TeXDict begin 17 16 bop 523 448 a FW(reason)23 b(for)h(this)h(is)g
+(that,)f(as)h(sho)n(wn)e(in)h(the)h Fa(type)p FW(-e)o(xample,)c(the)o
+(y)j(do)f(not)h(necessarily)g(come)f(from)523 548 y(abstractions,)18
+b(rather)h(the)o(y)f(can)h(be)g(chosen)f(arbitrarily)-5
+b(.)18 b(In)h(ef)n(fect)f(we)i(do)f FP(not)h FW(e)o(xpect)e(that)h
+(there)g(is)h(an)523 648 y(encoding)f(of)i(nominal)f(uni\002cation)g
+(problems)f(into)i(higher)n(-order)d(pattern)i(uni\002cation)g
+(problems)g(at)523 747 y(all.)648 847 y(W)-7 b(e)17 b(conjecture)e
+(that)h(some)g(of)g(the)h(problems)d(mentioned)h(abo)o(v)o(e)g(can)h
+(indeed)f(be)h(o)o(v)o(ercome)e(in)i(an)523 946 y(encoding)22
+b(of)i(nominal)f(uni\002cation)g(problems)g(into)h FQ(\025\033)s
+FW(\227a)h(calculus)f(with)h(de-Bruijn)e(indices)h(and)523
+1046 y(e)o(xplicit)f(substitutions)g(presented)f(in)h([7].The)f
+(resulting)h FQ(\025\033)k FW(uni\002cation)c(problems)f(can)h(be)g
+(solv)o(ed)523 1146 y(using)d(the)h(algorithm)e(presented)g(in)i(that)f
+(paper;)g(ho)n(we)n(v)o(er)e(it)k(is)f(not)f(clear)h(ho)n(w)f(to)h
+FP(decode)e FW(the)h(nom-)523 1245 y(inal)i(uni\002cation)e(solution)g
+(from)h(the)g(solution)g(of)g(the)g(encoded)f(problem)f(in)j
+FQ(\025\033)s FW(.)g(The)g(encoding)d(we)523 1345 y(ha)n(v)o(e)d(in)h
+(mind)f(is)h(dif)n(ferent)e(from)h(the)g(`pre-cooking')d(operation)h
+(in)j([7])f(and)g(is)i(rather)d(complicated:)h(it)523
+1445 y(depends)i(on)i(speci\002c)g(features)f(of)h(de-Bruijn)e(indices)
+i(and)f(e)o(xplicit)g(substitutions.)g(W)-7 b(e)21 b(shall)f(sk)o(etch)
+523 1544 y(the)g(encoding)e(belo)n(w)-5 b(.)648 1644
+y(Gi)n(v)o(en)21 b(a)i(nominal)e(uni\002cation)g(problem,)f(all)j(its)g
+(atoms)f(are)g(encoded)f(as)i(natural)e(numbers,)g(for)523
+1743 y(e)o(xample)e FQ(a)24 b FO(!)g FM(1)p FW(,)d FQ(b)i
+FO(!)h FM(2)d FW(and)f(so)h(on.)f(This)h(mapping)e(will)j(be)e(\002x)o
+(ed)g(in)h(an)g FP(en)m(vir)l(onment)p FW(,)e(which)h(we)523
+1843 y(write)g(as)h(a)g(sequence,)e(for)g(e)o(xample)1854
+1943 y FQ(a:b:c:d:e:f)1200 b FW(\(18\))523 2087 y(in)19
+b(which)f(the)h(position)f(determines)g(the)h(de-Bruijn)e(inde)o(x)h
+(of)h(an)f(atom.)h(As)g(hinted)f(abo)o(v)o(e,)f(v)n(ariables)523
+2186 y(need)27 b(to)h(be)f(`dependent')e(on)i(all)h(atoms)g(occurring)d
+(in)j(a)g(problem\227in)d FQ(\025\033)32 b FW(we)c(use)g(e)o(xplicit)f
+(sub-)523 2286 y(stitutions)h(for)g(this.)g(Gi)n(v)o(en)g(the)g(v)n
+(ariable)f FQ(X)35 b FW(and)27 b(the)i(en)m(vironment)c(in)j(\(18\),)f
+(the)h(encoding)e(leads)523 2386 y(to)1453 2485 y FO(j)p
+FQ(X)7 b FO(j)1575 2497 y FI(a:b:c:d:e:f)1900 2485 y
+FO(7!)23 b FQ(X)7 b FM([1)p FQ(:)p FM(2)p FQ(:)p FM(3)p
+FQ(:)p FM(4)p FQ(:)p FM(5)p FQ(:)p FM(6)p FQ(:)20 b FO(")2557
+2451 y FJ(6)2594 2485 y FM(])523 2629 y FW(where)e FM([1)p
+FQ(:)p FM(2)p FQ(:)p FM(3)p FQ(:)p FM(4)p FQ(:)p FM(5)p
+FQ(:)p FM(6)p FQ(:)j FO(")1221 2599 y FJ(6)1258 2629
+y FM(])e FW(is)h(an)f(e)o(xplicit)g(substitution.)f(One)h(reason)f(for)
+g(using)h(en)m(vironments)d(is)k(that)523 2729 y(permutations)e(can)i
+(be)g(encoded)f(as)i(permutations)d(on)i(en)m(vironments,)d(for)j(e)o
+(xample)1061 2903 y FO(j)p FM(\()p FQ(b)14 b(d)p FM(\))p
+FK(\001)p FQ(X)7 b FO(j)1367 2915 y FI(a:b:c:d:e:f)1692
+2903 y FO(7!)23 b(j)p FQ(X)7 b FO(j)1920 2915 y FI(a:d:c:b:e:f)2245
+2903 y FO(7!)23 b FQ(X)7 b FM([1)p FQ(:)p FM(4)p FQ(:)p
+FM(3)p FQ(:)p FM(2)p FQ(:)p FM(5)p FQ(:)p FM(6)p FQ(:)20
+b FO(")2902 2869 y FJ(6)2939 2903 y FM(])j FQ(:)523 3077
+y FW(The)g(other)f(is)j(that)e(the)o(y)f(help)h(to)h(de\002ne)e(the)i
+(encoding)d(for)h(abstractions.)h(Because)g(of)g(the)g(w)o(ay)h(ab-)523
+3176 y(stractions)17 b(are)g(represented)e(by)i(de-Bruijn)f(indices,)g
+(the)h(en)m(vironment)d(needs)j(to)g(be)g(modi\002ed)f(when-)523
+3276 y(e)n(v)o(er)24 b(an)h(encoding)e(deals)i(with)g(an)g
+(abstraction.)f(F)o(or)g(e)o(xample)g(the)h(abstraction)f
+FQ(c:e:X)31 b FW(is)26 b(encoded)523 3376 y(as)1025 3475
+y FO(j)p FQ(c:e:X)7 b FO(j)1268 3487 y FI(a:b:c:d:e:f)1593
+3475 y FO(7!)23 b FQ(\025)p FO(j)p FQ(e:X)7 b FO(j)1931
+3487 y FI(c:a:b:)p FL(\017)p FI(:d:e:f)2310 3475 y FO(7!)23
+b FQ(\025\025)p FO(j)p FQ(X)7 b FO(j)2634 3487 y FI(e:c:a:b:)p
+FL(\017)p FI(:d:)p FL(\017)p FI(:f)523 3619 y FW(where)20
+b(the)g(tw)o(o)g FO(\017)h FW(indicate,)e(roughly)f(speaking,)h(that)h
+FQ(c)h FW(and)e FQ(e)i FW(are)f(inaccessible)g(under)e(the)j
+FQ(\025)p FW(s.)648 3719 y(As)k(remark)o(ed)d(abo)o(v)o(e,)h(the)h
+(reason)g(why)g(we)g(cannot)g(see)h(ho)n(w)f(to)g(encode)f(nominal)g
+(uni\002cation)523 3819 y(problems)18 b(into)h(higher)n(-order)d
+(pattern)j(problems)f(has)h(to)h(do)e(with)i(our)f(freshness)f
+(constraints.)h(In)g FQ(\025\033)523 3918 y FW(there)24
+b(is)h(no)f FP(dir)m(ect)h FW(equi)n(v)n(alent)e(for)g(freshness)h
+(constraints,)f(b)n(ut)i(Franc)-32 b(\270)s(ois)24 b(Pottier)g(has)h
+(pointed)e(out)523 4018 y(to)f(us)h(that)f(freshness)f(is)i(e)o
+(xpressible)e(implicitly)h(in)g FQ(\025\033)k FW(and)c(hence)f(that)h
+(an)g(encoding)e(of)i(freshness)523 4117 y(problems)d(is)i(possible.)f
+(Consider)g FQ(b)j FM(#?)g FQ(X)7 b FW(,)20 b(which)g(can)g(be)g(solv)o
+(ed)g(if)g FQ(X)27 b FW(is)22 b(not)e(instantiated)f(with)i(a)523
+4217 y(term)f(containing)e FQ(b)j FW(freely)-5 b(.)18
+b(The)i(same)h(can)f(be)g(achie)n(v)o(ed)e(in)j FQ(\025\033)j
+FW(with)c(the)h(problem)1750 4391 y FQ(X)30 b FM(=?)23
+b FQ(Y)18 b FM([1)p FQ(:)p FM(3)p FQ(:)k FO(")2259 4357
+y FJ(3)2296 4391 y FM(])523 4565 y FW(where)17 b FQ(Y)37
+b FW(is)18 b(a)h(fresh)e(v)n(ariable)f(\(not)h(occurring)e(in)j(the)g
+(nominal)e(uni\002cation)h(problem\).)e(Note)i(that)h(the)523
+4664 y(e)o(xplicit)k(substitution)g FM([1)p FQ(:)p FM(3)p
+FQ(:)28 b FO(")1428 4634 y FJ(3)1464 4664 y FM(])c FW(ensures)f(that)g
+(no)f(matter)h(what)g(is)g(substituted)g(for)f FQ(Y)42
+b FW(the)23 b(resulting)523 4764 y(term)d(cannot)f(mention)g(the)h
+(de-Bruijn)f(inde)o(x)g FM(2)p FW(,)h(which)g(corresponds)e(to)i
+FQ(b)g FW(being)g(fresh)f(for)h FQ(X)7 b FW(.)648 4864
+y(The)25 b(main)h(obstacle)f(to)i(relating)e(nominal)f(uni\002cation)h
+(problems)g(with)h(uni\002cation)f(problems)523 4963
+y(in)20 b FQ(\025\033)25 b FW(is)c(the)f(decoding)e(of)i(solutions.)g
+(Consider)f(the)h(nominal)f(uni\002cation)g(problem)1641
+5137 y FQ(X)29 b FO(\031)p FM(?)23 b FQ(a:Y)42 b(;)37
+b(b:Z)28 b FO(\031)o FM(?)c FQ(X)p eop end
+%%Page: 18 4
+TeXDict begin 18 17 bop 523 448 a FW(which)29 b(has)h(the)f(solution)g
+FM(\()p FQ(b)p FM(#)p FQ(Y)5 b(;)14 b FM([)p FQ(X)46
+b FM(:=)40 b FQ(a:Y)5 b(;)14 b(Z)45 b FM(:=)40 b(\()p
+FQ(a)14 b(b)p FM(\))p FK(\001)o FQ(Y)19 b FM(]\))p FW(.)30
+b(Using)f(the)h(encoding)d(sk)o(etched)523 548 y(abo)o(v)o(e)19
+b(\(with)h(the)g(en)m(vironment)d FQ(a:b)p FW(\),)j(one)f(obtains)h
+(the)g(follo)n(wing)f(uni\002cation)g(problem)f(in)i
+FQ(\025\033)1033 731 y(X)7 b FM([1)p FQ(:)p FM(2)p FQ(:)21
+b FO(")1325 696 y FJ(2)1362 731 y FM(])i(=?)g FQ(\025)14
+b(Y)19 b FM([1)p FQ(:)p FM(3)p FQ(:)j FO(")1881 696 y
+FJ(3)1918 731 y FM(])h FQ(;)37 b(\025)14 b(Z)6 b FM([2)p
+FQ(:)p FM(1)p FQ(:)22 b FO(")2366 696 y FJ(3)2403 731
+y FM(])h(=?)g FQ(\025)14 b(X)7 b FM([1)p FQ(:)p FM(2)p
+FQ(:)22 b FO(")2931 696 y FJ(2)2968 731 y FM(])h FQ(:)523
+913 y FW(One)30 b(solution)g(for)g(this)h(problem)d(is)k
+FM([)p FQ(X)48 b FM(:=)42 b FQ(\025)14 b(Y)2063 883 y
+FL(0)2086 913 y FM([1)p FQ(:)42 b FO(")2258 883 y FJ(3)2295
+913 y FM(])p FQ(;)27 b(Z)48 b FM(:=)42 b FQ(Y)2669 883
+y FL(0)2693 913 y FM([)p FO(")2758 883 y FJ(1)2794 913
+y FM(])p FQ(;)28 b(Y)61 b FM(:=)42 b FQ(Y)3173 883 y
+FL(0)3197 913 y FM([1)p FQ(:)f FO(")3368 883 y FJ(2)3405
+913 y FM(]])31 b FW(in)523 1013 y(which)e FQ(Y)823 983
+y FL(0)876 1013 y FW(is)i(a)f(fresh)g(v)n(ariable.)e(While)i(in)g(this)
+g(case)g(the)g(solution)f(seems)h(to)g(correspond)d(to)j(the)523
+1112 y(solution)d(of)h(the)g(nominal)f(uni\002cation)g(problem)f(abo)o
+(v)o(e,)g(so)j(f)o(ar)f(we)g(ha)n(v)o(e)g(been)f(unable)g(to)h(gi)n(v)o
+(e)g(a)523 1212 y(decoding)22 b(that)i(allo)n(ws)g(us)g(to)g(relate)g
+(in)g FP(e)o(very)g FW(case)h(a)f(solution)f(produced)f(by)h
+(uni\002cation)g(in)h FQ(\025\033)k FW(to)523 1312 y(a)i(solution)f
+(produced)d(by)k(the)f(nominal)f(uni\002cation)g(algorithm.)g(As)i(can)
+g(be)f(seen,)g(the)h(freshness)523 1411 y(constraint)20
+b FQ(b)p FM(#)p FQ(Y)40 b FW(is)23 b(completely)c(embedded)g(into)i
+(the)h(solution)e FQ(\025\033)26 b FW(produces.)19 b(T)-7
+b(o)21 b(mak)o(e)g(it)h(e)o(xplicit)523 1511 y(again)d(seems)i(to)f(be)
+g(non-tri)n(vial)f(in)h(more)f(complicated)g(cases)i(and)e(remains)h
+(future)f(w)o(ork.)648 1611 y(Ho)n(we)n(v)o(er)m(,)i(e)n(v)o(en)i(if)h
+(the)g(encoding)d(into)j(and)f(decoding)f(from)g FQ(\025\033)29
+b FW(can)23 b(be)h(made)f(to)h(w)o(ork,)f(there)523 1710
+y(are)d(at)h(least)g(tw)o(o)f(reasons)g(which)g(justify)g(a)g(direct)g
+(algorithm:)581 1876 y FO(\017)41 b FW(The)27 b(calculations)e(in)m(v)n
+(olv)o(ed)g(in)i(using)f(it)h(to)g(encode)f(a)h(nominal)e
+(uni\002cation)h(problem,)e(\002nd)j(a)664 1976 y(solution)h(in)i
+FQ(\025\033)j FW(and)c(then)g(decode)e(it)j(back)f(to)g(a)h(`nominal')d
+(solution)h(are)h(lik)o(ely)g(to)g(be)g(v)o(ery)664 2076
+y(complicated.)g(Gi)n(v)o(en)h(the)h(relati)n(v)o(e)g(simplicity)f(of)h
+(the)g(nominal)f(uni\002cation)g(algorithm,)f(one)664
+2175 y(w)o(ould)20 b(certainly)f(hope)g(not)h(to)g(ha)n(v)o(e)g(to)g
+(use)h(this)f(roundabout)d(method)i(in)h(practice.)581
+2275 y FO(\017)41 b FW(An)15 b(encoding/decoding)c(does)k(not)g(gi)n(v)
+o(e)f(an)o(y)h(insight)f(into)h(the)h(properties)d(of)i(the)h
+FO(\031)p FW(-equi)n(v)n(alence.)664 2374 y(In)25 b(the)g(paper)f(we)h
+(ha)n(v)o(e)f(sho)n(wn)h(that)g FO(\031)g FW(is)h(indeed)e(an)g(equi)n
+(v)n(alence)f(relation)h(and)h(that)g(it)h(is)g(re-)664
+2474 y(spected)20 b(by)g(possibly-capturing)c(substitutions.)k(Both)g
+(f)o(acts)h(are)f(non-ob)o(vious.)3058 2444 y Fy(3)523
+2640 y FW(The)h(conclusion)g(we)h(dra)o(w)f(from)g(this)h(is)h(that)e
+(solving)g(nominal)g(uni\002cation)f(problems)h(by)g(the)h(uni-)523
+2740 y(\002cation)d(algorithm)f(described)h(in)h([7])f(might)g(be)g
+(possible,)g(b)n(ut)h(this)g(is)h FP(no)e FW(substitute)h(for)f(ha)n
+(ving)f(the)523 2839 y(direct)i(algorithm)e(we)j(presented.)p
+523 5050 473 4 v 558 5105 a Fu(3)606 5137 y FV(See)e
+FB(http)j FF(:)g FS(==)p FB(www)p FS(:)p FB(cl)p FS(:)p
+FB(cam)p FS(:)p FB(ac)p FS(:)p FB(uk)p FS(=)1607 5129
+y Fw(\030)1664 5137 y FB(cu200)p FS(=)p FB(Unification)h
+FV(for)c(the)g(Isabelle)g(proof)g(scripts.)p eop end
+%%Trailer
+
+userdict /end-hook known{end-hook}if
+%%EOF
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/index.html Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,118 @@
+<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.0 Transitional//EN"
+"http://www.w3.org/TR/REC-html40/loose.dtd">
+<HEAD>
+<TITLE>Nominal Unification</TITLE>
+<BASE HREF="http://www4.in.tum.de/~urbanc/Unification">
+</HEAD>
+
+<BODY TEXT="#000000"
+ BGCOLOR="#4169E1"
+ LINK="#0000EF"
+ VLINK="#51188E"
+ ALINK="#FF0000">
+
+<TABLE WIDTH="100%"
+ BGCOLOR="#4169E1"
+ BORDER="0"
+ FRAME="border"
+ CELLPADDING="10"
+ CELLSPACING="2"
+ RULES="all">
+
+<!-- right column -->
+<TR>
+<TD BGCOLOR="#FFFFFF" WIDTH="75%">
+<H2>Nominal Unification [<a HREF="http://www4.in.tum.de/~urbanc/Unification/nomu-tcs.ps">ps</a>]</H2>
+
+<A HREF="http://www4.in.tum.de/~urbanc/">Christian Urban</A>,
+<A HREF="http://www.cl.cam.ac.uk/~amp12/">Andrew Pitts</A>,
+Murdoch Gabbay
+<p>
+ We present a generalisation of first-order unification to the
+ practically important case of equations between terms involving
+ <i>binding operations</i>. A substitution of terms for variables
+ solves such an equation if it makes the equated terms
+ <i>alpha-equivalent</i>, i.e. equal up to renaming bound names.
+ For the applications we have in mind, we must consider the simple,
+ textual form of substitution in which names occurring in terms may
+ be captured within the scope of binders upon substitution. We are
+ able to take a "nominal" approach to binding in which bound
+ entities are explicitly named (rather than using nameless,
+ de Bruijn-style representations) and yet get a version of this
+ form of substitution that respects alpha-equivalence and
+ possesses good algorithmic properties. We achieve this by adapting
+ an existing idea and introducing a key new idea. The existing
+ idea is terms involving explicit substitutions of names for names,
+ except that here we only use <i>explicit permutations</i>
+
+ (bijective substitutions). The key new idea is that the
+ unification algorithm should solve not only equational problems,
+ but also problems about the <i>freshness</i> of names for
+ terms. There is a simple generalisation of the classical
+ first-order unification algorithm to this setting which retains
+ the latter's pleasant properties: unification problems involving
+ alpha-equivalence and freshness are decidable; and solvable
+ problems possess most general solutions.
+<BR><BR>
+
+In an <a HREF="http://www4.in.tum.de/~urbanc/Unification/app-nomu.ps">appendix</a> we
+discuss some issues about the relationship between nominal unification and higher-order
+pattern unification.
+
+<BR><BR><BR><BR>
+All results in the paper have been verified in
+<A HREF="http://www.cl.cam.ac.uk/Research/HVG/Isabelle/">Isabelle</A> (2004 version).
+Below are the theory files.
+
+<ul>
+<li><A HREF="Unification/Swap.thy">Swap.thy</A>
+ <br /> amongst other facts proves that swapping is a bijection (on atoms)
+<li><A HREF="Unification/Atoms.thy">Atoms.thy</A>
+ <br /> facts about atoms occurring in swappings
+<li><A HREF="Unification/Terms.thy">Terms.thy</A>
+ <br /> defines terms, occurs check and the notion of subterms
+<li><A HREF="Unification/Disagreement.thy">Disagreement.thy</A>
+ <br /> proves various facts about disagreement sets
+<li><A HREF="Unification/Fresh.thy">Fresh.thy</A>
+ <br /> defines the freshness relation and shows facts about its behaviour under swapping
+<li><A HREF="Unification/PreEqu.thy">PreEqu.thy</A>
+ <br /> defines the relation capturing the notion of alpha-equivalence (on open terms)
+ and proves a long lemma by mutual induction over the depth of terms which
+ is needed to show that the relation is an equivalence relation
+<li><A HREF="Unification/Equ.thy">Equ.thy</A>
+ <br /> proves various facts about the equivalence relations
+<li><A HREF="Unification/Substs.thy">Substs.thy</A>
+ <br /> defines substitutions and composition of substitutions, and establishes
+ some facts of substitution and our equivalence relation
+<li><A HREF="Unification/Mgu.thy">Mgu.thy</A>
+ <br /> defines the notion of unification problems and reduction rules over sets
+ of such problems; proves that every reduction leading to the empty set
+ produces an mgu
+<li><A HREF="Unification/Termination.thy">Termination.thy</A>
+ <br /> shows that every reduction reduces a (well-founded) measure, thus
+ proves that every reduction sequence must terminate
+<li><A HREF="Unification/Unification.thy">Unification.thy</A>
+ <br /> proves that all solvable problems reduce only to the empty set
+ (the "ok" configuration which provides an mgu) and all unsolvable
+ problems reduce to a "fail" configuration
+ (for which no unifier exists)
+</uL><BR>
+The old Isabelle-2002 files can be downloaded
+<A HREF="http://www4.in.tum.de/~urbanc/Unification/nomu-2002.tgz">here</A>.
+<BR><BR>
+A "quick and dirty" implementation of nominal unification in
+<A HREF="http://caml.inria.fr/">Ocaml</A> can be
+downloaded
+<A HREF="http://www4.in.tum.de/~urbanc/Unification/unification.ml">elsewhere</A>.
+
+
+<BR><BR><BR><BR>
+
+</TABLE>
+<P><!-- Created: Tue Mar 4 00:23:25 GMT 1997 -->
+<!-- hhmts start -->
+Last modified: Thu Feb 28 20:24:23 CET 2008
+<!-- hhmts end -->
+<a href="http://validator.w3.org/check/referer">[Validate this page.]</a>
+</BODY>
+</HTML>
Binary file Unification/nomu-2002.tgz has changed
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/nomu-tcs.ps Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,5507 @@
+%!PS-Adobe-2.0
+%%Creator: dvips(k) 5.92b Copyright 2002 Radical Eye Software
+%%Title: nomu-tcs.dvi
+%%Pages: 29
+%%PageOrder: Ascend
+%%BoundingBox: 0 0 596 842
+%%DocumentFonts: Times-Bold Times-Roman CMR10 CMR8 Times-Italic CMMI10
+%%+ CMMI12 CMSY10 CMTT12 CMR12 CMSY8 Courier CMBSY10 CMMI8 CMTI12 CMEX10
+%%+ CMTI10 CMMIB10 EUSM10 CMR6 CMMI6 CMSY6 CMTT8 CMBSY7
+%%EndComments
+%DVIPSWebPage: (www.radicaleye.com)
+%DVIPSCommandLine: dvips nomu-tcs.dvi -o nomu-tcs.ps
+%DVIPSParameters: dpi=600, compressed
+%DVIPSSource: TeX output 2004.04.09:1442
+%%BeginProcSet: texc.pro
+%!
+/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/id 0 N/rw 0 N/rc 0 N/gp 0 N/cp 0 N/G 0 N/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]/id Ci N/rw Cw 7 add 8 idiv string N/rc 0 N/gp 0 N/cp 0 N{
+rc 0 ne{rc 1 sub/rc X rw}{G}ifelse}imagemask restore}B/G{{id gp get/gp
+gp 1 add N A 18 mod S 18 idiv pl S get exec}loop}B/adv{cp add/cp X}B
+/chg{rw cp id gp 4 index getinterval putinterval A gp add/gp X adv}B/nd{
+/cp 0 N rw exit}B/lsh{rw cp 2 copy get A 0 eq{pop 1}{A 255 eq{pop 254}{
+A A add 255 and S 1 and or}ifelse}ifelse put 1 adv}B/rsh{rw cp 2 copy
+get A 0 eq{pop 128}{A 255 eq{pop 127}{A 2 idiv S 128 and or}ifelse}
+ifelse put 1 adv}B/clr{rw cp 2 index string putinterval adv}B/set{rw cp
+fillstr 0 4 index getinterval putinterval adv}B/fillstr 18 string 0 1 17
+{2 copy 255 put pop}for N/pl[{adv 1 chg}{adv 1 chg nd}{1 add chg}{1 add
+chg nd}{adv lsh}{adv lsh nd}{adv rsh}{adv rsh nd}{1 add adv}{/rc X nd}{
+1 add set}{1 add clr}{adv 2 chg}{adv 2 chg nd}{pop nd}]A{bind pop}
+forall N/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: 8r.enc
+% File 8r.enc as of 2002-03-12 for PSNFSS 9
+%
+% 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
+%
+% Authors: S. Rahtz, P. MacKay, Alan Jeffrey, B. Horn, K. Berry, W. Schmidt
+%
+% 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 + extra characters from Lucida + Euro.
+%
+% 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.
+%
+% (4) Remaining positions left undefined are for use in (hopefully)
+% upward-compatible revisions, if someday more characters are generally
+% available.
+%
+% (5) hyphen appears twice for compatibility with both ASCII and Windows.
+%
+% (6) /Euro is assigned to 128, as in Windows ANSI
+%
+/TeXBase1Encoding [
+% 0x00 (encoded characters from Adobe Standard not in Windows 3.1)
+ /.notdef /dotaccent /fi /fl
+ /fraction /hungarumlaut /Lslash /lslash
+ /ogonek /ring /.notdef
+ /breve /minus /.notdef
+% These are the only two remaining unencoded characters, so may as
+% well include them.
+ /Zcaron /zcaron
+% 0x10
+ /caron /dotlessi
+% (unusual TeX characters available in, e.g., Lucida Bright)
+ /dotlessj /ff /ffi /ffl
+ /.notdef /.notdef /.notdef /.notdef
+ /.notdef /.notdef /.notdef /.notdef
+ % very contentious; it's so painful not having quoteleft and quoteright
+ % at 96 and 145 that we move the things normally found there down to here.
+ /grave /quotesingle
+% 0x20 (ASCII begins)
+ /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 % rubout; ASCII ends
+% 0x80
+ /Euro /.notdef /quotesinglbase /florin
+ /quotedblbase /ellipsis /dagger /daggerdbl
+ /circumflex /perthousand /Scaron /guilsinglleft
+ /OE /.notdef /.notdef /.notdef
+% 0x90
+ /.notdef /.notdef /.notdef /quotedblleft
+ /quotedblright /bullet /endash /emdash
+ /tilde /trademark /scaron /guilsinglright
+ /oe /.notdef /.notdef /Ydieresis
+% 0xA0
+ /.notdef % nobreakspace
+ /exclamdown /cent /sterling
+ /currency /yen /brokenbar /section
+ /dieresis /copyright /ordfeminine /guillemotleft
+ /logicalnot
+ /hyphen % Y&Y (also at 45); Windows' softhyphen
+ /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: f7b6d320.enc
+% Thomas Esser, Dec 2002. public domain
+%
+% Encoding for:
+% cmb10 cmbx10 cmbx12 cmbx5 cmbx6 cmbx7 cmbx8 cmbx9 cmbxsl10
+% cmdunh10 cmr10 cmr12 cmr17cmr6 cmr7 cmr8 cmr9 cmsl10 cmsl12 cmsl8
+% cmsl9 cmss10cmss12 cmss17 cmss8 cmss9 cmssbx10 cmssdc10 cmssi10
+% cmssi12 cmssi17 cmssi8cmssi9 cmssq8 cmssqi8 cmvtt10
+%
+/TeXf7b6d320Encoding [
+/Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi /Omega
+/ff /fi /fl /ffi /ffl /dotlessi /dotlessj /grave /acute /caron /breve
+/macron /ring /cedilla /germandbls /ae /oe /oslash /AE /OE /Oslash
+/suppress /exclam /quotedblright /numbersign /dollar /percent /ampersand
+/quoteright /parenleft /parenright /asterisk /plus /comma /hyphen
+/period /slash /zero /one /two /three /four /five /six /seven /eight
+/nine /colon /semicolon /exclamdown /equal /questiondown /question /at
+/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 /bracketleft /quotedblleft /bracketright /circumflex /dotaccent
+/quoteleft /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 /endash /emdash /hungarumlaut /tilde /dieresis /suppress
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /space
+/Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi /.notdef
+/.notdef /Omega /ff /fi /fl /ffi /ffl /dotlessi /dotlessj /grave /acute
+/caron /breve /macron /ring /cedilla /germandbls /ae /oe /oslash /AE
+/OE /Oslash /suppress /dieresis /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+] def
+
+%%EndProcSet
+%%BeginProcSet: aae443f0.enc
+% Thomas Esser, Dec 2002. public domain
+%
+% Encoding for:
+% cmmi10 cmmi12 cmmi5 cmmi6 cmmi7 cmmi8 cmmi9 cmmib10
+%
+/TeXaae443f0Encoding [
+/Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi /Omega
+/alpha /beta /gamma /delta /epsilon1 /zeta /eta /theta /iota /kappa
+/lambda /mu /nu /xi /pi /rho /sigma /tau /upsilon /phi /chi /psi
+/omega /epsilon /theta1 /pi1 /rho1 /sigma1 /phi1 /arrowlefttophalf
+/arrowleftbothalf /arrowrighttophalf /arrowrightbothalf /arrowhookleft
+/arrowhookright /triangleright /triangleleft /zerooldstyle /oneoldstyle
+/twooldstyle /threeoldstyle /fouroldstyle /fiveoldstyle /sixoldstyle
+/sevenoldstyle /eightoldstyle /nineoldstyle /period /comma /less /slash
+/greater /star /partialdiff /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 /flat /natural /sharp /slurbelow
+/slurabove /lscript /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 /dotlessi /dotlessj /weierstrass /vector
+/tie /psi /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/space /Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi
+/.notdef /.notdef /Omega /alpha /beta /gamma /delta /epsilon1 /zeta /eta
+/theta /iota /kappa /lambda /mu /nu /xi /pi /rho /sigma /tau /upsilon
+/phi /chi /psi /tie /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef
+] def
+
+%%EndProcSet
+%%BeginProcSet: bbad153f.enc
+% Thomas Esser, Dec 2002. public domain
+%
+% Encoding for:
+% cmsy10 cmsy5 cmsy6 cmsy7 cmsy8 cmsy9
+%
+/TeXbbad153fEncoding [
+/minus /periodcentered /multiply /asteriskmath /divide /diamondmath
+/plusminus /minusplus /circleplus /circleminus /circlemultiply
+/circledivide /circledot /circlecopyrt /openbullet /bullet
+/equivasymptotic /equivalence /reflexsubset /reflexsuperset /lessequal
+/greaterequal /precedesequal /followsequal /similar /approxequal
+/propersubset /propersuperset /lessmuch /greatermuch /precedes /follows
+/arrowleft /arrowright /arrowup /arrowdown /arrowboth /arrownortheast
+/arrowsoutheast /similarequal /arrowdblleft /arrowdblright /arrowdblup
+/arrowdbldown /arrowdblboth /arrownorthwest /arrowsouthwest /proportional
+/prime /infinity /element /owner /triangle /triangleinv /negationslash
+/mapsto /universal /existential /logicalnot /emptyset /Rfractur /Ifractur
+/latticetop /perpendicular /aleph /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 /union /intersection
+/unionmulti /logicaland /logicalor /turnstileleft /turnstileright
+/floorleft /floorright /ceilingleft /ceilingright /braceleft /braceright
+/angbracketleft /angbracketright /bar /bardbl /arrowbothv /arrowdblbothv
+/backslash /wreathproduct /radical /coproduct /nabla /integral
+/unionsq /intersectionsq /subsetsqequal /supersetsqequal /section
+/dagger /daggerdbl /paragraph /club /diamond /heart /spade /arrowleft
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/minus /periodcentered /multiply /asteriskmath /divide /diamondmath
+/plusminus /minusplus /circleplus /circleminus /.notdef /.notdef
+/circlemultiply /circledivide /circledot /circlecopyrt /openbullet
+/bullet /equivasymptotic /equivalence /reflexsubset /reflexsuperset
+/lessequal /greaterequal /precedesequal /followsequal /similar
+/approxequal /propersubset /propersuperset /lessmuch /greatermuch
+/precedes /follows /arrowleft /spade /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+] def
+
+%%EndProcSet
+%%BeginProcSet: 09fbbfac.enc
+% Thomas Esser, Dec 2002. public domain
+%
+% Encoding for:
+% cmsltt10 cmtt10 cmtt12 cmtt8 cmtt9
+/TeX09fbbfacEncoding [
+/Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi
+/Omega /arrowup /arrowdown /quotesingle /exclamdown /questiondown
+/dotlessi /dotlessj /grave /acute /caron /breve /macron /ring /cedilla
+/germandbls /ae /oe /oslash /AE /OE /Oslash /visiblespace /exclam
+/quotedbl /numbersign /dollar /percent /ampersand /quoteright /parenleft
+/parenright /asterisk /plus /comma /hyphen /period /slash /zero /one
+/two /three /four /five /six /seven /eight /nine /colon /semicolon /less
+/equal /greater /question /at /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 /bracketleft /backslash /bracketright
+/asciicircum /underscore /quoteleft /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 /braceleft /bar /braceright
+/asciitilde /dieresis /visiblespace /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /space /Gamma /Delta /Theta /Lambda /Xi /Pi
+/Sigma /Upsilon /Phi /Psi /.notdef /.notdef /Omega /arrowup /arrowdown
+/quotesingle /exclamdown /questiondown /dotlessi /dotlessj /grave /acute
+/caron /breve /macron /ring /cedilla /germandbls /ae /oe /oslash /AE
+/OE /Oslash /visiblespace /dieresis /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+] def
+
+%%EndProcSet
+%%BeginProcSet: 10037936.enc
+% Thomas Esser, Dec 2002. public domain
+%
+% Encoding for:
+% cmbsy10
+%
+/TeX10037936Encoding [
+/minus /periodcentered /multiply /asteriskmath /divide /diamondmath
+/plusminus /minusplus /circleplus /circleminus /circlemultiply
+/circledivide /circledot /circlecopyrt /openbullet /bullet
+/equivasymptotic /equivalence /reflexsubset /reflexsuperset /lessequal
+/greaterequal /precedesequal /followsequal /similar /approxequal
+/propersubset /propersuperset /lessmuch /greatermuch /precedes /follows
+/arrowleft /arrowright /arrowup /arrowdown /arrowboth /arrownortheast
+/arrowsoutheast /similarequal /arrowdblleft /arrowdblright /arrowdblup
+/arrowdbldown /arrowdblboth /arrownorthwest /arrowsouthwest /proportional
+/prime /infinity /element /owner /triangle /triangleinv /negationslash
+/mapsto /universal /existential /logicalnot /emptyset /Rfractur /Ifractur
+/latticetop /perpendicular /aleph /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 /union /intersection
+/unionmulti /logicaland /logicalor /turnstileleft /turnstileright
+/floorleft /floorright /ceilingleft /ceilingright /braceleft /braceright
+/angbracketleft /angbracketright /bar /bardbl /arrowbothv /arrowdblbothv
+/backslash /wreathproduct /radical /coproduct /nabla /integral
+/unionsq /intersectionsq /subsetsqequal /supersetsqequal /section
+/dagger /daggerdbl /paragraph /club /diamond /heart /spade /arrowleft
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /space
+/minus /periodcentered /multiply /asteriskmath /divide /diamondmath
+/plusminus /minusplus /circleplus /circleminus /.notdef /.notdef
+/circlemultiply /circledivide /circledot /circlecopyrt /openbullet
+/bullet /equivasymptotic /equivalence /reflexsubset /reflexsuperset
+/lessequal /greaterequal /precedesequal /followsequal /similar
+/approxequal /propersubset /propersuperset /lessmuch /greatermuch
+/precedes /follows /arrowleft /spade /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+] def
+
+%%EndProcSet
+%%BeginProcSet: 74afc74c.enc
+% Thomas Esser, Dec 2002. public domain
+%
+% Encoding for:
+% cmbxti10 cmff10 cmfi10 cmfib8 cmti10 cmti12 cmti7 cmti8cmti9 cmu10
+%
+/TeX74afc74cEncoding [
+/Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi /Omega
+/ff /fi /fl /ffi /ffl /dotlessi /dotlessj /grave /acute /caron /breve
+/macron /ring /cedilla /germandbls /ae /oe /oslash /AE /OE /Oslash
+/suppress /exclam /quotedblright /numbersign /sterling /percent
+/ampersand /quoteright /parenleft /parenright /asterisk /plus /comma
+/hyphen /period /slash /zero /one /two /three /four /five /six /seven
+/eight /nine /colon /semicolon /exclamdown /equal /questiondown /question
+/at /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 /bracketleft /quotedblleft /bracketright /circumflex /dotaccent
+/quoteleft /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 /endash /emdash /hungarumlaut /tilde /dieresis /suppress
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /space
+/Gamma /Delta /Theta /Lambda /Xi /Pi /Sigma /Upsilon /Phi /Psi /.notdef
+/.notdef /Omega /ff /fi /fl /ffi /ffl /dotlessi /dotlessj /grave /acute
+/caron /breve /macron /ring /cedilla /germandbls /ae /oe /oslash /AE
+/OE /Oslash /suppress /dieresis /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+/.notdef /.notdef /.notdef /.notdef /.notdef /.notdef /.notdef
+] def
+
+%%EndProcSet
+%%BeginProcSet: texps.pro
+%!
+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
+%%BeginFont: CMBSY7
+%!PS-AdobeFont-1.1: CMBSY7 001.000
+%%CreationDate: 1992 Oct 22 12:18:11
+% Computer Modern fonts were designed by Donald E. Knuth.
+% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
+11 dict begin
+/FontInfo 7 dict dup begin
+/version (001.000) readonly def
+/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
+/FullName (CMBSY7) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Bold) readonly def
+/ItalicAngle -14.035 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMBSY7 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 1 /periodcentered put
+readonly def
+/FontBBox{0 -927 1542 750}readonly def
+/UniqueID 5032008 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5F086C44216EBE57F4BA37B479BF1E5A5139D8
+91F3E6DDA157B25D359C5E7FE4CFB264DF1707BB6497D3E074CFF95D9FD269B5
+0B1566D1161680C46D1548BBF384EF8653AF29FD474EA2336C876979AC00AE18
+DC87DC0DDD3164B96FC6C3ED826EBAAED383BB3EC5044BA84A0426B87ED04C9F
+4B3EDFC734C241D9B3D2321619F870FE68BB4BA7C060DBE8FBD12F641E7709F2
+C430491944B78B1E59580798F1B40CAC4D59DE95941217EF1A861DDB0156D5C3
+349AAB13FEFF6C646F6401550F5853BC09B267A6C63639228DF55BE60A99E4A5
+308C616892AA0DC96ADB7CD7AB8AEFA859F69B587B61930596A46A905661E4BA
+DAB5E1CF15C94CF060B7FA600B17162AEC2DEE64A156B3F87248E7A7F88C9154
+8C494273B33483BEDF0BEA4DF6A19941F52AA04717623ACBDE926B4851ED05A0
+28698A1C5AE63A46EF473A4F3DCAF3E73C4FAF0C1077EE7A6504074C0D77947F
+940B16425B3F5834763732F26D3385774A1CACA70C07F58887A0301D1BE530A7
+D7AC00A0664617A0CBA9F6281FA4B9168DBA3C1EDCF915778351E6BD8A9CE7E5
+3E56F2FFE0969E1CFFC83F07C01A3873EE1CE4E124565E8F493BE4FAA0A5D099
+A116CEE4EC6C8CB2E93B42771FFF67680A56501A201E12AFDA8448ABE80BEB9B
+80428F48753C4EFB174B693C69DB81CEF0A0B75C53A9D5B4C5F26FA58059A324
+8D4E9D4E9C54AB8F9C21CD66B9B259F9C797559384A653DD43ED4B9C2110BB5B
+C3A6370921186AEE29EB4E62793ECD96935C3D9D89DC6AAECFB745BABDB570FD
+E7E6248B6ED9C09EDA896ECFCAEECE8C1E8BE20BDE6F3558EF5A32ED390ABF86
+3A585DF34F2B8B9567778BF51A1BF9C1018AEEF42FF1F9AAB1F9F73627F6C7DE
+12747A5031EDFF0C8BBB61D651344F0D188BDFDAFE3CCE9916646C67C437A80B
+F00FCB77E47839B2E76E72333A9E4BB4734434BC32E2263F7B27289D59C17B58
+C262B8D8AE6A95FBC5A3D9F72BB299CA31A6E15E647AC3E4975EBEB7C1B7F562
+A238F35ED7C1787094BE1AC142EA716D2784431BCD49E9FB3C070C5F12D0F1DD
+70AC849C3F18153C05C67EF30CB0AB03B748BF8CBC4B80C9E9F79E960E8C4308
+940655DE1405914E2D32B06EADFDFA08B9B7B2F9B88902FC1A6DC36F968F2D33
+5860D1D1D1768392BB9CFA
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMTT8
+%!PS-AdobeFont-1.1: CMTT8 1.0
+%%CreationDate: 1991 Aug 20 16:46:05
+% 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 (CMTT8) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle 0 def
+/isFixedPitch true def
+end readonly def
+/FontName /CMTT8 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 0 /.notdef put
+readonly def
+/FontBBox{-5 -232 545 699}readonly def
+/UniqueID 5000830 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5F0187316F83DDE3E2D27FCDF6C5CE4F95B6EE
+3317BD91B7921F3039DD35FEA387D5CFB6C6E9DC84C178F3432994FC7FAC6E5A
+ED41A1E2EBA350178FBFEB45944511731BA827167DDAC238FC69A5486B995477
+C469E2E27493B0B711DF8E267D3D5613B450011921685147114106C9472580BD
+F531022F6DF5432B2A4EBC51A8032C7F9689B6FA942D849B29709631613DA68D
+4DF7B6F059A19304F40A3C3580CE3B51D79D42984194D4F178801720892FB6E7
+61FF43C63F9256B5E9F4227B1378222BAAD4D52C77462DF01892220E11129C16
+6C9E45BB9F01ED7C1AD5D8B4D72BE0E12969AFEA90FEF170603CDB91CB243173
+B19A56084D10293B80A35275F41BF78A054DDC98F4A1FFF592463D944960FB31
+6BE5F03960F9B1F213CBCC7FD448657FE388F10104D42B0715FC9571CC60CF23
+C72560CBB8835A0CA208FE06676B3B48B093CB7FB2C0C53AF17EC5B372A9771B
+BFD52FFB7062B4FE0106A01A2A1A1DD4EF5C8C7623EC9324A2CB3B402FCC1FCE
+52BFC8662F8A39D5F1B41C97E7CE34E16AC28A1E94007AEA7D4C519399F1B7A9
+48FA7DDB671067244F09C29F95DD60668223F45BBDA8B1C452E930A9F3F341C5
+351D59EA87462FFB30277D3B24E2104D4AAB873BB2B16DA5B23BEE25BE2C8128
+C4CF2F4F438A4E520CD932BAC455BF8775C27AEA6C73EED3EB2F8DB5E356AE27
+41B35C8AEFE73C4CD6A591AAE4F45762EBD6D3636C03F08C552BBFD0A13D11D5
+491F8369B4BAB8ED9D6F1DE7DB7AFD383986C4338D3AA71C9AF2B8A0955CFD86
+0345F16D9798B25156DDF826A7CB6A0CC4CB43078BEBD3E499DA95562A08EED9
+7CA27B7A0CE3FA7EBDAA87A6042231493B69CFEA7AE1992BCCF33B408CD50C7B
+8665EF678BB3B42D17D703B63DD8B76E3ECA117D3DE545A88A49AD0CB49D1B5A
+C920F89299DEBA6CA0B38D0131DE51CC2652D597386B659849699B7AB3465F93
+256D9CCF0A07891D6F310750D72ACFE1A498583FE2C167B057FB3EA6EC057AF4
+0A96B2399719B8902403E7833B7EB6F117C3F96FB7ADC2490E680A7ABA9BEA84
+49BD1D4307AB6F6401F913F85D9DDF0B414F1FF79BF7C5788A5EFF60CEA9B2F7
+5436EADECABB4C56C2E6BDFB6DB26AF884D590E545ECBDFD7425DB89C6B078AE
+FABE6479DC4F184445D71C052BCCEDD9A85AB9C94431E2BC305564F43C9BD36C
+BEB52E69ED948933A791CA47FC35625B97BF321BA9118DE302F526C7728194D2
+69622B146974D1D5105DC891279E8B0F7BF7CDEED7005AACDF972745E6290F2E
+38EFA0EA88D22A366FDC299883554B30D1A9C0EC5F4BEBBB72BF069E84EE3460
+1FEABC64557779EA9871F92473F25DB12E53EB4E0067D72EF60BF798DD4B96C4
+A051F4877F9040EA9881602F814996038FD60CB721B2BC34DB5115DCCBE09322
+43C75F2148B8CED5D250932CF91235019CF81300388EA235D0B0CD180A8D7ADD
+FE8EEDBAC7828A02B86A150FC415179B8083B3AD2D8C438E19CBC1C110E132CE
+7ABD9C28BD953636EE20FE72F56DE320C2E85690F15B774B89A0046F4EFC2D2D
+160FF39BFF8E71F2E5096F2B9464257703DA24538BB4764A43AD58A1C66CDEBE
+570D84E2EFC62A29274AF650A0843276C6A1EB4257704655EE1D11BBE8B8A0FB
+7AA7E93971EEF39F3EF2DC11B795D10581F89850C8E691C759B266E4DFDA8006
+AFDE411B07E822021E27BB6608F9206F239DA5E3F743D8780428809AE6649CC8
+628853E96C15B56223DC454EAC9750BB27457FBE40C34149848E11320BF06883
+1A1BA06C3A2BCD716E52AC4FA9DF97F82203CEC315A046C215A7A474A7BE6CA2
+D837D6CA08651D247FE2132CF89A530EC5CE24309CD42A63B38BD42B4B9DF104
+64A32984AC2DC3B4285D54815EC8D0702A1CAE2D0B1B42BE55E56EABC131D8C4
+282C11C95997CA005109CF71D8D277A468E23D101DB285250AEC26A778A2F3B4
+007A88C1B7FE3B4CB4DCE318824700203E4555816ADEFFD13A1C4D2A3485D1C2
+681482F0606BA2936543A623E121D140B70AAA459B0060FEABC5DF65C2470511
+07277A4E5B30B31A1DF870B571EC276A8EC80A11794DC5F1D444218690D89DCA
+4501F78C16A3118BF6A557F5A1A30E38A52145A545BC101DF51A51DE563001FC
+DB2773064DE8EB005CEF2D1F822E85B889104DB7DFB72D7E93068495531315EA
+C4983BFA1186ACE72A45E709B846D2D5861AF4745D360B8E39F14E37575041C4
+F72697AF00EC75F46C4021E4B03A6E1FE26326667B08F6FB584F6B7E0133D5FC
+DC9CB6C5B67BB5C7C52D115D189B796C
+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 0 /.notdef put
+readonly def
+/FontBBox{-4 -948 1329 786}readonly def
+/UniqueID 5000816 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052F09F9C8ADE9D907C058B87E9B6964
+7D53359E51216774A4EAA1E2B58EC3176BD1184A633B951372B4198D4E8C5EF4
+A213ACB58AA0A658908035BF2ED8531779838A960DFE2B27EA49C37156989C85
+E21B3ABF72E39A89232CD9F4237FC80C9E64E8425AA3BEF7DED60B122A52922A
+221A37D9A807DD01161779DDE7D5FC1B2109839E5B52DFB7605D7BA557CC35D6
+49F6EB651B83771034BA0C39DB8D426A24543EF4529E2D939125B5157482688E
+9045C2242F4AFA4C489D975C029177CD6497EACD181FF151A45F521A4C4043C2
+1F3E76EF5B3291A941583E27DFC68B9211105827590393ABFB8AA4D1623D1761
+6AC0DF1D3154B0277BE821712BE7B33385E7A4105E8F3370F981B8FE9E3CF3E0
+007B8C9F2D934F24D591C330487DDF179CECEC5258C47E4B32538F948AB00673
+F9D549C971B0822056B339600FC1E3A5E51844CC8A75B857F15E7276260ED115
+C5FD550F53CE5583743B50B0F9B7C4F836DEF1B263DB260571DCE3DFEC496E93
+CB244233CC5176296AB217F6487CB687C770B7AE548E41A54FED956042C976A4
+1D8AEA37F4E2AA260AAB5A2DD1005B8CAA0D43321E35B382D05E781FDCDD7C5C
+A4C8C5A15687AAEAB6387886992D3B232CFB4AF8322D95533D88DF326E5583C7
+5461B2998CE19A734626C656D2A74325ECDEA1EE8D67BC2E51DE040B195CF7C1
+57079434566EA74388C69D5DCDD00AC70DCB4A4727737AE645165783970E40B9
+A10FABB254B14066486084E2858FC4215951A36834BD88E18F8C333DE8AFD753
+EE1D18FC802F0E1A6F314D6846C507E1B859E8E84CD087A152BC7F446402ADE1
+C9E8FF1CFBA9B476EE319FE7B07ADD454B9F112544771BE4EF3DF1F9E05398CA
+63899ACA33397ABCF4B550FCCA3E0FDDB6A6BA7CCDFE160C21650AEEB0965AFC
+349AA45C53CDFCFE139247FD96DC11DCBA8BC7BE1730E0B812C972F522457F20
+5F772F260228B8F8699F99FECC311C63027472CBD7CD1D4960130845995B46BF
+41EF0CD00A3DBD818DB9A8FBE17CD2C67B8B2F2DCCE568603719D010EAEC7907
+83C1C96B15D05078EE4D305BB0EDF9C6017508AB49925D174D573D7E1BB55E58
+D23B2B99D0A38E6D2D28B3A9D5FC1367D3CFD59B93A2871B2970B772EA245541
+31CE9798FFD04D0DE6A654D6F6B63CEA582AA955FCB42EC094DCC8D346CB7DBB
+5EC069A1A35B55AF285E5920B9B2B82C9802DA2A0C41E5218BD4DDF7AC4BF78B
+8220A07CA271EB4F9F06E912E5A7A5DD6B3431E364CFBC5D654DB6F2F121A4DF
+15D1FED5EAD5C998BC9AC56F8CED6BB609FBD1827B67F1F75F177C24FCFBD551
+A362
+0000000000000000000000000000000000000000000000000000000000000000
+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 0 /.notdef put
+readonly def
+/FontBBox{11 -250 1241 750}readonly def
+/UniqueID 5087381 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D03D6A8F05B47AF95EF28A9C561DBDC98C47CF5
+5250011D19E9366EB6FD153D3A100CAA6212E3D5D93990737F8D326D347B7EDC
+4391C9DF440285B8FC159D0E98D4258FC57892DDF0342CA1080743A076089583
+6AD6FB2DC4C13F077F17789476E48402796E685107AF60A63FB0DE0266D55CF1
+8D0AD65B9342CB686E564758C96164FFA711B11C1CE8C726F3C7BB1044BBD283
+9AA4675747DF61E130A55E297CA5F0182A3F12F9085AF2F503481071724077A9
+387E27879A9649AD5F186F33500FAC8F7FA26634BDCE1221EC0ED0E359E5EA5E
+6166526FEB90C30D30099FBDC1BC2F9B62EFEEC48345160804AA98F8D0AA54B7
+A480E715426651865C8E444EDB798C7E11040AF6E5A7ED1888653C6DBF5E6169
+70BCD9C063B63B561EF165BF3AF11F8E519F37C6FDA2827685739DE2C48B5ADE
+EE84F067D704D4511DBFA49E166D543CFD9ECD7417055D8A827F51E087CD2927
+BAFC7E6CFBD70B0FE969F890A11149D3D44D422C3370495DA9951AEE7253A49F
+3A9444C8CD9158D84117299F7F2332FEB0F94E6ED8BC7AA789A3219BC2F227D3
+3B5BC75FB53B55D72AF4A6A7BB613FA235B11BB37D059FD87127CEF73D5B3FBF
+9F91ABAD78BD9240BD9525EBA78095EA0BDB25D1A19E876F292882EAD5619D46
+D20317A345D931F4FF4EAE6216C27044CBA525E3B917CEA25A04C120466C4B93
+FC720E6BA832A06CCA0A3916CEF0968D49085AEBD243C41A448289A6F05CE3F5
+79148DC112A3CC7E8FF810B8C1A09E05F496C0F1EBA334E42E05C376C98F5F69
+C06C71BFC0A2F3AC9951CFBB143C66FB84F9C4ED27DF70869352D61BD5E11508
+0797B87C76B5E6B51A001F9114E9D74CDA9B6803C24E8EC6A520CD68D0AFEF8A
+8EFCC86F74968B12657914526FE1C8B985C8FA5ADCE4A7F4A0977728E72CD181
+36F21C1B58A90A39AF99709302F31C525C51E02EE451508E850E7CC83C774E76
+3BF2BEE485884EF7B64D60A2CB34048A0C303985D01289B8B7AEF537546FB751
+E59F2AF706CC0B1451BE22050359481CCFFF77B3ACE7E374332C3A90408FBBB0
+E95F5B0FE53FECCE71AB47BA6845315649AD9FD28DB6CE342CFB80C010068A67
+C828AB77E7E3C12EC120BCFDA14ADB66E0913E28680BD4E5E5B169A58094B8E3
+9F6412F4F9B8545B3F5502398EBD1537828CE46867C6D346536CA5EB6F831085
+55C836EBACA7459820EF742BC3030E6D228B3D93011CB192F78FB6636AC87A24
+2CC8DBD38488AF51F3BB05AE68EDFAD6C95C758577D505F73E3A62D1F36F79C3
+DBE89941AB91B06F6237F60D408A23D2ED8E6F792090AA753F829AEC0E4AA8C3
+E7D5DE8DB0589DD06C360FE9E59998F68B89311A3A67A4C82E6256D9BE8D0DF5
+A6871AEA069269D3EA2BA981CAC84B4029122C2F439219DBA789530A34670EA1
+F4A7BE23CB7813AE9AE6BDB2E43A5BAFB75976F4E64520FC45406EC4176BC009
+01671311E0F374BDE1BE661C8D7D8951CCE5F5ACB6F4E710E9229E6F385180EE
+2775B63B37A94253D0A9C72CB6CD1F81EC6F988AF7623AB19F4E57C9A97E1DCB
+395A93F9E8EAA653F483649DAD587ED4030E31F6CAAD1776345E3D31C72C535F
+C2CB9A10921C9D0A9835B8C543FCD105810E7ADE67904EEB7A8D1D446A1C82AB
+C29E8D709BF7C860D70DEF69619209FC513482C8FE738BB68DBD8A61B9F5CF0A
+9560FF54E8B13F761A635E1B514D1D617355E2ECB7512BEA699CF64122FC5F3B
+243791E57EADDE4CD5C725404D7FD3E4E1AE86CC58E69FF3AB90C9531096C37B
+46D25AE7010D61DCC6B70BEEF4EF50727BD0143704D4EB138AB88F5CA81A8965
+5AB0054EA8D5A2ACD314C443DF26FB86BF24D70A7E3979A0D708D8E9C32EB3AE
+BC3FEA5E244035BF71AFE68D634BC6D9BAB22382A27A45A2263E49778874DBC7
+4E41B615D234515656166C8EDB3A34D81D559EA93707E1FABD5CD7B80E27407E
+EBCE4F346F726EFB68F34E5210FBEAC4E06A8082CE0AA1B5861A1D1D1895AC91
+025E0D90FB1DB1A49E0E0F571704C570E030ECA683C738FDC8E3E1EECDB33677
+7227A9E9CDF6FC83D0AB79186E7BC8047E88D0DC122627F44483EE672A4259F5
+ADD85ACEAD03B70F65CE7AED9595A1C5217E2B3D2669A3E1DDC4F6E27F6DDABF
+E5103FCE7D56CA120AFB9F0EF57D90EE8C55023FA2557353D8A1C0915C7388C7
+EA1A5B2AB218897E6A35AF771764E84D53672B18103CB950470EDF90F100C721
+B6AC684842655BF6B377FFFBDD632D7CC06C9FB4BDCC80E2FBF9F57B8CB85C3D
+B2A4050FFC4A17B0048CE4291D25B3C80BF802F2A6DB5AC61A2057F11E2E2823
+A489C1413972E764A25225C4279D23515D32A55ADF86612AA2E5491C6B08C6
+0000000000000000000000000000000000000000000000000000000000000000
+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 0 /.notdef put
+readonly def
+/FontBBox{-20 -250 1193 750}readonly def
+/UniqueID 5000789 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF4E9D2405B169CD5365D6ECED5D768D66D6C
+68618B8C482B341F8CA38E9BB9BAFCFAAD9C2F3FD033B62690986ED43D9C9361
+3645B82392D5CAE11A7CB49D7E2E82DCD485CBA17D1AFFF95F4224CF7ECEE45C
+BFB7C8C77C22A01C345078D28D3ECBF804CDC2FE5025FA0D05CCC5EFC0C4F87E
+CBED13DDDF8F34E404F471C6DD2E43331D73E89BBC71E7BF889F6293793FEF5A
+C9DD3792F032E37A364C70914843F7AA314413D022AE3238730B420A7E9D0CF5
+D0E24F501451F9CDECE10AF7E14FF15C4F12F3FCA47DD9CD3C7AEA8D1551017D
+23131C09ED104C052054520268A4FA3C6338BA6CF14C3DE3BAF2EA35296EE3D8
+D6496277E11DFF6076FE64C8A8C3419FA774473D63223FFA41CBAE609C3D976B
+93DFB4079ADC7C4EF07303F93808DDA9F651F61BCCF79555059A44CBAF84A711
+6D98083CEF58230D54AD486C74C4A257FC703ACF918219D0A597A5F680B606E4
+EF94ADF8BF91A5096A806DB64EC96636A98397D22A74932EB7346A9C4B5EE953
+CB3C80AA634BFC28AA938C704BDA8DC4D13551CCFE2B2784BE8BF54502EBA9AF
+D49B79237B9C56310550BC30E9108BB06EAC755D6AA4E688EFE2A0AAB17F20FE
+00CD0BFF1B9CB6BDA0FA3A29A3117388B6686657A150CE6421FD5D420F4F7FB5
+B0DAA1BA19D638676E9CF159AC7325EF17B9F74E082BEF75E10A31C7011C0FFA
+99B797CE549B5C45238DD0FADD6B99D233AC69282DF0D91EA2DBD08CE0083904
+A6D968D5AE3BD159D01BDFF42D16111BC0A517C66B43972080D9DD4F3B9AE7FB
+11B035CE715C1218B2D779761D8D7E9DEBE277531BD58F313EBD27E33BEF9DC5
+50C7821A8BBC3B9FDF899D7EAA0B94493B97AFEAC503EB5ED7A7AB62FBEFBEFA
+AD08ADD6648E354B05E584C970AACB6E015F56F3482678504279F10E01CDA446
+6F6E039511BE0576EC00ACDBBDE8DAD593121F1D855D175D616711DB93645AB6
+03227673BAF4CCF593EB7EF23AFEB7B2CB7AB053C26530E564BC01C2F74A3D1A
+4C1279D91F1EF1383E34E8218B3A9918341FE9E346EB3EBFE627E7518D6AB787
+8DC227C39FAF27D969769756B98A0A078058A6D04DBC5A18344ED6EC71C98E00
+7996DD675B953CBDA5AAECA2AB8819E1FFFE61F2A2C48F13B70F8334E9A2CDAB
+4603C68292AD8A93E29A04D262AFE205293FC10AD70F360DC07DC03E00CD12A1
+5213BFEA869BE17CB485E81907A09B54E11FDBD10D0F96A93331A1C67F567735
+D8103CDD80521A2ECBDFD5300A52A807A0CDFC1682A8152B969A4BB552C7F749
+B1CF7DF736594BF66F8BDA9AD556CDE9B1A7DB9D925BE2A6126ECFC9B5840DA9
+8DC33CE9E081A081726EA759DB1CD8D17674EC2EBD19ECB3EDA97CE76919BBF6
+FD8D8EDEA1CC16A552224B6B237450E93F953A3E5A831F9D56FF8A56EE011489
+F11F4770859118A61EB7F9D552E3AF26BF4C061D136A25969FA4F98524DE2B29
+37048AF15112C33E0A209BBE2BB660CF3690CB633122473A607E181FCAB6684A
+4EEF2C8B40869EBA5F3A9DC7384DA16DF2EE07986F86C00F0780117B1A6CEBA0
+0345A6B9F709B463B1AADEDB5A76FCD5D08EDECBB366C248EDFB02EF6EF4D5FE
+8EC92B011698C86A817B703C9E08511938C5E824C2A78E8F20D59F6F38471D6D
+C886487193C3742947E4DF31A01F89512E74344EA515FE8F3F0B5B2D0EAEAEB5
+79EE7DD825B87AB1BE8AF331DF5F087A8E61D9A606D684CA34BDD91BA76FE680
+C29D2AB77B8A513C259CEFB5CBA26071E55E0C5B17DB0A9E4476F79C84F2676F
+79B8C7AB8CCC4EBABFAB37BC961DC469FF21F3B212506E0CE29947219C2D2375
+1FEC3209D22EF95488F584655D7389E90204DD39E0C5938EB49107B49EA6343B
+D6ADC868BF393850973F912AA46ED3D4E11362DFAF19AC312F216FD5960FBA17
+8C5F89F51AE749F73150F077DADBA20BEE3A1B7DA363F34CF6F477B411EB5129
+1DBA86E8D7AFD4916A10FE2113FF2177E31F9456DB39984F2A123D51D26B36E3
+BE132D008D884BA2A44FF04031BA21A47AD823B57BE67B44116E7A7003EBC2E6
+97F17FFE8B65AD781E0AD24885FA94CADF751D2DB2BCC33DA1CC7C49843EBA
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: EUSM10
+%!PS-AdobeFont-1.1: EUSM10 2.1
+%%CreationDate: 1992 Nov 20 17:36:44
+% Euler fonts were designed by Hermann Zapf.
+% 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 (EUSM10) readonly def
+/FamilyName (Euler) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle 0 def
+/isFixedPitch false def
+end readonly def
+/FontName /EUSM10 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 85 /U put
+readonly def
+/FontBBox{-8 -194 963 738}readonly def
+/UniqueID 5031988 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF4E9D2405B169DE483FC3C5DBA587E58D683
+9E11948C1A3E8B5360ADE57410E9F910E561F679EBFBEB0730125606CD072EB1
+0A975FC5186A70BA3CAA77642D606B57469AD72289DA911758623D8C66DCB000
+A804D330FA2310AB619846B77495D97E7065EFF7E9863C2B64CF640C06301DE3
+EB4EE29B52BEA3204AF4CCEA9C444F88AD777F296CE544EE4A1E88E093DB16EE
+35D8C6473B422A716796536C1347DEF0115B8E31BBDD72D574FF3800E34CFA96
+9322738BF68A0EC2D78FB0CAB62512B656837C6184CA2E03BB30EC2396428702
+0EFE3DD99016DFF0BB433068C23FF079C1EA255E1510957CA6FA2B6307274252
+0E255FEBD892EC4434A09579FBE9BF52000882448823D29A95205EBAC8E4CDAB
+252C9C21441B4E3616AE348DE997113556FD7D4A14256BE54BA2EFF1F19A01A6
+4824F691EBBFA99BF593940B4A376C6FF81E89B0E672EEF95D12E8E962873673
+A80D036D1687EA75A40ECCD264BB3228450B425867FE2CE54BDA69C86A2210D8
+3B62AF190E473AE7F045B7DCC33B892F8009272D22B951B29B74690CC2FB8740
+0C1A7B0327C30AD0CF027A06E8CBE8CBE7C2224BA8E0FE349B5BF30033D027C2
+6EF80F00AC06A2AC1CE7C237ECB5F3616EE0C1CDD06A47F952F8F455172C4191
+59CC7C7D9378CB4CF4F7FA17D9FDBC94024D0323F115206DDBDEE9C97DB4A29F
+174BBD203ED0FDA1BFF0AFF74781F045E3D3F23AC926F472734F1E3C858FA60C
+1C5CC4DECF246DFCBBD6E7FDA25AA240952DA6118ECF0EA4D3F589A81E541018
+43E874DBE23F815D1728DE8699A74BE80B2F63553EA85A982FFAA0679AC682ED
+A87E0015ABA520C142F0F21214AA560C249278AD062087755F734029D3A9342C
+A45F4894A7AB991A660F4597ECEDF4D95E9866BD9ED2BA1BFCB7B4F60BBBA440
+C91605E3B358E1F81250DEF10D877374E9F61F161231B704F67DB722D50172F6
+A1B5C0E10718A5F8A8F7A0FF8C63CB49460AB402064F1C141E2BD409534AA315
+9801843C2F833633B222AE74F93A1A74B27935E9E133ED7E1C7FD93187C3E4E7
+B26DC68E12F835BCF5D24EAE7FFB4BE275AB5940B39D829C85F21B68C8B8757E
+889BB1AA153FDFDA189EE0840B591F2A9C9765641B560C177200EF108B7FCF5A
+6B3A42E3A61E22FE4771CD6BEB64E15640A52905DA14F4FAA6BD086244DA9593
+757DB7A4770E04BAE20A030E328A3DE2F6C9726F857373A0AA578F83FDFC0D41
+FD20887CA66B59C0F001AE122EAE21C92F5D54844EC4B33E0D09C248FD0A3E58
+9C865B9F37211292FF2BDC19E178469FDA43175560D02E69EF7C0F2E9B51E77A
+B90DD9B214A2B640B6F24AAEDB415F21BD2B81CF9996DE91042230BC20639EF6
+F2E0ED12EFEB6991C484992C5AFF42FC2F4CB67406DB8277EE0D743CAEEDECB6
+C1F662FFDFB8147D00131C20709598FAF2F16D656BAE8321B1F10B852738F436
+00FA43B92FADF4C987484716026CA895FF550062E81CB64C1EC9F97E7626D454
+CFD1F657E4802AD5834851315AF4C97A5C
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMMIB10
+%!PS-AdobeFont-1.1: CMMIB10 1.100
+%%CreationDate: 1996 Jul 23 07:54:00
+% 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 (CMMIB10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Bold) readonly def
+/ItalicAngle -14.04 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMMIB10 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 0 /.notdef put
+readonly def
+/FontBBox{-15 -250 1216 750}readonly def
+/UniqueID 5087392 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D004B836D34E88C25F6CE738846C8E2E59A2BCF
+4ACF80A26D78872E9343A0537BC3BD7715F32ACD958D5AAED865BFE129278935
+063A31C2634DE2F9077E0AAAAEB224466B779096D8E3FF0A12AD5157F6603DED
+1A82F3511359143311179080C510740B401C930C96270FD1AB3ECBCFEF5DE53F
+E846BAAE95828D5790922640EF8AB9D7CEBE7669FEA02B72F86872D3D8754A18
+A1629C40A7C00C956F140BC63362478279C36EE353638CD3E249897207A94504
+4400668C8E702058EBF7284C9BDF830A3FC79C7EE900CC4C3664F9767A237275
+CEE3671644A75F1E696DA906B4C66870DBE87F5B4A176920C078ADBE24F55C09
+3D18CDE21B5FBC1C6A8AB18E05EDBEF9D1C1C18B3E6377BA2A688579D4F708F9
+A5CF4F56C5E39E2726106E9713E638775E606464CD674E5DC25CE9A696A65806
+C8E9D206B421E246F18013ACC6C7B2985BA93B1B7D7745CCB25B09957F50128C
+B523A55ACA6A7A2A0193A536E590291ED9D577B527CAD0372E05BFCA1829FED1
+662D06144A5FFA628C587A4FA05B179F1A7E3B23B47765FDC054271A0DBF9C2B
+B4F4771F80D1F7AAD9024868C30DAD5CF728DB2A71D86D53B0E674996E8C01F7
+EF97B225A28872F8AD4752A466E078C2B020EB832F237CB9B5631EB2D2EDDB00
+709D3864CA3A6C3EF18085EAEABC011E9F35C9BE4B5D0B608361F329B5784DAC
+5557A602E9E3C204909D84DB988F0BAB914E87CD685C7DA55C5E0B9F0176184F
+FC39B570873BBF346A0D1DE3942DA05434949A65CE64D8BAB0A091C40F7FF47A
+4FC57CB4420221C7B3EB8B891044B5FB0227009F0F6028D3F28545E63C975F11
+6BCCB67C58A544AE706BA5309FE383C7789E0B88EA7691367C6FF15B28F3141F
+FE3736254383DA189E3F743B227B6A8092E37EC73FC92B713223F27F10156473
+E6413C306F5527B527904FBF5815EFE89BAE322ED6F5BF2631913A7B8F4DE1D4
+F5D5F2A026BCF9D84786CBF51F3C06E3C5387C4FA959D898968D1F9D1E4D314B
+C04991E9C14CDE77A3B93A517FCE4AFBD3C624CB348B0DDDF38E78A062801E91
+FDCE854591F1003EB5B6B12FE4D3B16130DC0B0EC7E4E39684457EBEB6CFFC47
+0A07F8E2901A074E53B32927C173E767AAE138A4B532F9182A0D60062BE39273
+4B9ECE65F4590FE8A177AB0DD5B0CCF3E3B1D62F36E5CF9C11CC4B07EC946DAF
+944420545A10E5631A0526DFEB781E0EE11CBC73610BFBF0FCC963A5D3EFACC0
+AAF0B187DCF3D55099AD9754CE10EF7AA265EB5067CDFC8E879647CEA93EC3E0
+E07D257B4295B0ED2254208F8A9B879708A8F95129D9953C1CEF8A0D45
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMTI10
+%!PS-AdobeFont-1.1: CMTI10 1.00B
+%%CreationDate: 1992 Feb 19 19:56:16
+% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
+11 dict begin
+/FontInfo 7 dict dup begin
+/version (1.00B) readonly def
+/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
+/FullName (CMTI10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle -14.04 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMTI10 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 0 /.notdef put
+readonly def
+/FontBBox{-163 -250 1146 969}readonly def
+/UniqueID 5000828 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D03D5993DFC0930297866E1CD0A319B6B1FD958
+9E3948FFB0B4E70F212EC976D65099D84E0D37A7A771C3101D6AD26A0513378F
+21EC3643079EECE0C9AB54B4772E5DCA82D0D4ACC7F42FB493AA04A3BF4A1BD6
+06ECE186315DBE9CFDCB1A0303E8D3E83027CD3AFA8F0BD466A8E8CA0E7164CF
+55B332FAD43482748DD4A1CB3F40CB1F5E67192B8216A0D8FE30F9F05BF016F5
+B5CC130A4B0796EE065495422FBA55BEE9BFD99D04464D987AC4D237C208FA86
+0B112E55CE7B3782A34BC22E3DE31755D9AFF19E490C8E43B85E17ECE87FA8B9
+1485831624D24F37C39BF9972D74E6EC4784727AC00B9C4A3AD3DA1C22BD6961
+7E0ADAF55422F22ACA5E4DCD4DF9FCD187A566B7FB661D0530454D0DD6C6C50A
+7A3875C6CBF8EC7769F32A1F3F7FC1C072BADEC97794D4E90E0035282A170402
+356E5A9CD9ABD80AC4342A5283E458A7269252F4541CBB6452B39ED54D336D0B
+19928E9CD1AB26AD83EB209E2EC75011A2643813053B5DBB0246097C4821B5F2
+C92554E9140BE35B2DBFCD98809A8EC9FC910FDE9E0D86457C70ACB056EBF90F
+244DC0A5BBD455E15D6E3180311D52CF50B0BF7D0A7F64F3A1821E0AEDBC2E7B
+AEB549FE1D51088C153799C6E089B5D5D65E1C4E2D2B430CDF1FFA23CCB25D95
+592943209E846E55B4CB54F6658CBA3C0B29796D69D0435D5431ABECF3448C15
+98CA2F36F3659E29AEB79355EC2ADF835CF0886C21B766B9DEBC3950B5B3B496
+2E06D980A8C60305B273232D4604F12632FB4F1B2F9703952C823C098543AED1
+CFB4ECF259A11985F0C944A57B5AFD853374FCF12305601200C2A393E2FC77FD
+F78C2BEAF545FB34D38EED3579B16A9724302E591A2F873A766ADD5572E32935
+1FC9784C66AF7E2E13D92D32C0C48F82D45BFACCC2BCAFA04B854CEC608FBCA9
+643028338C2CD69ACB53C7E6AC418FC6F1635EA2702668C3B98AA439066990B4
+CE74A7C22B992987625D2EE48EDB70865F9E4A1270E31FAC2D3E53A8439B7B53
+8B9EEA7FE7A10F17C39BB7B59B378048D647E7262FF36D2F50AD1585CF7A329D
+ACB0AA69E4B9237A38B34F678D425DDC223513DBB24647F26DC9AE1F4D32A65D
+D8BF0FB2E22E0F124BD5E2BCDD118AB22573E6F6F222B66E06B1A978C9351185
+5D92E88E5EFD2A2F0C7409E7EFAFCD2068698E376056B34D26EF1B9B3F8F4939
+8B78507CDB2784F367825A6FF7966F91A2484D81D88EB29C28C177F614C9CBE1
+9742FEFAE2EF8986837A44A7DA8ECD775F92A67B241F371B0C510DDB9FC04380
+40699EE66D74DB29E2A42542325DE900C6C3D5F2AC32856B2E8750F907434487
+CAA85339DFC689CF96323EEF28D4EB5EE0F89503D5E7BBE1BE59E2731F7CC5FB
+EA68224C8049A7657367AF2D50CA834854EFAA771B90D92397A137D8F54C31DA
+6F2911B0806BC00DF1123C0C48095FD5B31BF2BD2DDD55A6AC84C7C7C406D5B4
+6D0BCD5FE8E3EE46CA12B533AD713F8D50C53D95B489599443D342DA57291846
+EEF677D1B88EF3EE1BA20C0BD3445E9F1384AD7C487CA413C486C95FED100038
+814926A1BC8B9CD3C6073F80096255FACDEA9C2F50B38E9FEE14053A2788D98B
+C49EF4098B
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMR10
+%!PS-AdobeFont-1.1: CMR10 1.00B
+%%CreationDate: 1992 Feb 19 19:54:52
+% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
+11 dict begin
+/FontInfo 7 dict dup begin
+/version (1.00B) readonly def
+/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
+/FullName (CMR10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle 0 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMR10 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 0 /.notdef put
+readonly def
+/FontBBox{-251 -250 1009 969}readonly def
+/UniqueID 5000793 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF7158F1163BC1F3352E22A1452E73FECA8A4
+87100FB1FFC4C8AF409B2067537220E605DA0852CA49839E1386AF9D7A1A455F
+D1F017CE45884D76EF2CB9BC5821FD25365DDEA6E45F332B5F68A44AD8A530F0
+92A36FAC8D27F9087AFEEA2096F839A2BC4B937F24E080EF7C0F9374A18D565C
+295A05210DB96A23175AC59A9BD0147A310EF49C551A417E0A22703F94FF7B75
+409A5D417DA6730A69E310FA6A4229FC7E4F620B0FC4C63C50E99E179EB51E4C
+4BC45217722F1E8E40F1E1428E792EAFE05C5A50D38C52114DFCD24D54027CBF
+2512DD116F0463DE4052A7AD53B641A27E81E481947884CE35661B49153FA19E
+0A2A860C7B61558671303DE6AE06A80E4E450E17067676E6BBB42A9A24ACBC3E
+B0CA7B7A3BFEA84FED39CCFB6D545BB2BCC49E5E16976407AB9D94556CD4F008
+24EF579B6800B6DC3AAF840B3FC6822872368E3B4274DD06CA36AF8F6346C11B
+43C772CC242F3B212C4BD7018D71A1A74C9A94ED0093A5FB6557F4E0751047AF
+D72098ECA301B8AE68110F983796E581F106144951DF5B750432A230FDA3B575
+5A38B5E7972AABC12306A01A99FCF8189D71B8DBF49550BAEA9CF1B97CBFC7CC
+96498ECC938B1A1710B670657DE923A659DB8757147B140A48067328E7E3F9C3
+7D1888B284904301450CE0BC15EEEA00E48CCD6388F3FC3A307A1A4455646E1F
+B73047F2C36AB0DE64B2660795139BF3E491E7329C6804B62E7607010C423008
+37F93F3851E472BD8F6D8DD9DFC1577C04A9531F8BEF1154F701F3FE438FD25D
+0B9C4EB4E2132F3094A166D7D2D3D877B0447947E929A44135D567DFC946AC97
+E6EDD4B4ED1204D5D2402C351212FC96C0EACEF4222FD9A271D0C3DB27CEA5D1
+81D7E6BEDC9CE8937288DBC4F96661C812D31312FDBD0E680C6EEF544CCF26CD
+B6FF82F3378CB493CA2527D51AA87E3A80D75383603152011530B8ABD2294D0E
+60C55FDEB881EF5212133C9FD7C7BBC8675CE57FAE119EF37404C17535023E73
+FD32A1EF0C659555FBFE6D9582A7DD2A35828301408B25709FA7CC2A6258B759
+24199D0B7A5547947DC1BB9D24782DA5E78890DA1B1B82FD8797EEEB29846B6D
+B84DEFD0C7F2FDA61D1056E8D567AFE76F964B9F717BD287A94CF9CFEA668F9B
+6380CAE5E75DF3B8B749E7052BFF9D02D39133E848BA45FBD6DA80780B7F07B1
+5BA28781C1F3745C86A375AEA1A0FDAA7D46CF5216ADCB21A84A41FC78B26E3D
+7A2C72B64493E5043685A5A66E5C53F98E62AE673E6D288CD2A3095CF4B93D27
+73F4C10B37F259AD713BC6DA67172869F9E4359B75F27C0256E0ECC8F09A260B
+FCE7C040D10575084BC081821FB75A61F7E12B131FEDC0B252830ECD89C6A6CC
+70C7533AA07A8F4604FEAAF4A452A9017E5CEE18ED0DCBCF97785515A78B7A54
+EABB57F4F771336D985C0C360B2F5A5877FD00B5C99F6CEA624002296A1EFE64
+6E370390407746B9D3B77BF54FD908650F54EDCFF07F1FDC473ABE8F145314B4
+872D158E88EE3F59090FB93A5747EDA57889E69C9114F441822FB770581BB0AB
+68253F075B952533D5C2B1F3E25CE13CE1961D18A72345C9DE917268939A7077
+AB632E33BB67BFBB8E4B55321E0BDD5848E516BE23C494EDDD23ECCC58122E19
+4DF11B8E0EC7921DC518CB144363F76E6B7E8D6F7A316B014ECC0696385F8BAA
+52A45CED5C4C0FDDB01868A95850B5AEA735AE3FD63C0CAAF6FE863A7C454C81
+7414BF1C8774A26A892F24A373DD2657FF1A7DD924FD08F60B0FA332A0AA735A
+E42F2542E0CDFC20F1E2239F8FF4DCF88C10217175E46CA827D283DE125EDED4
+52C353D573BD1EE7F076D261720A124DE83795976CFEBD2414C2FF8AB761EB63
+387D6D99128B5BF1895E61814A6628ECB6DCE1428637CB1D228FCDC455D924F3
+73590C61BCBFC65AACC5E7900CEBC70897095A960D5717B028EB7AD8CB7BC8E5
+DB996EEA5205602BEE0D53586DE1736A7A3093307D333437CD724D8064C2A1DE
+FEC7F02EF5299036BAF5A17F6BC6B2385E430DCF30B7C5AB395AF40C8EEFE4E8
+E8F73ECC88AB829C95C96A345D59F7011DDF9AFD77DB33FFCDC03C052C7459EC
+4065D48574216CBA90BFA6B73F0963A217F054F71846C7D6D22414BAACB511EC
+B3D5CEDDC599F610FACE80DFDBAF04FBE72318D619B5D693E52FC6317801C2CB
+2363CA7FDC7D3864DA5CC01055B4B681BCF63FBEF43758DE58200AC5B41003A3
+1AA4D3E2C0A056D20F7B44ED1B5CF7592206656527380E5484D9D71D6813182D
+D478603F75DA560D36A7279BF65B8462E3B7DD2217214CD5A1A565C371D3C7AA
+A9243D5F281FD5431E1A821A8D63B8EC801F8EC1AE9290C745CB9058C33027EC
+573C7FDFE8DD8BDE8805046B418D2E21A3CC3B56CEF1F919DA1DD6B6169E52F2
+4D5B92EAA2EA854874A6AED31D5BC88A5701FEB114367C2340F268EB2ED69BF1
+4F792BFBA92520C562C8DCAC2807EB726FED9CEADC14C7A202191F0EA735380F
+42F04E6B578D80AD35E8CABC1E8FF7033AAFC3A3FC7CD3151E566767AB19B745
+69649A18223524CABCFC0282C3300AB7B48F46DD82C4A60778379DB573AC0431
+D5F6F5EE98AE202933787B589790A31AF611C1DF720B6ADA6BC239A59134DC30
+547DA55366EF694471338DCF77717D5949BB0E8789D77D7E36182DF642C924A2
+07E925EA4109B6C6C09355228BF9BE741F7310A292D1C75A10E6C1807196880A
+B5486540163F88FD975FB054F371C01FDF83E2D59D897138D6403B1612C72B5C
+7B90042F423F6B2AC1587C449090E08F25385A79AC4435D13C12728B2877F58B
+2C390DED557EFE21AF7272D00DF6696C5E0BB7FBDF0518BD6D587166FA6F7088
+2606217866FF5187418D1824F372E14D1FDE095B4DE8903025B7964692E22371
+809B3B7EADFC13D6912434F14CC610CD435E94700EA6F7952ED64C7B1A71E184
+6FC45B08F69F737E4DE1002F873E2DEE7947E509D481EFD17B408F9E45B50B14
+338D2DD9208E6351DA8F53CE8D110FD6FF7C318DBD670AD5C20D8F65625BF89B
+3E4C3CAFE69549A45677793FC568F45C5727D859AEF9E31631DC03BA171D30F6
+487E65AFAB16F109593A68608EDB5FB7BF5B8B22F27910BEBE72F04B19E224E9
+BA9EEEE0EC8582EA87E366A78C90BCD345C68988D09816794AE90789E6DB942B
+A05A589C1698F7E403974494979C390C989FF8D76BB359BAD5210F1E829848D7
+EA27770CCCA5FFC909EE182ACD1B89DEA1C40B69580B1D21C7EA77B1340C7E0A
+EA4845C2A319AD5D1BCBCB2F6CB1CB79B3E902FA8554836986BC812ABE2A8E7B
+EADBF231192077F29889857052D491A53B1015204D52F3D8E9623C803D52AC7B
+3587E1D076BDDC8A5B7D4F385A05F34E70FF84478109C98AF835B3BE03B989B9
+159C601F2286FAAFA3C6D7462EB895964C544D37BEB8350469DD83C0CF3E87CE
+8AA8958623553311DAB5DA3DBA0121467908538545D9F031AFB5C02A6A97CDBF
+63B48B09BE325781786A936547A04C2936D26F48F029259E19B8F57C6866F153
+8B4E9571E7
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMSY10
+%!PS-AdobeFont-1.1: CMSY10 1.0
+%%CreationDate: 1991 Aug 15 07:20:57
+% 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 (CMSY10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle -14.035 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMSY10 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 0 /.notdef put
+readonly def
+/FontBBox{-29 -960 1116 775}readonly def
+/UniqueID 5000820 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052F09F9C8ADE9D907C058B87E9B6964
+7D53359E51216774A4EAA1E2B58EC3176BD1184A633B951372B4198D4E8C5EF4
+A213ACB58AA0A658908035BF2ED8531779838A960DFE2B27EA49C37156989C85
+E21B3ABF72E39A89232CD9F4237FC80C9E64E8425AA3BEF7DED60B122A52922A
+221A37D9A807DD01161779DDE7D31FF2B87F97C73D63EECDDA4C49501773468A
+27D1663E0B62F461F6E40A5D6676D1D12B51E641C1D4E8E2771864FC104F8CBF
+5B78EC1D88228725F1C453A678F58A7E1B7BD7CA700717D288EB8DA1F57C4F09
+0ABF1D42C5DDD0C384C7E22F8F8047BE1D4C1CC8E33368FB1AC82B4E96146730
+DE3302B2E6B819CB6AE455B1AF3187FFE8071AA57EF8A6616B9CB7941D44EC7A
+71A7BB3DF755178D7D2E4BB69859EFA4BBC30BD6BB1531133FD4D9438FF99F09
+4ECC068A324D75B5F696B8688EEB2F17E5ED34CCD6D047A4E3806D000C199D7C
+515DB70A8D4F6146FE068DC1E5DE8BC570370531B9EAA95D650E5C23803BF5C2
+75082A9BF73D0F7781F7541D64AFD7C99AF409D4E2AC42AAD5E7E86A6728D134
+0A9E2596DB05C582C3CD6360A7224419C99D6D63D5A03767AA63C1FB127F2279
+6954A5C1EF3C8E03DF1993DFCB40FF5CECCBC498F584A60B549F9AEE28C7EF4B
+8CA07284608D9E404A1E020845EEF5D393502C680F135D7CD7F1799B288D6585
+A77116CBB7195E4878BA2A6FB4357ABDB89FE940218335675846DB18868A9658
+F864B1C438AE6915FED20CC3B3D962688F2164C457EC13BF26FC61A44616F51D
+DAB0FBBF3C98676F7304B5455CFD65C6ED21E582AE8B78F1C125E2ECC603264C
+8E9FB3B50411FD6841462547085A6D3AC4CBBAD10256CCBA55C42F96A5E87169
+CDBB40425B321CA17301122CA7AE90856685EDE8D4704BB2BEAC91F21DAC869A
+B64D1FE0B556727B875474C6C5BE660E82BDA7BBE25DEDF6E59F50D001A04BBD
+31949F48B9A908E4F74FF36D64DCDCD30940FAE4063AB32FAE56BCC8B996E0F2
+5ACF051FF643C3058CC411CC3FD22DD50ED43678CCD2B9C06E1BEE5E686DF3F0
+E118322E30040BFC8AABC8C28B02FD1B7066DD7276BE23309C2AB3BC9BB2143E
+0F975F4985AA2299853DA5091AAD2D020FA8B10395055F0B865BD998A59F62BD
+DB8F6B8331B3C7FD5DFD51F2203CE018C9A5E227BF58FD9DE9CED50617EEAACB
+6D2F782763D019E192E75A4B5F3926A27E25076AEF787623542B6DFEB7FD03BC
+C8DEBB242D43EAF7F2BE14DB97A19CEA9750F83A0242AFE2F7BD980F4F84B1B9
+6825964503D6DD07FDE3A5C9B164446939A3AEFB5BEFCA4378452E96F54697ED
+3397D66C7C1335DDCE8A506A6AE2B814233A86D07E525144A16F893AE18B9109
+ECD1DBD22E41EBD0CBA00D36E7DA95C9B618442D334EE34DFEE461153FC05B6C
+65345748836A807FD360ABD721F7B17F8F84E9107FC7AB79E2485C692EFD19E1
+E45AB951588BC4065FBAB9FF0DB5CA542A2169A6974281C5D124A825820ADE16
+FB335B7FF5DFE9C41CC767CD93A1664DFA6B4C2984CF604B5DDC02C74C8C5752
+36596E8E546DB447A98BDCA476E85E92DCD9B80C85EB341547BC13E31C968AD2
+93663F93CAEBECB1F4F1D75D27873620E79A6A135C97219637E197E5E6AAC9F1
+47C833C5DB688992A57D5FB1C34B4C121ECF9EF5156E1BE4A1BC20C66B4E0072
+FC928AA1CEF326B749F5D53E195083500DBDE786A988FAF8416F4B6C3BAE0F9D
+52D917726A9603A7F40F739B95086B4DF288FB06675AAA2EA3170498C04038DC
+6E978A18F766A4F788D0B5B5458FFCC73AE7FED1E1C04D7E661D44089DAB6987
+66B6561C345F4984F3F157E1CCED47ED5631413B76474AC227F51E070DB3DBC6
+9491979366EC67BF5B58AAE64536791DDF799ACB35B5150120BCE6039B9E2B3C
+10A519113758AB57234ED0A65C93F5FE31F43D9A68BF854B74119ED124AD73BD
+16867EDC7EC72D90F6A9D5F92C9DF39C4E656D10CEEDECF90BCE0A2FC0D0F061
+A4BEB82CB5B80BDE0B88299FFC951C81B0A911A640CDD5DE95F49DA12DF4F450
+4DFC6BABED6B763AAFE899AE391C95EA11BB8651D4CC8F726DDF691D68A29778
+C2E0126C7D43045437AA04551D3678D3B0B489C4F011DDCE04EA99EE049F1841
+7FAED59BEF7472AA50F1A907689D55EAA79CD318D0CB64B5DE013AD23FCCD985
+DA0D9FEDFA0E5FE24A74FD4004022BD6336B4504DC748170F9ECCFD748AD0DEB
+B1927F19412A1CE3839CE17CFDD9D5B75A00508B3F834E89D661CA49CD0B2644
+E46BAD173B124969FDA2D2997DF28EA53FCDFA8EF90A58CA8A1189EFFDE1C05D
+B309373DD98A0E6AE0A78D9010A9213DD5C00445F9F3AD02D7D27772D5B3FE74
+51915B6E3A2D62F00A30CC12D63309D7D1AAAC1B3F63CC1B4A7F9F86A9F7F501
+DC057E9F9D496338FFAD12FDF008AAF5F7468F5732CACC68551CFB79A81354E2
+CFA6DADB0D8C1B55A3E76177F22876396240FC028BCC244CD579C3EDD4CC1827
+BD83D19CE2D601AC195FEF8DC83EA390CA8994A1B023A3DD687FA1F34309A5AF
+BD413E7FBA11C14EAE026361A73B15A73C7D996F4277EFB81CBF47B45E42DE92
+5BB41B1831CB5244B8FDA84ADC0A385B5B0AF56071DC5328B4EBF8AD10794635
+62A25213ED5712720ADC12AC9C9949ACCC70B0BF203F5B20CFB472695165737B
+CE476E46A90CFEF299D9CACEB0C201C6275647B1C13122AEA330474907F82898
+2A973D904FCA51DC76CE776D22E29AA16A34D5F5726B48FFCE740F6DD32005C8
+18FEA3FCD32A40A21F6D10ADCFAD8D71680FD430CBB25DA32A9F8190ED1DB32D
+A860D10063E80F6C35E0AB67BDB6A48D9BCDE145B0393265DE3264A7E6728B0B
+30B4AA7A7187786697D14D3A872D0018C1A6E4C507129E80751000025B55EF8A
+FC82E6D7178B4D88AAD67BBD5805B78F07238C092D4F66984D9E2B30DFFDF499
+323AC5F0D66C45EC56A22DD28C8575E2E0A79AD1D999215D31402143643B45F0
+DE3522A9B77D1C7A3A0952939FF49E850FC5732D8F3DE9FE2C5894180EA559C8
+9D765620005B53AAC916C622752BB5B89B235B4AFF5847D79DA903DC63A61CB4
+9CC15243CD91A5F630FADAA2B18DD53D47C46D03F6DF6D9F9714F5A6276F4AC5
+881889E6978BDE4B668966F50268A8AE3544E5EC0C180134660265ED3ED0056D
+CC6C8C5B1131F09B123ED3A30F4F66B9D543281F1DA102F1AF10102BE7DBF169
+0529DDAB30A9D3F3229BEF61635DF44920A81AE0F51A94CDE613D71F0A390A5B
+6BCBB27A0B5D903B5B36EDE600DD337D7F8D033556BA91EA4B9E219CA06951DC
+B22D5D48CBB3350DEC250430F610379866ADABD259ECE2690E013B586381A2E8
+37B7F0329E49B4C77634B145AC536D31E4DC898D004A444B91BE1773F27511A6
+EA602CE23131009BB2EF282AE6192FA96EC332F8C367B7774B67E656B7D9369B
+30C103179C5C4014517FA61E67BA37EF2045F7DB75B8BC35574727D7D2B0AE35
+2201033C51849454205F5F03FB4F7013E501E939A6575E0A88AD8F5AF850046A
+00002BF6BEB09160F4810BCF3AFFF7DF9E7489C0BAAC3DC8382562A649138FBD
+8C03984BDCB50246A5CA5ABBA1DA3B52E17D8646BCA357CF8F953D6E544D4086
+1922FF9B526BC4CD28CB4400992CD2D171B3A1E530F27886B70D1A09C8AB416B
+ECB957171D0422C8A3E22845328E15D8B4DEBFF2EACAEC2B7B87A246EC52E3FC
+AA9395C7969ACB53C0208F761AC7952AC750403FFE9AD505562B52C37035BF02
+4B69679BEC2D99FCE838148AC18A9C7B66614F07219C27C798709B7EB547E458
+409F298F8D785145030DE7B04D0B0CF063D8A52EAA8445988A3D56A84755F473
+7CC7137451275A5277A67B0319493CDC9FC3BF8A044B49120A4DB24427FDAFA9
+605D2F694C14252181C916322CFF65E89F1204B763BFD71429AF8453420757B5
+9628B1E88FF6DC865F4A2C8F2A5B56C546FDC866418A8C4E2F553E280BDCC690
+770CE5C795674A0BF1D91322F9ED93865C156FC0F26FB37DB11A19BF933C5D4C
+E06230764344BEA35CA0C373D6BD355B3BB33431E870558DDAB64E9753BF51D9
+84ADDE6DA3E7153EE1CB62D87B581C6EFE8CCD5ABA12DBA85F88B576148F3B90
+5D4D5247A654F0B9661B3DFB7E64108B1CF0FDB54053A4C8C7DB1518001934F8
+FCE735E41E8DB3B0D262491B8F99226EA9A7D53F57C45E96D923A4F5DB70AEAF
+C2319396CBB34361FC3127778A9532AD6DDA383C054EC66CDC2FC4B721D91EDF
+CDC41741107FA1F2B1790D1301413EDD316E150CA78FA6064F37D6E861E9C7F7
+DC68D5BED811E591E943EC6075BC28BA2255C3A8AF78C0FFAA66E419F13308F4
+6E9DDAA5BCC2392EFB7DC428D4D0C02DE85D7D817B2AC7BF147FE116D78B553A
+8E5D97242F4CB876E8714F3129A738A97C87FA74195FB25632D7B9F1150989C4
+FB47E8F53FA3532C0C1C80B218D9CBBB3DBF4C8E2EE3F678B52A8938B11B8A1A
+3F5A5A12A9DEA229F6F46395074344853E0C8C5DFCB75AB1CEE296F4D0DDA749
+460E4A7D3509D652CF699BFDA5C0CEDA5FA37B14FE27E63E14C8B086265114A1
+5890D4A698A926C556609A2945C0A0FB1191052674731318F40DD5389134D0A9
+28174D9E62E6666C0F3D5404CC02D3496E09E24C3C6569F9E0F571B10882FB51
+189BE89BB40D4785C77B46B5B1A2B2100804169DBD745EC72B0CD11747BAAE0A
+2630741C3D9FCB0895AB91FFADD9E0417DBED8FE3A9CD1934F027CF89E0953BC
+4E4EEADDE09BF03864666976A9E2767F55D28E27FCEDF94A89893DB833AB55B0
+C032380A851D80C78F2ECB63A2BAD5D3DDB02439FC64A5DEC2C0C1B41BF0760F
+0CF6D654898119F9AFE2C18412AB91650F6582756A3FE4CA590BF665B3BF2CD3
+79EDDC1B8D3DAEBB8E5AD3E6DFDAC57265E2E4B293BC95D561734B14959B889C
+0A336506ACE8BADE4470D37F5777FE74399AF45A2FBF18140473080B787549EF
+DCE93E30D2679353D50D5EC7A9A9A8373D74A59C5A01FAA002CFD455DA63F61C
+C09BD1DE5155FA9DE2F42B0CA08ABC391B2B5A431470E3040D2F15A65E0464
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMBSY10
+%!PS-AdobeFont-1.1: CMBSY10 1.00
+%%CreationDate: 1992 Jul 23 21:21:18
+% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
+11 dict begin
+/FontInfo 7 dict dup begin
+/version (1.00) readonly def
+/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
+/FullName (CMBSY10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Bold) readonly def
+/ItalicAngle -14.035 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMBSY10 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 0 /.notdef put
+readonly def
+/FontBBox{-27 -940 1332 825}readonly def
+/UniqueID 5000762 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF17EFB0FFE6C69FEBA8389DCC1923D30683D
+A8CD93F7195D5A07BA2F18CB3FD5FFEDA4D83BF758062134D84AC0100187A6CD
+1F80F5DC15B47D73F69655445AD218A8AD78C16EF96F385C9E2D46F8A330C7B5
+A859EB0610C78FC5CE39715A1C5458D30498C0A339504A74C7E8F84B3DEC1516
+B3ABAA0A06DEDCD5F9FEAA5AC4AE8D5A5BA5EC0B64784454F58049E13467D705
+8F13A22BDED5F93EDDCAB7A1886A5168D25B120F8BBCC23546BC7398D4E3EC17
+138921404C390EB84C3CC243C0FF3DEC9EBFFF3DEA73365F1E4BC2F3AB911B2F
+780946F4F6F49935A54EF955D9894FEB37239C896CF98240162F6A6E9677EA24
+06BEE1F04463C033047F7F972C560213C7A02BFEE5AE5AE5BF72377CED942A6D
+8059E59CF03CD6782BD34BC02AA4FD1BA25A5CBE32569D7FED28EFB4C0F5F7C8
+6DADC1A047CB514E19B36A84D4DB390FFE5B841C390666FE27C712E23E22FC84
+A8670626E8B72700B9EE9F06F2121264C1CF69FEEC3E20897D0D9057032830FE
+A18A4BA2AD5CE10EE4FED4BB9E2A9C06965779827D7CBA93926793A7161454E3
+C5AC6A3AAEB75EC64556142508DE6E37B71058F8B97C1A9B4CEBF74FBD2D6D84
+F5DAA2B04AD30B313070B33789935E83DB470FAB8EC65165679F247964BD0C20
+78291B6E13C29E8B86429C1B90C396729D6BDE4CCF24BE000390D798DA73BBEC
+AC5C9B1AC19B2C660CF1CDEC05289F6CAEF0E43465E3627DE26670BAA825429B
+4B8FE57928267D5EBE38C5BF93F90304EB89DE120F81362FB5A3D374AB25B33C
+D03A8E9E176E41C964625E58A65EA958EF2B089933C06B71E29249A96D5A2395
+DE687A0C60B837B5657B90F8642A27B037E4FFFA82343351B7C36566DB55E543
+704DF628D0D6C4A672B6BF5C32E797279E72EEFD88551A3B581C615C3D9A11C8
+270ECE7BBDE9ED6DAAE1E81635A51F046840086FC9FFE90840982501EACE70FB
+3495CA202A5F29CA2A4F56C99CE83F882A551087BC666D0A90C14A4AC08F5158
+A2903B69BA116FEF3715532F5E441037A44D2648D62E14A3569E9D57ED99D92A
+85DA381440E32FFF9546B9BFD2B14508D42F198C975076E2269C8B2BBF1AE20E
+C463B0EBE68BEF1F29F27E86E7600E0A7ECF879B5350A8B74101625D3DDDAC09
+083BCA5E10DACF
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMEX10
+%!PS-AdobeFont-1.1: CMEX10 1.00
+%%CreationDate: 1992 Jul 23 21:22:48
+% Copyright (C) 1997 American Mathematical Society. All Rights Reserved.
+11 dict begin
+/FontInfo 7 dict dup begin
+/version (1.00) readonly def
+/Notice (Copyright (C) 1997 American Mathematical Society. All Rights Reserved) readonly def
+/FullName (CMEX10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle 0 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMEX10 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 56 /bracelefttp put
+dup 57 /bracerighttp put
+dup 58 /braceleftbt put
+dup 59 /bracerightbt put
+dup 60 /braceleftmid put
+dup 61 /bracerightmid put
+dup 62 /braceex put
+dup 88 /summationdisplay put
+readonly def
+/FontBBox{-24 -2960 1454 772}readonly def
+/UniqueID 5000774 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF5B8CAC6A7BEB5D02276E511FFAF2AE11910
+DE076F24311D94D07CACC323F360887F1EA11BDDA7927FF3325986FDB0ABDFC8
+8E4B40E7988921D551EC0867EBCA44C05657F0DC913E7B3004A5F3E1337B6987
+FEBC45F989C8DC6DC0AD577E903F05D0D54208A0AE7F28C734F130C133B48422
+BED48639A2B74E4C08F2E710E24A99F347E0F4394CE64EACB549576E89044E52
+EABE595BC964156D9D8C2BAB0F49664E951D7C1A3D1789C47F03C7051A63D5E8
+DF04FAAC47351E82CAE0794AA9692C6452688A74A7A6A7AD09B8A9783C235EC1
+EA2156261B8FB331827145DE315B6EC1B3D8B67B3323F761EAF4C223BB214C4C
+6B062D1B281F5041D068319F4911058376D8EFBA59884BA3318C5BC95684F281
+E0591BC0D1B2A4592A137FF301610019B8AC46AE6E48BC091E888E4487688350
+E9AD5074EE4848271CE4ACC38D8CBC8F3DB32813DDD5B341AF9A6601281ABA38
+4A978B98483A63FCC458D0E3BCE6FD830E7E09B0DB987A6B63B74638FC9F21A5
+8C68479E1A85225670D79CDDE5AC0B77F5A994CA700B5F0FF1F97FC63EFDE023
+8135F04A9D20C31998B12AE06676C362141AAAA395CDEF0A49E0141D335965F2
+FB4198499799CECCC8AA5D255264784CD30A3E8295888EFBC2060ADDD7BAC45A
+EEEECDFF7A47A88E69D84C9E572616C1AC69A34B5F0D0DE8EE4EDF9F4ADE0387
+680924D8D5B73EF04EAD7F45977CA8AD73D4DD45DE1966A3B8251C0386164C35
+5880DD2609C80E96D1AB861C9259748E98F6711D4E241A269ED51FF328344664
+3AF9F18DCE671611DB2F5D3EA77EE734D2BED623F973E6840B8DAD1E2C3C2666
+DD4DD1C1C32F82156385E55FA0C948E3754B7B4883DDB996924249A3C980CA14
+FC56E075F7740213108F818042AF3B0ACDF856B449002CEB305B32D2829585DD
+3095946E0DC5C8763880ADCDF9771E90E192CBA58D81FA1C219094BB9D674112
+DCB4815502CA23BCE201C6755C7F3066372784BE0A61184F6DEC58305E675EFD
+EB24A7C934E67D1C02494B135ED5D53E9099CDE3F5507A2D5303C79CC8A40BD7
+FE9C68638DE98263DEB0E072A99B6FE5D9881A6E3ECA167CAB4919CA5B0E71E3
+17B55A30F4A4EB7600423DFDEF7F14551681716E81086CED4C87DDFF7F87E42E
+DAEBE51957D584114AC57AE4B36D46082FBB317B5018A3161546F714FC3DBAEB
+097FF37F609820A4D87C0F66379133CCC024A1E40ECD03854800880C1A78985C
+882C2F4B63BAD251898BCC685375E570DBA2C04D3E872E2BE6B9203CED0D1B26
+A32406B51B0EEF80E005F9F734EA1A4B3CDE85C822448F3CEC17D11362787B4C
+1C04C62B3365A9C659D15184ED588B4CFB1058533B750551C0DA0A1905D05464
+0F4675395D1D3D53F2696FE0C5CA4F722471A68E741386A22FAA72C4E8D9AAB8
+7F3238FDBAB8286F94F12393A99B6818B56E8D51BC2F983025DACD9B88DCCFEF
+500C61BA95CAEE78EE77A3EDDC18DF93A39C90B9B8A46C98645655592F74DBB6
+BBBFC2D2A391A1FB9FA7EB8514633877EB75BAABF214AB1B946A8A03D282C691
+F34CD79FDB2E07264212EAD27B08314B2F1568E6BC1A7F500E59FC06C80374D8
+2369A4DE0DD5C86591E9434DEF0712D2C94DAACE050342DE5EE1DD8F26E61487
+584ED5C920D3A2F9E847A1B95A1F447AC25BA73C51B8C49A8020CEA53FA2E798
+2A34B72678DF15A9D68D2081ECF22AD4C2D4B595573D3E5669A5BB6598C106E5
+7041D141FF0AEDCBA295B22F71611AD6D9389FF73CEB27E46B738B5358262C72
+AD1E374A239DFEDB66549B9BE1D07A8284AF976516CB56B467C76CA8F48C493D
+8E4D615F0693A725BA70995F127D54E43D6D9F2FCB1E943BBE3CF1BB50B3E704
+2E82E001D7A6ED0821F03CB265494034A80B457F9C2C5DCD76DF155D1F6E8F1C
+04C578A199CDFA4529F946EE5B3EFA100B63D478B4C1FE725E39221E98197202
+412FE717E5F3F497CAC904547C40ABABFAE91F27949DBFABC7CD202934B52C43
+A84A7D78F406B7AC6C08EF1BA039297D21421F56E7E11B9E6371893718AF99FA
+577613AACE689420859429DCC763A560E9D33EB916B86AA6FCB152E09716FCB7
+E501F99429E8E2E74E0928221513CA6ABEEED2E2787D7802EBF90CC56AA4F495
+40A8B0E7333A8F5E127AB78DD6E3D9B657325A789F66977C2CBF700AE95C73F9
+B24201C5B6C0167DFED53935A8C2C3663FA6A39BDF2D82765AA5D65C1EE869F4
+3E76139C9A83B3C281E943EBF473FD6AA33940EBF165DE3E216C5F2A4662C4DA
+8FA925408C0C610948186A51C654D822846577405B76A798877B03B72A14F4AE
+230E6DCA55112699A001AF1B18D754EC91F1DDC21263714CF6F12EAF5101EDB6
+55AAD1C7946CA0BF585B84DE0F4A3A331DAC2CA46B33EE3E39303D00CD7FFF5B
+95B93241312DFFE16F96A6F75AD4D2480ABA8A4274F149DCE287D10CAD57C712
+272AE67B59E4EC9EA31DBB4C052B03C7B89F
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMTI12
+%!PS-AdobeFont-1.1: CMTI12 1.0
+%%CreationDate: 1991 Aug 18 21:06:53
+% 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 (CMTI12) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle -14.04 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMTI12 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 0 /.notdef put
+readonly def
+/FontBBox{-36 -251 1103 750}readonly def
+/UniqueID 5000829 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D03D6A8F05B47AF95EF28A9C561DBDC98C47CF5
+525003F3DBE5BF07B2E83E66B7F97DDD7CE0EEB75A78BD9227BF359D002B6ADB
+8AC57A33FED4EF021A7085B1E2B933DE602F0FF71467ECD501744AE338AF29A0
+26F7D368AC6F25CCB882DB7B7343566192BD687E1349225982823027D3B66703
+3B0DB7A7E680A682B98023D39C7FAE81A5D5B867A0A66C8AA0DBC83B1596A84F
+0436AC6A7900B767BDCCE0060A4811003C79FDCC71D73F7F2D0A6675E93AD21A
+56B4CD8EF75EED3DE8C0A18BEBF7B9D1BE72504872D56EDB272F1E97FC726CB6
+68C85C713059DA19F6C2E0F3E12710A59B6FC4699AE883DE8C8615B7292AC25C
+D5714B6CFB14EF0EF11EB13009BEBA4F345A5D3D6D9926ABC2BAD7DB1328651E
+437BFB3C46DA7B62219660FC368CF3D3704DAD3AB461C28F711665BF484BF61C
+052093D231CA65618EA463D63E406ECE858D180A6C0589B2FEDC321371C28E77
+DE974D655DF5FF7D41ED01FE717D928A885F6FA6CFE4D2C0807F8E7F937916E0
+96EDD1A3BA67802B1F4A49100E75613BA0356D9DCBBAD4DAB3C59E70A47058F5
+2163D1730F0EE4D1F87C3A4AE723A23CFD7986FC4FBD399347E9F5946354E013
+D860FC446AFF0B0744F5DA27CC777C96ADB388D1E835DDCBE123FB517679B9B7
+EF696E091A9D51510BE264701A41C04FA8125A48F306ACA7A83E35D5BA0C296A
+BC594ECA2CB27E92FED95B595C21E5BF0DA724D40761CB377BDE5FB98C9D152D
+6C0DC98C4083E9656321BFC445CD6FCC142DEF16E27DD6FAD0B3185223B1A7D6
+779F39C70793184F2C3B721FD0AE6D8E063BD47804785DAEA74AF8C75483B713
+6506165055CD1430305941C5813EA8BDB4832B3A8BEDBA91DF61C9B419499D4D
+AF7C963D310E2C7E1A518F02ECB659C98DF4C85AB936BA60DCDD34876F39D496
+156BB1D255FBFF4BDAE741475D832B26C856987EDEACCA657E02B639B51D49E5
+90B07711F5671EE3238DB9D65DFD5E754AE2169D9496BBEA59461FE7727BC69D
+40A4BA901C57454C3141B4E97CB736BD097510DE9AAD0854A7DB55B1E0EEE202
+A86CD6D42CB7902121492F0BE9586F71C1D57ADD98B5186F06766B5E40809541
+24856D394AD5E63B25733FE09A8C3C666CC8A64FD99F499D40D6228A6238BC6F
+3A911B5115CF7CA1C35B7D2BCF64F1D2F493EF6DD1A35F5254DCE35890FC0C3D
+9559CDFEBE196E420722FEF2A069B21127ED777FCD8F16EC42F1BBFAE47C3C7D
+CF69CAE2BCE88BCE6D6209003DD375B7179E72FF5569D4BF61F48EA554EC93E2
+AA6B34F8C8C3D284F59163830163F79ECA604121F302B6AC2100DDEFCB16EC88
+7213D4C1A50FD6812F87E141794710B486F7FD1397BB9C30C74A71AE92AA3695
+F044A1915E1382FCB0014AC0D3E99B81EF0BD4BD76E0FA30145CC20192A8116C
+3F8118FDCF9C435204A62F2A1E83BA523EED32B5BA95700622C01F7F967EE294
+276BA0E7C69A0482950F448B979A41828D6B12B5E6AC43B858398EA55F73FBB6
+26891E76A1482C6DDAA342C1A3AA820AAC1AB929217ED7A1CDA01D60963E57B6
+1BD432BF4056F68CA42FDF5F18A7CFED79874BD4DBC6FD3B35C981509D418136
+A9270E2A10EBE47D765316AF230A6B163F9C6680182BACC5BE0A3058B76C26A2
+DB53917DFE062E25C7BAE60B76E10709EDCA549D15C0937F03C764F67987CA5A
+FD1448002CBF50EEDF9FB48420E428CD0ACEE827EF5BD787DDF8A1818DC179FC
+D456F5407CDC66D62E0CBFAEBE33ABEEF8E7DFD5537DB011982B6D4BBE97C603
+21FFC4CA284D58D691AD00828A45ED0977934283E2975C73E72F40CF56086FCD
+765A63351A467EC2934FD8F589AECC3C642F150609CD9D57D37156E2ABC312CB
+FB74EFBBD19ED616CBA9F4816C5C43124A471EE1878F43131020D9C412EB6616
+A127925DC2930C85664F69A5E3FD3F0E309BFD1B2C12EDA9C306E8388154F65D
+1CB5F6BD4EBF0CAFF85B71B918A6D484CA516A4E9D1622807134B70339670CD6
+6738B1FEF6ED55E4B34ABDFC895F12CD274BF400CB59252D2D89A53C2BBF3FE7
+6FE7CB1A22362F48895F708C0B9B8D337843A72571DED644CA95947CAA30E039
+FA4D92805AE2EB6038D47ABBEAB734B84DD328BE04AFAD44044FCF73ACBC872A
+94E0A3DF54E3BB5A1651369F49A0ACABF563D99EA8DD5FF8FDC25F1AEF22A778
+C3CCA8EFC27E8D5A09AAE53A75323B2242D7A90F3E95E79739FDE2EF0AB8CCED
+EDFAAC033B3F30C70896C5476F848C7725E7D41FEB138D57EAC8349B170ED1A9
+C7942F505B448638C41C42770C71AE85FAD42D28B70B6F239E4A22ED34AC2842
+218B9B7F2873C7F506F524F1C9694C0F9EAD60C373AF0FC4427B72FE05527369
+3B4E89E9754A0965F83765EA0BB93933FEBA181EE2B146A50EA42CFB2D7D86AE
+17592EB4800B0CFDB7023DBF884321799F0A17D31F563AAB1146C8E1A2266285
+7F534DF2570B0B11285570C6360F1EF95D3FD9F23316060374CA16161B2D20DE
+8A7B0E65491BBF78DF1D732949C7311EC4FFBB683349E15527750A9B27C4E417
+EAA25D8961DECFA7A2CA45AB2AF0F8E6B65DD8C6557BF244F8FCD3D290F9E881
+976CBE6F1088F8D0D77864C8D2E5D4C814CFC0DD061CE01CAAE307ABEA03F24D
+C3E7D0E8DE72804B3E809FFD5E1AEA18F30368D6F44CBCABBDF2E25EBF8F3B79
+07FA918238ADE92C689E221B04EDA99AC82A2DA5FB9E4B79F652D2B1C76A6761
+59E93AFC467D1C1D184B72D3321B45A0BCD7E3832547FD7E188F7AB8DCC8A31A
+21362FF0C9EDC038C9BD97FDC7FAC046EB43690F4A307259627D37B7AFFDFE13
+3323D22CE53D7ECE119CBD846E565317B661BDE2BABCB4C741A5E2928B21EC69
+EE0BBD2E95EB445CE7468A3C49AAA05CC4D71E5151B3E8FE4D083833A63BD1C0
+A2802E9D5245DD9DDBCF17CD152796E0B25C3E1E139ED27BC17B2254E869312D
+C1CD9E23B672E95081F19EB2A944D5ECE09418E462FDDE0BE52BEFFF750F9550
+84D97E7BD8C9BC2254966DB0076BA4AF9E4D4C17107954EBA5A443136C5E833F
+400080D557C81A74144CE2684C5B880325DD8F1F46CFCB6A2141886F72D458AB
+956EC1FA948112B45D0A6BA8D9109C615DD25D7D0E0E63CCF08FFD02E57D0168
+D392904E7F0105158885A3D27F4EE4BB1C852A4CA3CE642F8437A8E4349FB1CB
+4F7F395AC011354CF9DDED270E76941F5E154803ECF8071FE81F97D87D494235
+965C3A15592B3BC875CC61A96F7789FEE005295FBAF9FF791C3D786E9576A384
+C87F0C762618B1B5870C2545F357F17B9EA745DCEE514A0A7DF9245499280C8E
+98E3FE24327B75CEBCF4E767692D4DD670A5EB0BB4F285D57CCCFBD2836683D6
+65C7CE73EB96F1B371B7C8E256F3BDCB14A1732B28BAE44D1F2C8FD73B10BF9D
+5DFB41F493635A07D14D16064BBD5E136836E9C0AF7F0398B1C2B836EFFEA517
+71A5102BD3A2D43FA3BE41B1B0D728E7D3206D9CD112D419AB04515051899BBA
+CDA524535E3A139E60C2157F1AEEE84E3C6A7CBA48D47AE1EE23C8825B4A5651
+3FD5BA4E8D36D4BFCE02AD80A2801CA7184258B2A9E0A11B1E7AF262C17D17DC
+23527DDF7848B139B10F645B60022E755137BDCC72870FF2D962E0099D4087E3
+06B966FB88AF83C4A113728471048A8D2B1B043CD9B14BE2430CA6E41F4CC674
+10F14CEC66FC0032D9B0324A946909D957B8011002314080296A3D7270FD1EA9
+5E4796CBF1640BA3FED68739BA35CB4FCEED6FDDEAA49C56A697C55C11BB7CEA
+61351288A39ACAF3DE7532DF6F9236A35C42D7C6171053CD76AD990A1633329D
+1330160138953911C0A1AB1AE87BE040A6180A2B836553F71D4D6460B5E8AD7B
+5036FF166854C4BA98C8A3B9B9D7767D3D580D1163A93881ADFBBB6B6120DF4B
+11BD6401DB9AE68EB8EEDBC67D73D0681D97376778C192626FB25745B95DC03C
+F4E8F6E6A42C4FAD12194583DBEC5BC0798ECB6C4A5F4EF40AE1AAF7DB06EEB5
+99F2E0C6AC662D4874885E6A30CA6D735AB0268D56301681A14A6F94BC8D7B03
+86C61520D5FE8749000B1C2DB39884C7BC98ED8611A55F
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%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 0 /.notdef put
+readonly def
+/FontBBox{-24 -250 1110 750}readonly def
+/UniqueID 5087383 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D03D6A8F05B47AF95EF28A9C561DBDC98C47CF5
+5250011D19E9366EB6FD153D3A100CAA6212E3D5D93990737F8D326D347B7EDC
+4391C9DF440285B8FC159D0E98D4258FC57892DDF753642CD526A96ACEDA4120
+788F22B1D09F149794E66DD1AC2C2B3BC6FEC59D626F427CD5AE9C54C7F78F62
+C36F49B3C2E5E62AFB56DCEE87445A12A942C14AE618D1FE1B11A9CF9FAA1F32
+617B598CE5058715EF3051E228F72F651040AD99A741F247C68007E68C84E9D1
+D0BF99AA5D777D88A7D3CED2EA67F4AE61E8BC0495E7DA382E82DDB2B009DD63
+532C74E3BE5EC555A014BCBB6AB31B8286D7712E0E926F8696830672B8214E9B
+5D0740C16ADF0AFD47C4938F373575C6CA91E46D88DE24E682DEC44B57EA8AF8
+4E57D45646073250D82C4B50CBBB0B369932618301F3D4186277103B53B3C9E6
+DB42D6B30115F67B9D078220D5752644930643BDF9FACF684EBE13E39B65055E
+B1BD054C324962025EC79E1D155936FE32D9F2224353F2A46C3558EF216F6BB2
+A304BAF752BEEC36C4440B556AEFECF454BA7CBBA7537BCB10EBC21047333A89
+8936419D857CD9F59EBA20B0A3D9BA4A0D3395336B4CDA4BA6451B6E4D1370FA
+D9BDABB7F271BC1C6C48D9DF1E5A6FAE788F5609DE3C48D47A67097C547D9817
+AD3A7CCE2B771843D69F860DA4059A71494281C0AD8D4BAB3F67BB6739723C04
+AE05F9E35B2B2CB9C7874C114F57A185C8563C0DCCA93F8096384D71A2994748
+A3C7C8B8AF54961A8838AD279441D9A5EB6C1FE26C98BD025F353124DA68A827
+AE2AF8D25CA48031C242AA433EEEBB8ABA4B96821786C38BACB5F58C3D5DA011
+85B385124F409FE733070651C9B7A2D6E683A2BFB681080792962AE48EE5B3F8
+820BC9D61F438E027AC3EECE10A7FCDB529C856728E86A8025D8948EB0F9740D
+1BAB41670C54AE8925698A6A197B742196487461F7DCF7B0F53E94A3B37653DB
+5AFCEF96E05DB75B724F80154FC91C8C08313775EEB7ADA99399A20868A061F2
+E758C49383633ECDC5A408D25937F6556EEC505365FCE73745BE8578229C92A6
+12A4678795EC2A4407445222CFBD0837A31C421DB9B03F92FE3FA72193D1C65D
+F69E9CFD70B658B4F5F43129259D35A83360363E8841758AF61DF261CE0A82F8
+A5A0D3FC91A95C52C118B02386AF037DC75FC6A3C6D2D25FBB80C0EAAEDE0A90
+6DF9E23B18C238C8A4662687467EEA19643150ECFAB377FEDCE630B7D135E079
+6B7F0BB23926CC6BB9AD8F6E7F0C7705EFE760845C1E65F0AC6AE2B18A2B4CDD
+8DE291E4880F20FE65E0AD26A4C5B27B74BB4FCFBDB787E29869A8EBB6EC7ED5
+BB72CF73852F836B3BFC6355086F6A210B45056B9C3D8A0A2BEF49D119A6A7BE
+D6F617379DD0D6A467C416BBBDA2ED9DF4ADF5AE13E344DF608358C6DE2DAB4F
+989A2258E664FD0EDAC5C0E873EA27B64FE1DFAB4D03F1747F9FF11D75076BE6
+C7390F000890FACB9FAB28118A9867BF50CB37CFD50F7541C53DDE4E9E24E898
+E5F8013278BAF060F416A18BEB37D50410DE02817988F4DF20D557EF0A026946
+EDB6F2E3627F1E6AC26B178724B4343524E02E533491295EC532C4D896A38C81
+6EE508D6F9CDA8DDDE8372222BFD274D26A5969F0F38E622B0FB3967B5CF27C9
+AB9FAD864C26219A662985E2282EC6E3B920CD62828458F8A67AE7D26724ED66
+8F68B5492ED92F4E035AED8A9AD60FA22577DC7D7350A90D63F771964844713B
+AFD5C85ED0C3FDBC3F3C2DBAE3E578BA78A9B79E09132CB86BED3EE46CBA86D1
+784D3D1AA674AFF374EAE08628054F185917EA3A9F03701932C9CC44B6685EA5
+57C276AD9F7CDFF5E4C75374935B0FCF3BB87027FF75315181E02DAFFF2DEF22
+E1B01499E82D56CF546DD27273AD0D585F290C4F54115D7AD5BF508C20EFB0AD
+33A20B023EDA3F5D430946A3026DF418B626C5676D613B03CA34AF1F9DBA89E7
+A4E48A754BD9E03A7605EE37F20D162D4BBCC3734A037CC81188BABCF691016E
+7B9B9888742CD6D71C04F38762D1431A660773E553FA3B6CE5D42FC89AEDF9B2
+96D2DB8B65E94C425DC5BA65309CD1B247939D4FD70E9175524DA0A334DBDF2D
+72465A5A87018167CCD26DC491D7E9519C67D6D135C485C45DF9388F5747BD83
+A4B64490EB211AC050EC0C906FFC9687F5CA42875E9938112B7BDAC56546F9FC
+084099AA0864BF7C789CF2218CF7592059B02DFA9DB5600BC10CE4BC76C4056D
+DB7F88A52D82B050D82EC132F835D194A1C4EE98ACD2260D234FD25C9F5A1A20
+A666ED5C5A7D26E75B9069C860928739B5851460B37E73AE1AD568C071A88F3C
+A3EA67E6DBF60156F3DB90E2E390D5115EBBB918C67AAB291DE9B2E79E2E501E
+02547B4A9BF1112286740E9ED6FFEDD9781B3029C4FB95B480B86F532569784A
+61D8FE161C1B9A9E8A8DDA8A8411E244CA11C210B38D90C5803380DA4AF29382
+D31C7B40C022005B866BA920A50FEF87F7FCBF02EF5CB44EE712391D2D8A6028
+99E20908FDA20065E1F75D114CFC64D8EDCA860669961FEDC7E8883ACF6576C9
+256EB11023D756A7EE06DDAFF16B701633333EBD5A635ED4E320724F643169F2
+800CDC46F79DB8A69CC6DA7FB1E6B370EC8C9DAB9FB2C989E46A924BF183B92C
+A4C7C042E841232FE398586CC7EB55B14D4200B65353A27CCD066F2175F9206E
+553F6550A2A5CC8A701EC01CCA4EEECFDAA9E4065E6FDB8710CF42D80303B88B
+A7613784015CC08BFBC6524993479CF4FBCB223E714D2A47A34DAD4B0FF35147
+348D0A6FD407092D7B218106414BE9FABD0921104A2F3A8124AA5E3DE161EC2F
+70F86A50C9432A0935200505F56DD433514D0C94D72922C941C122FEA1051403
+D913325F733D9DEE3D7F3334C422E569A2ED3C47CB53BCFE21CDD39FF6D5137B
+FE1664C591846FFD8433DC240C95D90AFAB0DD29792EA8345BDFC9EF17230967
+9F0974FEC656E93F64EFE7589BE7FD159DC167E0DC1C46781EBBD9A51D48DF27
+462B4D1F491A4D671F0CAF11E7545F78F2E0624006CFE650BF435FC8FC2A7992
+5C43B7F51AF90DD5BE9A7B08744C681459D8078A61A14F556DE2926907DF7B6F
+C037CE01772DEDB0FFD81A298FEBDBD2196AF5EDA93AF7D5B5E96C4FCA5A6538
+B1A999956633854B5900DEA2BC4ED3D89185DE9EDACD18EDA620D4A6337853DC
+352C0BFDC6FA40B3FA7332B1189EA2372DB4346116D6C34F68F1B7422EEB5453
+5456AB3377FCB21B440322996A8F99EF4784167AA0316665B14A7862EB573310
+86ABBCEFE41B39ECDE38553380DC808E7E821F646FF5EB73776C25084FBA6432
+4A5BD37F30DCCCDFF78D50E6A94D8816BC6A3549BC407C2D9290778D44A6BDB8
+F2F2812BC145AD649586C9014C4BA89945CC0FFB99C311F7F60208600AA9A803
+DA78A3CA0A23A3F2B87FDE73D51310635DA106B4000FAFD8D7FC830EAA57DFAE
+5F2D9A08CC2D663A6E1FCF075ED17DD7569DDF2F0939C642989C85F475A9AA1B
+30D5B7ECBA17C9FC2367E97E37B79B23BFD20082811F75DBBC3714257D309AFC
+217A352E802B5EC49D19306526E27AAF4D6EC2D0CE4DB379D5FF9298F878A999
+06EBCD32E4E1336ACB9E6558FD63DB6924C8C604FB4C725AEC0D00CD27CB00A4
+48ECF6464EF9A6585D4C48C4DC5743FE01F77889934FF05CBF5A5AAA110A0D7A
+E3DC96EF096228D53EF63A752118CEE758E3B4EA072434C99CE5A80E34BA2328
+E8B4BF9915226EB470D2F07B6667E236C644D90BE033E9DD1FF1C8DBD118CDA1
+CB1A40478AB4907D1CD6AAAC0869ABD917C5E7B1A71FE79488FE16F48AFB22D8
+207FF6C4D6456EA280F1CF924671E4E11220EFEEC49E547D764FD6F7DC3B03D0
+38DC4EAC7B3D3E5161F3CF10E62531A8C88548AD04A4E64CAF6FC648238E6AF1
+C07188AE52FDA6CFB83F975BB9F78B94A600A82CAB470E333C5E3F66B43FDE5A
+ABAC621C1C158B0265D10C20ADDE8B3A91CEBFCED8864060265DFCF65FDF47A8
+FA5A55F153E43DD7D08D18C2BE1E13B9BD9D9DBA8679943B83749D7AC27A2210
+8AAFB25C425DE39C9EB74BD0615D4ECB7B2B88D01C694ACCC711BB687686CF83
+6530F6681C91CE58B14E7267E321BDA7878F2E843A8A8523BFDAFB66402223C3
+AAA823E5337C21B9B6BBF91FD2D16387DA0CF2E5B440510900692A1D941429E0
+68F62E5E8374D3D55D6ECBA044CC7DA50FD22F447C76EB1299249D6F2850122B
+9F54D7B21692CBAFD00DECAC0C775B745C7EBE1336564CC4D05208087C33F584
+188703D73B2C12D023F5F821706ABD93D7F9227682D871772477F83773F8FA78
+CA5B9A1A7B1663960D58B5673CCF9ED40F60B2BBEBF61AC5322AF56FB683C114
+677D292978E9EA4FE621E00CB0980E34A13ECEA5905A6150DDCCD826AB7AFE74
+DF462EA0FEDD0195DD34F073878F19C2981FC4B3D71E647594528FC7AACC5D18
+FDE91646D19289A9400A626BF395C7DE1CDB2996E3F9DE9880A433BC111D5B31
+8332DF2D4B5975D29319F75A764C31B6580210D082DA1EC3D23F674BCD955982
+19EDC2DEEE18979E3FCB9E308075D630E8D1DB9589FB82CCDBE7B5044740F857
+EB4ACF1964B91A654E85FE7C1D81205E238A62041825EC820F6073C0A6E4531A
+792CBAD8466CC895E8331A750209136C2B9645E8B011EE3F566759D644BAD46D
+4DF543EAD466DB7A94CA21A617B3FDD06062D984180E0FDC00B4E0E6462DA4A0
+B6009352F1DAE2C798B09228CB4DAF70EDE9F947406814EB1AC8347BB647C6B6
+71FC2C0564D8AC2337E43B257F42D2E5AB72E90D5CA6FC4691C974B85E1346E8
+3AEC45511086E40677D9379396ACAA0F18262DD2A13377F8DD93CA31AF8B34BA
+C34B09BF79239B8F7F29EBDB3C3D2B165B3024E013305585778A5D738FC6DAF1
+C61100A1286DF275E751E5315C4030A621723A39D41D418C5EFF417FD72814C4
+4E036349270C48C15D75AAE80CCF97FC484BE3027B49C237873505ADD46280C6
+00CAAD68E99B2391B062C085CBD8D0184E9BDB479A1E6490CA9E8D376ECA6919
+C4D8A169C75CBA9E709587579D6886BF96FFAFB682E6863E0BD138236388DAAF
+7D64E649204B7A1C6FC224D36C53E5DBF34B22B9644442156D8F5CFDE07CC59E
+1F1A867CDDFDA80D52E32DB46191107D8238C5E1A06BAD964548463AD011C0E4
+A44ADBAC88927E91923199A64A04B118965833D4FF4E1408AEBB086E0FC3418B
+75D644EA6430AA1E4950CDD3823B1D2363EABBFFDB23F58ECD990247DBA7E4A9
+0103ECDB8970478C23B5700C351E17B4D4BE1303E9320050FC651EAEA684D24D
+DA4D91BEC4B5D12EEE2EEBCDEB18C7951961A1202FE12FDEF1FE286CC2E3A0BF
+C9DB9C55FD16F0A7E8BDE50819891875657C2B867704FAE7E0EEE20DBB445CE3
+B7CB48D20BF8CC14ABB1310013472C9B919AED96B7FC4205C09FD98B5DDA07C1
+DEA9CDE6983E579C72B03506B10BC01770953BEE9BCC403BAD053C4822A85FB4
+245DB868CB4B82BB4185274ABC3DA78207E4811D5C8696277D02B5D80EC22DC6
+491FF010A925BE7ACF100A6261E04071E939D1A86B87C1C0C33BE53FE0E2DA1F
+0B96E84A761045
+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 0 /.notdef put
+readonly def
+/FontBBox{-30 -955 1185 779}readonly def
+/UniqueID 5000818 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052F09F9C8ADE9D907C058B87E9B6964
+7D53359E51216774A4EAA1E2B58EC3176BD1184A633B951372B4198D4E8C5EF4
+A213ACB58AA0A658908035BF2ED8531779838A960DFE2B27EA49C37156989C85
+E21B3ABF72E39A89232CD9F4237FC80C9E64E8425AA3BEF7DED60B122A52922A
+221A37D9A807DD01161779DDE7D5FC1B2109839E5B52DFBB2A7C1B5D8E7E8AA0
+5B10EA43D6A8ED61AF5B23D49920D8F79DAB6A59062134D84AC0100187A6CD1F
+80F5DDD9D222ACB1C23326A7656A635C4A241CCD32CBFDF8363206B8AA36E107
+1477F5496111E055C7491002AFF272E46ECC46422F0380D093284870022523FB
+DA1716CC4F2E2CCAD5F173FCBE6EDDB874AD255CD5E5C0F86214393FCB5F5C20
+9C3C2BB5886E36FC3CCC21483C3AC193485A46E9D22BD7201894E4D45ADD9BF1
+CC5CF6A5010B5654AC0BE0DA903DB563B13840BA3015F72E51E3BC80156388BA
+F83C7D393392BCBC227771CDCB976E9330253242123300FC673C8BCFE359BBD9
+53AD5F3D9983B036D9202C8FCC4FA88AF960E1E4992ED4D358A43D0B625E7B8F
+2CB5599132B9E31EC3421EBBF4BBEE085CFDD794E0B4B694DEC1FA0F4659EBB4
+8519781FC955ECFB19C20E94CACA623E755471ECCF677190C3C26822E97567AC
+2CD723A5F65DE62EF0E103658C4E94456BD12851C89BBB7D042FDAB8B6D01DE0
+7B930D1F63DCD59628371A2EA482EB65A6BAC43680C318427A70B844D313DF5D
+12D35D89990DDC7033205197EE96F3C0596AB0B69AD48366BBF98E8022AE9282
+B803D6E86273FF3C739031C6C8D90DBD12501A684A2B37AA83D5C9946CF8A671
+76B4B3A04AA295986963279D3B3A30E9F815CE728FD4DCAEE31A95C15FA0491B
+54F936970A892E2C72A09A834F40736240602C872CFDA6AF79468AFD10481A35
+5E6527A2792F33A15CD2EBA335139CFDA1F4F43B9F3646BE3DCDDA0C799ACEF0
+38B1A979ACAE33810F684A2D364F44A77C3B67085A51DF9192170244582F3201
+53004DB0A58D4952A9D6615553F92599238469AA282D6DABC50B0DA0FB785D84
+3F40FE2695F4338B905FEBC0F8B5FC0DA2A9B924C25BBE03CB87E0D4F7FB291F
+0E16C58370CEC38D1BE34CA95E305A158F902A7DE77CB0AF321518BBCC89BE81
+1867C8B603D935813902769CB2E7FAF7EABA0129F40156E5B35F8A698194464B
+BFBA5E4F12F6D6C336EF50B7EE05B2CB80B723CFD68623969253A157EE2B0E00
+CEAC662C7FFD809A69C633ED8A64DA0ACB640C74B44DBE77E2495E6C40CC44DC
+EEF07CB15CA136E69353A19A15815FC945DF68F9F0883D0920148EC50368A357
+E772E010664404D11AE2F57C62268F7D2432972CD870D10B2824EC7E9DD517FE
+3F69C034B754A575180A3465992D2B78A3EFE66941BE8A4A4D54AAE2BA022A1F
+5B5FA60658D9BE08F1004961C5F665274BBAC4C94F9D14BD56E911F421427863
+408D498BAC667AE754B80CF6D0E431EA431961E97BD4C7F270C64A99BC24F97A
+7882E5EB5A0371FAF10A16D13CAB75B23E068B8CD528A420D0C49527626F9484
+5269CA342605A606D420108564B66DEC54CC3EAA7D11E6A49A6522C7C7DE5731
+74A674E37DBA9223F9D27233388F892FED220811CB368D97046D529CAF13FC9E
+965A0DB8038B071C371675C2457B48E7FFD22A04486C07594E6DA373B2F03077
+E6F90F905F4E1702F4DD117DD7F6454B51BD634CC7763D5C007AD05875CB33C6
+0103D78D3D2BD8250DAA8E67C4F31F2B3DD1297B15F4DE37B0531CB58BC5019F
+9E7B6F9D1A683726BCC64F4772665AC3DD5CB75CCD44B95D3F5B3B687DE89B70
+762F6EBA2A961C74AF1DE7718B122B25409C9A75C8C2C097894ECFCBCF9D0348
+0A4B286938E345E1ACDC8DC473CC5B08317104462E0EF60B15EEEA321FDFC485
+F076087EF16578D1552C9D164A61AFB915545C7FA1F06599991DD5C55A565246
+67FFDE8AE5AB35634EE4ACDA69485F522A0EFF1F15D7488EBDCA0A3BF4B5A92D
+A13A5D3F02B55E6E83751F21F990AF1E7D4E1AD6AA7D6E46A374BB7757C2EDF7
+9C1D73D2151014798FB74829ABE5CBD7F42E8F115CF8F5252922768B39B50449
+55B26CB872ABC2F37203A14B871B9E1CC24A85E8236374697DB1BEA09C40A44B
+50F3D99F3B227F1A43D8F425FF2E83D55D1D6F636F4A4686825928EE89E8A295
+57DDE3C523F603C4AD47FD831B09D6AB3B90ABEF5293F0093B34A50011053BD5
+E9CCAF59722A38BE82009F5B55F6D663494DE9C60D56E6FB00AC1F7F34621AF7
+815C190238A4EF5E9DFFE36506502E24D0AE6246BF9559F4D095E26D366560A3
+328F31993B8F9F1B5B8E685273901237F150183D89E5D6647499F37A416E2FE5
+A3AD0B08741D1DEA18E73CA3E5276786512CBEE5BDFA6146A78A92885CA7098F
+2D892E738AE230EC834CEF5258D3C1CDE61CFC1346767C51466CCA684E7CEE32
+9BB5081348C9131154300A1370EF67F5BC9EF51E9764A14DB5291FF72D0821D8
+2D74FD12CB12E17B3B87384E01EFD8F4B48D691030FD220A7343C85F31D2B4AB
+3866B32BE68DA800ADDFAC77640763491BA9EB2A38B556AB660E83A2233E5BD1
+99DE36CDAAF848D9D3AABD0E552FCFD2E924CE2D8A8E8ED37468A4C6AC4A4603
+8B501CD0884C5D03FFA9D73D6485131C983ECE11F0BED9CED6EDEB14D7B070AB
+453EB129FA7FB31FFCFAB97BA677FC3603C1F85E34121EC8C31937978631EFF2
+04E80EE83EDCC0AF4F285B6DE20BD5C22E922A050C3FC5F756AD0A86B703C71C
+7A83A6BD77E41EAB2260DA0A8679EE46241190A0A0BBC32EE8A183EC0896988E
+5C52FD02B5772F55713ADA039464F9C3286E69DCE72B210FF5702BBF548967BF
+E9EFBCECF6D18BC47BA2FB658B8314C39CD60655C122C4C472BA45D17B7F339E
+FEE4E025310262CA496973E847138612A96B8695B9443E3C3CA6C736C7847C17
+8AC79F60062FFD7C93795129DA6A813F2BE35230CD27882FC2D2E673988B3171
+1A59DA52493CDE64AD80B5E646EFF62658DC16B1E658E774AC53E9BD9670808C
+58684F567CDF9FD8168F13A7A9647E750962E7033AE8B5C0D412D4B08FF853A4
+3873FC51C7A9A3F841FD863607634233427B34BC1CA3BBCE45B1F1EA6FD5B3F8
+153AAAAEA0B58D737B6FBD372DB9B49918A1FC0D6FE38640291D0AA2510C8AEC
+E0CD2A0A42D0B20F6731390F0BC3A1DF2B867390885CC89FCEAAEE182C4BCFF8
+5E6C0771996EC36B4354057AB0E9FA30F5C364B7A6EA958A06A47D7C4DD143C1
+2269AFF37CAF54351E1C88B34115438E6D613F6E27CD26DB65ED59C057B1B343
+13AF7D9A8908900616BC1CEA9219E550D4BCA97D93262C2C963EFCA33D5D1060
+617F72239B7D99A97E75361DB327CF433774DDA8FD32908B0D977A53A2604DDD
+5ADDF4871C7BA75B198489EE29ACDFDEE64366E9C4C5BED44094162281C5C5AD
+0EC815CDEF12C36BBD792BC6654E3F2B3D6EDADBB7A477E397079898DA51048D
+B23AB0242C7F113F9835B4626B6AA11B07A92AA8
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMR12
+%!PS-AdobeFont-1.1: CMR12 1.0
+%%CreationDate: 1991 Aug 20 16:38:05
+% 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 (CMR12) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle 0 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMR12 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 0 /.notdef put
+readonly def
+/FontBBox{-34 -251 988 750}readonly def
+/UniqueID 5000794 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF4E9D2405B169CD5365D6ECED5D768D66D6C
+68618B8C482B341F8CA38E9BB9BAFCFAAD9C2F3FD033B62690986ED43D9C9361
+3645B82392D5CAE11A7CB49D7E2E82DCD485CBA04C77322EB2E6A79D73DC194E
+59C120A2DABB9BF72E2CF256DD6EB54EECBA588101ABD933B57CE8A3A0D16B28
+51D7494F73096DF53BDC66BBF896B587DF9643317D5F610CD9088F9849126F23
+DDE030F7B277DD99055C8B119CAE9C99158AC4E150CDFC2C66ED92EBB4CC092A
+AA078CE16247A1335AD332DAA950D20395A7384C33FF72EAA31A5B89766E635F
+45C4C068AD7EE867398F0381B07CB94D29FF097D59FF9961D195A948E3D87C31
+821E9295A56D21875B41988F7A16A1587050C3C71B4E4355BB37F255D6B237CE
+96F25467F70FA19E0F85785FF49068949CCC79F2F8AE57D5F79BB9C5CF5EED5D
+9857B9967D9B96CDCF73D5D65FF75AFABB66734018BAE264597220C89FD17379
+26764A9302D078B4EB0E29178C878FD61007EEA2DDB119AE88C57ECFEF4B71E4
+140A34951DDC3568A84CC92371A789021A103A1A347050FDA6ECF7903F67D213
+1D0C7C474A9053866E9C88E65E6932BA87A73686EAB0019389F84D159809C498
+1E7A30ED942EB211B00DBFF5BCC720F4E276C3339B31B6EABBB078430E6A09BB
+377D3061A20B1EB98796B8607EECBC699445EAA866C38E02DF59F5EDD378303A
+0733B90E7835C0AAF32BA04F1566D8161EA89CD4D14DDB953F8B910BFC8A7F03
+5020F55EF8FC2640ADADA156F6CF8F2EB6610F7EE8874A26CBE7CD154469B9F4
+ED76886B3FB679FFDEB59BB6C55AF7087BA48B75EE2FB374B19BCC421A963E15
+FE05ECAAF9EECDF4B2715010A320102E6F8CCAA342FA11532671CEBAC7808842
+A680E4250DD66FFE6EBEAF7EDC45B5F1765BE4BE4563AB9399932FE758F01BB4
+ABA03C0ABEADD1E3BC6430DA54E9BB7CE8E0B9C43755BF7C8DAB08C3E064C7BA
+676348BE1E40EB030B29B8918B8B9B9C6422DF15891BC607868ADEF2A9966F27
+40D6CB4F3AA8E715A5AD53F83CA9A5307F749D7BF47FCA32FEF324C5D4BC494D
+DD4CFFE171B658F7C368A959DB70BF553CC32E1A1D4794529596125BDE09516B
+FBDC0EF45E0E7E8740481A03D8672AEBE59812724401EFBE3D6F225E5BEFE06A
+57CE7183B9C8D2EDB0CF134DDC6BE5CE05FBE1B2149C1F9E3CDBFE13A94DEE0D
+535145F0520FA5E8BBD9A4FBDCF821B005F814E1226F40FAD7590DB1D1D4FA2A
+9DC75D06C9BB5A168D535A5EA58D90299B92051301F3C23A17D85495EB9BFBAC
+C395FD5EDF9B4A65641CE5443756CE71F7F8727DBF2C90C62BFFE9F83310006B
+8E6F252CCF78AF2B3B93F8E86C115A621956A3E6CD67FC2864DF3B4F4E7DED18
+91EC2503E05396403D0322DED2972C5E76AB1FDF5A503CC51227C069DCAE7127
+24937D598A25377470D62C3441CFEA23257BCCCE4146B2CEBE0C72A8FE9DE1A0
+60C389B022B65DE4371A3D735344F6DC941FAEC4FB84FCCDD9D4C3C21085ABB4
+F816D0A93C7717FB08DAFDD30860A89DFDF14B347868251F774943A4F074DA18
+EBCD1D12440837D7DAA821460091232940B90A029AE9F5D24128C328335EE33B
+ED847B428AEC894DAAE93FEC51015085CD55D14FB43CCFCE3F59401B36065C6C
+348F7165980AB56DD6FBA1C3E0A77A6C4C8231C6D1E8E62CFC424746DDE1B1B2
+BE1AE4B944239E0E203FFE952F72BD1D3F6C25DCCE2B3F49030FCC3C72D94D4D
+61A7CD54E3E01EAADA5A128D409C164F5ED5ABBB9E845922248F405572ED695E
+5354F4B0D11C320475051405D488790D3897ACF237ACD6743E4901D4F425B1E5
+1BB473420C6EEBE25F67CF2DBBDFED2CBB5AD309FC4766B21159AF475A1CFDAC
+80FBCBDD1B3C1E6D554C460A1C113FAC7A13ECDC5DC97BD64C4B4104FC52CDB5
+0746A8A6E5FD2D8686C4B003263BEF16DED50283876144FE31288DF758055FE6
+8E52FE1F58F099E2D572A8CA07C2FEC43992773F7091218A4E2215F876C4F4B6
+E36D9C898D89F5343E1943CB748401E44699934A97FF332CAA115266C8368F01
+B04049FA3CF67D4102E49D803FE79CCB0D240FF69C3E8AA197C012225F8A9CAA
+AA5976F3047FA0820983E0C642F60D49EEB47AF4BF93425E2F13F53073366A74
+3BDD0AAA9FCBB80177232703181613D12A6114BC827D003923B43D0293BB6B90
+FDD4724AB6B30104199D7742FA5F020A0218200ED323482CAA2012294A590C43
+547FABD566C46C01FACDF88DB60BF4D8B3C017A4755EE076EEA8059B7D8ABBB8
+35C4A15875FD9EEF3A1EA350DE4B3E393CBC64437608AB396AD4F630D91C4F66
+EDBC0013E7D15B22FA62B4936FB6BECE464730B594D041C217D29B51FCB4289E
+79FE3BC9441F1C1D91FD6DE98639D207923497E0209FB56FD59E23D0A44DB80D
+76E82E78F6B44132C4051295731D3907F7B0BAF39B3BA5C73615C72CA1491482
+FD7BBB9113D4575D1A7F5C4830AEF9D20FE552EEF2A02BEF35CF7E8D39BE47E0
+C0907FF2E373BA034AA39D51029FF938AD3D1EAC4740F5E35ADB893E07807273
+65C18F7D0CFBCA8BB8600386C6294B9EBC5CFCA07A8BEF89387582DCB555CD6A
+6627F4F242AF73124DF0D40224A3D85DD88900F4BB8A20E8BB23F21AA31618E8
+CA1B639A42CBC5C991433E8B1413A7317D87AAB3828C85240D60976D6A1854F1
+4A84BB02EDA57662E001C773DBE5F3B5AC234EC18D83EFF6C5AABB16C4ACBE91
+F45A914AAAFC0DC42022A39AE77F1F8A7A8CD992368AA9C560E86C72806C70AC
+81DA1C74F597B3618E8488609CACF0AA351ABF1AE2890D86ED1F45F99CFCE700
+205FF90931523AE6871084DBC316EF77A597A6B2461533F3FD57EEBD25FFD63E
+5B93760EB7F6EED86AB37A2EC44873602A0E5934F7665A295350DFDDA09049A7
+5513D47FE8D8522C528369B6141F7CB278B41763F98AADADFEBBC328C829CE41
+64370CA3F408D6A40A445093EE52229D579B30E72A3A517588676C8BAC358157
+94A6079D23A8B88D6D072F026B58A9D6F7B440715828FAA57FFB80BBDF81B5EA
+B606FFF53DD4EA34B87B445F8E21F17114A51E1F1C9567A79F1606C1EDEA996A
+FA8F7C4C0D992CB41A43A0657309C370E0C402F3FD669C8B3672782F45E1FC05
+14097245702386542A23FA8C07CE4E1AAE9F4B67489CA82AFFA8C84C25F95CBE
+04B5761D45C7EF3AB503CC27654B37BAF54F3EA41F0E51B5159396532635D6B8
+14E3D17130C28573FFAE2ED6314DE0AF7880DDA3761A03941AEA016B9E84B32B
+0C57EEAC7DA17879606225DC4B518EC364A003F8F74F2A1D750BF9468CF8779C
+A44C9602E682E21DBC8BF6EA6BC9BCFB4749DAD9008B36BFCC71844BF1ADDB77
+83B2654E88D079245343EB57CC740027158C9B
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMTT12
+%!PS-AdobeFont-1.1: CMTT12 1.0
+%%CreationDate: 1991 Aug 20 16:45:46
+% 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 (CMTT12) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle 0 def
+/isFixedPitch true def
+end readonly def
+/FontName /CMTT12 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 0 /.notdef put
+readonly def
+/FontBBox{-1 -234 524 695}readonly def
+/UniqueID 5000833 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5F0364CD5660FE13FF01BC20148F9C480BCD0E
+C81D5BFC66F04993DD73F0BE0AB13F53B1BA79FE5F618A4F672B16C06BE3251E
+3BCB599BFA0E6041FBD558475370D693A959259A2699BA6E97CF40435B8E8A4B
+426343E145DF14E59028D4E0941AB537E34024E6CDE0EA9AF8038A3260A0358D
+D5B1DB53582F0DAB7ADE29CF8DBA0992D5A94672DFF91573F38D9BFD1A57E161
+E52DA1B41433C82261E47F79997DF603935D2A187A95F7A25D148FB3C2B6AA32
+6B982C32C6B25867871ED7B38E150031A3DE568C8D3731A779EAAF09AC5CE6C5
+A129C4147E56882B8068DF37C97C761694F1316AF93E33FF7E0B2F1F252735CE
+0D9F7BCE136B06EE967ABE0C8DF24DCBBF99874702ED252B677F407CB39678CC
+85DDFC2F45C552BA967E4158165ED16FECC4E32AC4D3B3EB8046DCDD37C92FDF
+F1F3710BB8EF5CA358ABACA33C7E5ACAD6BF5DC58BDFC3CF09BA2A38291D45A4
+C15FF1916FE2EC47FDC80911EB9C61F5D355BEDFC9DB17588547763AC5F0B1CC
+12D2FFB32E0803D37E3281DA9CE36C5433655526ACFB3A301C56FAB09DF07B5D
+048B47687348DEB96F3F9C53CE56DDD312B93D3918CD92AF53FB9461864D11B8
+0138918D0B1270C54873C4012CDE6F886DB11BCEA04B023EBB43E0D0A06BE725
+741D08B9DB688731A6C9886C15A83C28DADCC81385EA239E045E8F3670CE03DB
+9EE77ED067036595C9F3B1854343BE3A12E486B6E5A2F8AC44FA5378D28DCCEE
+306B0E283AA444423F9A4FF38E2B56DCF67A39CEB2C643DAE86865517D5D0371
+CB8797208ADEC637330A3A57902C9A88EDB75A7C16FA9850075D9F19578EC666
+1353CC1FC512D59DFF847ACCD04A1C36D8C678E99A0CAF0B7AA7DE226F220BA6
+6E824D881ED93192711FF2A2B6F20577C0BDA67126FED82E844451FFBBE43A78
+739AB1AC68AD49FC94008354A3B578BCCCD2D583C24805FCCB60B8E3B3DE6DB9
+4A90D26A563AC41C40BA1D40178F2FF6855E1D1F76A437143A46F7A6380109AC
+4A6F8B577DC4F10940B7545E4BE6DD9204C2D53F0A7FCFE15E6ABF21461AB7FD
+B904A363CDAC0E6B0C41298C8592C710AEA61C5795A6ECFE8BDF2AC93976FF11
+8A27DEDD539DE9ACC076946E7EC5CFC23F320F67F1F39FDDE1F5D4E31FEF6D1D
+AE4814DDBB384E2590EE915AD4B269108C16EE65FA1A1FFA6960C6592F6B5445
+50D53D4B071BB3B1D346EB4F8E9EA42A506B7F5420AF938D6BD5FF951C0C4EAD
+853725CA1D14AD14513076AD8B1E2B82FA4A43102DFD785243E620EC3698A7B0
+BB5AD9371A5D18FB0B9FDF1D6DCE03DCBD24BC4169F2A4F76C9B494179D8EDD2
+2514F33B71CE5D980642E2D0D2223CB8E56E3BF72866ACD19818DF93CB0517B6
+00D5C0DA6B151D88C492B2634FF7DC9C7A3123EF007DED1D78D1A657DABAFDED
+85CF6CEC398394902814A25A1536DF2EA701D31048F3DEBE306409B27270FDD9
+E857FB4A01F969F2B044D89A361609DCE5B499A813F98DF81781033ED3B287BA
+E00F0EF7319D784200F749EA14C80E1057393AD192E32FCC5602847887E31AD0
+94B999230CD63BF764BD0629DD011594A995C5898EAD9FFD5C413CB4F56C0820
+062C002027596F05E727622A08E7E05A96314563A7455460A300646459721371
+01F8EE8E4DE6486B275475B8D910F232041C824C3512A51A00421A2F3C681370
+70653D2CFBC920879A25F1DBFC6A497F795F1CC8EDD011C7ECA94D06EC530E2D
+F0331A47D27F6D3B5A0184387A033B7202A2275757D73564B2042CB0AD2E66CF
+408A6DCD456BE9DC79D8DF90FA2218C72BF2005BA9C0F72A4D104384F9443651
+61BCC6E44099C69168227DB5477EBDFE97B056BF85A61B295D76E5A947CA1F25
+585808BDC21F7D9E749FC35E5942DC479CF9507DEAB1A3AABE22A2AB1B276B59
+698C596553CCE435336E595DEB4188FEFFADC09C05F8FD7B8C0A60CDAB94C9B1
+330C4D4038530C62901BE50A955D400B8C38038ADC89474B46B096178629C2EF
+F09AA101C5D4A45D03EFCF6F3A368B7EDA6164DCE6AA5109D755634420172BB9
+C8A0D0718A2BBD50E559271A4E05D8135960160AC0795AAFE8D64EF684B7686C
+3B2F80434DC561EB83D8219D372F64168EDB02A1BEEB49F1F4FF666C4592899A
+3001F169A938547AC4AEF8CEAD874F45EC85AD197F5746B5C7456ABB3CD0A492
+3B523974C048E2AC727DCE81C9495B8C39A2DAA093D0969E1D2D65BCD899BE3C
+FB7FA34D180B66024900FA5600915E18D76F2CE7591D87817D94CF491C9E1DF0
+3AF08201169990CE357DB95B4A2A01F8E280369D9CC5D98079AAD9CBFAF9FF5E
+34D404C957EED42BE91455369AF2D02BB14C31C2DEC69B45E12D73FAFF5D60C9
+9CB4E047AD1CEFF656711F9CA8739019D5943DCD9D7F070386FA73073F56068A
+9FABDFBDA79ABDC708B7A348D67414B8206796CCB3BC6F7E068ABA90555BE6CF
+407764F375840F13A4B4CCF6B12CB015F593A977A6DD0188A583832F06A686C9
+2BCD8C828B2DD6007B8CD887DBAD24D4A403E1A4960A3F638F151767B84988AE
+BFC4823601F6E6570094D3A50ED465D2F0D0DAE4976D7C77DD32C1A87339E700
+7AEE6731763EEC7809C29B7EFC86F8AD677F072B5BFD53E2ED424C9CF5A92E19
+27A8DF7428FEE3C4236F3D6AF4A5041B85D71268866DA3C761CF3875F00633C3
+06EDD71E14E4ACED004D7CF649739C6519CA190129BAF9292A5DCD32C054E940
+D2496D270DDAF9AFDAB7BFB8EF2D064F1F8596E89056A214039F44E6DE153E1F
+493109758811AF6381C170B10875FC06E81468F56930172E67C43D137E8098DD
+E88C33C3BC3DBD23DD1E62984A2C1DBACDA12190515D11270F5AD372CD2AD103
+065C46590D9CAC2DBA3D2019332929098D0D22144F5FA0724260AE0F7D57054F
+8B2ABF0FFE003CFDCD267C0B2623AB972B4A1F6513F5F1E0A8B800D140BEA77F
+F0C9C46FC836B0D626A83B5B49E3DA3BB42BABA4D5AC84FF3612FE3B161012FB
+4B375D510195977BF4AB9E8545750057EF0C832D35CF530A784E8AEF56B4EB41
+6439D0EDA59B546F9652AEF13B623691CB7B3850A98A2FC0175C1DB00F905B92
+F9AE7256D77379CA51CBF8F9711126687EBD3FD10389D95D9786C72E91D2164F
+4FFB5C7CDF0580137358411B20E30B6313F5BABDB401BB449E2C232EA65A3CEE
+07054B6B681A07C99EBB27F0D5C2840D6F36345ABCFC4227D2DF242855F3240C
+036504489964B8BEB1884F51A2DE3AC4D420249DD26A7D095D2997AD97CE16BC
+DDF942D5A1031FAA7A47A3B19C48764C0F92C37D83917FB2E0A630EA7106EBAA
+795D30C7A6925A2B0CE21562EF9407BAA9B2541DA5E27B5D5FED7692DA3A1F6A
+CAEFDCDEB7A0EA1872489B8B3D35207C79066F821DE83218520FB7CB9EA386D9
+D7B4F5ADDCBE378E5CFA2C865D36DCE1EC95261B24B0E77BFD5D319D37C9EDD5
+D9E8CAD326AB15D5DC2697111E8942E995D289212F2DFC9AF434311E5E45AF5A
+5815BF4BBC668218853FBEBFA3FAB8538740F5FD704DAD0078190917498D4574
+C0A130EB89C1B0392F51811C09A26975738C77DC02742E42E3716B0BE137E4D2
+44CE2CD06A1F6D1103B3A7914AFA61979E74ACB0E9D755D9C5FBEC35C4EBBF48
+5244F338B9BC5AEE9026CDFDDAD0C88063DF27963EF31DDE1B835433FC3A7529
+726F93CBF249F7978D632D2806C894327EB5380904AD250DECD6D645C3767913
+F970808A9AE93E17C48F62667F88A3D2A8AA30EC2A2306F7817C1E19CDEA0001
+E9C5ECFAF9607DFFE1F65018F9492EBF28D4D2586231E987DB045551F40B49D8
+E5B0671F0A9437C2A1266F3F69A45D511FDDE55B68D201028F91D060E6A9E799
+2F092E69A7A06F399EC11A68EBB4950CFFF301F22541605328A5A72EC2A3051D
+57AA161E7EE53E2390FFE8253E705A834AF5BFC5A10B6730E70CDF50238EADC9
+C3612C8082E122252F3BB5E57EC5F0EC8A5261928AAFBD5CCA01B113BF8D3F07
+736F00480E0989AFE4BD305A1093FA7FC782B9660C374CB7E42E194F5CE95CC2
+A85115C183245F4BBB8591E827FD4E6319E11C5EBD01B33F54957B7B98D4DC6C
+29E62B36E6169677B7E1C274C21F5D5345929D907C173DB641FFCEF3115C9DD9
+8E33E83FD0A9F015CCE3DEEA8309A6039690CBCEC2E9EE5F84056F5D345D607D
+C5CE9D6CA932E7E87C28260BCB0287BEC328F52FF3299A5DE62DEC8AD58DEA3D
+695C3C910E8D26BE92B1C7BFD5E626E43A0B99B007A59BAF84C19AE0F2D6D47B
+A9FD012AFD1D0ADD4E24E2C719974D896AD4BE86B3B0E6B0F1C766EF56F99146
+F510FA63FFDB3B1C6E5C6CD67207C33816E1D645F6840D848C4C292530B7C6BF
+65C7D9C0CE5DF6A8FF3AB4376BEF07DFCCCF2CF06C7FBB4340CF234902E710CA
+ADEDC285BCCA4D759FB4E7840C599933AF4AAD8FA216FEFD8E68F9A5A7D838AF
+BA2E0831657BED7C672D1CC575DF29FAF0EEEF3A095F86BF51B531463D3E571A
+4CEE1225008DCEBCC332CA5B962591F52D22C9E01FC4F1912C8EB01F9E8A3B88
+64D651E3A46ED42FE5DAE72F324CFFF7DE025B69EC22E30BFCD699B243BB339E
+11DC361DB1CF16702AD628ECF7346125A6056320AD68E254D7DB2B615C4CEC76
+8D708A5A6CC907B1AABC85DC013207F0865B3A62A9DAA0021810D35FFC32B843
+083C1EF45BA5B8110A60B8A3DF4661236100BFCE351D222825BEA451CB05D6D7
+63DE40DDDBE3D3F5BEAF675B32291F04DB986015BA826D2A1DD8AB9119612446
+2F086EB283CC9D2C5960B2760975A5DC25BC300252BB90A6800C5BE349D09782
+DBD43EBC5A80FE1793898B1691D8958AA1D491FCA440312759BECF36F5C2CD39
+5D6DBD52D9658778516D86ABEF80AB609457136DB3EC59CFB47E5A374E568C95
+55C0B48903C785506A46E8B877A347AA4189EE02067CC0F5630BE279F78AD656
+2BAF25C3218164214FE8F4409F6280CECEA06F04FB87E54E685D91857D6876A4
+176C5F9519C4696D04F719C818F68BDDE2021FBF53E1557E813E3374B52507B6
+EE20D0387092D1A771B4887EF6071D26F54A36304B85685AB7C007FF650B50A1
+DF0A3B3CCD1A53E4AE1B21C21911A8094BC6446AF117D07580D568F3BFD2F3BB
+DADBD36A9AA31C367E4FB2539DF2AAF0DC7D4420CDFCE2EEB6E6924EA9CFC9EE
+091F6911D69E00D110040F773C2655B24E7F0754AF5252F24E7EBE2981A1E9D8
+0970F623E3B1FC1E7B8BB31237ED2496BBAD0A0B4D6A7B78EDABCD9BED0E03B2
+0EDC8F63613FC9AFD0FD19F35690CBFE3CDD42C5B3BA5159B111A5503C34B444
+0B2019621CE302FEF8E5A92A367E0EE70BA0086C763FAC7E62121F50168C69E1
+8FBD6D38EBF758680FEAED5273B64EAC04AC46157BE365B2275AFBCBABD81065
+6683BCB8FDEE461D3A2340B82A4B3B20D120B400A15D894FEE8A0227D2E9EB8D
+3233D6EED7B2EF5539DDF083F990FD487FAE6382B14DDE7777D6F0AB5AE95D36
+C8886F6E3E8BEED0D01CD515299E1293F3F47DC1652093335EA77C50047817D4
+EF1EE91C43EAAF79A10BCC
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMMI12
+%!PS-AdobeFont-1.1: CMMI12 1.100
+%%CreationDate: 1996 Jul 27 08:57: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 (CMMI12) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle -14.04 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMMI12 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 0 /.notdef put
+readonly def
+/FontBBox{-30 -250 1026 750}readonly def
+/UniqueID 5087386 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D03D6A8F05B47AF95EF28A9C561DBDC98C47CF5
+5250011D19E9366EB6FD153D3A100CAA6212E3D5D93990737F8D326D347B7EDC
+4391C9DF440285B8FC159D0E98D4258FC57892DCC57F7903449E07914FBE9E67
+3C15C2153C061EB541F66C11E7EE77D5D77C0B11E1AC55101DA976CCACAB6993
+EED1406FBB7FF30EAC9E90B90B2AF4EC7C273CA32F11A5C1426FF641B4A2FB2F
+4E68635C93DB835737567FAF8471CBC05078DCD4E40E25A2F4E5AF46C234CF59
+2A1CE8F39E1BA1B2A594355637E474167EAD4D97D51AF0A899B44387E1FD933A
+323AFDA6BA740534A510B4705C0A15647AFBF3E53A82BF320DD96753639BE49C
+2F79A1988863EF977B800C9DB5B42039C23EB86953713F730E03EA22FF7BB2C1
+D97D33FD77B1BDCC2A60B12CF7805CFC90C5B914C0F30A673DF9587F93E47CEA
+5932DD1930560C4F0D97547BCD805D6D854455B13A4D7382A22F562D7C55041F
+0FD294BDAA1834820F894265A667E5C97D95FF152531EF97258F56374502865D
+A1E7C0C5FB7C6FB7D3C43FEB3431095A59FBF6F61CEC6D6DEE09F4EB0FD70D77
+2A8B0A4984C6120293F6B947944BE23259F6EB64303D627353163B6505FC8A60
+00681F7A3968B6CBB49E0420A691258F5E7B07B417157803FCBE9B9FB1F80FD8
+CA0A265B570BA294792DD2FC75CE2C83DCC225B902551DBD11E687EAC6E85D2B
+02C28359A40AE66A6A6A8862CB17815B41E280313F0EFAA9981755611F7F683D
+35603984D60BB0C772054355A97A5E03C689E23B04DA79080CE4579CC90EF38B
+1A64CDD92B907AE83192C3C46C5FC40BB412F6656DC6349E6D29B5936DCE94CB
+98E3B465FFF7574095F57BB3750F1A525FDE28FB8FCA4F0DBD8BF16A3F1FD9A9
+CD52473E23A90E6632D39EBD39223004B2382EA7B1CE8E90BE77C69CB5B9F659
+009CF2105EBACE0CF9BA0295F1937CDBF7F718B6B85C984490D1E560E9989B89
+D7905696D5F64BFA604C5A93CA60A8C5C4A5AACDB6EF5562BF55E551CC048835
+68743F501BB40ADB8DEC07D5975369A720A6D249F9A5AC91B1ED3896F2C6874D
+6C2EA8F580296EE177A97F5C50FF92FB804C7EE832FEE09D4975CE752F2D3CF7
+66AF8F6748D64821EB2DEAD1C5E50275E46460F6DBAAB372244C84107B983B1A
+5C7863DDF3174A09E363E92A1EAC874E4BF7544FC321A653D2995F020A0BC9F8
+959C4ABDF4FE7C245CA0CEDD499A368E7D55BACD328699AA3E465C41C19E8164
+2B60DDF6CA62288E4E780F8614C7749C9406325C827EC8A67B3BA58728304A89
+D56F596F0E1E2A35E4D40D70D5990DD9B4CA9726CAC5224C853F2229F0B337D0
+C9BECED9D75C25F5CDFDE290494439ED79D6305E42CAAAB1CD158975FF45AFC5
+6539492DE05FBC7E5D5A2A65C07D4FBCBB8FB94BE05DFF2B3C136D05597DC315
+ACA725635AA768294B6319A6506C5CF42A65CA282CE66583F82EC2C242200FD1
+6C0FC15C0DCB4C59BBAE3A0C3B766E84C51DFE03A5AB4EEDFFA3A48F17115F0E
+B2991D13D9CDF6E18FCCE668886402A6534071AAE50D5ACDCB2163996299FA25
+E75395AAA60265F5F9C4FF8CAF1AE0D867CEDDB932C05F604C97248C54B5D684
+D444577E6AA72F0BD08EADBB9EBE4B1EBF7091C0DDC1AC4A2265D10F6A7CEC42
+FC81D1F4CBEE94ECDB043188968A9B764661916B832E25CE49A8B04A9DEAAB1B
+74EEB7B908245CDDB7A093C1AA3078E881AF16CD559F65C3FB88475B20A342A7
+D2D6623578765F8E3699C303C9830CA8AD726BD1E09FDC90F12AC14EF20726B9
+72C731CE019699113C009C09E30DB38B1B301F105C93E25C7CFA6A6EEA12BCD8
+F321C3B2A0DBFC4809AFEB576CEDDBFDADEB99A3D372154801D0D805E4A7DB68
+DF780AC0F0FDEEA064AFBAAAF136F6723174414F05A93B7AE16CB40DA8C9DA8D
+0B4A7435075F9D69574A0C84EF7D0C0A78D24BCAFAD53F634FE29D5697CD01F9
+91022223D03A5F3EFC2F3D86EC4BD4B3929DC04C21A98E1D513948B7DE7D312D
+2517A932BE08DB51ACD11007AA194D3CD05F37F438576A773263485C12E6155D
+B7DB67F081430DB01F383937068CD274C150E1B8AA9E57D02C1F40DD537B72E1
+5DED842BAD6EDCB212D5ADB51AFEEF9E7582E79A147D57BF7E3139F4B3C8123D
+2F5C17982B8B744C9ED014BCEDA12390611DC885626AFC8538C1ADEEF2966B29
+BB6401A4DB94840169872A715D0016C6EFAE68F28D0FDAC830342469AA657EA4
+1F29AA9ED11B9B867C62289B740F40EB38E93B738E6D63328D3A9074A17A2C09
+4A3E7A78C1C8EECD5DEC2A474476C58009BD295F9D9E809E7D9D355B321A0C93
+E5B850FC4FD48BE9E990937F73D6FD830B08FBB36B93FAA9042DD30F34D3D73E
+9D0470832920A206ED263B64366AFDCEF245A6F1471D408EC3DF742F207A6C03
+E8529C48F18F3F45B380A17E6E463455270B3804D1580B3B3C912E99F6BDB738
+90310431B5235914531465DDFA1C281FF1B14EA956538AEF8E3C834935A5AA59
+35BEA495EEEF110F88E32436635110BFD83E87566EB86F195D178DB1AA17A658
+430FC620AF24EAFA7E04B1ADD4D278F02116933E00553D48F6C5309A317F5CC8
+23D267BB4FC6E690C02A85559FC1395372F20B5E94A540D1DD3A0FAC30D7A9C1
+1AF7BD0CA0746C30FCCA90F18C77D925D42B519BC6A8A7953C6B8F39D574798A
+F4748F74C50112E74092385CDB2671C13F560AF894695B0AE2E6CA4921D9AA39
+D5A05EEFA079306C31ACCAD6C93C0612CEAFD042F73E9B914409EFE5EB437B9C
+91B2826084E5EF932FDF538E2DA4CEFDB262C03BFE9029B3C1B13DF4808B1C1D
+5A20E2925A91C134089A9F7420AE7228FE6809263DC9910DF38D8C619D6DA89F
+EDB6DE40C208D5C1635A6AA6ADF21ADE74FB9C8161C71B04AC1D980B0E53CFB5
+07BAC208F07CA9B23B0EFE63927395B6602A95FE14E435F4A37805CBBA6F9B03
+037A58B5FE08E76A322E1D5163D55885DE4C2048C76FFE51ACAEFB5926BBD6E0
+DA9826016BF125909B215037FDB8E383AF4102A6E71A28CF70B642F2B3CAAECC
+6C0B89195F5CEA39FD16FF9C75A185A22F042614513829268295CF429D4E4F2A
+FB638A78265B8835EDB72EABA1F9A2D95D817F9EB7D9897F01260E957A7175B4
+C9EB4D27F2E53EFD6A9743244D04A8E7F05714F6F9F7169020E6E567F3928070
+B865352893D17A0B71826154CA2714C1291BC04F60D62D5F7AB09B8A13D11F63
+9C91CDD8496A128EAB280A341CED4918C091271694CE819432E7372333C188B5
+C5AA47A39B7740B93D1569D16F23F4E2DDB59158B3CA94B2E04CFBB4E8C8BEAB
+0F3C872C815227428ED415F4CE2CC9A1DEB3C082DC9368F2FB7681C056662EF8
+3FA26C2CD267ADD8AD2923062D2AC80052DA6ACCD4283F27D9D0C5513BC7F3CE
+DD6C771F8E50EC29BE58484968832F5C1F163F753702B577DA808488DEF8DB90
+04CFFB431163D169BE5311F528E6D5A48F5576AEE3D0B7210FFEFDCCFD6B9CD1
+1B437FEA6A3FE0B563794EA424C59F8664A1E08DFFC4F94C3189B8D07CAE7F15
+A49FDEDA1D4D6AC6E90C42A960C469498459DDD922971E8DEDF1079E15FE1F46
+DFD74B0D697846D4E373838AB59EF1D82D7D739F3204E5230782D9B2BBB5C6A5
+8B41C50707B53FB29005896880F04CB2351AD1013C258DF77E9056193AE22780
+3D422BB21A55C488244B73C1A4A7EA54AB24614E1507258686C3BB3C9D461EBB
+491C211503068B7EBA700000305757FAEE10251CECE5DFC98C355CBF29753649
+FC90089F1E12E8EF350D06A8051A15E59790212519B5C83C8A0829D35B5CE001
+65BADB38FAB29F8F159DCCC0C0D8D049C35F787AD7312783F9F3A38A30242A41
+331C327903EA3911786FED5AC82933A4388FD65D7A4AF9B626595D6DD8C36547
+278ECC6A008EA12CF589EE3B310AB8BB4112BF8CC872E053DA11F406FE6EA79D
+57886D1716E84F5CEBF78F7268CB4D32A27175A6B60BEB1181A87727831C1B34
+965458F141151034EAD12FA27E9B9D9D8856CB6B7B828B836792899BA0CAE999
+899FFA1499F7D04B9CCBC3BFAAEAD54B4032337E3E6C072128C9F9696C76CDD4
+81E5227E69F0DD18E2D266635A11F1C2B2E935BD1CA02ED6DF48E51B67659D9A
+45E7351DA161E55FD4D75EBC90361485A352D07BFB8F7F43E609C46040D122AB
+3845BB3D82DB86109DDE04E69925F82185892AE63390867A77D02B9BD6ED20B5
+14CE83662C2AACF619158B5985291D9585CB6F68CE9065F60A2B6A4AC0896F55
+E99F078952898DACC3DBD2E0548A3503803A91478AFC6E7CC4A0F2BEAABA4106
+20C860EEEFDB2ADCA7FE1FDEEDEFD040E769C6011E95209C8978E13350CA9D32
+B2C591BE1FCB5274C13454EE2B702E9904C29CFC05D57DAE08B9F5783D68F8B1
+ED657083C3F61D25100CFF02488EB8BE9EFE3A464D39F2FC92B701B286AE293F
+9DE7047431B97150EF0C1922A9F6842CC733CC58AF67299D635EC62F5D8AA537
+20CA49D4EE3256E9A30A6175D2F4CA07DC0D44AD2F41802F4775E508762ACFA6
+A2C33DB8F2A10C49DF4AE412200A56BF9035A7A4769E27B5002BBA697E7F3EB5
+599C6D065BC86D514902A83D2799234B3C1C9C667CB8257DA1EADB70CD739432
+B14B751DF7685573CB4B59E74D87AD8779879F42213CA7A6953A98A57402BBC2
+CCF1BCA977BC7462872347DB109E836A03D5BD956B0559D792F8F061786E71D5
+B66C0CD112D4002321B5C52B1C6EA8B86AED1FECC35CE62755551A0A26B3AC8A
+B80739D91220DFA13EA96CBB87CDA2693DD486D17E36CE79460196AAB90CC9E5
+731111BBBB2AD93D06E333941494044158E4ED44BCA2FD7D72CF35C5AF94567E
+878F3E74E668A17E5451A87F4076021D9BDA111674BD9A19D22EAC5EF915D038
+03E3E2F58769B23E41CFE21B967491EE251C07E66F855A5FD1DABE027D415580
+08A7F1469FC267772030E54CAC214F69D5BB99B48B85688E5A364D09D2EACE4B
+97DFD477EB843EF65A8302177C6D3A91A2EE8D1149E9F9FD09C2523A9F7C4AB8
+7B1249CF351CF15FE45DC9A918DAD2487EAC151A5F9B5470B0E91C2EEBDD86A6
+3B51AFAACA719DBC358FCC896F872886239D08B2E8E497A3334427D4A5B2C398
+DEC67337A6BCFBFECFCF1A20D921AABB4259A81AFE89B4B580C1C87EDD8FBE24
+F60C75D407221CE179408672D4BF159AAF41FFAC4FE92A23C11A6DD8A1BF2559
+B611FEA32F34B5E79A7FA81873494EA81ECCFCC3C3208AAF526E1FD72EA6B4F6
+395D8902DD0C37184ABED0C5A90851FB606197A2CC23ADA0B8BF20BBCFDAB7E4
+979ED53877C932686DD43756C50295A6048C8BDBF064585A613B556B97832495
+A84C403188AEB3F9D1983F98E8CC781330EE164A6B1870F1E231D3E32FDD3B40
+0C5EFD4125F34963FA1BE919F0C6D9588674140F1077558854B96ABA372D8E7C
+09E28C87CDD0A98944BBBF747489E145EFE86982F40C8EA7530B5A91961D9698
+855A949436D5E5CC3CC8ED1EBBB83ACE2522FF09742FD41D6B5F249CAF451404
+016C4206F2C0D2298D1A6ED64565DA36BB937D98DF80BA79E222EA057765FBBE
+96BE2BBD2F0C62B937612513B832EC7B07B3628E7B645211DA2E9BA62D5021B2
+916A596AC9340262A6823031CF7BA1832BCA926899EAE37B05B5DFE95E8BB9E0
+7CDC935F7A34613A9BD964576DB9992BEE8A18F3EC6D5B5AEDEF63DE5C67DF33
+1ECEAC31FB2A328A0E04706C972D67D6C471FB4BF24D0C2296BB4E98A9E92E28
+7F8198050A5DBAC2B06710C7F59571EF8BF1AB029CCE9C5B294B57421E6BC839
+0514976305ADBE6658C6E679FD72220085A5E4F4422063262459B4A240B555EC
+E48987762ADB3B41E64B8BB3A667CEFC519B18AEBADE13EB92B93FDD429AB398
+3FF7D45011BAE7E64C561D59085282B4E262AFB408F0F337D120554D1131EADF
+CBD9BF5DA98622EBC26D5413A632529407E5AB3F6A36338721FC8195A551F1A0
+2BA9F14B9540B1ACD633662D65D0D2857CCCF5C6EADFA081892107A1F0095617
+971E1CEEF8AD80D7D97F6D6528D73371D51A1D50E2D4B0ABCCEEEEF213EAE057
+34E7ADB3948A9C28D9582E51A211CFB426943C980ED2276FEC39DE88FB5DFF6C
+B5B25376237D8D97D24BD313C5450E8B0D2AC7F8CCA11428E2ADAF00C49F5691
+B9206D77353908933412D4C66C3920B2D097ECDE735C6C7346D23B92A18D005C
+A2D40A8AE89DC0DF20EC2BA62FA047327C89027170369A2A556A41BCA1FEF6AB
+E1991DC453E46BFEC9A272F4DA2E5CCF8E33BC0F718422BAE65D913E5B72D0E8
+083B1CE88FB8889CBD38B0BF788FB57724F8A58AB57AE5E32F9CA84E5C50EE60
+C0B94C3295D9085759BBBF1C7F1D9A4B01D832CAF556467AD054D0F7387B560D
+C4457FE90921F7B26011FB52C3E5F6D796819FCB2265971636B459D3CDD2D2CF
+2BD0686A5BCAC27A2175DE8489FD1C63524C6D56E2943B95EEF2AC331020C243
+7714A5A486F10D689EDB5ECD530BBB8CA492A46DDA77F8FE5EF3697C8D950470
+5985974925EB9AD2B5D7D5BD5CA31F5363452E2417BAB34D13F55F66AE145F6D
+E97EFAA58A17AC512DD5386587AA2E5B0B2D55F325CAE7D6B22CC80DF165FE84
+E5B577CFC07A9EC7B8D60BACF393BED56AC64E1827880240467C2C17C942F467
+F9DA255E9EAD277003E5607BD1FE24AE6400FC922C24083C8FD68B6D93A3997B
+CF16F3B167DC3AA5DAF2F68C386B781019D665E2A894DAAC96875007E44505F2
+38A44F198FD58A8F21961FE57A7122546D5F6EAB32D1903183A35F75579515F5
+8961CDA85B2E0C542105D15140BBED60730C2146DACD0A484E6A4672D0CDD5ED
+028DD6BE49361BC87EBE7AF022FD60DC1CB7E39574C6C19DBF7758E3C60BB1F0
+ABAC9DBADF4D9416555630D540133CFB0E371FB8A268F5B055713D00FA48FB41
+F9F1E586E5DCFEA4CE89209BD4EA76ADDCA1DC3FC10C874FA4194801156666A2
+BD71A0D59764B595EE4C6B197C08EED7579B85B5800FACCC85D7165DCFDA5FC8
+FF25E237B7E19004524E7052007171B0DDEA436250351A5BB0CAD870D59D308C
+88347C7F5B4B9C61FEF57CE6ECD702202861C05062A355C63AFAE51A4A53E682
+235F23142506EF797F86349DD9D1B744BB5C101BD76B65FE9AB98D835B35ED29
+AA4667062BDA956FA6DC687FD0980FCE3E7D8CA35DF28686FDF0B66D00B7B17B
+880A9BB199AA4415EABE5EE3C7ABBB298AD252FA077F5D499F865D7C3BB7428C
+F2A57E23D62E3E6408C41B0A25DB06C8ABA05F7DE811C9827D977DCA19B4589D
+C7E33F0DAEB66AE57F31C1444CBA850B19E0C1F4B909A228DF6228B0EF4A630B
+7ABF89866837A740C3B84321255FEC85E3D8818829EA4268D56C7A043BE05B35
+662E4FCDFCC9D2931944977E6505AB0E6B6F35BFB12823FCA4FFE700B73C1FBA
+A217EF67F1422A62DC24AA3D6907382EDEE0688CB40DB54315132B209825B621
+0AB665D5CF74DF684FD7E06201B404218C91BF62012F230F1D49F365CE60E66B
+609B054EEFF4C271E9CB3C8634AF050D24583AE5400807ED15416940AC70892E
+E2A26AF83EFC90768B72AB256B399D782512622FE84D305AB089A1571FAD33F4
+D2C82D492549EB70034032898AF950BD8159FAC7F9578A8F0D0183FCBE85683E
+FBEFC3C013FA42CA534151C30A0FBC43BD37DFB2B2EBED0D86E1B643AB49C78C
+B8A8FE79AEBCF9DE2C577499C895232D210A36A8136D4BF29F04FBFC3CA9CEAB
+0DA5B89AA84438B619D92743F1B2D402108746C7A9B9AED068BCC8CD3FE8CA85
+9BE65CFB9DCD8A7876EE15E091F318CE4BA5072EA842285809C20FC0165B6988
+4159CDC475B2B03B318FB442A425871A99C4DE7150685A600377B623EAA013B1
+61C2B3E478D09E210FB7518C2ABDBF0201F03669C0D354AC24D1A2F82659978C
+A7A4FDBA9C66C380F67EBDBC20EDDC4124BB1E9BE310024551CD7FAB6D748E6A
+51005EB82A52DAC7747C4D2D589DDA746407865524E1119D4114FDE37130C51E
+17A6A0CC1B2674B5283E8529B22B32A87D515E4962EE4CDAF8053DFC83310121
+01E0608E5E47BACDFEEA8A26AB5DB40D2FD00941B28BBBB1F548E359922F57C9
+DECE7F283036F828256AC28894A39E73287A3A8AD0FA351FF28C937798A089A8
+5E59FB7446B9503E726724F88070979D95AB07A09C2335E0A5A3F3EA49673C68
+B273E056AB7FCE5EFBA00939A57606926457A7731472EB77E80E1BA182555E48
+CAD63759D59B732CDB8944932D2DF1352890D8A1F22C28D25DD7EC543282489E
+52CC8039133728490D60D24B06BC5A2D76E4155E24DB5C0880F771187815E5C2
+03421093FE423D9C00F95630D25B9FF5FF43500CC8C4CA4BB70A0C002B6E4DB6
+A9E693F6E710A6F94176DD1722BF60E6EAF7A96A67290B0DA2D2B538008C4192
+13223F71D14EDE00E1C1B0CD1362B706E6CD4E3FF79FDFAB0AE7E906AFAD5E29
+093210FA9C129EF6CF9B6A39F879FE40DF5973518E88CAFB57CB3056EFF120C4
+2FE1CD8193C7F177C68F24A00006ABC4868FBB1FBB3A3E2B0063E2178B5BC4FA
+99977259BA94A75A657986483C425168340A7CD62C45D203882B3375861104BB
+46A26F492B1B5B2AB757BA6C5ECC8611E6BECBA2CBE4EE2EC6E6581122E9D05A
+706A3C9A9BE9DFECDD8E7052B188CD21B6D80E52A9BF7D13726D230660DEA9E3
+3DACE3480A24FD847908A52F9E8D7602DA473FD216AAD2827D22CF909AB21DEF
+3556BED769D8F0CB0579C095607723DB09E30E14803DE7DC635842F400069FB4
+A846D7C4D83E52BDBA554121290F76AE1200F2900E0F90BFC81EF640DDC111C5
+6B97F67BC0AD181F031A659B90FF37574869C2A6159B063B6B4D130084051406
+250FEDE2162EC5C27A3CAFA5F9BA0EDE1D2A1B23A4B72C1C3C010D27CEC514CE
+9A773DA82A2D4A4A95DDF3B66D1C074ED8D39CF3484E1259420D43DEF233BA08
+BB9595D9FF785A830EFA86EA451CB47D4A5529AFC3853090CDD68B6B144F8DE9
+45667F91417843B6C599C02355E650D7C73EEC61156D6EEB3BB310B1FBB444C3
+3B0974A987AFD8115FB870F9EAE580DB8080EF5EF316F73A7071D0DDED532F88
+AF2CFCDDBC680EB5DD708FB070C44C54DE8F60946E0A637651F73A349B41974F
+CEF49787E442768431F9CF24936E3EA1644A701237A486AE1C676D3FDA92D0D0
+BA8B4EFDA8F4E2CA9BF95E915338EF0B71266022C5CFC338777607A38E0DACDB
+BFC1A46C1AD9216925333651F28CF857BB10022A585646152238FC1DA6A16F08
+983EC45B808B60012C1ED1CAAAEA90513CCD663DB1DB1E75ADC05F1AF51D31BF
+BDDEAF4B6C2A6A775D74BE5CB1043471E38304A02CC73301F5FC71D02A412829
+F1AE76628C0BE7006CBE98554AB631330A13B327799B83CD8C12D652C693ACF0
+242F727E4AC4E1EEE393904A72C48EEEAA45BB830B578643FAB0C227F6EEAFC1
+5295E99902A3DC49603BC567D3C0EE700DE1C01BB5721BE9B112AF5BEF20F4E8
+DE3E34640FDFCDEC5B8C3523756030B13B1896FF20B718BC2F59EFB980609129
+9517B94BCD0972CCBE74BECBC818A2BDFE69C8285774C4E8274605AEDAC944FC
+5B310159F5D8013C00C9ED1C6A34495E3923F92827150634F018FCF91978E5DF
+DFE920B20B1605358D6447F9C0660D03DA8D86E98711465AC82DE46C9E20EE82
+29755663A47FA70A59022342121396380C92C729AD26288830D12BF23DA68FFC
+1D84C22AC369E0424EA3D3F1CEC67E97C3BF268FA906DCEBBA1766562389D568
+CD390101FE0E089250FCE948ABEFF7E929384DF332D825195790B560E48A904B
+F16FE927422981CF50FF2E4C8AC390B36188F11CF80052DB5288232542D80BBF
+BE42CCC5C19CF944BF25F33756728B61CFE5121DEE2727DA82FD1F4F87CE4969
+3ACB1D9D7C1B205C1D499AF032D0B10D93269296B60B91A3542EC535DF279514
+2482D18AE45F6DCC0EBACD44BCA8E157548A618313C215F1F92E13B52A87D4A0
+2A28BADFC1108E221E5F7B5C744EE8CE33015CBAC4FFC2D73E5C23635467A3D4
+3B9D83E37115F60EF19F043A527C0E2A91149CE69FF10977E60D7E6DB95302FF
+050A6A63B99A3E47A30EF26D394B95E6088EE398BD0220D0FA9E8AFD9D046FBA
+089FAEEFC1FC84601C1678C2CA4AB6D8E2E1B19921AB25BBB7DAD0A29232A36E
+CA59D24A4832815AD346019D0BACADF30084DAC4AD99D33B41A7C133B1B0ED78
+153245F50A51FB230A324805DDA67747A147362041CAFD2F957312BBEF554E4F
+4D90F13068744FFA4B3728579E224249DE3692BEA1B4F238CD82E6D0D809984D
+2A9149DF0A028675659C10D0FF75BA77CBD0054A034122DFF0ADD9DB7C46511F
+0113B972BFB88C1DAF11F901C6E6B0D407566F026F3FC4A703511C6B0E65051F
+FA7FD005BE83A05138F2F1D24B2BE78D48C01ABC6F195779B6900643FC8AFB8B
+6FB2F0F1C09CC07F80A1F62FAB45A7BA874FD6A899F9E11C5155AAB1BB1BFBE8
+EE8A5BAFE8F2EAD4A2983ED2542AAEFE0BB1CE85DB67316BD7B05A8C7503ADB6
+221BA9587108C6E7C7A6130C0CB111DBE7E4D6A71E6D8AA4C160F3BACBE9574B
+3932BF539AF8ABE02C42E2FC9BA9E99CB8635CE83555147F903218BD64972B64
+AA75FB6B2327345805A9B464351ECC153588CE54413A5282A1315FF3285E9390
+80BCECB7C22710261E08537F6C33824520522230D20BAC10EF4B7F62B1CB2AC3
+3FECF85811977E91B75A0580638B60ADD8B0C9F1895B84BE79D2022E9A2FAB8E
+2F7B0086CE955F5E3562AB6F29AB5FB1C2C93CD4CBF1F71439E4C2FABFEF483A
+FE4B2CC1651EE52359AA1627D0BC21E94D1481857A8BCD37FB1413828C99761B
+858865D8EF52056A021139559E8D02A63829A0C6C3637F3133571E892EE2CAAF
+861CD99ACC70F0F2ABB9124380D3FD9F7F1BC0A2C4F03B1E963BD56EC633FA06
+B6B69DB25EDC510C3A3D4B05825DAFCDABCC2E86EF47396D673E98D7A21460A4
+1E97106A29D01736D7ECDE1ECB4A734ACCBB4519BB1F553A0BC76A019846656C
+11314F256E02AF9076ECE402B54DFF02B24EA882D601E4513E3E42A7254D0DD5
+DDE61CDD878D251B09EB8D946967349AC708C2CA49A1DE36B7B090865F5EFFC2
+E4975FE03E457CA278E0A5D71E0C794955BEDF0EB08A41A278EAE2DD8A29E505
+0475DC6A9B960FB489DADE582158905E3CEC6D2ADAD84E5FBA48F39E4CB3914D
+00B6FD106B0BD666B451DE1E305885B21B47A271F2CBE1D0484DA381D10396E2
+9C105FF20694FD1955F35E394526756DE8FC6C1BAC431A2D69A8C82B3854112F
+D9DDA32C6EB4D44B312DED7B0B75BE95C6675547AAE74B4E89E528A72DA5C1B5
+7EE26BD3062703E9EA1D388804331482E04927DBD05157B1896E01BD1DB8836B
+C265561D244B3E1BC4929681C61274173E8A8313E44364963AA3ED8144A9AD75
+8899C9EC8A3921D13FB3AE0CB24B44167ADCAB8D23EC12567D07B2C882ABA823
+230D2F0F2620FB6FC182AD6C26618F6744B880DFEBA2A6C8A29589AFDB3E1383
+B9928C728A82EDF2D35B06C8E0A90862F90C56558342D5DCB9285FB0D61C9AB8
+26814F5EBA1EF483F52506CD21305D277EFEAD08B004EDE0E7D6B5576975E496
+775F6C2457B4E233D3D638339C996AD60C1E47701E02B8502D5C3C092A526013
+02F912F9F25CB9566625CCBF957955E9F92F09853EAA842A83870B12813FD6DA
+C2243E273CDE50DB9523905C29BEE6321A1D7719B9F63AEAB36CB04BFE378D2C
+794A95FF387333008402119EB82F0CA1A21151B9524A66A3C81008ADC05ED8FE
+B9A583CF80BFC7215774508993D0EE46DB13603EB587124DD0ADCBF300012017
+472C19F111332547559B10773FAF9E7DC9EDEF2B44ADE1E25A72A2C22F14C709
+9BA62F470561EB4E2518ECE2E5C27FCDACD63ED67109EB0BE3A5DF06FB91554A
+1AC155C990E299D230B006CB40C881E6D3D11642C8B050C87B71CA76B532B1B4
+6B00DE4C5AE3CA8C29BA9972FC4885A1B0F93F898493E9C213D0E411C3563529
+A89E7CFA0AEBE73A6F1CA1851C7C529A9AF79621CD69A47144CEB77CE9D6820A
+6F9CF55FDDCE90D61EE37AD3120A1D83159A54DA67994FA120312AC6D0AA2F04
+B01E7B59BBEBF7670EC00296706235B72478F2618C5AB064A80C852AFD3CE8EB
+2DACBE43926F60D21D8DAB9C9EBF3FAE1550BB4FA2094D5070FEB8BE18DAD57D
+44A87E339425F94350E51CD3779968E7CDDC0F05750086D69A86100967D6D1E7
+908EE83F31857A27C0C725901D6634B74B4E0453C27361372D4C78EFFC245AB0
+8E59A9CF786F57B0081946634BB17ABCC31DBFDCD123652323FF192430631D92
+A994339DB0BF3E94554CF953139A23F0010CB8391E2F96690A1A17248AD43B72
+F4D2FF8B06990F191001B02A717EB27A50FB5E5C4F1A681C9FABE2E19FB8B391
+A94E233574A2B880E43D364646B254C6463F4E8FA104FC0BCBA60D8E000F9EDB
+BBEAD68A42CB1D669659DDD59AE8E16B
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+%%BeginFont: CMMI10
+%!PS-AdobeFont-1.1: CMMI10 1.100
+%%CreationDate: 1996 Jul 23 07:53:57
+% 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 (CMMI10) readonly def
+/FamilyName (Computer Modern) readonly def
+/Weight (Medium) readonly def
+/ItalicAngle -14.04 def
+/isFixedPitch false def
+end readonly def
+/FontName /CMMI10 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 0 /.notdef put
+readonly def
+/FontBBox{-32 -250 1048 750}readonly def
+/UniqueID 5087385 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA0529731C99A784CCBE85B4993B2EEBDE
+3B12D472B7CF54651EF21185116A69AB1096ED4BAD2F646635E019B6417CC77B
+532F85D811C70D1429A19A5307EF63EB5C5E02C89FC6C20F6D9D89E7D91FE470
+B72BEFDA23F5DF76BE05AF4CE93137A219ED8A04A9D7D6FDF37E6B7FCDE0D90B
+986423E5960A5D9FBB4C956556E8DF90CBFAEC476FA36FD9A5C8175C9AF513FE
+D919C2DDD26BDC0D99398B9F4D03D5993DFC0930297866E1CD0A319B6B1FD958
+9E394A533A081C36D456A09920001A3D2199583EB9B84B4DEE08E3D12939E321
+990CD249827D9648574955F61BAAA11263A91B6C3D47A5190165B0C25ABF6D3E
+6EC187E4B05182126BB0D0323D943170B795255260F9FD25F2248D04F45DFBFB
+DEF7FF8B19BFEF637B210018AE02572B389B3F76282BEB29CC301905D388C721
+59616893E774413F48DE0B408BC66DCE3FE17CB9F84D205839D58014D6A88823
+D9320AE93AF96D97A02C4D5A2BB2B8C7925C4578003959C46E3CE1A2F0EAC4BF
+8B9B325E46435BDE60BC54D72BC8ACB5C0A34413AC87045DC7B84646A324B808
+6FD8E34217213E131C3B1510415CE45420688ED9C1D27890EC68BD7C1235FAF9
+1DAB3A369DD2FC3BE5CF9655C7B7EDA7361D7E05E5831B6B8E2EEC542A7B38EE
+03BE4BAC6079D038ACB3C7C916279764547C2D51976BABA94BA9866D79F13909
+95AA39B0F03103A07CBDF441B8C5669F729020AF284B7FF52A29C6255FCAACF1
+74109050FBA2602E72593FBCBFC26E726EE4AEF97B7632BC4F5F353B5C67FED2
+3EA752A4A57B8F7FEFF1D7341D895F0A3A0BE1D8E3391970457A967EFF84F6D8
+47750B1145B8CC5BD96EE7AA99DDC9E06939E383BDA41175233D58AD263EBF19
+AFC0E2F840512D321166547B306C592B8A01E1FA2564B9A26DAC14256414E4C8
+42616728D918C74D13C349F4186EC7B9708B86467425A6FDB3A396562F7EE4D8
+40B43621744CF8A23A6E532649B66C2A0002DD04F8F39618E4F572819DD34837
+B5A08E643FDCA1505AF6A1FA3DDFD1FA758013CAED8ACDDBBB334D664DFF5B53
+95601766758D28AA39D6FF7952C9CF80458BACD8BE9F676B2D79F89AF3AECA49
+FF0EAD54B11C7245097449D1931B719E828B09C0401C011AE1DCBDB313BBC8C5
+48130AF24A0969DFB53344C8A7C4F0ADEB3C8C87312571DB58E804DF5E773235
+0DF5A07883F33603F1B1FE26A37124F0D26B911BBC0F5D294609395C92AA347C
+655BB711D63FEFD58CC0C617CD3BD9879A68E2C3B912C476CC561D9D007D66A7
+FB2006D9EA7027B292445002420703C8A6E0419D3AB9ECD7796861ECA6F39E44
+234E86830CDCC75E06AA1D4ABBFC66A23D25D432ACB057B1C37F481CAEEFE071
+548358246B1443338BEB45C86BDA17C2B66938DB0A2B63BA76CCC63090764C00
+5CDCA0A6F87E06DE95069C7855DFD334D509117033218BCB84D1313E21EE3323
+349275E9A904B639B654FF72A9A78DA234EB42FF76EC262E6E0FD37E66ED3A5E
+57C99C0AC2A051D2D2737052D85A1A9EF3607F14CA8134304FC7E3CD8F976808
+7226B96AC9D105A92568F1B27627E1FB3FBBE3AA27407AF8CF00B1BB30F5E800
+73DB225E2AA5C4F83F4E3C6F11BDEC3899436499334C6603BD8083A6B322AFA2
+78C674A79D148517FDAD94A67CC46E079940CC82B6AFB6ADF73AE1E6FC5E33C5
+791940C072CBC565BC100A3ADFB14AAF4A030CCDCA1AA2EBBB83872B6D909B4E
+43686A71A987944CB71AC052823FFAC41725525D4DD7258F9AFCD753D81756FD
+2A9050F4B765E258BDBBC5261B3CD7735E0CB80A60CDB4704063E50E195D090A
+1565D9271CCFEA85C95483A9DFF9C3ED8043D5F7346A25EF015A53266BD31A6E
+3F0F2C5E6C9AA71BE40EB7D6958F6517EDEBE27595E418EC7EC8B7388490A0BC
+0FE992798E466C843E39D582E1BB1F0FA3C94FBF687980688F62C949A4CA13C8
+88AC84004B7753BC38283CB2C7EBD8B0A454D21868D7D099EDBEF0A718639408
+059ED88D834AE22C5396AB5B43F2D6B9F5D79611849CCCD1F9490984337F3C18
+43F0B5343BA99AC7DD2EFDFBA09CB720569FD473A96C8B4F40E597818FD3E6A4
+B211CC867382B1CE44736327D4130133E260C2556336FD12FDFD743012BB8798
+587B4A491FAFC622ECC6B32390852AA419253E6522DA0F9BEBB675460F9DF0BA
+494DE6F876768609BA7ADC213EF25847689AE890C3B35853097673A9EC89AE2C
+716A6F1F35FD69BA6846CEE3709E06820F9D6DEB78831C14BCF82A20508DE487
+8A9CCF42BEA79279B7180652BCA734EC8F657F4A8BD9015BE61F8BD7620F1031
+D96D1CE49AEEA9DA2E4A0364C685FB146C2F0183672E7804E9AE51C4003B8431
+B17A7BFB4C094FF83804B9F9BAE23DB6081EF1C799E33708927126A933D7CBD1
+311D3A21F10F2E65FFD2EDD324F524FF593720C30E0A1A9ED3544595D18F9E82
+885A0AA4B2A8D33D623D453DED34B72588ECA222A54C05E97DAEA8E2A11D0EE7
+EB5E4BD5F43CCF045E09892DEE140AA2694975AC9866EC303AEA2912133DF6C7
+24B2093BE47D28389FBAE9E2B4BE669CDEE068148243DC9D3ED36A9AE449B7B5
+95DB5815D7642CE1406A73FB777D6E3B4E76B195BFE7DE8997895A12EA03E3C3
+9326686B5A6E3A3D977DB8C5E2AEB5BCD5A1A076C750D9C7D1096948771DA578
+58897A9685EBBCCEA69C886BC304844DD8C20A398611BC105EE2A785CFD1474F
+B47C1FC5AC8421C2454AEC68494A2DEE8F9671DECFBDC69C5CC72950BE6E0ED9
+C69F11B65EF7BC65AA9A281DE4B9C468AA4B9C7ACD76FCFA08AAA8932A644735
+EB888350AC067BC06E41805FF544A33C10615C4060757A974388438988A309A6
+4B12A05503F5CE5DAE4ADD212DB7D9E2877CD8AC9B3C9C9AF2B44A91429E35D6
+BD77395B3A2D83D316212FBDD67A7100FCF8EAD473DC86DB2F07F71109335674
+08A86E262847A6891C51BFB0DEF5EB561782261A45C590505EAE8FEDEC8F576B
+404ADF163D330388081B4001ADDFFEAFE47E67B39857D35BF0AB1C2178F17C37
+468ED2B7E82DE2B3CE84A581FEDF443132E22630D06B8E8B1AAB390DE6B8E4D4
+3C3FF490226D850D6CCDF24A401297DF0A9737EDA9CA827B7F0F36E096ECBB73
+9E2DB239DE71F57A65E30D21E9D1A761BA216C7D9E7DDDBEE046CEC49E3B7B51
+918173EFC47B477DEA7FD81543DDE87E53943BBDDBC0872B1776E8264688C2CB
+5572236D87DD7434C18804F3337C42B390E939E8788890EF607CCB2388EFC322
+FA577F1A6D73562D0F23FBE061B0153C5A02DD0EC3F428EC6FF9443E6E92AB6D
+A3E82C259279948AF7D63FEED3328E3FD5F1D09755A9B00E3FED86626590D39E
+FB027BC1FC2963EBEA9002B483E142FF28811646D3E7EEA85256B329CBC6F859
+2E80464EB00C95AC41B2F9C9B1437D1B7156664AF7D347C3F1398EEC3244CEDA
+C27CBBB44C6EB707B35383E4BCF02E8BCFAB5BB4FF1CB19C4AA719FF129E7000
+C18C315CEB710842B6A1EAD9A0C8EB8002A1A2BD82A6DA9BB24A367095B9A4D5
+2890C62ECFAC7BFEDD0AE71F679690CB8631FB182426DADA6DEDFCB0925A2669
+3C22E7FB30F9156B03022CDB1C92FD7D5C71D4A4A4DB16F7C57FAA22A4745BE2
+900D2C1FFD557AB45AE34E6BAA68FDAA943A92FD0738D73D49FB363377481E84
+67EEBB2B1B2027B4044343797E9ADDE6E7F92FAAD9795B5D48D2B892BEDABC7C
+22552F70716CDF9BB17BABEC606C657DAF3B4093863A7440DB3B6DAB06CEFA4B
+2302B25D9ECF7B5AF6EFDE56B107D4F6CCB5CF618054983E0B38DBB27B567646
+E281C1DA9C72472C30063D65302AE24E63256FBB84BB1218A505FFCF64872FAD
+ED6D94E3D46DDFF42D70F5A6656417516E44
+0000000000000000000000000000000000000000000000000000000000000000
+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 0 /.notdef put
+readonly def
+/FontBBox{-36 -250 1070 750}readonly def
+/UniqueID 5000791 def
+currentdict end
+currentfile eexec
+D9D66F633B846A97B686A97E45A3D0AA052A014267B7904EB3C0D3BD0B83D891
+016CA6CA4B712ADEB258FAAB9A130EE605E61F77FC1B738ABC7C51CD46EF8171
+9098D5FEE67660E69A7AB91B58F29A4D79E57022F783EB0FBBB6D4F4EC35014F
+D2DECBA99459A4C59DF0C6EBA150284454E707DC2100C15B76B4C19B84363758
+469A6C558785B226332152109871A9883487DD7710949204DDCF837E6A8708B8
+2BDBF16FBC7512FAA308A093FE5CF4E9D2405B169CD5365D6ECED5D768D66D6C
+68618B8C482B341F8CA38E9BB9BAFCFAAD9C2F3FD033B62690986ED43D9C9361
+3645B82392D5CAE11A7CB49D7E2E82DCD485CBA1772CE422BB1D7283AD675B65
+48A7EA0069A883EC1DAA3E1F9ECE7586D6CF0A128CD557C7E5D7AA3EA97EBAD3
+9619D1BFCF4A6D64768741EDEA0A5B0EFBBF347CDCBE2E03D756967A16B613DB
+0FC45FA2A3312E0C46A5FD0466AB097C58FFEEC40601B8395E52775D0AFCD7DB
+8AB317333110531E5C44A4CB4B5ACD571A1A60960B15E450948A5EEA14DD330F
+EA209265DB8E1A1FC80DCD3860323FD26C113B041A88C88A21655878680A4466
+FA10403D24BB97152A49B842C180E4D258C9D48F21D057782D90623116830BA3
+9902B3C5F2F2DD01433B0D7099C07DBDE268D0FFED5169BCD03D48B2F058AD62
+D8678C626DC7A3F352152C99BA963EF95F8AD11DB8B0D351210A17E4C2C55AD8
+9EB64172935D3C20A398F3EEEEC31551966A7438EF3FEE422C6D4E05337620D5
+ACC7B52BED984BFAAD36EF9D20748B05D07BE4414A63975125D272FAD83F76E6
+10FFF8363014BE526D580873C5A42B70FA911EC7B86905F13AFE55EB0273F582
+83158793B8CC296B8DE1DCCF1250FD57CB0E035C7EDA3B0092ED940D37A05493
+2EC54E09B984FCA4AB7D2EA182BCF1263AA244B07EC0EA901C077A059F709F30
+4384CB5FA748F2054FAD9A7A43D4EA427918BD414F766531136B60C3477C6632
+BEFE3897B58C19276A301926C2AEF2756B367319772C9B201C49B4D935A8267B
+041D6F1783B6AEA4DAC4F5B3507D7032AA640AAB12E343A4E9BDCF419C04A721
+3888B25AF4E293AACED9A6BDC78E61DA1C424C6503CC1885F762B92DCAD2A47B
+6B9A8DBEF5C9099E3D87CF31254B8FE1F02269893F6ED28045486321A2C89D4C
+401189A7E2BB7D6930806A94FD603E6F41A29BAB742D3B39EE5D466F01EAE670
+ECE883C1A03319790D5485E89B210AA2DFB079A06037B590231088141C188224
+41B36511D172B85BAC8DC4A5D8923DDFCCC7FB9E7139041A3987364398C6C4BD
+906A152728000BC1BFCE959CC41A31C5F856D1BDCC9670044E649159603C3686
+06119C8A797D2B4CCD12980C936C8D6448AD5F636E500B7BCA12A228713F0C8C
+39271F7439E90049AC0BED962697D62B4FC578AB76B8ED3435466FFB06D0AB33
+99E38FEF539575C80A0A17B3234211136543862D0D08575ACDB45D4A9CBF03B4
+4EB9BC13D8285413C3CF8FE3483C5AD80C24405EB4F303AE4B53A840D8551717
+998778B4E0150C6E2903DA6B66675F5FE94E4E459A8BE4F1115BF77891002D98
+95D6CDFAEF4BE5742C180B9EFD08C447D622697B2DC3D4EEF36146000873341C
+C217B24F9CDCD7FCBBB737B9A1F9BF56B8691CFE3461E2D3AEEDE48635D3E88A
+7262691A41E07FB73F9F78FA6D4DC3C7183F3494A24FB2B832DFA7D8512EEE18
+CBBEAEDAAED794D087C950AD664BBB4A29C09BA4A2B39793494A6D9F3AFF11D5
+643EF3CE2D55D900B8A388440E9D842C7029262842398AB9B03E89CCA05CA341
+4F13583239AFCE7DBBB890F37A443465D2A05A302167008364C6B767CFF820E4
+C49B152DF2425389B70353A66A18C0DD524DA5748F82CE00650FE1D7F8072587
+6FE107E243BC593D46EE71968FF2642811F9AB469E8ABBCB546E60EC161A10A8
+C58F38D4CDF75319AFD62E478C9CC4D344C31768C0AA2F5A961B68BD7453B8B8
+E39E07BE521C9AC9CA1483358850175E6FAF41E2E5A417B83A81DB59BCB7E1A2
+1430C4505D383B21B83E51BBE0951D8C10F70DA67DCFA7CD923C1FBA6A30F384
+60986F1405C92B3CDB376B83C4829F6A1331070798D832C1AEC4F3DDE65F6093
+1BDB3D4F5A32015DAD95BEB0965C52D42492B3C4DBD5F7308E349A65C56E220B
+5C6B9F7146596C3E34F56612FB0DC2EB5EBD03E6DDDD34B3BCB88546E616C388
+B33F8116C6780D766B93BD07261118700120B21B2C1C5839D96656F3A83166E6
+873FAB12C093EBE91D5BE747F2A40E2693A7A6A839F622B24CA1AA2400C3D615
+F930A678F92FFBF61395EF917BE8E0FCF625BBD4C4CDC3A3A113AF556C34C137
+6DAFAA08DFE2E05C5B50E3522EC0CE5BF62E6670969B11B6010685253BE2D38F
+43A30648E37FE431E144BA820B5EEEAD6E5E8BCF943EB26BD59BF70CD69FF6E5
+BEF8A70815D591B4848962B1D9FCBA27DD8F14E80B16DA66104EE470A5829668
+401F537283ECBD11585A0CEDA647ED58003599FDD2B4A6833BA63AB7982A82D4
+475D2B398079D6170EADF48793A334611ABE0010E3FBCB38185C0409E70A86FB
+57D283A31176696DE8ED59B20605C4B9F3105F50FD362E9ABDC63BEDBE943859
+4FE15ED4B2D958A3CCBCFE6B3A0A4B6DE4997FAB76A2E31AB780DDF0291118AE
+3061AA19BAA87E7A31BE9BA9947C3F4D5E7BB1DF7587BE990F3E888547AE7A3B
+4C9906D15B9C1EB902219D7F920BEC3C3FDA2B2C6411776033B3F82153A3F72A
+C35F6FC174E5F25651D097D172112BE48F882D092DF106E9E6121DB4EBA0D1F3
+1A7E9DDD48D1E7A08AF7BCE12F94AFBDF788C1E3F611AC8B85A981A198C06211
+BB334D67280B3E981F090E93FAECE611C519ABB22E283033713312F9D5258BB1
+36C51A59BB216F2D752493A42C23218CBF62D844DAB3FCF38D0A01544F061F7C
+0E3CABF49097B5332A5C22A086566E3654352C3BDD609EB81C6222F5377E0AB1
+9B51F024F9A961A25C27759E2EB94E896378B189AE69CAA1A1CEDFE8BC5DF441
+536637C800FFF0ABC3DCDAE48BB9B1B98B8C1B3EB1CECABA5E9C56811EEBD94E
+7363493E005F730AED8DC4391BDC2F2B6B9CAB02437A9233BE2B3EA9C7A20FA3
+2D5CDEFF1332C0196CA6A5043DA277D95305526B4B993D8797E802FC95284AED
+D2F39E49462314FDD677169691D77EF89FF593FB70A7094F73E461148A4D309D
+596BD21562FC266E915F09187590156AFEADDA857210DB29EF5ADFAC01D109B3
+0B8F1DBE1AD241752117650CEC4CCD5DD47FCAA50F772759B2D2B97A75C68500
+A6775AD82109E38E40813FE41695D93B1D7EB8170A4C4BB08BAC9FDCC818C1A1
+ABBA03842FE42812EBCD0DEA9D37867BD1056F34A86067F9FF1E4559B1BC0DB8
+D05B7E59F3C6F8D0494148CE1411E0E15065ED1CE4AA0B1F4613DFFDBC55E08C
+6A99AE501DB257D8446F42CB04AFF7D32BC61C6746A8A82E14EAE912550CB4AD
+3D734E8EECBDC6FE06858017C139C3DD48B21CFA9FF32B713F5BFF5EBBBF7A51
+2E9040F038DEF98B8A41E0A7D13129C70005B39C780942755AB01C60F7DB4ECB
+3A363B4897EE5E5ADABF06512B8EB62FEC96DB7BD5A40493A71F88EEB5704965
+95981F7C78D2519DF1FE0FE6EA7121E14984B86FAF4CC4D6A610EC3D4C09F009
+D6A1DCCC0400BC3FF557EFDDD8557E071F619D9016B0EEDA30DF51EF3C188534
+0E2178DB5AF5E709432604CCEDF1533E678B0B8E03AAA58E35481579AC244D86
+B1EABAEC15D2B160B385779338D3738A17467C4EDB75EA0B9ABFF35485B5887C
+65D05EC069E155E4F1A3F98CFBDA0F868E20DD1090C51278D1CB3625530DDFFD
+28217B4304D577CE3EC3F74A42C5665AA40514A08DE2B590EDBA6470AA13400E
+04822649CE07453764CF737429BF36226FF147FE1EE0BBFEB04D84F2C0C11FBD
+2C7AA90CA94A8056E9CF45B93C204D21BACEBC5DF69C71D2AE65FC0D70F3BDF0
+9EC2285CA3EA182E50997B31503665D0F26EB4814C408D851FA7B8CA0A2A6799
+1C1119B911DBFE9C12E8A4182287F412FB4B382A7485165EC6629935C128F063
+785FEFE52E1EDDA51BBDA48C4E64DF9E04A1C12A0C4700E26D1FE61CED5D75C2
+D4BF8D0CC925BF50C69820C2FD3C3EF37C1AA100AD2D838824F5DDC2DC861D67
+48CF15EA2014475F236DAE3ECCB23363ADF3FB97849DF394FFC5062642E89F7F
+D057989D509A61F896DDC6036910B4D159087CF9139961B417ACAFA20F4AEB37
+B7447C9CB920708F65A674C0C23829A7181743460930A08310CFCA1FF4110BBB
+6E8AA33D0A67B365F26C74A372717BECC2F7D5A98E591DDF389E875834AA4099
+CBF8320005B2DE88D643B578EAC66A4CE4E157C5BF620A61B30B6526CD129B00
+157AA321C6A4F1A2A1AE6B10F9AAD30AFE39FDBEAB00ACBA4CD2FC345A350FB6
+94FE9EEAEEA0A71021F42B902C211C32B9F6968F6EF098C1A1662FFF10CECB40
+B219D915DEB1D6ADB07E475D82E6A6631A370328B4C25DE35B0B8E6EA1A981D9
+12823E2B2822DD9930BAA09B41781ABD8AB66705864E5FB3534AFB834932677A
+C56BC76CAEAD09833FC7E9248D9C86CFCA8A26F02B5A92C005FA08D35B1B4BD7
+D431E0CD7052A9741EAB928D2694AF94693525C82036CC951F951778B791F717
+662D25C6A91E66457310DF68CBC579818FA8DD6AE35FF96C3E1FF0CE4218C440
+9CA8E53256CF78BB85834978F435919BCBAE9DC7E9BABD9B6B0D60D29BB28873
+B3F089DC022E08A09B6C468E90210148A4A9152EC230DC8D7A1415B199C92211
+D09BBCB2B38D47A8BA2837279C5862D90C7897B1C4DF3DED8B353B68481FBC9A
+F9220D98A4747088434A2E576BA266D8DE6421B773831F637944348CD12C3B78
+C1F9FCFD0975318DCF06FB621ED1D1AF36A09B3511271DF1B9009527EDF2C9ED
+E97F76A3EB070C2F586662CF44AD631D85A09E3E11A32F24830EAA6B148F8C3E
+131407DDEBC04F78C1BCCBD59BB646F278F32216918A7DB4C322F527726887E9
+A5EEFDEC5F0915D70247CDEC0E6563800B9894F72895D20C8C19D137BF48E54B
+B02A58F4943CAC6A1AD120A1009DA6CBC9C4FF6E264C714AA529ECD30E64308C
+41B29D416AA381F3F7FFAB311389AE51A0CCCA6EED45F27001DB9AE68E827159
+90B64316DAE0467A1A518360469658C8713CA72DA0DC5DCBBBD4004C64C98D3E
+50F4752CDB7C85556D4F3DDB3A50A0C77884C7808BC5F9718302BEF4D9904473
+6F25C221B1B1D9AD73E187CF47B786A7FF29FC3843D11E77312ADBF799706084
+1A44A21873815221AFF18593750515745D05A681189D65BF8412C64938CDE022
+B230A41D0B5D143D75B2A4AED9210270B6D5B7EA4B8BBEBEACA0AD53C85F66FE
+E7AED9EAB9D930BBD6802683965EFE9E99E427DB4986AFE4C63F0429389FBA41
+F588EC41010D14B04FEA34DF800CF8138DB21CEB194D40DF97B9CBD24392A524
+06E1A813BBCBB1D44737E9F2844D7A12E56989113AD2E7E3F83F8D2E84678CE6
+2CA043ECBB799E064558132615B285D7C48C9FAC99490A593EE20C15768B2043
+62226C1750374F5F75971BAC962D862DA4764C5AD03532EEE7370193F095FBD1
+6980C098AEB5CF08F1FE63FBD7FA377282096B13DB71941F5DD90BCF642AAD8B
+04813C462D9FA42ECEA19707E7BC6506B3002A42DCE293AC81583E5AB07E681B
+2A79FCB1A08E9ACAC517E35578D1F466A60552DBB994D0573F2842C279C53424
+E19C62ACA5BEBE45E11DCDF80D7D33F73676FB112B06CF1BD4FA5F06DC7EADE2
+76AF40BD241CD757E4BA428889D7B7536F6BB53C79FB1891F5EBD56030A9E6F6
+BEAF410AA8BF02FBD0E861AF94967B47193A9029D80450FE7DAFB18CEC856892
+36F2B330802133A1EC
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%EndFont
+TeXDict begin 39158280 55380996 1000 600 600 (nomu-tcs.dvi)
+@start /Fa 254[23 1[{}1 66.4176 /CMBSY7 rf /Fb 137[35
+3[35 1[35 1[35 7[35 4[35 97[{ TeX09fbbfacEncoding ReEncodeFont }6
+66.4176 /CMTT8 rf /Fc 141[51 54[32 10[18 48[{
+ TeXbbad153fEncoding ReEncodeFont }3 49.8132 /CMSY6 rf
+/Fd 145[38 54 74[29 6[36 27[{ TeXaae443f0Encoding ReEncodeFont }4
+49.8132 /CMMI6 rf /Fe 200[30 30 1[30 2[30 30 48[{
+ TeXf7b6d320Encoding ReEncodeFont }5 49.8132 /CMR6 rf
+/Ff 170[71 85[{}1 99.6264 /EUSM10 rf /Fg 153[18 29 33
+100[{ TeXBase1Encoding ReEncodeFont }3 66.4176 /Times-Italic
+rf /Fh 244[76 11[{ TeXaae443f0Encoding ReEncodeFont }1
+99.6264 /CMMIB10 rf /Fi 136[55 1[55 55 55 55 1[55 55
+55 55 55 55 1[55 55 1[55 55 1[55 1[55 11[55 26[55 7[55
+1[55 55 55 46[{ TeXBase1Encoding ReEncodeFont }23 90.9091
+/Courier rf /Fj 140[37 14[46 100[{ TeX74afc74cEncoding ReEncodeFont }2
+90.9091 /CMTI10 rf /Fk 192[43 1[71 19[35 35 4[76 35[{
+ TeXf7b6d320Encoding ReEncodeFont }5 90.9091 /CMR10 rf
+/Fl 141[76 8[35 35 7[56 41[0 3[61 24[71 25[{
+ TeXbbad153fEncoding ReEncodeFont }7 90.9091 /CMSY10
+rf /Fm 254[29 1[{ TeX10037936Encoding ReEncodeFont }1
+90.9091 /CMBSY10 rf /Fn 167[120 25[74 74 74 74 74 74
+74 56[{}8 83.022 /CMEX10 rf /Fo 134[33 1[48 33 33 18
+26 22 1[33 33 33 52 18 2[18 33 1[22 29 33 29 33 29 46[33
+33 3[22 3[22 22 40[{ TeXBase1Encoding ReEncodeFont }25
+66.4176 /Times-Roman rf /Fp 135[45 1[45 2[40 2[50 50
+1[80 3[30 3[45 50 29[64 4[72 65[{ TeX74afc74cEncoding ReEncodeFont }11
+99.6264 /CMTI12 rf /Fq 139[25 5[43 62 3[24 6[30 37 8[58
+7[45 20[20 20 23[33 6[40 5[41 3[35 4[40 45 11[{
+ TeXaae443f0Encoding ReEncodeFont }16 66.4176 /CMMI8
+rf /Fr 254[32 1[{ TeX10037936Encoding ReEncodeFont }1
+99.6264 /CMBSY10 rf /Fs 134[60 60 60 60 1[60 1[60 1[60
+60 60 60 60 2[60 2[60 60 3[60 8[60 3[60 5[60 60 5[60
+4[60 60 6[60 10[60 60 60 60 2[60 60 4[60 35[{
+ TeXBase1Encoding ReEncodeFont }30 99.6264 /Courier rf
+/Ft 141[59 8[27 27 35 35 42[35 3[0 4[47 1[19 6[71 7[71
+7[55 23[20 55{ TeXbbad153fEncoding ReEncodeFont }14 66.4176
+/CMSY8 rf /Fu 162[27 1[27 26[76 46 1[76 2[27 8[49 5[76
+1[38 38 4[81 34[61{ TeXf7b6d320Encoding ReEncodeFont }12
+99.6264 /CMR12 rf /Fv 135[51 51 51 51 51 51 51 1[51 51
+51 51 51 2[51 51 51 51 51 51 2[51 46[51 51 49[{
+ TeX09fbbfacEncoding ReEncodeFont }21 99.6264 /CMTT12
+rf /Fw 141[83 7[28 39 39 50 50 5[61 2[66 66 66 31[50
+4[0 3[66 6[61 1[100 7[100 5[77 1[77 9[50 50 11[77 28
+77{ TeXbbad153fEncoding ReEncodeFont }23 99.6264 /CMSY10
+rf /Fx 107[55 55 24[39 44 44 66 44 50 28 39 39 50 50
+50 50 72 28 2[28 50 50 28 44 50 44 50 50 12[55 50 1[72
+61 6[33 72 72 61 1[72 2[61 5[33 33 4[50 1[50 50 50 2[25
+33 25 2[33 33 33 36[50 2[{ TeXBase1Encoding ReEncodeFont }49
+99.6264 /Times-Italic rf /Fy 138[45 25 1[30 1[45 45 45
+71 25 2[25 45 1[30 40 45 40 45 40 12[56 51 15[61 1[66
+18[23 46[{ .167 SlantFont TeXBase1Encoding ReEncodeFont }21
+90.9091 /Times-Roman rf /Fz 134[48 55 3[35 46 3[47 58
+85 3[33 2[48 45 51 42 41 51 7[57 81 1[57 5[63 2[94 66
+5[63 2[70 74 73 2[76 49 1[27 27 11[49 11[46 5[42 55 1[55
+1[48 1[57 3[48 2[43 1[55 62 11[{ TeXaae443f0Encoding ReEncodeFont }39
+99.6264 /CMMI12 rf /FA 87[33 16[100 50 1[44 44 10[33
+13[44 50 50 72 50 50 28 39 33 50 50 50 50 78 28 50 28
+28 50 50 33 44 50 44 50 44 3[33 1[33 3[94 1[72 61 55
+66 72 55 72 72 89 61 72 39 33 72 72 55 61 72 66 66 72
+1[44 3[28 28 50 50 50 50 50 50 50 50 50 50 28 25 33 25
+2[33 33 33 5[33 15[28 13[55 55 2[{ TeXBase1Encoding ReEncodeFont }80
+99.6264 /Times-Roman rf /FB 87[33 45[44 50 50 72 50 55
+33 39 44 55 55 50 55 83 28 55 1[28 55 50 33 44 55 44
+55 50 9[100 1[72 66 55 72 78 61 78 72 1[66 2[39 2[61
+66 72 72 1[72 7[50 50 50 50 50 50 50 50 50 50 1[25 33
+3[33 33 22[28 14[55 2[{ TeXBase1Encoding ReEncodeFont }58
+99.6264 /Times-Bold rf /FC 139[33 13[45 4[48 8[75 28[25
+25 32[52 3[53 8[51 58 11[{ TeXaae443f0Encoding ReEncodeFont }10
+90.9091 /CMMI10 rf /FD 75[30 11[30 17[45 1[40 40 24[40
+45 45 66 45 45 25 35 30 45 45 45 45 71 25 45 25 25 45
+45 30 40 45 40 45 40 3[30 1[30 56 66 1[86 66 66 56 51
+61 66 51 66 66 81 56 66 35 30 66 66 51 56 66 61 61 66
+5[25 25 45 45 45 45 45 45 45 45 45 45 25 23 30 23 2[30
+30 30 71 20[25 5[30 8[51 2[{ TeXBase1Encoding ReEncodeFont }81
+90.9091 /Times-Roman rf /FE 139[30 35 40 14[40 51 45
+31[66 65[{ TeXBase1Encoding ReEncodeFont }7 90.9091 /Times-Bold
+rf /FF 134[40 40 61 40 45 25 35 35 45 45 45 45 66 25
+2[25 45 45 25 40 45 40 45 45 11[66 2[56 1[56 1[61 2[61
+1[30 2[56 2[61 1[56 6[30 12[30 23 44[{ TeXBase1Encoding ReEncodeFont }
+35 90.9091 /Times-Italic rf /FG 157[39 35 3[20 1[20 14[44
+2[25 2[46 4[53 1[33 1[55 2[20 2[35 35 35 35 35 35 35
+35 6[27 27 4[59 35[{ TeXf7b6d320Encoding ReEncodeFont }22
+66.4176 /CMR8 rf /FH 157[46 42 97[{ TeXf7b6d320Encoding ReEncodeFont }2
+83.022 /CMR10 rf /FI 134[60 1[86 1[60 33 47 40 2[60 60
+4[33 60 2[53 60 53 60 53 11[86 4[66 2[106 5[86 3[80 1[86
+65[{ TeXBase1Encoding ReEncodeFont }21 119.552 /Times-Roman
+rf /FJ 139[48 4[72 80 120 40 2[40 5[64 1[72 11[104 6[104
+75[80 2[{ TeXBase1Encoding ReEncodeFont }11 143.462 /Times-Bold
+rf end
+%%EndProlog
+%%BeginSetup
+%%Feature: *Resolution 600dpi
+TeXDict begin
+%%PaperSize: A4
+ end
+%%EndSetup
+%%Page: 1 1
+TeXDict begin 1 0 bop 1176 315 a FJ(Nominal)34 b(Uni\002cation)563
+776 y FI(Christian)c(Urban)1350 732 y FH(a)1426 776 y
+FI(Andre)m(w)g(Pitts)2071 732 y FH(a)2146 776 y FI(Murdoch)f(Gabbay)
+3006 732 y FH(b)1018 948 y FG(a)1057 981 y FF(Univer)o(sity)d(of)d
+(Cambridg)o(e)o(,)i(Cambridg)o(e)o(,)g(UK)1394 1094 y
+FG(b)1437 1127 y FF(INRIA,)d(P)-7 b(aris,)23 b(F)-5 b(r)o(ance)p
+166 1417 3288 4 v 166 1546 a FE(Abstract)166 1752 y FD(W)e(e)31
+b(present)j(a)e(generalisation)k(of)c(\002rst-order)i(uni\002cation)h
+(to)d(the)g(practically)j(important)f(case)f(of)166 1865
+y(equations)j(between)f(terms)f(in)l(v)n(olving)i FF(binding)g(oper)o
+(ations)p FD(.)g(A)c(substitution)37 b(of)d(terms)g(for)f(v)n(ari-)166
+1978 y(ables)d(solv)o(es)f(such)h(an)e(equation)j(if)d(it)h(mak)o(es)g
+(the)g(equated)h(terms)f FC(\013)p FF(-equivalent)p FD(,)j(i.e.)c
+(equal)h(up)g(to)166 2091 y(renaming)21 b(bound)h(names.)d(F)o(or)g
+(the)h(applications)k(we)18 b(ha)n(v)o(e)j(in)e(mind,)h(we)f(must)h
+(consider)i(the)e(simple,)166 2204 y(te)o(xtual)25 b(form)f(of)f
+(substitution)k(in)d(which)g(names)g(occurring)i(in)e(terms)g(may)f(be)
+h(captured)h(within)g(the)166 2317 y(scope)f(of)f(binders)i(upon)f
+(substitution.)j(W)-7 b(e)22 b(are)h(able)h(to)f(tak)o(e)h(a)e
+(\223nominal\224)j(approach)h(to)d(binding)i(in)166 2430
+y(which)e(bound)h(entities)h(are)d(e)o(xplicitly)j(named)e(\(rather)i
+(than)e(using)h(nameless,)g(de)e(Bruijn-style)j(rep-)166
+2543 y(resentations\))30 b(and)d(yet)f(get)h(a)f(v)o(ersion)i(of)e
+(this)h(form)f(of)g(substitution)k(that)d(respects)h
+FC(\013)p FD(-equi)n(v)n(alence)166 2656 y(and)23 b(possesses)i(good)f
+(algorithmic)h(properties.)g(W)-7 b(e)21 b(achie)n(v)o(e)j(this)f(by)g
+(adapting)i(tw)o(o)d(e)o(xisting)i(ideas.)166 2769 y(The)h(\002rst)f
+(one)i(is)f(terms)g(in)l(v)n(olving)j(e)o(xplicit)f(substitutions)i(of)
+c(names)g(for)g(names,)h(e)o(xcept)g(that)f(here)166
+2882 y(we)g(only)h(use)g FF(e)n(xplicit)h(permutations)i
+FD(\(bijecti)n(v)o(e)e(substitutions\).)j(The)25 b(second)i(one)f(is)f
+(that)h(the)g(uni-)166 2994 y(\002cation)h(algorithm)g(should)g(solv)o
+(e)g(not)f(only)g(equational)j(problems,)e(b)n(ut)f(also)h(problems)g
+(about)g(the)166 3107 y FF(fr)m(eshness)j FD(of)e(names)h(for)f(terms.)
+g(There)h(is)e(a)h(simple)h(generalisation)j(of)c(classical)i
+(\002rst-order)g(uni-)166 3220 y(\002cation)i(problems)h(to)e(this)g
+(setting)i(which)f(retains)g(the)g(latter')-5 b(s)32
+b(pleasant)i(properties:)g(uni\002cation)166 3333 y(problems)23
+b(in)l(v)n(olving)i FC(\013)p FD(-equi)n(v)n(alence)g(and)d(freshness)i
+(are)d(decidable;)k(and)d(solv)n(able)h(problems)g(pos-)166
+3446 y(sess)h(most)g(general)h(solutions.)166 3717 y
+FF(K)m(e)m(y)e(wor)m(ds:)47 b FD(Abstract)25 b(Syntax,)f(Alpha-Con)l(v)
+o(ersion,)j(Binding)d(Operations,)i(Uni\002cation)p 166
+3823 V 166 4194 a FB(1)99 b(Intr)n(oduction)166 4537
+y FA(Decidability)28 b(of)i(uni\002cation)e(for)i(equations)f(between)g
+(\002rst-order)h(terms)f(and)h(algorithms)166 4658 y(for)20
+b(computing)e(most)g(general)i(uni\002ers)g(form)f(a)h(fundamental)e
+(tool)h(of)h(computational)d(logic)166 4778 y(with)26
+b(man)o(y)f(applications)g(to)h(programming)f(languages)g(and)i
+(computer)n(-aided)e(reasoning.)166 4898 y(Ho)n(we)n(v)o(er)l(,)e(v)o
+(ery)i(man)o(y)e(potential)h(applications)f(f)o(all)i(outside)e(the)i
+(scope)f(of)h(\002rst-order)g(uni-)166 5019 y(\002cation,)39
+b(because)h(the)o(y)e(in)l(v)n(olv)o(e)g(term)h(languages)g(with)g
+(binding)f(operations)g(where)i(at)166 5139 y(the)d(v)o(ery)f(least)h
+(we)g(do)g(not)f(wish)h(to)f(distinguish)f(terms)h(dif)n(fering)g(up)h
+(to)g(the)f(renaming)166 5259 y(of)e(bound)e(names.)i(There)g(is)f(a)h
+(lar)n(ge)g(body)f(of)g(w)o(ork)h(studying)e(languages)h(with)f
+(binders)166 5380 y(through)h(the)g(use)g(of)h(v)n(arious)e
+Fz(\025)p FA(-calculi)i(as)f(term)g(representation)g(languages,)g
+(leading)g(to)189 5712 y Fy(Accepted)25 b(for)e(publication)k(in)d
+(Theoretical)h(Computer)g(Science.)p eop end
+%%Page: 2 2
+TeXDict begin 2 1 bop 166 83 a Fx(higher)n(-or)l(der)34
+b(uni\002cation)g FA(algorithms)f(for)i(solving)f(equations)g(between)h
+Fz(\025)p FA(-terms)f(mod-)166 203 y(ulo)23 b Fz(\013)q(\014)6
+b(\021)t FA(-equi)n(v)n(alence.)21 b(Ho)n(we)n(v)o(er)l(,)h(higher)n
+(-order)i(uni\002cation)e(is)h(technically)g(complicated)166
+324 y(without)k(being)h(completely)f(satisf)o(actory)g(from)h(a)h
+(pragmatic)f(point)f(of)h(vie)n(w)-6 b(.)27 b(The)i(reason)166
+444 y(lies)38 b(in)f(the)h(dif)n(ference)g(between)h(substitution)c
+(for)j(\002rst-order)h(terms)e(and)h(for)g Fz(\025)p
+FA(-terms.)166 565 y(The)31 b(former)h(is)f(a)h(simple)e(operation)h
+(of)g(te)o(xtual)g(replacement)g(\(sometimes)f(called)h
+Fx(gr)o(aft-)166 685 y(ing)g FA([1],)h(or)f Fx(conte)n(xt)h
+(substitution)c FA([2,)k(Sect.)g(2.1]\),)f(whereas)h(the)g(latter)f
+(also)g(in)l(v)n(olv)o(es)f(re-)166 805 y(namings)d(to)i(a)n(v)n(oid)f
+(capture.)h(Capture-a)n(v)n(oidance)g(ensures)f(that)g(substitution)e
+(respects)j Fz(\013)q FA(-)166 926 y(equi)n(v)n(alence,)c(b)n(ut)g(it)h
+(complicates)f(higher)n(-order)h(uni\002cation)f(algorithms.)f
+(Furthermore)i(it)166 1046 y(is)d(the)g(simple)f(te)o(xtual)g(form)g
+(of)i(substitution)c(rather)j(than)g(the)g(more)g(complicated)f
+(capture-)166 1166 y(a)n(v)n(oiding)j(form)i(which)f(occurs)g(in)g(man)
+o(y)g(informal)g(applications)f(of)h(\223uni\002cation)g(modulo)166
+1287 y Fz(\013)q FA(-equi)n(v)n(alence\224.)k(F)o(or)h(e)o(xample,)f
+(consider)h(the)g(follo)n(wing)e(schematic)h(rule)i(which)e(might)166
+1407 y(form)j(part)g(of)g(the)g(inducti)n(v)o(e)e(de\002nition)h(of)i
+(a)f(binary)g(e)n(v)n(aluation)e(relation)i Fw(+)g FA(for)g(the)g(e)o
+(x-)166 1528 y(pressions)24 b(of)h(an)g(imaginary)e(functional)h
+(programming)g(language:)1330 1758 y Fv(app)q Fu(\()p
+Fv(fn)h Fz(a:Y)5 b(;)17 b(X)8 b Fu(\))55 b Fw(+)g Fz(V)p
+1304 1803 931 4 v 1304 1889 a Fv(let)26 b Fz(a)i Fu(=)g
+Fz(X)k Fv(in)26 b Fz(Y)76 b Fw(+)56 b Fz(V)2290 1806
+y FA(.)3338 1838 y(\(1\))166 2145 y(Here)27 b Fz(X)8
+b FA(,)25 b Fz(Y)47 b FA(and)26 b Fz(V)48 b FA(are)26
+b(meta)n(v)n(ariables)f(standing)f(for)j(unkno)n(wn)d(programming)g
+(language)166 2266 y(e)o(xpressions.)k(The)h(binders)g
+Fv(fn)c Fz(a:)p Fu(\()p Fw(\000)p Fu(\))30 b FA(and)f
+Fv(let)d Fz(a)36 b Fu(=)g Fz(X)d Fv(in)25 b Fu(\()p Fw(\000)p
+Fu(\))30 b FA(may)f(v)o(ery)g(well)g(capture)166 2386
+y(free)34 b(occurrences)g(of)g(the)f(v)n(ariable)g(named)g
+Fz(a)g FA(when)g(we)h(instantiate)e(the)h(schematic)g(rule)166
+2507 y(by)g(replacing)h(the)g(meta)n(v)n(ariable)e Fz(Y)56
+b FA(with)32 b(an)i(e)o(xpression.)f(F)o(or)g(instance,)h(using)e(the)i
+(rule)166 2627 y(scheme)25 b(in)f(a)h(bottom-up)e(search)j(for)f(a)g
+(proof)g(of)1393 2873 y Fv(let)h Fz(a)i Fu(=)f(1)e Fv(in)g
+Fz(a)56 b Fw(+)f Fu(1)1111 b FA(\(2\))166 3120 y(we)25
+b(w)o(ould)f(use)h(a)g(substitution)d(that)i(does)h(in)l(v)n(olv)o(e)e
+(capture,)i(namely)1304 3366 y Fu([)p Fz(X)36 b Fu(:=)28
+b(1)o Fz(;)17 b(Y)49 b Fu(:=)28 b Fz(a;)17 b(V)49 b Fu(:=)27
+b(1])166 3613 y FA(in)i(order)g(to)g(unify)f(the)h(goal)f(with)h(the)f
+(conclusion)g(of)h(the)g(rule)g(\(1\)\227generating)f(the)h(ne)n(w)166
+3733 y(goal)h Fv(app)q Fu(\()p Fv(fn)25 b Fz(a:a;)17
+b Fu(1\))38 b Fw(+)f Fu(1)30 b FA(from)g(the)g(hypothesis)e(of)i
+(\(1\).)g(The)h(problem)e(with)g(this)h(is)f(that)166
+3853 y(in)22 b(informal)g(practice)i(we)f(usually)e(identify)h(terms)g
+(up)h(to)f Fz(\013)q FA(-equi)n(v)n(alence,)g(whereas)h(te)o(xtual)166
+3974 y(substitution)f(does)j(not)f(respect)h Fz(\013)q
+FA(-equi)n(v)n(alence.)f(F)o(or)h(e)o(xample,)f(up)h(to)f
+Fz(\013)q FA(-equi)n(v)n(alence,)g(the)166 4094 y(goal)1403
+4241 y Fv(let)h Fz(b)k Fu(=)e(1)e Fv(in)g Fz(b)56 b Fw(+)f
+Fu(1)1121 b FA(\(3\))166 4417 y(is)34 b(the)f(same)h(as)g(\(2\).)g(W)-8
+b(e)34 b(might)f(think)g(\(erroneously!\))h(that)f(the)h(conclusion)f
+(of)h(rule)g(\(1\))166 4537 y(is)i(the)f(same)h(as)g
+Fv(let)26 b Fz(b)49 b Fu(=)f Fz(X)33 b Fv(in)25 b Fz(Y)70
+b Fw(+)48 b Fz(V)58 b FA(without)35 b(changing)g(the)h(rule')-5
+b(s)35 b(hypothesis\227)166 4658 y(after)c(all,)f(if)g(we)h(are)g
+(trying)e(to)h(mak)o(e)g Fz(\013)q FA(-equi)n(v)n(alence)f(disappear)i
+(into)e(the)h(infrastructure,)166 4778 y(then)d(we)g(must)f(be)h(able)g
+(to)f(replace)i(an)o(y)e Fx(part)i FA(of)f(what)g(we)g(ha)n(v)o(e)g
+(with)f(an)h(equi)n(v)n(alent)e(part.)166 4898 y(So)31
+b(we)g(might)e(be)h(tempted)g(to)g(unify)g(the)g(conclusion)f(with)h
+(\(3\))h(via)f(the)g(te)o(xtual)g(substitu-)166 5019
+y(tion)25 b Fu([)p Fz(X)35 b Fu(:=)28 b(1)p Fz(;)17 b(Y)49
+b Fu(:=)27 b Fz(b)q(;)17 b(V)49 b Fu(:=)27 b(1])p FA(,)f(and)g(then)f
+(apply)g(this)g(substitution)e(to)i(the)h(hypothesis)d(to)166
+5139 y(obtain)28 b(a)h(wrong)f(goal,)h Fv(app)p Fu(\()p
+Fv(fn)d Fz(a:b;)17 b Fu(1\))35 b Fw(+)g Fu(1)p FA(.)28
+b(Using)g Fz(\025)p FA(-calculus)g(and)h(higher)n(-order)g(uni\002-)166
+5259 y(cation)k(sa)n(v)o(es)f(us)g(from)h(such)g(slopp)o(y)f(thinking,)
+f(b)n(ut)h(at)h(the)g(e)o(xpense)f(of)h(ha)n(ving)f(to)h(mak)o(e)166
+5380 y(e)o(xplicit)22 b(the)g(dependence)i(of)f(meta)n(v)n(ariables)f
+(on)h(bindable)f(names)h(via)g(the)g(use)g(of)g(function)1773
+5712 y(2)p eop end
+%%Page: 3 3
+TeXDict begin 3 2 bop 166 83 a FA(v)n(ariables)24 b(and)h(application.)
+e(F)o(or)i(e)o(xample,)f(\(1\))h(w)o(ould)f(be)h(replaced)g(by)g
+(something)e(lik)o(e)1310 289 y Fv(app)17 b Fu(\()p Fv(fn)g
+Fz(\025a:F)31 b(a)p Fu(\))17 b Fz(X)62 b Fw(+)56 b Fz(V)p
+1310 334 1001 4 v 1369 425 a Fv(let)18 b Fz(X)24 b Fu(\()p
+Fz(\025a:F)30 b(a)p Fu(\))56 b Fw(+)f Fz(V)1109 b FA(\(4\))166
+647 y(or)l(,)25 b(modulo)e Fz(\021)t FA(-equi)n(v)n(alence)1371
+853 y Fv(app)17 b Fu(\()p Fv(fn)g Fz(F)d Fu(\))j Fz(X)63
+b Fw(+)55 b Fz(V)p 1371 898 798 4 v 1468 984 a Fv(let)18
+b Fz(X)24 b(F)69 b Fw(+)56 b Fz(V)2224 901 y FA(.)3338
+984 y(\(5\))166 1206 y(No)n(w)26 b(goal)h(\(3\))g(becomes)g
+Fv(let)17 b Fu(1)g Fz(\025b:b)32 b Fw(+)g Fu(1)26 b FA(and)h(there)g
+(is)g(no)g(problem)f(unifying)f(it)i(with)f(the)166 1326
+y(conclusion)i(of)h(\(5\))h(via)f(a)g(capture-a)n(v)n(oiding)g
+(substitution)d(of)j Fu(1)h FA(for)f Fz(X)8 b FA(,)29
+b Fz(\025c:c)g FA(for)h Fz(F)43 b FA(and)29 b Fu(1)166
+1447 y FA(for)c Fz(V)d FA(.)166 1668 y(This)36 b(is)h(all)f(v)o(ery)h
+(\002ne,)g(b)n(ut)f(the)h(situation)e(is)i(not)f(as)h(pleasant)f(as)h
+(for)g(\002rst-order)h(terms:)166 1788 y(higher)n(-order)43
+b(uni\002cation)g(problems)f(can)h(be)h(undecidable,)e(decidable)h(b)n
+(ut)g(lack)g(most)166 1909 y(general)h(uni\002ers,)f(or)g(ha)n(v)o(e)g
+(such)g(uni\002ers)g(only)g(by)g(imposing)e(some)i(restrictions)f([3];)
+166 2029 y(see)36 b([4])g(for)g(a)g(surv)o(e)o(y)f(of)g(higher)n
+(-order)h(uni\002cation.)f(W)-8 b(e)36 b(started)g(out)f(w)o(anting)g
+(to)g(com-)166 2149 y(pute)42 b(with)g(binders)f(modulo)g
+Fz(\013)q FA(-equi)n(v)n(alence,)g(and)h(someho)n(w)f(the)h(process)g
+(of)h(making)166 2270 y(possibly-capturing)31 b(substitution)f
+(respectable)k(has)f(led)g(to)f(function)h(v)n(ariables,)f(applica-)166
+2390 y(tion,)24 b(capture-a)n(v)n(oiding)g(substitution)e(and)j
+Fz(\014)6 b(\021)t FA(-equi)n(v)n(alence.)23 b(Does)h(it)h(ha)n(v)o(e)f
+(to)h(be)g(so?)f(No!)166 2611 y(F)o(or)43 b(one)g(thing,)f(se)n(v)o
+(eral)g(authors)h(ha)n(v)o(e)g(already)g(noted)g(that)f(one)h(can)h
+(mak)o(e)f(sense)g(of)166 2731 y(possibly-capturing)35
+b(substitution)f(modulo)i Fz(\013)q FA(-equi)n(v)n(alence)f(by)i(using)
+f Fx(e)n(xplicit)h(substitu-)166 2852 y(tions)21 b FA(in)h(the)h(term)f
+(representation)g(language:)g(see)g([1,5\2269].)g(Compared)h(with)e
+(those)h(w)o(orks,)166 2972 y(we)36 b(mak)o(e)g(a)h(number)f(of)g
+(simpli\002cations.)e(First,)i(we)g(\002nd)g(that)g(we)g(do)g(not)g
+(need)g(to)g(use)166 3093 y(function)24 b(v)n(ariables,)g(application)f
+(or)i Fz(\014)6 b(\021)t FA(-equi)n(v)n(alence)23 b(in)h(our)g
+(representation)g(language\227)166 3213 y(lea)n(ving)d(just)f(binders)g
+(and)i Fz(\013)q FA(-equi)n(v)n(alence.)d(Secondly)-6
+b(,)21 b(instead)f(of)i(using)e(e)o(xplicit)g(substitu-)166
+3333 y(tions)26 b(of)h(names)g(for)h(names,)e(we)i(use)f(only)f(the)h
+(special)g(case)h(of)f Fx(e)n(xplicit)g(permutations)e
+FA(of)166 3454 y(names.)33 b(The)g(idea)g(of)g(using)f
+(name-permutations,)f(and)i(in)f(particular)h(name-sw)o(appings,)166
+3574 y(when)g(dealing)g(with)g Fz(\013)q FA(-con)l(v)o(ersion)f(w)o(as)
+i(described)f(in)g([10])h(and)f(there)h(is)f(gro)n(wing)f(e)n(vi-)166
+3694 y(dence)k(of)f(its)f(usefulness)g(\(see)i([11\22613],)f(for)g(e)o
+(xample\).)f(When)i(a)f(name)g(substitution)d(is)166
+3815 y(actually)f(a)g(permutation,)f(the)h(function)f(it)h(induces)g
+(from)g(terms)g(to)f(terms)h(is)g(a)g(bijection;)166
+3935 y(this)g(bijecti)n(vity)f(gi)n(v)o(es)g(the)i(operation)g(of)g
+(permuting)e(names)i(v)o(ery)g(good)f(logical)h(proper)n(-)166
+4056 y(ties)21 b(compared)g(with)f(name)h(substitution.)d(Consider)j
+(for)h(e)o(xample)e(the)h Fz(\013)q FA(-equi)n(v)n(alent)e(terms)166
+4176 y Fv(fn)e Fz(a:b)j FA(and)g Fv(fn)d Fz(c:b)p FA(,)j(where)g
+Fz(a)p FA(,)f Fz(b)h FA(and)g Fz(c)f FA(are)h(distinct.)e(If)i(we)f
+(apply)g(the)h(substitution)c Fu([)p Fz(b)3141 4170 y
+Ft(7!)3212 4176 y Fz(a)p Fu(])k FA(\(re-)166 4296 y(naming)f(all)h
+(free)h(occurrences)f(of)h Fz(b)f FA(to)g(be)g Fz(a)p
+FA(\))g(to)g(them)f(we)i(get)e Fv(fn)f Fz(a:a)i FA(and)g
+Fv(fn)d Fz(c:a)p FA(,)j(which)g(are)166 4417 y(no)27
+b(longer)g Fz(\013)q FA(-equi)n(v)n(alent.)f(Thus)g(renaming)h
+(substitutions)d(do)k(not)e(respect)i Fz(\013)q FA(-equi)n(v)n(alence)
+166 4537 y(in)20 b(general,)h(and)g(an)o(y)f(uni\002cation)h(algorithm)
+e(using)h(them)g(needs)h(to)f(tak)o(e)h(e)o(xtra)f(precautions)166
+4658 y(to)j(not)f(inadv)o(ertently)f(change)i(the)g(intended)f(meaning)
+h(of)g(terms.)f(The)h(traditional)f(solution)166 4778
+y(for)27 b(this)f(problem)g(is)g(to)g(introduce)g(a)h(more)g
+(complicated)e(form)i(of)g(renaming)f(substitution)166
+4898 y(that)e(a)n(v)n(oids)f(capture)h(of)h(names)e(by)h(binders.)g(In)
+g(contrast,)f(the)h(simple)f(operation)h(of)g(name-)166
+5019 y(permutation)29 b(respects)h Fz(\013)q FA(-equi)n(v)n(alence;)f
+(for)i(e)o(xample,)e(applying)g(the)h(name-permutation)166
+5139 y Fu(\()p Fz(a)17 b(b)p Fu(\))30 b FA(that)e(sw)o(aps)h(all)g
+(occurrences)h(of)g Fz(a)f FA(and)g Fz(b)h FA(\(be)g(the)o(y)e(free,)i
+(bound)f(or)g(binding\))f(to)h(the)166 5259 y(terms)f(abo)o(v)o(e)g(gi)
+n(v)o(es)f Fv(fn)17 b Fz(b:a)29 b FA(and)g Fv(fn)17 b
+Fz(c:a)p FA(,)29 b(which)f(are)h(still)f Fz(\013)q FA(-equi)n(v)n
+(alent.)e(W)-8 b(e)29 b(e)o(xploit)e(such)166 5380 y(good)38
+b(properties)g(of)h(name-permutations)d(to)i(gi)n(v)o(e)g(a)h
+(conceptually)e(simple)h(uni\002cation)1773 5712 y(3)p
+eop end
+%%Page: 4 4
+TeXDict begin 4 3 bop 166 83 a FA(algorithm.)166 303
+y(In)22 b(addition)e(to)h(the)g(use)g(of)h(e)o(xplicit)e
+(name-permutations,)g(we)h(also)g(compute)g(symbolically)166
+423 y(with)26 b(predicates)g(e)o(xpressing)f Fx(fr)l(eshness)g
+FA(of)h(names)g(for)h(terms.)e(Such)i(predicates)f(certainly)166
+544 y(feature)43 b(in)e(pre)n(vious)g(w)o(ork)h(on)f(binding)g(\(for)h
+(e)o(xample,)f(Qu-Prolog')-5 b(s)41 b Fs(not)p 3030 544
+30 4 v 35 w(free)p 3305 544 V 35 w(in)166 664 y FA(predicate)26
+b([8],)g(the)f(notion)f(of)i(\223algebraic)g(independence\224)g(in)f
+([14,)g(De\002nition)g(3],)g(and)h(the)166 785 y
+(\223non-occurrence\224)31 b(predicates)f(of)g([15]\).)g(But)h(once)f
+(again,)f(the)h(use)g(of)g(such)g(a)g(freshness)166 905
+y(predicate)d(based)g(upon)f(name)h Fx(swapping)f FA(rather)h(than)f
+(renaming,)g(which)h(arises)g(naturally)166 1025 y(from)33
+b(the)g(w)o(ork)g(reported)g(in)g([10,16],)f(gi)n(v)o(es)g(us)h(a)g
+(simpler)f(theory)h(with)f(good)h(algorith-)166 1146
+y(mic)28 b(properties.)f(It)h(is)g(easy)g(to)f(see)i(why)e(there)h(is)g
+(a)g(need)g(for)h(computing)d(with)h(freshness,)166 1266
+y(gi)n(v)o(en)i(that)h(we)h(tak)o(e)g(a)g(\223nominal\224)f(approach)h
+(to)f(binders.)g(\(In)h(other)f(w)o(ords)g(we)h(use)g(con-)166
+1386 y(crete)25 b(v)o(ersions)e(of)i(binding)e(and)h
+Fz(\013)q FA(-equi)n(v)n(alence)g(in)g(which)g(bound)f(entities)h(are)h
+(named)f(e)o(x-)166 1507 y(plicitly)-6 b(,)26 b(rather)i(than)g(using)e
+(de)i(Bruijn-style)f(representations,)g(as)h(for)g(e)o(xample)f(in)h
+([1,7].\))166 1627 y(A)i(basic)h(instance)e(of)i(our)f(generalised)g
+(form)g(of)h Fz(\013)q FA(-equi)n(v)n(alence)e(identi\002es)h
+Fv(fn)25 b Fz(a:X)39 b FA(with)166 1748 y Fv(fn)25 b
+Fz(b:)p Fu(\()p Fz(a)17 b(b)p Fu(\))p Fr(\001)q Fz(X)33
+b FA(pro)o(vided)25 b Fz(b)h Fx(is)e(fr)l(esh)i(for)h
+Fz(X)8 b FA(,)25 b(where)h(the)f(subterm)f Fu(\()p Fz(a)17
+b(b)p Fu(\))p Fr(\001)p Fz(X)34 b FA(indicates)24 b(an)i(e)o(x-)166
+1868 y(plicit)h(permutation\227namely)e(the)i(sw)o(apping)g(of)h
+Fz(a)f FA(and)h Fz(b)p FA(\227w)o(aiting)f(to)g(be)h(applied)f(to)g
+Fz(X)8 b FA(.)166 1988 y(W)-8 b(e)30 b(write)f(\223)p
+Fz(b)g FA(is)g(fresh)g(for)h Fz(X)8 b FA(\224)29 b(symbolically)e(as)i
+Fz(b)36 b Fu(#)g Fz(X)8 b FA(;)29 b(the)g(intended)f(meaning)h(of)g
+(this)166 2109 y(relation)g(is)g(that)g Fz(b)h FA(does)f(not)g(occur)g
+(free)i(in)e(an)o(y)f(\(ground\))h(term)g(that)g(may)g(be)h
+(substituted)166 2229 y(for)23 b Fz(X)8 b FA(.)23 b(If)g(we)g(kno)n(w)f
+(more)h(about)f Fz(X)31 b FA(we)23 b(may)f(be)h(able)g(to)f(eliminate)g
+(the)h(e)o(xplicit)e(permuta-)166 2350 y(tion)26 b(in)g
+Fu(\()p Fz(a)17 b(b)p Fu(\))p Fr(\001)p Fz(X)8 b FA(;)26
+b(for)h(e)o(xample,)e(if)i(we)f(kne)n(w)g(that)g Fz(a)31
+b Fu(#)g Fz(X)k FA(holds)25 b(as)i(well)f(as)h Fz(b)k
+Fu(#)g Fz(X)8 b FA(,)26 b(then)166 2470 y Fu(\()p Fz(a)17
+b(b)p Fu(\))p Fr(\001)p Fz(X)33 b FA(can)25 b(be)g(replaced)g(by)g
+Fz(X)8 b FA(.)166 2690 y(It)24 b(should)e(already)i(be)g(clear)h(from)e
+(these)h(simple)e(e)o(xamples)h(that)g(in)h(our)g(setting)e(the)i
+(appro-)166 2810 y(priate)c(notion)f(of)h(term-equality)f(is)h(not)g(a)
+g(bare)h(equation,)e Fz(t)28 b Fw(\031)g Fz(t)2430 2774
+y Ft(0)2454 2810 y FA(,)20 b(b)n(ut)g(rather)g(a)h(hypothetical)166
+2931 y(judgement)j(of)g(the)h(form)1597 3051 y Fw(r)j(`)f
+Fz(t)h Fw(\031)g Fz(t)1999 3010 y Ft(0)3338 3051 y FA(\(6\))166
+3200 y(where)37 b Fw(r)g FA(is)g(a)g Fx(fr)l(eshness)f(en)l(vir)l
+(onment)r FA(\227a)f(\002nite)i(set)f Fw(f)p Fz(a)2317
+3215 y FG(1)2407 3200 y Fu(#)50 b Fz(X)2619 3215 y FG(1)2658
+3200 y Fz(;)17 b(:)g(:)g(:)f(;)h(a)2928 3215 y Fq(n)3025
+3200 y Fu(#)50 b Fz(X)3237 3215 y Fq(n)3284 3200 y Fw(g)37
+b FA(of)166 3320 y(freshness)25 b(assumptions.)d(F)o(or)j(e)o(xample)
+1055 3539 y Fw(f)p Fz(a)j Fu(#)g Fz(X)r(;)17 b(b)28 b
+Fu(#)g Fz(X)8 b Fw(g)27 b(`)h Fv(fn)d Fz(a:X)36 b Fw(\031)28
+b Fv(fn)d Fz(b:X)782 b FA(\(7\))166 3758 y(is)27 b(a)g(v)n(alid)f
+(judgement)f(of)i(our)g Fx(nominal)f(equational)f(lo)o(gic)p
+FA(.)h(Similarly)-6 b(,)26 b(judgements)f(about)166 3879
+y(freshness)g(itself)f(will)g(tak)o(e)h(the)f(form)1571
+4098 y Fw(r)k(`)g Fz(a)f Fu(#)h Fz(t)g(:)1290 b FA(\(8\))166
+4317 y(T)-8 b(w)o(o)34 b(e)o(xamples)f(of)i(v)n(alid)e(freshness)h
+(judgements)f(are)i Fw(f)p Fz(a)46 b Fu(#)f Fz(X)8 b
+Fw(g)45 b(`)h Fz(a)f Fu(#)h Fv(fn)25 b Fz(b:X)60 b FA(and)166
+4438 y Fw(;)28 b(`)f Fz(a)h Fu(#)g Fv(fn)d Fz(a:X)8 b
+FA(.)166 4658 y(The)20 b(freshness)f(en)l(vironment)f
+Fw(r)i FA(in)f(judgements)f(of)i(the)f(form)g(\(6\))h(and)g(\(8\))f(e)o
+(xpresses)g(fresh-)166 4778 y(ness)28 b(conditions)f(that)h(an)o(y)g
+(te)o(xtual)f(substitution)f(of)j(terms)f(for)g(v)n(ariables)g(must)f
+(respect)i(in)166 4898 y(order)k(for)f(the)g(right-hand)g(side)g(of)g
+(the)g(judgement)f(to)h(be)g(v)n(alid)f(after)i(substitution.)d(This)
+166 5019 y(e)o(xplicit)37 b(use)h(of)g(freshness)f(mak)o(es)h(the)g
+(operation)f(of)i(te)o(xtual)d(substitution)g(respect)i(our)166
+5139 y(generalised)29 b(form)g(of)h Fz(\013)q FA(-equi)n(v)n(alence.)e
+(F)o(or)h(e)o(xample,)g(if)g(we)h(were)g(na)m(\250)-30
+b(\021v)o(ely)28 b(to)h(re)o(gard)g(the)166 5259 y(terms)23
+b Fv(fn)i Fz(a:X)32 b FA(and)23 b Fv(fn)j Fz(b:X)32 b
+FA(as)23 b Fz(\013)q FA(-equi)n(v)n(alent,)f(then)h(applying)f(for)i(e)
+o(xample)e(the)h(capturing)166 5380 y(substitution)j
+Fu([)p Fz(X)44 b Fu(:=)35 b Fz(a)p Fu(])30 b FA(or)f
+Fu([)p Fz(X)44 b Fu(:=)35 b Fz(b)p Fu(])30 b FA(results)e(into)h(tw)o
+(o)f(terms)h(that)g(are)g Fx(not)i Fz(\013)q FA(-equi)n(v)n(alent)1773
+5712 y(4)p eop end
+%%Page: 5 5
+TeXDict begin 5 4 bop 166 83 a FA(an)o(ymore.)22 b(\(A)h(similar)e
+(observ)n(ation)g(partly)h(moti)n(v)n(ates)f(the)h(w)o(ork)h(in)f
+([17].\))h(Ho)n(we)n(v)o(er)l(,)e(if)i(we)166 203 y(assume)e
+Fz(a)27 b Fu(#)h Fz(X)h FA(and)22 b Fz(b)28 b Fu(#)g
+Fz(X)h FA(as)21 b(in)g(\(7\),)g(then)g(all)g(problematic)f
+(substitutions)e(are)k(ruled)e(out.)166 324 y(In)k(this)e(w)o(ay)i(we)g
+(obtain)f(a)g(v)o(ersion)g(of)h Fz(\013)q FA(-equi)n(v)n(alence)e
+(between)i(terms)f(with)f(v)n(ariables)h(that)166 444
+y(is)g(respected)g(by)g(te)o(xtual)e(substitutions)f(\(see)k(Lemma)e
+(2.14)h(belo)n(w\),)f(unlik)o(e)g(the)h(traditional)166
+565 y(notion)h(of)g Fz(\013)q FA(-equi)n(v)n(alence.)166
+897 y Fx(Summary)166 1237 y FA(W)-8 b(e)26 b(will)f(represent)h
+(languages)g(in)l(v)n(olving)e(binders)h(using)g(the)g(usual)h(notion)e
+(of)i(\002rst-order)166 1357 y(terms)34 b(o)o(v)o(er)f(a)h(man)o
+(y-sorted)f(signature,)g(b)n(ut)h(with)f(certain)h(distinguished)e
+(constants)h(and)166 1478 y(function)26 b(symbols.)f(These)h(gi)n(v)o
+(e)f(us)i(terms)f(with:)f(distinguished)f(constants)i(naming)f(bind-)
+166 1598 y(able)35 b(entities,)f(that)g(we)i(call)e Fx(atoms)p
+FA(;)g(terms)h Fz(a:t)g FA(e)o(xpressing)f(a)h(generic)g(form)g(of)g
+Fx(binding)166 1719 y FA(of)g(an)f(atom)g Fz(a)h FA(in)f(a)h(term)f
+Fz(t)p FA(;)g(and)h(terms)f Fz(\031)t Fr(\001)o Fz(X)43
+b FA(representing)34 b(an)g(e)o(xplicit)f Fx(permutation)g
+Fz(\031)166 1839 y FA(of)d(atoms)e(w)o(aiting)h(to)g(be)g(applied)g(to)
+h(whate)n(v)o(er)e(term)h(is)h(substituted)d(for)j(the)f(v)n(ariable)g
+Fz(X)8 b FA(.)166 1959 y(Section)20 b(2)h(presents)e(this)h
+(term-language)g(together)f(with)h(a)h(syntax-directed)e(inducti)n(v)o
+(e)f(def-)166 2080 y(inition)25 b(of)i(the)g(pro)o(v)n(able)e
+(judgements)h(of)h(the)f(form)h(\(6\))g(and)g(\(8\))g(which)f(for)h
+Fx(gr)l(ound)h FA(terms)166 2200 y(\(i.e.)36 b(ones)g(with)f(no)g(v)n
+(ariables\))h(agrees)g(with)g(the)f(usual)h(notions)e(of)i
+Fz(\013)q FA(-equi)n(v)n(alence)f(and)166 2320 y(\223not)26
+b(a)g(free)h(v)n(ariable)e(of)5 b(\224.)26 b(Ho)n(we)n(v)o(er)l(,)f(on)
+h(open)f(terms)h(our)f(judgements)g(dif)n(fer)g(from)h(these)166
+2441 y(standard)h(notions.)f(Section)i(3)f(considers)g(uni\002cation)g
+(in)g(this)g(setting.)f(Solving)h(equalities)166 2561
+y(between)h(abstractions)f Fz(a:t)34 b Fw(\031)q Fu(?)f
+Fz(a)1385 2525 y Ft(0)1409 2561 y Fz(:t)1471 2525 y Ft(0)1522
+2561 y FA(entails)28 b(solving)e(both)h(equalities)g
+Fz(t)34 b Fw(\031)q Fu(?)f(\()p Fz(a)17 b(a)3130 2525
+y Ft(0)3153 2561 y Fu(\))p Fr(\001)p Fz(t)3258 2525 y
+Ft(0)3310 2561 y FA(and)166 2682 y(freshness)37 b(problems)f
+Fz(a)51 b Fu(#?)g Fz(t)1297 2645 y Ft(0)1320 2682 y FA(.)37
+b(Therefore)h(our)f(general)h(form)f(of)g Fx(nominal)f(uni\002cation)
+166 2802 y(pr)l(oblem)23 b FA(is)h(a)g(\002nite)h(collection)e(of)h
+(indi)n(vidual)e(equality)i(and)g(freshness)g(problems.)f(Such)h(a)166
+2922 y(problem)e Fz(P)37 b FA(is)23 b(solv)o(ed)f(by)h(pro)o(viding)e
+(not)i(only)f(a)i(substitution)c Fz(\033)28 b FA(\(of)23
+b(terms)g(for)g(v)n(ariables\),)166 3043 y(b)n(ut)31
+b(also)g(a)h(freshness)g(en)l(vironment)e Fw(r)i FA(\(as)g(abo)o(v)o
+(e\),)f(which)g(together)g(ha)n(v)o(e)h(the)f(property)166
+3163 y(that)25 b Fw(r)i(`)h Fz(\033)t Fu(\()p Fz(t)p
+Fu(\))g Fw(\031)h Fz(\033)t Fu(\()p Fz(t)976 3127 y Ft(0)999
+3163 y Fu(\))c FA(and)g Fw(r)j(`)g Fz(a)g Fu(#)g Fz(\033)t
+Fu(\()p Fz(t)1751 3127 y Ft(00)1794 3163 y Fu(\))d FA(hold)f(for)h
+(each)h(indi)n(vidual)d(equality)h Fz(t)k Fw(\031)p Fu(?)g
+Fz(t)3430 3127 y Ft(0)166 3284 y FA(and)i(freshness)g
+Fz(a)38 b Fu(#?)g Fz(t)1030 3247 y Ft(00)1102 3284 y
+FA(in)30 b(the)g(problem)f Fz(P)14 b FA(.)30 b(Our)g(main)f(result)h
+(with)f(respect)i(to)e(uni\002ca-)166 3404 y(tion)20
+b(is)h(that)g Fx(solvability)e(is)i(decidable)g(and)g(that)f(solvable)g
+(pr)l(oblems)g(possess)g(most)h(g)o(ener)o(al)166 3524
+y(solutions)30 b FA(\(for)i(a)g(reasonably)g(ob)o(vious)d(notion)i(of)h
+(\223most)f(general\224\).)h(The)g(proof)g(is)f(via)g(a)166
+3645 y(uni\002cation)i(algorithm)f(that)h(is)g(v)o(ery)g(similar)g(to)g
+(the)h(\002rst-order)g(algorithm)e(gi)n(v)o(en)g(in)h(the)166
+3765 y(no)n(w-common)23 b(transformational)g(style)h([18].)h(\(See)h
+([19,)f(Sect.)g(2.6])g(or)g([20,)g(Sect.)g(4.6])g(for)166
+3885 y(e)o(xpositions)k(of)j(this.\))f(Section)h(4)f(considers)h(the)f
+(relationship)g(of)g(our)h(v)o(ersion)f(of)h(\223uni\002-)166
+4006 y(cation)26 b(modulo)f Fz(\013)q FA(-equi)n(v)n(alence\224)g(to)g
+(e)o(xisting)g(approaches.)h(Section)g(5)g(assesses)g(what)g(has)166
+4126 y(been)f(achie)n(v)o(ed)f(and)h(the)f(prospects)h(for)g
+(applications.)166 4458 y Fx(Quiz)166 4799 y FA(T)-8
+b(o)29 b(appreciate)i(the)e(kind)g(of)h(problem)e(that)h(nominal)g
+(uni\002cation)g(solv)o(es,)f(you)h(might)f(lik)o(e)166
+4919 y(to)h(try)g(the)g(follo)n(wing)f(quiz)h(about)g(the)g
+Fz(\025)p FA(-calculus)g([21])h(before)f(we)h(apply)f(our)g(algorithm)
+166 5039 y(to)24 b(solv)o(e)g(it)g(at)h(the)g(end)g(of)g(Section)f(3.)
+166 5259 y Fx(Assuming)i Fz(a)i Fx(and)g Fz(b)g Fx(ar)l(e)g(distinct)e
+(variables)p FA(,)g(is)i(it)f(possible)f(to)i(\002nd)f
+Fz(\025)p FA(-terms)h Fz(M)3062 5274 y FG(1)3101 5259
+y Fz(;)17 b(:)g(:)g(:)f(;)h(M)3414 5274 y FG(7)166 5380
+y FA(that)24 b(mak)o(e)h(the)g(follo)n(wing)e(pairs)h(of)h(terms)f
+Fz(\013)q FA(-equi)n(v)n(alent?)1773 5712 y(5)p eop end
+%%Page: 6 6
+TeXDict begin 6 5 bop 206 83 a FA(\(1\))50 b Fz(\025a:\025b:)p
+Fu(\()p Fz(M)764 98 y FG(1)821 83 y Fz(b)p Fu(\))125
+b FA(and)f Fz(\025b:\025a:)p Fu(\()p Fz(a)17 b(M)1753
+98 y FG(1)1794 83 y Fu(\))206 203 y FA(\(2\))50 b Fz(\025a:\025b:)p
+Fu(\()p Fz(M)764 218 y FG(2)821 203 y Fz(b)p Fu(\))125
+b FA(and)f Fz(\025b:\025a:)p Fu(\()p Fz(a)17 b(M)1753
+218 y FG(3)1794 203 y Fu(\))206 324 y FA(\(3\))50 b Fz(\025a:\025b:)p
+Fu(\()p Fz(b)17 b(M)822 339 y FG(4)862 324 y Fu(\))125
+b FA(and)f Fz(\025b:\025a:)p Fu(\()p Fz(a)17 b(M)1753
+339 y FG(5)1794 324 y Fu(\))206 444 y FA(\(4\))50 b Fz(\025a:\025b:)p
+Fu(\()p Fz(b)17 b(M)822 459 y FG(6)862 444 y Fu(\))125
+b FA(and)f Fz(\025a:\025a:)p Fu(\()p Fz(a)17 b(M)1763
+459 y FG(7)1803 444 y Fu(\))166 679 y FA(If)28 b(it)g(is)f(possible)g
+(to)g(\002nd)h(a)g(solution)e(for)j(an)o(y)e(of)h(these)f(four)h
+(problems,)f(can)h(you)g(describe)166 800 y(what)36 b(all)g(possible)e
+(solutions)h(for)h(that)f(problem)h(are)h(lik)o(e?)e(\(The)i(answers)f
+(are)h(gi)n(v)o(en)d(in)166 920 y(Example)24 b(3.8.\))166
+1423 y FB(2)99 b(Nominal)25 b(equational)g(logic)166
+1779 y FA(W)-8 b(e)31 b(tak)o(e)g(a)g(concrete)g(approach)g(to)f(the)h
+(syntax)e(of)i(binders)f(in)g(which)h(bound)f(entities)f(are)166
+1899 y(e)o(xplicitly)d(named.)i(Furthermore)h(we)g(do)f(not)g(assume)g
+(that)f(the)i(names)f(of)g(bound)g(entities)166 2019
+y(are)23 b(necessarily)f(v)n(ariables)f(\(things)g(that)h(may)f(be)h
+(substituted)e(for\),)j(in)f(order)g(to)g(encompass)166
+2140 y(e)o(xamples)29 b(lik)o(e)g(the)h Fz(\031)t FA(-calculus)f([22],)
+h(in)f(which)h(the)f(restriction)g(operator)h(binds)f(channel)166
+2260 y(names)g(and)h(these)f(are)i(quite)e(dif)n(ferent)g(from)h(names)
+f(of)h(unkno)n(wn)e(processes.)h(Names)h(of)166 2380
+y(bound)25 b(entities)g(will)g(be)h(called)g Fx(atoms)p
+FA(.)f(This)h(is)f(partly)h(for)g(historical)f(reasons)h(\(stemming)166
+2501 y(from)32 b(the)f(w)o(ork)h(by)f(the)h(second)g(tw)o(o)f(authors)g
+([10]\))h(and)g(partly)f(to)h(indicate)f(that)g(the)h(in-)166
+2621 y(ternal)26 b(structure)h(of)f(such)g(names)g(is)g(irrele)n(v)n
+(ant)g(to)g(us:)g(all)g(we)h(care)g(about)f(is)g(their)g(identity)166
+2742 y(\(i.e.)i(whether)g(or)g(not)f(one)h(atom)g(is)f(the)h(same)g(as)
+g(another\))g(and)g(that)f(the)h(supply)f(of)h(atoms)166
+2862 y(is)c(ine)o(xhaustible.)166 3097 y(Although)i(there)h(are)i(se)n
+(v)o(eral)d(general)i(frame)n(w)o(orks)f(in)f(the)i(literature)f(for)g
+(specifying)g(lan-)166 3217 y(guages)j(with)f(binders,)g(not)h(all)f
+(of)h(them)g(meet)f(the)h(requirements)f(mentioned)g(in)h(the)f(pre-)
+166 3338 y(vious)g(paragraph.)h(Use)g(of)h(the)e(simply)g(typed)g
+Fz(\025)p FA(-calculus)h(for)g(this)f(purpose)h(is)f(common;)166
+3458 y(b)n(ut)35 b(as)g(discussed)f(in)h(the)g(Introduction,)f(it)h
+(leads)g(to)f(a)i(problematic)e(uni\002cation)h(theory)-6
+b(.)166 3579 y(Among)29 b Fx(\002r)o(st-or)l(der)j FA(frame)n(w)o
+(orks,)d(Plotkin')-5 b(s)29 b(notion)g(of)h Fx(binding)f(signatur)l(e)g
+FA([23,24],)g(be-)166 3699 y(ing)f(unsorted,)h(equates)f(names)h(used)g
+(in)f(binding)g(with)g(names)h(of)g(v)n(ariables)f(standing)g(for)166
+3819 y(unkno)n(wn)22 b(terms;)h(so)h(it)f(is)g(not)g(suf)n(\002ciently)
+g(general)h(for)g(us.)g(A)g(\002rst-order)g(frame)n(w)o(ork)f(that)166
+3940 y(does)h(meet)f(our)h(requirements)f(is)g(the)h(notion)e(of)i
+Fx(nominal)e(alg)o(ebr)o(as)g FA(in)i([15].)f(The)h Fx(nominal)166
+4060 y(signatur)l(es)19 b FA(that)i(we)g(use)g(in)f(this)g(paper)h(are)
+h(a)f(mild)e(\(b)n(ut)i(practically)f(useful\))h(generalisation)166
+4181 y(of)27 b(nominal)e(algebras)i(in)g(which)f(name-abstraction)g
+(and)h(pairing)f(can)h(be)g(mix)o(ed)f(freely)h(in)166
+4301 y(arities)h(\(rather)g(than)g(insisting)e(as)i(in)g([15])g(that)g
+(the)g(ar)n(gument)g(sort)f(of)i(a)f(function)f(symbol)166
+4421 y(be)e(normalised)f(to)g(a)h(tuple)f(of)h(abstractions\).)166
+4656 y FB(De\002nition)f(2.1.)40 b FA(A)23 b Fx(nominal)f(signatur)l(e)
+h FA(is)g(speci\002ed)g(by:)g(a)h(set)g(of)f Fx(sorts)g(of)g(atoms)f
+FA(\(typical)166 4777 y(symbol)28 b Fz(\027)6 b FA(\);)31
+b(a)f(disjoint)e(set)i(of)g Fx(sorts)e(of)i(data)f FA(\(typical)g
+(symbol)g Fz(\016)t FA(\);)g(and)h(a)h(set)e(of)h Fx(function)166
+4897 y(symbols)25 b FA(\(typical)h(symbol)f Fz(f)11 b
+FA(\),)26 b(each)h(of)g(which)f(has)g(an)g Fx(arity)g
+FA(of)g(the)h(form)f Fz(\034)42 b Fw(!)30 b Fz(\016)t
+FA(.)c(Here)h Fz(\034)166 5018 y FA(ranges)d(o)o(v)o(er)f(\(compound\))
+g Fx(sorts)g FA(gi)n(v)o(en)g(by)g(the)h(grammar)g Fz(\034)39
+b Fu(::=)28 b Fz(\027)34 b Fw(j)27 b Fz(\016)32 b Fw(j)27
+b Fu(1)h Fw(j)f Fz(\034)k Fw(\002)19 b Fz(\034)39 b Fw(j)28
+b(h)p Fz(\027)6 b Fw(i)p Fz(\034)39 b FA(.)166 5138 y(Sorts)26
+b(of)g(the)h(form)f Fw(h)p Fz(\027)6 b Fw(i)p Fz(\034)37
+b FA(classify)26 b(terms)g(that)f(are)i(binding)e(abstractions)g(of)i
+(atoms)e(of)h(sort)166 5258 y Fz(\027)31 b FA(o)o(v)o(er)24
+b(terms)g(of)g(sort)g Fz(\034)11 b FA(.)25 b(W)-8 b(e)25
+b(will)e(e)o(xplain)h(the)g(syntax)g(and)g(properties)g(of)g(such)h
+(terms)e(in)h(a)166 5379 y(moment.)1773 5712 y(6)p eop
+end
+%%Page: 7 7
+TeXDict begin 7 6 bop 166 83 a FB(Example)31 b(2.2.)43
+b FA(Here)31 b(is)f(a)g(nominal)f(signature)g(for)h(e)o(xpressions)f
+(in)h(a)g(small)f(fragment)h(of)166 203 y(ML)24 b([25]:)638
+416 y(sort)g(of)h(atoms:)267 b Fp(vid)638 530 y FA(sort)24
+b(of)h(data:)340 b Fp(exp)638 645 y FA(function)24 b(symbols:)97
+b Fv(vr)29 b Fu(:)e Fp(vid)38 b Fw(!)27 b Fp(exp)1454
+759 y Fv(app)i Fu(:)f Fp(exp)f Fw(\002)c Fp(exp)33 b
+Fw(!)27 b Fp(exp)1454 874 y Fv(fn)i Fu(:)e Fw(h)p Fp(vid)10
+b Fw(i)p Fp(exp)33 b Fw(!)27 b Fp(exp)1454 988 y Fv(lv)i
+Fu(:)e Fp(exp)h Fw(\002)23 b(h)p Fp(vid)9 b Fw(i)p Fp(exp)33
+b Fw(!)28 b Fp(exp)1454 1102 y Fv(lf)h Fu(:)e Fw(h)p
+Fp(vid)10 b Fw(i)p Fu(\(\()p Fw(h)p Fp(vid)f Fw(i)p Fp(exp)c
+Fu(\))23 b Fw(\002)f Fp(exp)6 b Fu(\))27 b Fw(!)h Fp(exp)33
+b FA(.)166 1322 y(The)27 b(function)f(symbol)f Fv(vr)j
+FA(constructs)e(terms)g(of)h(sort)g Fp(exp)32 b FA(representing)27
+b(v)n(alue)f(identi\002ers)166 1442 y(\(named)e(by)g(atoms)f(of)h(sort)
+f Fp(vid)10 b FA(\);)24 b Fv(app)h FA(constructs)e(application)g(e)o
+(xpressions)f(from)i(pairs)g(of)166 1562 y(e)o(xpressions;)29
+b Fv(fn)p FA(,)h Fv(lv)h FA(and)g Fv(lf)g FA(construct)e(terms)h
+(representing)g(respecti)n(v)o(ely)f(function)g(ab-)166
+1683 y(stractions)k(\()p Fv(fn)25 b(x)j Fu(=)p Fz(>)g
+Fv(e)p FA(\),)34 b(local)f(v)n(alue)g(declarations)h(\()p
+Fv(let)26 b(val)f(x)j Fu(=)g Fv(e1)d(in)h(e2)f(end)q
+FA(\))34 b(and)166 1803 y(local)28 b(recursi)n(v)o(e)e(function)h
+(declarations)g(\()p Fv(let)f(fun)g(f)f(x)j Fu(=)g Fv(e1)d(in)h(e2)f
+(end)q FA(\).)j(The)f(arities)h(of)166 1924 y(the)f(function)g(symbols)
+e(specify)j(which)f(are)h(binders)f(and)g(in)g(which)g(w)o(ay)g(their)g
+(ar)n(guments)166 2044 y(are)32 b(bound.)f(F)o(or)h(e)o(xample,)f(in)g
+(the)h(e)o(xpression)e(\()p Fv(let)c(fun)g(f)f(x)j Fu(=)f
+Fv(e1)f(in)f(e2)h(end)p FA(\))32 b(there)g(is)166 2164
+y(a)h(binding)e(occurrence)i(of)g(the)f(v)n(alue)g(identi\002er)g
+Fv(f)h FA(whose)f(scope)h(is)f(both)f(of)i Fv(e1)g FA(and)f
+Fv(e2)q FA(;)166 2285 y(and)i(a)g(binding)e(occurrence)i(of)g(the)g(v)n
+(alue)f(identi\002er)g Fv(x)h FA(whose)f(scope)h(is)f(just)g
+Fv(e1)p FA(.)h(These)166 2405 y(binding)27 b(scopes)h(are)i
+(re\003ected)f(by)g(the)f(ar)n(gument)g(sort)h(of)f(the)h(function)e
+(symbol)g Fv(lf)q FA(.)i(This)166 2525 y(kind)i(of)h(speci\002cation)g
+(of)g(binding)e(scopes)i(is)f(of)h(course)g(a)g(feature)h(of)f
+Fx(higher)n(-or)l(der)f(ab-)166 2646 y(str)o(act)24 b(syntax)h
+FA([26],)g(using)g(function)f(types)h Fz(\027)6 b Fw(!)p
+Fz(\034)37 b FA(in)25 b(simply)e(typed)i Fz(\025)p FA(-calculus)g
+(where)h(we)166 2766 y(use)g(abstraction)g(sorts)g Fw(h)p
+Fz(\027)6 b Fw(i)p Fz(\034)11 b FA(.)26 b(W)-8 b(e)27
+b(shall)f(see)h(that)f(the)g(latter)g(ha)n(v)o(e)g(much)g(more)g
+(elementary)166 2887 y(\(indeed,)f(\002rst-order\))g(properties)f
+(compared)h(with)f(the)h(former)-5 b(.)166 3116 y FB(De\002nition)40
+b(2.3.)49 b FA(Gi)n(v)o(en)38 b(a)i(nominal)e(signature,)h(we)g(assume)
+g(that)g(there)h(are)g(countably)166 3236 y(in\002nite)21
+b(and)h(pairwise)f(disjoint)e(sets)j(of)f Fx(atoms)g
+FA(\(typical)g(symbol)f Fz(a)p FA(\))i(for)g(each)g(sort)f(of)h(atoms)
+166 3357 y Fz(\027)6 b FA(,)27 b(and)f Fx(variables)g
+FA(\(typical)g(symbol)f Fz(X)8 b FA(\))26 b(for)h(each)g(sort)f(of)h
+(atoms)e Fz(\027)34 b FA(and)26 b(each)h(sort)f(of)h(data)166
+3477 y Fz(\016)t FA(.)36 b(The)h Fx(terms)e FA(o)o(v)o(er)h(a)g
+(nominal)f(signature)h(and)g(their)g(sorts)f(are)i(inducti)n(v)o(ely)d
+(de\002ned)j(as)166 3597 y(follo)n(ws,)23 b(where)j(we)f(write)f
+Fz(t)k Fu(:)g Fz(\034)36 b FA(to)25 b(indicate)f(that)g(a)i(term)e
+Fz(t)h FA(has)g(sort)f Fz(\034)11 b FA(.)166 3826 y FB(Unit)25
+b(v)o(alue)50 b Fw(hi)27 b Fu(:)h(1)p FA(.)166 3947 y
+FB(P)o(airs)49 b Fw(h)p Fz(t)510 3962 y FG(1)550 3947
+y Fz(;)17 b(t)629 3962 y FG(2)668 3947 y Fw(i)27 b Fu(:)h
+Fz(\034)831 3962 y FG(1)893 3947 y Fw(\002)23 b Fz(\034)1035
+3962 y FG(2)1075 3947 y FA(,)h(if)h Fz(t)1245 3962 y
+FG(1)1312 3947 y Fu(:)j Fz(\034)1409 3962 y FG(1)1474
+3947 y FA(and)d Fz(t)1678 3962 y FG(2)1745 3947 y Fu(:)j
+Fz(\034)1842 3962 y FG(2)1882 3947 y FA(.)166 4067 y
+FB(Data)50 b Fz(f)27 b(t)h Fu(:)g Fz(\016)t FA(,)c(if)h
+Fz(f)36 b FA(is)24 b(a)h(function)f(symbol)f(of)i(arity)g
+Fz(\034)39 b Fw(!)27 b Fz(\016)i FA(and)c Fz(t)j Fu(:)g
+Fz(\034)11 b FA(.)166 4188 y FB(Atoms)50 b Fz(a)27 b
+Fu(:)h Fz(\027)6 b FA(,)25 b(if)g Fz(a)g FA(is)g(an)g(atom)f(of)h(sort)
+f Fz(\027)6 b FA(.)166 4308 y FB(Atom-abstraction)50
+b Fz(a:t)29 b Fu(:)e Fw(h)p Fz(\027)6 b Fw(i)p Fz(\034)11
+b FA(,)26 b(if)e Fz(a)i FA(is)e(an)h(atom)f(of)h(sort)f
+Fz(\027)32 b FA(and)24 b Fz(t)k Fu(:)g Fz(\034)11 b FA(.)166
+4428 y FB(Suspension)51 b Fz(\031)t Fr(\001)p Fz(X)35
+b Fu(:)28 b Fz(\034)11 b FA(,)21 b(if)g Fz(\031)32 b
+Fu(=)27 b(\()p Fz(a)1414 4443 y FG(1)1471 4428 y Fz(b)1512
+4443 y FG(1)1551 4428 y Fu(\)\()p Fz(a)1678 4443 y FG(2)1735
+4428 y Fz(b)1776 4443 y FG(2)1815 4428 y Fu(\))17 b Fw(\001)g(\001)g
+(\001)e Fu(\()p Fz(a)2092 4443 y Fq(n)2156 4428 y Fz(b)2197
+4443 y Fq(n)2244 4428 y Fu(\))21 b FA(is)f(a)h(\002nite)g(list)f(whose)
+h(elements)266 4549 y Fu(\()p Fz(a)355 4564 y Fq(i)400
+4549 y Fz(b)441 4564 y Fq(i)469 4549 y Fu(\))27 b FA(are)h(pairs)f(of)g
+(atoms,)f(with)g Fz(a)1559 4564 y Fq(i)1614 4549 y FA(and)h
+Fz(b)1826 4564 y Fq(i)1882 4549 y FA(of)g(the)g(same)f(sort,)h(and)g
+Fz(X)35 b FA(is)26 b(a)h(v)n(ariable)g(of)266 4669 y(sort)d
+Fz(\034)11 b FA(,)25 b(where)h Fz(\034)36 b FA(is)24
+b(either)h(a)g(sort)g(of)f(data)h(or)g(a)g(sort)g(of)g(atoms)f(\(i.e.)g
+Fz(\034)40 b Fu(::=)27 b Fz(\027)34 b Fw(j)28 b Fz(\016)t
+FA(\).)166 4898 y(Recall)20 b(that)f(e)n(v)o(ery)g(\002nite)g
+(permutation)f(can)i(be)f(e)o(xpressed)g(as)h(a)f(composition)f(of)h
+(sw)o(appings)166 5019 y Fu(\()p Fz(a)255 5034 y Fq(i)300
+5019 y Fz(b)341 5034 y Fq(i)370 5019 y Fu(\))p FA(;)28
+b(the)g(list)f Fz(\031)33 b FA(of)28 b(pairs)h(of)f(atoms)g(occurring)g
+(in)g(a)h(suspension)e(term)h Fz(\031)t Fr(\001)p Fz(X)36
+b FA(speci\002es)28 b(a)166 5139 y(\002nite)34 b(permutation)f(of)h
+(atoms)g(w)o(aiting)f(to)h(be)g(applied)g(once)g(we)h(kno)n(w)e(more)h
+(about)g(the)166 5259 y(v)n(ariable)e Fz(X)40 b FA(\(by)32
+b(substituting)d(for)k(it,)e(for)i(e)o(xample\).)e(W)-8
+b(e)33 b(represent)f(\002nite)g(permutations)166 5380
+y(in)24 b(this)g(w)o(ay)h(because)g(it)f(is)g(really)g(the)h(operation)
+f(of)g(sw)o(apping)g(which)g(plays)g(a)h(fundamen-)1773
+5712 y(7)p eop end
+%%Page: 8 8
+TeXDict begin 8 7 bop 166 83 a FA(tal)33 b(r)8 b(\210)-41
+b(ole)33 b(in)g(the)f(theory)-6 b(.)33 b(Since,)g(semantically)f
+(speaking,)g(sw)o(apping)g(commutes)g(with)g(all)166
+203 y(term-forming)f(operations,)h(we)h(can)g(normalise)e(terms)h(in)l
+(v)n(olving)f(an)i(e)o(xplicit)e(sw)o(apping)166 324
+y(operation)24 b(by)h(pushing)e(the)i(sw)o(ap)f(in)h(as)g(f)o(ar)g(as)g
+(it)f(will)g(go,)h(until)e(it)h(reaches)i(a)f(v)n(ariable)f(\(ap-)166
+444 y(plying)g(the)h(sw)o(apping)g(to)f(atoms)h(that)g(it)g(meets)f(on)
+h(the)h(w)o(ay\);)f(the)g(terms)g(in)g(De\002nition)f(2.3)166
+565 y(are)38 b(all)f(normalised)e(in)i(this)f(w)o(ay)-6
+b(,)37 b(with)f(e)o(xplicit)f(permutations)h(\223piled)g(up\224)h(in)g
+(front)g(of)166 685 y(v)n(ariables)30 b(gi)n(ving)g(what)h(we)g(ha)n(v)
+o(e)g(called)g Fx(suspensions)p FA(.)f(In)h(case)h(the)f(permutation)e
+Fz(\031)36 b FA(in)30 b(a)166 805 y(suspension)23 b(is)i(the)f(empty)g
+(list,)g(we)h(just)f(write)g Fz(X)33 b FA(for)25 b Fz(\031)t
+Fr(\001)p Fz(X)8 b FA(.)166 1038 y FB(De\002nition)38
+b(2.4.)48 b FA(The)37 b Fx(permutation)e(action)p FA(,)h
+Fz(\031)t Fr(\001)p Fz(t)p FA(,)h(of)g(a)h(\002nite)f(permutation)e(of)
+j(atoms)e Fz(\031)166 1159 y FA(on)31 b(a)h(term)f Fz(t)h
+FA(is)f(de\002ned)h(as)f(in)g(Figure)h(1,)f(making)f(use)i(of)f(the)h
+(follo)n(wing)d(notations.)h(The)166 1279 y(composition)36
+b(of)i(a)h(permutation)d Fz(\031)42 b FA(follo)n(wed)37
+b(by)h(a)h(sw)o(ap)f Fu(\()p Fz(a)17 b(b)p Fu(\))38 b
+FA(is)g(gi)n(v)o(en)e(by)i(list-cons,)166 1399 y(written)25
+b Fu(\()p Fz(a)17 b(b)p Fu(\))30 b(::)f Fz(\031)t FA(.)d(\(Note)f(that)
+h(we)g(apply)f(permutations)f(to)h(terms)g(on)h(the)f(left,)h(and)f
+(hence)166 1520 y(the)20 b(order)g(of)g(the)g(composition)e(is)h(from)h
+(right)f(to)h(left.\))g(The)g(composition)d(of)j Fz(\031)k
+FA(follo)n(wed)19 b(by)166 1640 y(another)j(permutation)f
+Fz(\031)1050 1604 y Ft(0)1096 1640 y FA(is)h(gi)n(v)o(en)f(by)h
+(list-concatenation,)f(written)h(as)g Fz(\031)2772 1604
+y Ft(0)2795 1640 y Fu(@)p Fz(\031)t FA(.)g(The)h Fx(identity)166
+1761 y FA(permutation)j(is)g(gi)n(v)o(en)g(by)h(the)g(empty)f(list)g
+Fu([])p FA(;)h(and)g(the)g Fx(in)l(ver)o(se)g FA(of)g(a)g(permutation)f
+(is)g(gi)n(v)o(en)166 1881 y(by)f(list)e(re)n(v)o(ersal,)h(written)h
+(as)g Fz(\031)1273 1845 y Ft(\000)p FG(1)1367 1881 y
+FA(.)166 2114 y(Permutation)h(actions)g(ha)n(v)o(e)g(e)o(xcellent)g
+(logical)g(properties)g(\(stemming)f(from)h(the)g(f)o(act)h(that)166
+2234 y(the)o(y)f(are)i(bijections\).)e(W)-8 b(e)28 b(e)o(xploit)d
+(these)i(properties)g(in)g(our)g(de\002nition)f(of)h
+Fz(\013)q FA(-equi)n(v)n(alence)166 2355 y(for)g(terms)e(o)o(v)o(er)h
+(a)h(nominal)e(signature,)g(which)h(is)g(respected)h(by)f(substitution)
+d(of)k(terms)f(for)166 2475 y(v)n(ariables)33 b(e)n(v)o(en)h(though)f
+(the)h(latter)g(may)g(in)l(v)n(olv)o(e)f(capture)h(of)g(atoms)g(by)g
+(binders.)f(T)-8 b(o)34 b(do)166 2595 y(so)e(we)h(will)f(need)h(to)f
+(mak)o(e)h(use)g(of)f(an)h(auxiliary)f(relation)g(of)h
+Fx(fr)l(eshness)f FA(between)h(atoms)166 2716 y(and)k(terms,)f(whose)g
+(intended)g(meaning)g(is)g(that)g(the)g(atom)g(does)h(not)f(occur)h
+(free)g(in)f(an)o(y)166 2836 y(substitution)23 b(instance)i(of)h(the)f
+(term.)g(As)h(discussed)e(in)h(the)h(Introduction,)e(our)h(judgements)
+166 2957 y(about)j(term)f(equi)n(v)n(alence)g(\()p Fz(t)34
+b Fw(\031)h Fz(t)1374 2920 y Ft(0)1397 2957 y FA(\))29
+b(need)f(to)g(contain)f(hypotheses)f(about)i(the)g(freshness)g(of)166
+3077 y(atoms)k(with)g(respect)h(to)f(v)n(ariables)g(\()p
+Fz(a)43 b Fu(#)g Fz(X)8 b FA(\);)32 b(and)h(the)g(same)f(goes)h(for)g
+(our)f(judgements)166 3197 y(about)25 b(freshness)g(itself)g(\()p
+Fz(a)30 b Fu(#)f Fz(t)p FA(\).)d(Figure)f(2)h(gi)n(v)o(es)e(a)i
+(syntax-directed)e(inducti)n(v)o(e)g(de\002nition)166
+3318 y(of)h(equi)n(v)n(alence)f(and)g(freshness)h(using)f(judgements)f
+(of)i(the)g(form)1115 3564 y Fw(r)i(`)h Fz(t)g Fw(\031)g
+Fz(t)1517 3523 y Ft(0)1740 3564 y FA(and)199 b Fw(r)28
+b(`)f Fz(a)h Fu(#)g Fz(t)166 3810 y FA(where)40 b Fz(t)g
+FA(and)g Fz(t)743 3773 y Ft(0)806 3810 y FA(are)g(terms)f(of)h(the)f
+(same)h(sort)f(o)o(v)o(er)f(a)i(gi)n(v)o(en)e(nominal)h(signature,)f
+Fz(a)i FA(is)166 3930 y(an)34 b(atom,)g(and)g(the)g Fx(fr)l(eshness)f
+(en)l(vir)l(onment)i Fw(r)f FA(is)f(a)i(\002nite)f(set)g(of)g
+Fx(fr)l(eshness)f(constr)o(aints)166 4050 y Fz(a)c Fu(#)g
+Fz(X)8 b FA(,)26 b(each)g(speci\002ed)g(by)f(an)g(atom)g(and)h(a)g(v)n
+(ariable.)e(Rule)i(\()p Fw(\031)p FA(-suspension\))f(in)g(Figure)h(2)
+166 4171 y(mak)o(es)e(use)h(of)g(the)g(follo)n(wing)e(de\002nition.)166
+4404 y FB(De\002nition)36 b(2.5.)46 b FA(The)35 b Fx(disa)o(gr)l
+(eement)f(set)i FA(of)f(tw)o(o)f(permutations)f Fz(\031)39
+b FA(and)34 b Fz(\031)2909 4368 y Ft(0)2967 4404 y FA(is)h(the)f(set)h
+(of)166 4540 y(atoms)24 b Fp(ds)8 b Fu(\()p Fz(\031)t(;)17
+b(\031)727 4504 y Ft(0)750 4540 y Fu(\))815 4487 y Fo(def)820
+4540 y Fu(=)32 b Fw(f)p Fz(a)27 b Fw(j)h Fz(\031)t Fr(\001)o
+Fz(a)g Fw(6)p Fu(=)g Fz(\031)1444 4504 y Ft(0)1467 4540
+y Fr(\001)p Fz(a)p Fw(g)p FA(.)166 4773 y(Note)38 b(that)f(e)n(v)o(ery)
+g(disagreement)g(set)h Fp(ds)8 b Fu(\()p Fz(\031)t(;)17
+b(\031)1856 4737 y Ft(0)1879 4773 y Fu(\))38 b FA(is)f(a)h(subset)f(of)
+h(the)g Fx(\002nite)f FA(set)h(of)g(atoms)166 4893 y(occurring)29
+b(in)f(either)h(of)g(the)g(lists)f Fz(\031)33 b FA(and)c
+Fz(\031)1712 4857 y Ft(0)1735 4893 y FA(,)g(because)h(if)e
+Fz(a)i FA(does)e(not)h(occur)g(in)g(those)f(lists,)166
+5014 y(then)i(from)g(Figure)h(1)f(we)h(get)g Fz(\031)t
+Fr(\001)o Fz(a)39 b Fu(=)f Fz(a)g Fu(=)g Fz(\031)1819
+4977 y Ft(0)1842 5014 y Fr(\001)p Fz(a)p FA(.)31 b(T)-8
+b(o)30 b(illustrate)f(the)i(use)f(of)h(disagreement)166
+5134 y(sets,)24 b(consider)h(the)f(judgement)909 5380
+y Fw(f)p Fz(a)k Fu(#)g Fz(X)8 b(;)17 b(c)27 b Fu(#)h
+Fz(X)8 b Fw(g)27 b(`)h Fu(\()p Fz(a)17 b(c)p Fu(\)\()p
+Fz(a)g(b)p Fu(\))p Fr(\001)o Fz(X)36 b Fw(\031)28 b Fu(\()p
+Fz(b)17 b(c)p Fu(\))p Fr(\001)p Fz(X)57 b(:)1773 5712
+y FA(8)p eop end
+%%Page: 9 9
+TeXDict begin 9 8 bop 166 3 3288 4 v 166 1022 4 1020
+v 791 332 a Fu([])p Fr(\001)p Fz(a)1028 280 y Fo(def)1032
+332 y Fu(=)104 b Fz(a)327 628 y Fu(\(\()p Fz(a)454 643
+y FG(1)510 628 y Fz(a)561 643 y FG(2)601 628 y Fu(\))27
+b(::)h Fz(\031)t Fu(\))p Fr(\001)p Fz(a)1028 576 y Fo(def)1032
+628 y Fu(=)1212 404 y Fn(8)1212 479 y(>)1212 504 y(>)1212
+528 y(<)1212 678 y(>)1212 703 y(>)1212 728 y(:)1286 488
+y Fz(a)1337 503 y FG(1)1528 488 y FA(if)c Fz(\031)t Fr(\001)p
+Fz(a)k Fu(=)g Fz(a)1938 503 y FG(2)1286 632 y Fz(a)1337
+647 y FG(2)1528 632 y FA(if)c Fz(\031)t Fr(\001)p Fz(a)k
+Fu(=)g Fz(a)1938 647 y FG(1)1286 777 y Fz(\031)t Fr(\001)p
+Fz(a)100 b FA(otherwise)2430 188 y Fz(\031)t Fr(\001)p
+Fw(hi)2698 135 y Fo(def)2702 188 y Fu(=)k Fw(hi)2237
+368 y Fz(\031)t Fr(\001)o Fw(h)p Fz(t)2401 383 y FG(1)2441
+368 y Fz(;)17 b(t)2520 383 y FG(2)2559 368 y Fw(i)2698
+316 y Fo(def)2702 368 y Fu(=)104 b Fw(h)p Fz(\031)t Fr(\001)p
+Fz(t)3047 383 y FG(1)3086 368 y Fz(;)17 b(\031)t Fr(\001)p
+Fz(t)3256 383 y FG(2)3295 368 y Fw(i)2321 549 y Fz(\031)t
+Fr(\001)p Fu(\()p Fz(f)27 b(t)p Fu(\))2698 496 y Fo(def)2702
+549 y Fu(=)104 b Fz(f)27 b Fu(\()p Fz(\031)t Fr(\001)p
+Fz(t)p Fu(\))2318 730 y Fz(\031)t Fr(\001)p Fu(\()p Fz(a:t)p
+Fu(\))2698 677 y Fo(def)2702 730 y Fu(=)104 b(\()p Fz(\031)t
+Fr(\001)p Fz(a)p Fu(\))p Fz(:)p Fu(\()p Fz(\031)t Fr(\001)p
+Fz(t)p Fu(\))2229 910 y Fz(\031)t Fr(\001)o Fu(\()p Fz(\031)2416
+874 y Ft(0)2439 910 y Fr(\001)p Fz(X)8 b Fu(\))2698 858
+y Fo(def)2702 910 y Fu(=)104 b(\()p Fz(\031)t Fu(@)p
+Fz(\031)3114 874 y Ft(0)3137 910 y Fu(\))p Fr(\001)p
+Fz(X)p 3450 1022 V 166 1025 3288 4 v 1075 1162 a FD(Fig.)22
+b(1.)h(Permutation)i(action)g(on)f(terms,)f FC(\031)s
+Fm(\001)q FC(t)p FD(.)p 166 1267 V 166 3748 4 2482 v
+637 1441 445 4 v 637 1525 a Fl(r)i(`)f(hi)i(\031)f(hi)1123
+1464 y FD(\()p Fl(\031)p FD(-unit\))1653 1395 y Fl(r)g(`)g
+FC(t)1868 1409 y FG(1)1932 1395 y Fl(\031)g FC(t)2061
+1362 y Ft(0)2061 1419 y FG(1)2191 1395 y Fl(r)g(`)g FC(t)2406
+1409 y FG(2)2471 1395 y Fl(\031)g FC(t)2600 1362 y Ft(0)2600
+1419 y FG(2)p 1653 1439 987 4 v 1739 1525 a Fl(r)g(`)f(h)p
+FC(t)1988 1539 y FG(1)2028 1525 y FC(;)15 b(t)2101 1539
+y FG(2)2141 1525 y Fl(i)25 b(\031)g(h)p FC(t)2365 1492
+y Ft(0)2365 1550 y FG(1)2405 1525 y FC(;)15 b(t)2478
+1492 y Ft(0)2478 1550 y FG(2)2518 1525 y Fl(i)2680 1460
+y FD(\()p Fl(\031)p FD(-pair\))494 1693 y Fl(r)25 b(`)g
+FC(t)g Fl(\031)g FC(t)863 1660 y Ft(0)p 424 1713 531
+4 v 424 1800 a Fl(r)g(`)g FC(f)g(t)g Fl(\031)g FC(f)f(t)932
+1767 y Ft(0)997 1734 y FD(\()p Fl(\031)p FD(-function)i(symbol\))2052
+1693 y Fl(r)f(`)g FC(t)g Fl(\031)g FC(t)2421 1660 y Ft(0)p
+1979 1713 539 4 v 1979 1800 a Fl(r)g(`)g FC(a:t)g Fl(\031)g
+FC(a:t)2494 1767 y Ft(0)2559 1736 y FD(\()p Fl(\031)p
+FD(-abstraction-1\))741 1968 y FC(a)g Fl(6)p Fk(=)g FC(a)958
+1935 y Ft(0)1072 1968 y Fl(r)g(`)g FC(t)g Fl(\031)g Fk(\()p
+FC(a)15 b(a)1554 1935 y Ft(0)1578 1968 y Fk(\))p Fm(\001)p
+FC(t)1675 1935 y Ft(0)1789 1968 y Fl(r)25 b(`)g FC(a)g
+Fk(#)g FC(t)2178 1935 y Ft(0)p 741 2010 1461 4 v 1190
+2097 a Fl(r)g(`)g FC(a:t)g Fl(\031)g FC(a)1647 2064 y
+Ft(0)1671 2097 y FC(:t)1729 2064 y Ft(0)2243 2033 y FD(\()p
+Fl(\031)p FD(-abstraction-2\))p 338 2314 400 4 v 338
+2394 a Fl(r)g(`)g FC(a)g Fl(\031)g FC(a)779 2337 y FD(\()p
+Fl(\031)p FD(-atom\))1350 2265 y Fk(\()p FC(a)g Fk(#)g
+FC(X)7 b Fk(\))26 b Fl(2)f(r)50 b FD(for)24 b(all)52
+b FC(a)25 b Fl(2)g Fj(ds)8 b Fk(\()p FC(\031)s(;)15 b(\031)2620
+2232 y Ft(0)2644 2265 y Fk(\))p 1350 2307 1330 4 v 1685
+2394 a Fl(r)25 b(`)g FC(\031)s Fm(\001)p FC(X)32 b Fl(\031)25
+b FC(\031)2209 2361 y Ft(0)2233 2394 y Fm(\001)p FC(X)2721
+2329 y FD(\()p Fl(\031)p FD(-suspension\))p 659 2687
+427 4 v 659 2772 a Fl(r)g(`)g FC(a)h Fk(#)f Fl(hi)1128
+2710 y FD(\()p Fk(#)p FD(-unit\))1663 2650 y Fl(r)g(`)g
+FC(a)g Fk(#)g FC(t)2052 2664 y FG(1)2182 2650 y Fl(r)g(`)g
+FC(a)h Fk(#)f FC(t)2572 2664 y FG(2)p 1663 2687 948 4
+v 1831 2772 a Fl(r)g(`)g FC(a)g Fk(#)g Fl(h)p FC(t)2255
+2786 y FG(1)2295 2772 y FC(;)15 b(t)2368 2786 y FG(2)2408
+2772 y Fl(i)2652 2709 y FD(\()p Fk(#)p FD(-pair\))1215
+2933 y Fl(r)25 b(`)f FC(a)i Fk(#)f FC(t)p 1180 2971 459
+4 v 1180 3050 a Fl(r)g(`)g FC(a)g Fk(#)g FC(f)f(t)1680
+2992 y FD(\()p Fk(#)p FD(-function)i(symbol\))p 422 3263
+463 4 v 422 3343 a Fl(r)f(`)g FC(a)g Fk(#)g FC(a:t)926
+3286 y FD(\()p Fk(#)p FD(-abstraction-1\))1794 3218 y
+FC(a)h Fl(6)p Fk(=)f FC(a)2012 3185 y Ft(0)2126 3218
+y Fl(r)g(`)g FC(a)g Fk(#)g FC(t)p 1794 3256 721 4 v 1912
+3343 a Fl(r)g(`)g FC(a)g Fk(#)g FC(a)2316 3310 y Ft(0)2339
+3343 y FC(:t)2556 3279 y FD(\()p Fk(#)p FD(-abstraction-2\))731
+3515 y FC(a)g Fl(6)p Fk(=)g FC(a)948 3482 y Ft(0)p 638
+3552 428 4 v 638 3639 a Fl(r)g(`)f FC(a)i Fk(#)f FC(a)1042
+3606 y Ft(0)1107 3575 y FD(\()p Fk(#)p FD(-atom\))1682
+3516 y Fk(\()p FC(\031)1772 3483 y Ft(\000)p FG(1)1867
+3516 y Fm(\001)p FC(a)g Fk(#)g FC(X)7 b Fk(\))26 b Fl(2)f(r)p
+1682 3559 693 4 v 1767 3639 a(r)g(`)g FC(a)g Fk(#)g FC(\031)s
+Fm(\001)q FC(X)2417 3581 y FD(\()p Fk(#)p FD(-suspension\))p
+3450 3748 4 2482 v 166 3751 3288 4 v 1092 3889 a(Fig.)e(2.)g(Inducti)n
+(v)o(e)i(de\002nition)g(of)f Fl(\031)e FD(and)i Fk(#)p
+FD(.)166 4073 y FA(This)g(holds)f(by)i(applying)e(rule)i(\()p
+Fw(\031)p FA(-suspension\))f(in)g(Figure)h(2,)f(since)h(the)f
+(disagreement)g(set)166 4193 y(of)h(the)g(permutations)e
+Fu(\()p Fz(a)17 b(c)p Fu(\)\()p Fz(a)g(b)p Fu(\))24 b
+FA(and)h Fu(\()p Fz(b)17 b(c)p Fu(\))25 b FA(is)f Fw(f)p
+Fz(a;)17 b(c)p Fw(g)p FA(.)166 4417 y FB(Remark)32 b(2.6)f(\(Fr)n
+(eshness)h(en)l(vir)n(onments\).)46 b FA(Note)31 b(that)g(the)f
+(freshness)h(en)l(vironment)f(on)166 4537 y(the)d(left-hand)f(side)h
+(of)g(judgements)e(in)i(the)g(rules)f(in)h(Figure)g(2)g(does)f(not)h
+(change)g(from)f(hy-)166 4658 y(potheses)d(to)f(conclusion.)g(So)i(in)f
+(the)g(same)g(w)o(ay)g(that)g(we)h(assume)e(v)n(ariables)h(ha)n(v)o(e)g
+(attached)166 4778 y(sorting)i(information,)f(we)i(could)f(dispense)g
+(with)g(the)g(use)h(of)g(freshness)f(en)l(vironments)f(en-)166
+4898 y(tirely)d(by)g(attaching)f(the)h(freshness)g(information)f
+(directly)h(to)g(v)n(ariables.)f(Ho)n(we)n(v)o(er)l(,)g(we)i(\002nd)166
+5019 y(the)31 b(use)g(of)g(freshness)g(en)l(vironments)f(more)h(ele)o
+(gant)f(\(for)i(one)f(thing,)f(without)g(them)g(tw)o(o)166
+5139 y(v)n(ariables)24 b(with)g(the)h(same)g(name)g(b)n(ut)f(dif)n
+(ferent)g(freshness)h(information)f(w)o(ould)g(ha)n(v)o(e)g(to)h(be)166
+5259 y(re)o(garded)30 b(as)h(dif)n(ferent\).)f(The)o(y)g(also)g(mak)o
+(e)g(life)h(simpler)e(when)i(we)f(come)h(on)f(to)g(nominal)166
+5380 y(uni\002cation)24 b(problems)g(and)h(their)f(solutions)f(in)h
+(the)h(ne)o(xt)f(section.)1773 5712 y(9)p eop end
+%%Page: 10 10
+TeXDict begin 10 9 bop 166 83 a FA(Belo)n(w)27 b(we)g(sk)o(etch)f(a)i
+(proof)e(that)h Fw(\031)g FA(is)g(an)g(equi)n(v)n(alence)e(relation.)i
+(At)f(\002rst)h(sight)f(this)g(prop-)166 203 y(erty)k(might)e(be)i
+(surprising)e(considering)g(the)i(\223unsymmetric\224)e(de\002nition)h
+(of)h(the)f(rule)h(\()p Fw(\031)p FA(-)166 324 y(abstraction-2\).)25
+b(Ho)n(we)n(v)o(er)e(it)i(holds)f(because)i(of)f(the)g(good)g(logical)f
+(properties)h(of)g(the)g(rela-)166 444 y(tion)30 b Fw(\031)h
+FA(with)f(respect)h(to)f(permutation)f(actions.)h(Although)f(reasoning)
+h(about)g Fw(\031)h FA(is)f(rather)166 565 y(pleasant)k(once)h(equi)n
+(v)n(alence)e(is)h(pro)o(v)o(ed,)f(establishing)f(it)i(\002rst)h(is)f
+(rather)g(trick)o(y\227mainly)166 685 y(because)29 b(of)f(the)h(lar)n
+(ge)g(number)f(of)g(cases,)h(b)n(ut)f(also)g(because)h(se)n(v)o(eral)f
+(f)o(acts)g(needed)h(in)f(the)166 805 y(proof)21 b(are)i
+(interdependent.)1180 769 y FG(1)1258 805 y FA(W)-8 b(e)22
+b(\002rst)g(sho)n(w)e(that)h(permutations)f(can)i(be)g(mo)o(v)o(ed)e
+(from)h(one)166 926 y(side)28 b(of)g(the)g(freshness)g(relation)g(to)g
+(the)g(other)g(by)g(forming)f(the)i(in)l(v)o(erse)e(permutation,)g(and)
+166 1046 y(that)d(the)h(freshness)g(relation)f(is)g(preserv)o(ed)h
+(under)g(permutation)e(actions.)166 1266 y FB(Lemma)i(2.7.)206
+1486 y Fx(\(1\))50 b(If)41 b Fw(r)28 b(`)g Fz(a)f Fu(#)h
+Fz(\031)t Fr(\001)p Fz(t)42 b Fx(then)f Fw(r)28 b(`)f
+Fz(\031)1500 1450 y Ft(\000)p FG(1)1594 1486 y Fr(\001)p
+Fz(a)h Fu(#)g Fz(t)g Fx(.)206 1606 y(\(2\))50 b(If)41
+b Fw(r)28 b(`)g Fz(\031)t Fr(\001)o Fz(a)g Fu(#)g Fz(t)42
+b Fx(then)f Fw(r)28 b(`)f Fz(a)h Fu(#)g Fz(\031)1688
+1570 y Ft(\000)p FG(1)1782 1606 y Fr(\001)p Fz(t)g Fx(.)206
+1727 y(\(3\))50 b(If)41 b Fw(r)28 b(`)g Fz(a)f Fu(#)h
+Fz(t)42 b Fx(then)f Fw(r)28 b(`)g Fz(\031)t Fr(\001)o
+Fz(a)g Fu(#)g Fz(\031)t Fr(\001)p Fz(t)g Fx(.)166 2059
+y FB(PR)m(OOF)-11 b(.)49 b FA(\(1\))27 b(and)g(\(2\))h(are)g(by)f
+(routine)g(inductions)f(on)h(the)g(structure)g(of)g Fz(t)p
+FA(,)h(using)e(the)h(f)o(act)166 2179 y(that)g Fz(\031)t
+Fr(\001)p Fz(a)33 b Fu(=)g Fz(b)28 b FA(if)n(f)f Fz(a)34
+b Fu(=)e Fz(\031)1067 2143 y Ft(\000)p FG(1)1162 2179
+y Fr(\001)o Fz(b)d FA(;)e(\(3\))h(is)f(a)h(consequence)g(of)g(\(2\))g
+(and)f(the)h(f)o(act)g(that)f(permuta-)166 2300 y(tions)d(are)h
+(bijections)f(on)g(atoms.)p 3382 2300 4 68 v 3386 2236
+60 4 v 3386 2300 V 3445 2300 4 68 v 166 2632 a(According)h(to)h(the)f
+(de\002nition)g(of)h(the)f(permutation)g(action)g(gi)n(v)o(en)f(in)h
+(Figure)h(1,)g(if)g(we)g(push)166 2752 y(a)38 b(permutation)f(inside)g
+(a)i(term,)e(we)i(need)f(to)g(apply)f(the)h(permutation)f(to)h(all)f
+(atoms)h(we)166 2873 y(meet)e(on)g(the)g(w)o(ay)-6 b(.)35
+b(Suppose)h(we)h(apply)e(tw)o(o)h(distinct)e(permutations,)h(say)h
+Fz(\031)k FA(and)c Fz(\031)3292 2836 y Ft(0)3315 2873
+y FA(,)g(to)166 2993 y(a)d(term)f Fz(t)p FA(,)g(then)g(in)g(general)h
+Fz(\031)t Fr(\001)o Fz(t)g FA(and)f Fz(\031)1583 2957
+y Ft(0)1607 2993 y Fr(\001)o Fz(t)h FA(are)g(not)f Fz(\013)q
+FA(-equi)n(v)n(alent\227the)d(disagreement)j(set)166
+3113 y Fz(ds)p Fu(\()p Fz(\031)t(;)17 b(\031)463 3077
+y Ft(0)485 3113 y Fu(\))36 b FA(characterises)g(all)f(atoms)g(which)g
+(potentially)f(lead)h(to)h(dif)n(ferences.)f(Ho)n(we)n(v)o(er)l(,)166
+3234 y(if)30 b(we)g(assume)g(that)g(all)f(atoms)h(in)f
+Fz(ds)p Fu(\()p Fz(\031)t(;)17 b(\031)1709 3198 y Ft(0)1732
+3234 y Fu(\))30 b FA(are)h(fresh)f(for)g Fz(t)p FA(,)h(then)e(we)i(can)
+f(infer)g(that)g(the)166 3354 y(permutation)19 b(actions)g(produce)h
+(equi)n(v)n(alent)e(terms.)h(This)h(is)f(made)h(precise)g(in)g(the)g
+(follo)n(wing)166 3474 y(lemma.)166 3694 y FB(Lemma)31
+b(2.8.)44 b Fx(Given)30 b(any)h Fz(\031)j Fx(and)c Fz(\031)1488
+3658 y Ft(0)1511 3694 y Fx(,)h(if)f Fw(r)38 b(`)g Fz(a)g
+Fu(#)g Fz(t)31 b Fx(holds)e(for)h(all)g Fz(a)38 b Fw(2)g
+Fz(ds)p Fu(\()p Fz(\031)t(;)17 b(\031)3166 3658 y Ft(0)3189
+3694 y Fu(\))p Fx(,)30 b(then)166 3815 y Fw(r)e(`)f Fz(\031)t
+Fr(\001)p Fz(t)h Fw(\031)g Fz(\031)683 3779 y Ft(0)706
+3815 y Fr(\001)p Fz(t)p Fx(.)166 4147 y FB(PR)m(OOF)-11
+b(.)49 b FA(By)25 b(induction)g(on)g(the)g(structure)h(of)f
+Fz(t)p FA(,)h(for)g(all)f Fz(\031)30 b FA(and)25 b Fz(\031)2541
+4111 y Ft(0)2590 4147 y FA(simultaneously)-6 b(,)23 b(using)166
+4267 y(the)e(f)o(act)h(about)f(disagreement)g(sets)g(that)g(for)h(all)f
+(atoms)f Fz(a;)d(b)p FA(,)22 b(if)f Fz(a)28 b Fw(2)g
+Fz(ds)p Fu(\()p Fz(\031)t(;)17 b Fu(\()p Fz(\031)t Fr(\001)o
+Fz(b)51 b(\031)3090 4231 y Ft(0)3113 4267 y Fr(\001)p
+Fz(b)p Fu(\))28 b(::)g Fz(\031)3393 4231 y Ft(0)3416
+4267 y Fu(\))166 4388 y FA(then)d Fz(a)i Fw(2)h Fz(ds)p
+Fu(\()p Fz(\031)t(;)17 b(\031)832 4352 y Ft(0)855 4388
+y Fu(\))p FA(.)p 3382 4388 V 3386 4324 60 4 v 3386 4388
+V 3445 4388 4 68 v 166 4720 a(An)27 b(e)o(xample)g(of)g(this)g(lemma)f
+(is)h(that)g Fw(r)33 b(`)f Fz(\031)t Fr(\001)p Fu(\()p
+Fz(a)17 b(b)p Fu(\))p Fr(\001)p Fz(t)32 b Fw(\031)h Fu(\()p
+Fz(\031)t Fr(\001)p Fz(a)50 b(\031)t Fr(\001)p Fz(b)p
+Fu(\))p Fr(\001)p Fz(\031)t Fr(\001)p Fz(t)27 b FA(is)g(a)h(v)n(alid)e
+(judge-)166 4840 y(ment,)e(because)h(the)g(disagreement)f(set)h
+Fz(ds)p Fu(\()p Fz(\031)t Fu(@\()p Fz(a)17 b(b)p Fu(\))g
+Fz(;)32 b Fu(\()p Fz(\031)t Fr(\001)p Fz(a)50 b(\031)t
+Fr(\001)p Fz(b)p Fu(\))28 b(::)f Fz(\031)t Fu(\))e FA(is)g(empty)-6
+b(.)p 166 4943 299 4 v 166 5007 a FG(1)257 5040 y FD(In)107
+b(addition)j(some)d(further)i(simple)g(properties)h(of)d(permutations)j
+(and)e(dis-)166 5153 y(agreement)j(sets)g(need)f(to)f(be)h(established)
+j(\002rst.)c(A)f(machine-check)o(ed)114 b(proof)166 5266
+y(of)i FF(all)g FD(results)h(using)g(the)g(theorem)f(pro)o(v)o(er)h
+(Isabelle)h(can)e(be)g(found)h(at)166 5379 y Fi(http://www.cl.c)o(am)o
+(.a)o(c.)o(uk/)o(us)o(er)o(s/)o(cu)o(200)o(/U)o(ni)o(fi)o(ca)o(tio)o(n)
+p FD(.)1748 5712 y FA(10)p eop end
+%%Page: 11 11
+TeXDict begin 11 10 bop 166 83 a FA(The)25 b(ne)o(xt)f(lemma)g(sho)n
+(ws)f(that)i Fw(\031)g FA(respects)g(the)g(freshness)f(relation.)166
+317 y FB(Lemma)h(2.9.)41 b Fx(If)25 b Fw(r)j(`)f Fz(a)h
+Fu(#)g Fz(t)d Fx(and)g Fw(r)i(`)h Fz(t)g Fw(\031)g Fz(t)1818
+281 y Ft(0)1842 317 y Fx(,)c(then)h Fw(r)j(`)f Fz(a)h
+Fu(#)g Fz(t)2510 281 y Ft(0)2533 317 y Fx(.)166 770 y
+FB(PR)m(OOF)-11 b(.)49 b FA(Routine)24 b(induction)f(on)i(the)f
+(de\002nition)g(of)h Fw(\031)h FA(using)d(Lemma)i(2.7.)p
+3382 770 4 68 v 3386 706 60 4 v 3386 770 V 3445 770 4
+68 v 166 1224 a(F)o(or)j(sho)n(wing)d(transiti)n(vity)g(of)j(the)f
+(relation)g Fw(\031)p FA(,)h(it)f(will)f(be)i(necessary)g(to)f
+(de\002ne)h(a)g(measure)166 1344 y(that)c(counts)g(all)h(term)f
+(constructors)g(occurring)h(in)f(a)i(term.)166 1578 y
+FB(De\002nition)g(2.10.)40 b FA(The)25 b Fx(size)g FA(of)f(a)i(term)e
+Fz(t)h FA(is)g(the)f(natural)h(number)f Fw(j)p Fz(t)p
+Fw(j)h FA(de\002ned)g(by:)1184 1844 y Fw(j)p Fz(\031)t
+Fr(\001)p Fz(X)8 b Fw(j)p Fz(;)17 b Fw(j)p Fz(a)p Fw(j)p
+Fz(;)g Fw(jhij)1774 1792 y Fo(def)1778 1844 y Fu(=)32
+b(1)1368 2012 y Fw(j)p Fz(a:t)p Fw(j)p Fz(;)17 b Fw(j)p
+Fz(f)26 b(t)p Fw(j)1774 1959 y Fo(def)1778 2012 y Fu(=)32
+b(1)22 b(+)g Fw(j)p Fz(t)p Fw(j)1420 2179 y(jh)p Fz(t)1522
+2194 y FG(1)1561 2179 y Fz(;)17 b(t)1640 2194 y FG(2)1680
+2179 y Fw(ij)1774 2126 y Fo(def)1778 2179 y Fu(=)32 b(1)22
+b(+)g Fw(j)p Fz(t)2118 2194 y FG(1)2158 2179 y Fw(j)f
+Fu(+)h Fw(j)p Fz(t)2368 2194 y FG(2)2408 2179 y Fw(j)166
+2539 y FA(Notice)d(that)g(the)g(size)g(of)h(a)f(term)g(is)g(preserv)o
+(ed)g(under)h(permutation)d(actions)i(\(i.e.)g Fw(j)p
+Fz(\031)t Fr(\001)p Fz(t)p Fw(j)27 b Fu(=)h Fw(j)p Fz(t)p
+Fw(j)p FA(\))166 2660 y(and)d(respected)g(by)f(the)h(relation)f
+Fw(\031)i FA(in)e(the)h(sense)f(that)h(if)g Fw(r)i(`)h
+Fz(t)g Fw(\031)g Fz(t)2591 2624 y Ft(0)2639 2660 y FA(then)d
+Fw(j)p Fz(t)p Fw(j)i Fu(=)h Fw(j)p Fz(t)3121 2624 y Ft(0)3144
+2660 y Fw(j)p FA(.)166 2894 y FB(Theor)n(em)f(2.11)d(\(Equi)o(v)o
+(alence\).)43 b Fw(r)28 b(`)f(\000)h(\031)g(\000)e Fx(is)e(an)h
+(equivalence)f(r)l(elation.)166 3347 y FB(PR)m(OOF)-11
+b(.)49 b FA(Re\003e)o(xi)n(vity)19 b(is)h(by)h(a)g(simple)e(induction)h
+(on)g(the)h(structure)f(of)h(terms.)f(T)m(ransiti)n(vity)166
+3467 y(is)k(by)h(an)g(induction)e(on)h(the)h(size)g(of)g(terms:)e(a)i
+(slight)f(complication)f(is)h(that)g(man)o(y)g(subcases)166
+3588 y(need)e(to)f(be)h(analysed)f(\(for)h(e)o(xample)f(\002)n(v)o(e)g
+(subcases)h(when)f(dealing)g(with)g(abstractions\))g(and)166
+3708 y(also)26 b(that)g(transiti)n(vity)e(needs)j(to)f(be)g(sho)n(wn)g
+(by)g(mutual)f(induction)g(with)h(the)g(f)o(act)h(that)f
+Fw(\031)h FA(is)166 3829 y(preserv)o(ed)e(under)f(permutation)g
+(actions,)g(that)g(is)837 4076 y(gi)n(v)o(en)f(an)o(y)k
+Fz(\031)t(;)45 b FA(if)27 b Fw(r)h(`)f Fz(t)h Fw(\031)g
+Fz(t)1870 4035 y Ft(0)1922 4076 y FA(then)f Fw(r)g(`)h
+Fz(\031)t Fr(\001)p Fz(t)g Fw(\031)g Fz(\031)t Fr(\001)p
+Fz(t)2705 4035 y Ft(0)2756 4076 y Fz(:)555 b FA(\(9\))166
+4323 y(W)-8 b(e)36 b(illustrate)e(the)i(proof)f(of)h(transiti)n(vity)d
+(for)j(the)f(case)i(when)e Fw(r)48 b(`)f Fz(a)2767 4338
+y FG(1)2807 4323 y Fz(:t)2869 4338 y FG(1)2957 4323 y
+Fw(\031)h Fz(a)3133 4338 y FG(2)3172 4323 y Fz(:t)3234
+4338 y FG(2)3310 4323 y FA(and)166 4444 y Fw(r)28 b(`)f
+Fz(a)416 4459 y FG(2)456 4444 y Fz(:t)518 4459 y FG(2)585
+4444 y Fw(\031)i Fz(a)742 4459 y FG(3)781 4444 y Fz(:t)843
+4459 y FG(3)908 4444 y FA(hold,)24 b(with)g Fz(a)1388
+4459 y FG(1)1428 4444 y FA(,)g Fz(a)1528 4459 y FG(2)1593
+4444 y FA(and)h Fz(a)1813 4459 y FG(3)1877 4444 y FA(all)g(distinct)e
+(atoms,)h(and)h(we)g(ha)n(v)o(e)f(to)h(pro)o(v)o(e)166
+4564 y Fw(r)j(`)f Fz(a)416 4579 y FG(1)456 4564 y Fz(:t)518
+4579 y FG(1)585 4564 y Fw(\031)i Fz(a)742 4579 y FG(3)781
+4564 y Fz(:t)843 4579 y FG(3)883 4564 y FA(.)20 b(By)h(the)f(\()p
+Fw(\031)p FA(-abstraction-2\))h(rule)g(we)f(can)h(infer)g(from)f(the)g
+(assumptions)166 4685 y(the)25 b(follo)n(wing)e(f)o(acts:)803
+4954 y(\(i\))155 b Fw(r)27 b(`)h Fz(t)1286 4969 y FG(1)1353
+4954 y Fw(\031)g Fu(\()p Fz(a)1547 4969 y FG(1)1604 4954
+y Fz(a)1655 4969 y FG(2)1694 4954 y Fu(\))p Fr(\001)p
+Fz(t)1799 4969 y FG(2)2075 4954 y FA(\(ii\))119 b Fw(r)28
+b(`)f Fz(a)2566 4969 y FG(1)2633 4954 y Fu(#)h Fz(t)2777
+4969 y FG(2)803 5135 y FA(\(iii\))99 b Fw(r)27 b(`)h
+Fz(t)1286 5150 y FG(2)1353 5135 y Fw(\031)g Fu(\()p Fz(a)1547
+5150 y FG(2)1604 5135 y Fz(a)1655 5150 y FG(3)1694 5135
+y Fu(\))p Fr(\001)p Fz(t)1799 5150 y FG(3)2075 5135 y
+FA(\(i)n(v\))99 b Fw(r)28 b(`)f Fz(a)2566 5150 y FG(2)2633
+5135 y Fu(#)h Fz(t)2777 5150 y FG(3)166 5380 y FA(Belo)n(w)d(we)g(gi)n
+(v)o(e)e(the)i(steps)f(that)g(pro)o(v)o(e)g Fw(r)k(`)f
+Fz(a)1826 5395 y FG(1)1866 5380 y Fz(:t)1928 5395 y FG(1)1995
+5380 y Fw(\031)i Fz(a)2152 5395 y FG(3)2191 5380 y Fz(:t)2253
+5395 y FG(3)2293 5380 y FA(.)1748 5712 y(11)p eop end
+%%Page: 12 12
+TeXDict begin 12 11 bop 166 126 a FA(\(a\))106 b Fw(r)28
+b(`)f Fu(\()p Fz(a)670 141 y FG(1)726 126 y Fz(a)777
+141 y FG(2)817 126 y Fu(\))p Fr(\001)p Fz(t)922 141 y
+FG(2)989 126 y Fw(\031)h Fu(\()p Fz(a)1183 141 y FG(1)1239
+126 y Fz(a)1290 141 y FG(2)1330 126 y Fu(\)\()p Fz(a)1457
+141 y FG(2)1513 126 y Fz(a)1564 141 y FG(3)1604 126 y
+Fu(\))p Fr(\001)p Fz(t)1709 141 y FG(3)2740 126 y FA(by)c(\(iii\))h
+(and)g(IH)g(\(9\))166 307 y(\(b\))100 b Fw(r)28 b(`)f
+Fz(t)616 322 y FG(1)683 307 y Fw(\031)i Fu(\()p Fz(a)878
+322 y FG(1)934 307 y Fz(a)985 322 y FG(2)1024 307 y Fu(\)\()p
+Fz(a)1151 322 y FG(2)1208 307 y Fz(a)1259 322 y FG(3)1298
+307 y Fu(\))p Fr(\001)p Fz(t)1403 322 y FG(3)2255 307
+y FA(by)c(\(i\),)g(\(a\))g(and)g(IH)g(\(transiti)n(vity\))166
+488 y(\(c\))106 b Fz(ds)p Fu(\(\()p Fz(a)606 503 y FG(1)662
+488 y Fz(a)713 503 y FG(2)752 488 y Fu(\)\()p Fz(a)879
+503 y FG(2)935 488 y Fz(a)986 503 y FG(3)1026 488 y Fu(\))17
+b Fz(;)33 b Fu(\()p Fz(a)1230 503 y FG(1)1286 488 y Fz(a)1337
+503 y FG(3)1377 488 y Fu(\)\))27 b(=)h Fw(f)p Fz(a)1685
+503 y FG(1)1724 488 y Fz(;)17 b(a)1819 503 y FG(2)1859
+488 y Fw(g)1038 b FA(by)25 b(de\002nition)166 668 y(\(d\))100
+b Fw(r)28 b(`)f Fz(a)632 683 y FG(1)699 668 y Fu(#)h(\()p
+Fz(a)897 683 y FG(2)954 668 y Fz(a)1005 683 y FG(3)1044
+668 y Fu(\))p Fr(\001)p Fz(t)1149 683 y FG(3)2361 668
+y FA(by)c(\(ii\),)h(\(iii\))f(and)h(Lemma)f(2.9)166 849
+y(\(e\))106 b Fw(r)28 b(`)f Fz(a)632 864 y FG(1)699 849
+y Fu(#)h Fz(t)843 864 y FG(3)1804 849 y FA(by)c Fu(\()p
+Fz(a)2017 864 y FG(2)2073 849 y Fz(a)2124 864 y FG(3)2164
+849 y Fu(\))p Fr(\001)p Fz(a)2285 864 y FG(1)2352 849
+y Fu(=)k Fz(a)2507 864 y FG(1)2546 849 y FA(,)d(\(d\))g(and)g(Lemma)f
+(2.7\(i\))166 1029 y(\(f\))117 b Fw(r)28 b(`)f Fu(\()p
+Fz(a)670 1044 y FG(1)726 1029 y Fz(a)777 1044 y FG(2)817
+1029 y Fu(\)\()p Fz(a)944 1044 y FG(2)1000 1029 y Fz(a)1051
+1044 y FG(3)1091 1029 y Fu(\))p Fr(\001)o Fz(t)1195 1044
+y FG(3)1263 1029 y Fw(\031)h Fu(\()p Fz(a)1457 1044 y
+FG(1)1513 1029 y Fz(a)1564 1044 y FG(3)1604 1029 y Fu(\))p
+Fr(\001)p Fz(t)1709 1044 y FG(3)2220 1029 y FA(by)c(\(c\),)h(\(i)n
+(v\),)g(\(e\))g(and)g(Lemma)f(2.8)166 1210 y(\(g\))100
+b Fw(r)28 b(`)f Fz(t)616 1225 y FG(1)683 1210 y Fw(\031)i
+Fu(\()p Fz(a)878 1225 y FG(1)934 1210 y Fz(a)985 1225
+y FG(3)1024 1210 y Fu(\))p Fr(\001)p Fz(t)1129 1225 y
+FG(3)2244 1210 y FA(by)c(\(b\),)g(\(f\))g(and)g(IH)g(\(transiti)n
+(vity\))166 1390 y(\(h\))100 b Fw(r)28 b(`)f Fz(a)632
+1405 y FG(1)672 1390 y Fz(:t)734 1405 y FG(1)801 1390
+y Fw(\031)h Fz(a)957 1405 y FG(3)997 1390 y Fz(:t)1059
+1405 y FG(3)2162 1390 y FA(by)c(\(e\),)i(\(g\))f(and)g(\()p
+Fw(\031)p FA(-abstraction-2\))166 1621 y(The)g(other)g(cases)g(are)h
+(by)f(similar)e(ar)n(guments.)i(Symmetry)f(is)g(then)h(by)g(a)g
+(routine)f(induction)166 1741 y(on)h(the)f(de\002nition)g(of)h
+Fw(\031)g FA(using)f(Lemma)g(2.8)h(and)g(transiti)n(vity)-6
+b(.)p 3382 1741 4 68 v 3386 1677 60 4 v 3386 1741 V 3445
+1741 4 68 v 166 2082 a(No)n(w)23 b(it)h(is)f(relati)n(v)o(ely)g
+(straightforw)o(ard)g(to)h(obtain)f(the)h(follo)n(wing)e(properties)i
+(of)g(our)g(equi)n(v-)166 2202 y(alence)h(relation)g(with)f(respect)h
+(to)f(permutation)g(actions.)166 2422 y FB(Cor)n(ollary)g(2.12.)206
+2643 y Fx(\(1\))50 b Fw(r)28 b(`)f Fz(t)h Fw(\031)g Fz(\031)798
+2606 y Ft(\000)p FG(1)893 2643 y Fr(\001)o Fz(\031)t
+Fr(\001)p Fz(t)1050 2606 y Ft(0)1126 2643 y Fx(if)c(and)h(only)f(if)52
+b Fw(r)28 b(`)g Fz(t)f Fw(\031)i Fz(t)2088 2606 y Ft(0)2139
+2643 y Fx(.)206 2763 y(\(2\))50 b Fw(r)28 b(`)f Fz(t)h
+Fw(\031)g Fz(\031)t Fr(\001)p Fz(t)865 2727 y Ft(0)941
+2763 y Fx(if)c(and)h(only)f(if)52 b Fw(r)28 b(`)f Fz(\031)1758
+2727 y Ft(\000)p FG(1)1853 2763 y Fr(\001)p Fz(t)g Fw(\031)i
+Fz(t)2088 2727 y Ft(0)2139 2763 y Fx(.)206 2883 y(\(3\))50
+b(Given)32 b(any)h Fz(\031)j Fx(and)c Fz(\031)1151 2847
+y Ft(0)1174 2883 y Fx(,)h(if)48 b Fw(r)42 b(`)f Fz(\031)t
+Fr(\001)p Fz(t)h Fw(\031)g Fz(\031)1909 2847 y Ft(0)1932
+2883 y Fr(\001)p Fz(t)50 b Fx(then)32 b(for)f(all)h Fz(a)h
+Fx(in)f Fz(ds)p Fu(\()p Fz(\031)t(;)17 b(\031)3030 2847
+y Ft(0)3052 2883 y Fu(\))32 b Fx(we)i(have)372 3004 y
+Fw(r)28 b(`)f Fz(a)h Fu(#)g Fz(t)p Fx(.)166 3345 y FB(PR)m(OOF)-11
+b(.)49 b FA(\(i\))38 b(follo)n(ws)f(immediately)g(from)h(Lemma)g(2.8)h
+(and)f(transiti)n(vity;)e(\(ii\))i(follo)n(ws)166 3465
+y(from)i(\(9\))h(and)f(\(i\);)g(and)g(\(iii\))g(is)g(by)g(a)g(routine)g
+(induction)f(on)h(the)g(structure)g(of)g Fz(t)h FA(using)166
+3585 y(Lemma)24 b(2.9.)p 3382 3585 V 3386 3522 60 4 v
+3386 3585 V 3445 3585 4 68 v 166 3927 a(The)33 b(main)f(reason)h(for)g
+(using)f(suspensions)f(in)h(the)h(syntax)f(of)g(terms)h(is)f(to)g
+(enable)h(a)g(def-)166 4047 y(inition)d(of)j Fx(substitution)c(of)i
+(terms)h(for)f(variables)g FA(that)h(allo)n(ws)f(capture)h(of)g(free)h
+(atoms)e(by)166 4167 y(atom-abstractions)i(while)i(still)e(respecting)i
+Fz(\013)q FA(-equi)n(v)n(alence.)e(The)i(follo)n(wing)e(lemma)i(es-)166
+4288 y(tablishes)24 b(this.)f(First)i(we)g(gi)n(v)o(e)f(some)g
+(terminology)f(and)i(notation)e(for)i(term-substitution.)166
+4508 y FB(De\002nition)37 b(2.13.)47 b FA(A)36 b Fx(substitution)d
+Fz(\033)40 b FA(is)c(a)g(sort-respecting)f(function)g(from)h(v)n
+(ariables)f(to)166 4628 y(terms)28 b(with)h(the)f(property)h(that)f
+Fz(\033)t Fu(\()p Fz(X)8 b Fu(\))35 b(=)g Fz(X)i FA(for)29
+b(all)g(b)n(ut)f(\002nitely)h(man)o(y)f(v)n(ariables)g
+Fz(X)8 b FA(.)28 b(W)-8 b(e)166 4749 y(write)20 b Fp(dom)6
+b Fu(\()p Fz(\033)t Fu(\))20 b FA(for)g(the)g(\002nite)f(set)h(of)g(v)n
+(ariables)f Fz(X)27 b FA(satisfying)19 b Fz(\033)t Fu(\()p
+Fz(X)8 b Fu(\))27 b Fw(6)p Fu(=)g Fz(X)8 b FA(.)20 b(If)g
+Fp(dom)7 b Fu(\()p Fz(\033)t Fu(\))20 b FA(con-)166 4869
+y(sists)28 b(of)h(distinct)f(v)n(ariables)g Fz(X)1269
+4884 y FG(1)1308 4869 y Fz(;)17 b(:)g(:)g(:)f(;)h(X)1608
+4884 y Fq(n)1684 4869 y FA(and)29 b Fz(\033)t Fu(\()p
+Fz(X)2035 4884 y Fq(i)2063 4869 y Fu(\))36 b(=)f Fz(t)2283
+4884 y Fq(i)2341 4869 y FA(for)29 b Fz(i)36 b Fu(=)g(1)p
+Fz(::n)p FA(,)29 b(we)h(sometimes)166 4989 y(write)25
+b Fz(\033)k FA(as)1204 5110 y Fz(\033)j Fu(=)27 b([)p
+Fz(X)1502 5125 y FG(1)1569 5110 y Fu(:=)h Fz(t)1735 5125
+y FG(1)1775 5110 y Fz(;)17 b(:)g(:)g(:)e(;)i(X)2074 5125
+y Fq(n)2149 5110 y Fu(:=)27 b Fz(t)2314 5125 y Fq(n)2362
+5110 y Fu(])p Fz(:)872 b FA(\(10\))166 5259 y(W)-8 b(e)37
+b(write)f Fz(\033)t Fu(\()p Fz(t)p Fu(\))g FA(for)h(the)f(result)g(of)g
+Fx(applying)f(a)h(substitution)e Fz(\033)40 b FA(to)c(a)h(term)f
+Fz(t)p FA(;)g(this)f(is)h(the)166 5380 y(term)g(obtained)g(from)g
+Fz(t)h FA(by)f(replacing)g(each)h(suspension)e Fz(\031)t
+Fr(\001)p Fz(X)44 b FA(in)36 b Fz(t)h FA(\(as)g Fz(X)44
+b FA(ranges)37 b(o)o(v)o(er)1748 5712 y(12)p eop end
+%%Page: 13 13
+TeXDict begin 13 12 bop 166 83 a Fp(dom)7 b Fu(\()p Fz(\033)t
+Fu(\))p FA(\))39 b(by)h(the)f(term)g Fz(\031)t Fr(\001)p
+Fz(\033)t Fu(\()p Fz(X)8 b Fu(\))39 b FA(got)g(by)g(letting)f
+Fz(\031)43 b FA(act)d(on)f(the)g(term)g Fz(\033)t Fu(\()p
+Fz(X)8 b Fu(\))40 b FA(using)e(the)166 203 y(de\002nition)27
+b(in)g(Figure)h(1.)g(F)o(or)g(e)o(xample,)e(if)i Fz(\033)37
+b Fu(=)c([)p Fz(X)41 b Fu(:=)33 b Fw(h)p Fz(b;)17 b(Y)22
+b Fw(i)p Fu(])27 b FA(and)h Fz(t)33 b Fu(=)g Fz(a:)p
+Fu(\()p Fz(a)17 b(b)p Fu(\))p Fr(\001)q Fz(X)8 b FA(,)27
+b(then)166 324 y Fz(\033)t Fu(\()p Fz(t)p Fu(\))h(=)g
+Fz(a:)p Fw(h)p Fz(a;)17 b Fu(\()p Fz(a)g(b)p Fu(\))p
+Fr(\001)p Fz(Y)k Fw(i)p FA(.)k(Gi)n(v)o(en)e(substitutions)f
+Fz(\033)29 b FA(and)c Fz(\033)2169 288 y Ft(0)2193 324
+y FA(,)g(and)f(freshness)h(en)l(vironments)f Fw(r)166
+444 y FA(and)h Fw(r)418 408 y Ft(0)441 444 y FA(,)g(we)g(write)869
+747 y(\(a\))100 b Fw(r)1162 705 y Ft(0)1213 747 y Fw(`)28
+b Fz(\033)t Fu(\()p Fw(r)p Fu(\))199 b FA(and)g(\(b\))100
+b Fw(r)27 b(`)h Fz(\033)k Fw(\031)c Fz(\033)2728 705
+y Ft(0)3288 747 y FA(\(11\))166 1049 y(to)i(mean,)f(for)i(\(a\),)f
+(that)g Fw(r)1119 1013 y Ft(0)1179 1049 y Fw(`)37 b Fz(a)g
+Fu(#)h Fz(\033)t Fu(\()p Fz(X)8 b Fu(\))30 b FA(holds)f(for)h(each)g
+Fu(\()p Fz(a)38 b Fu(#)f Fz(X)8 b Fu(\))37 b Fw(2)h(r)30
+b FA(and,)f(for)i(\(b\),)166 1169 y(that)24 b Fw(r)k(`)g
+Fz(\033)t Fu(\()p Fz(X)8 b Fu(\))27 b Fw(\031)h Fz(\033)955
+1133 y Ft(0)978 1169 y Fu(\()p Fz(X)8 b Fu(\))25 b FA(holds)f(for)h
+(all)f Fz(X)36 b Fw(2)28 b Fp(dom)7 b Fu(\()p Fz(\033)t
+Fu(\))22 b Fw([)g Fp(dom)7 b Fu(\()p Fz(\033)2601 1133
+y Ft(0)2624 1169 y Fu(\))p FA(.)166 1430 y FB(Lemma)28
+b(2.14)g(\(Substitution\).)44 b Fx(Substitution)25 b(commutes)j(with)f
+(the)h(permutation)d(action:)166 1551 y Fz(\033)t Fu(\()p
+Fz(\031)t Fr(\001)p Fz(t)p Fu(\))j(=)f Fz(\031)t Fr(\001)p
+Fu(\()p Fz(\033)t Fu(\()p Fz(t)p Fu(\)\))p Fx(.)e(Substitution)d(also)i
+(pr)l(eserves)h Fw(\031)g Fx(and)f Fu(#)i Fx(in)e(the)h(following)e
+(sense:)206 1812 y(\(1\))50 b(if)24 b Fw(r)535 1776 y
+Ft(0)586 1812 y Fw(`)k Fz(\033)t Fu(\()p Fw(r)p Fu(\))d
+Fx(and)f Fw(r)k(`)f Fz(t)h Fw(\031)g Fz(t)1494 1776 y
+Ft(0)1518 1812 y Fx(,)d(then)f Fw(r)1847 1776 y Ft(0)1898
+1812 y Fw(`)k Fz(\033)t Fu(\()p Fz(t)p Fu(\))f Fw(\031)i
+Fz(\033)t Fu(\()p Fz(t)2422 1776 y Ft(0)2445 1812 y Fu(\))p
+Fx(;)206 1932 y(\(2\))50 b(if)24 b Fw(r)535 1896 y Ft(0)586
+1932 y Fw(`)k Fz(\033)t Fu(\()p Fw(r)p Fu(\))d Fx(and)f
+Fw(r)k(`)f Fz(a)h Fu(#)g Fz(t)p Fx(,)d(then)g Fw(r)1844
+1896 y Ft(0)1895 1932 y Fw(`)i Fz(a)h Fu(#)g Fz(\033)t
+Fu(\()p Fz(t)p Fu(\))p Fx(.)166 2612 y FB(PR)m(OOF)-11
+b(.)49 b FA(The)19 b(\002rst)h(sentence)g(follo)n(ws)e(by)i(a)g
+(routine)f(induction)f(on)i(the)f(structure)h(of)g Fz(t)p
+FA(.)g(The)166 2733 y(second)25 b(follo)n(ws)e(by)h(induction)g(on)g
+(the)h(de\002nition)f(of)h Fw(\031)g FA(and)g Fu(#)g
+FA(using)f(Lemma)g(2.8.)p 3382 2733 4 68 v 3386 2669
+60 4 v 3386 2733 V 3445 2733 4 68 v 166 3413 a(W)-8 b(e)38
+b(claim)f(that)g(the)g(relation)g Fw(\031)h FA(de\002ned)g(in)f(Figure)
+h(2)g(gi)n(v)o(es)d(the)j(correct)g(notion)e(of)i Fz(\013)q
+FA(-)166 3533 y(equi)n(v)n(alence)d(for)h(terms)f(o)o(v)o(er)g(a)h
+(nominal)e(signature.)h(This)g(is)h(reasonable,)g(gi)n(v)o(en)e(Theo-)
+166 3654 y(rem)24 b(2.11)g(and)g(the)g(f)o(act)g(that,)g(by)f
+(de\002nition,)g(it)h(satis\002es)g(rules)g(\()p Fw(\031)p
+FA(-abstraction-1\))g(and)g(\()p Fw(\031)p FA(-)166 3774
+y(abstraction-2\).)34 b(Further)i(e)n(vidence)f(is)f(pro)o(vided)g(by)h
+(the)g(follo)n(wing)e(proposition,)g(which)166 3894 y(sho)n(ws)d(that)i
+(for)g(ground)f(terms)g Fw(\031)h FA(agrees)g(with)f(the)h(follo)n
+(wing)e(more)h(traditional)g(de\002ni-)166 4015 y(tion)24
+b(of)h Fz(\013)q FA(-equi)n(v)n(alence.)166 4276 y FB(De\002nition)c
+(2.15)f(\(Na)954 4275 y(\250)957 4276 y(\021v)o(e)h Fh(\013)p
+FB(-equi)o(v)o(alence\).)37 b FA(De\002ne)21 b(the)f(binary)h(relation)
+f Fz(t)28 b Fu(=)2966 4291 y Fq(\013)3043 4276 y Fz(t)3078
+4240 y Ft(0)3122 4276 y FA(between)166 4396 y(terms)h(o)o(v)o(er)f(a)i
+(nominal)e(signature)h(to)g(be)h(the)f(least)g(sort-respecting)g
+(congruence)g(relation)166 4517 y(satisfying)i Fz(a:t)43
+b Fu(=)818 4532 y Fq(\013)909 4517 y Fz(b:)p Fu([)p Fz(a)1055
+4511 y Ft(7!)1127 4517 y Fz(b)p Fu(])p Fz(t)33 b FA(whene)n(v)o(er)f
+Fz(b)h FA(is)f(an)h(atom)f(\(of)h(the)f(same)h(sort)f(as)g
+Fz(a)p FA(\))h(not)f(oc-)166 4637 y(curring)j(at)h(all)f(in)g(the)h
+(term)f Fz(t)p FA(.)h(Here)g Fu([)p Fz(a)1624 4631 y
+Ft(7!)1695 4637 y Fz(b)p Fu(])p Fz(t)g FA(indicates)f(the)h(result)f
+(of)g(replacing)h(all)f(free)166 4757 y(occurrences)26
+b(of)f Fz(a)g FA(with)f Fz(b)h FA(in)g Fz(t)p FA(.)166
+5019 y FB(Pr)n(oposition)39 b(2.16)f(\(Adequacy\).)50
+b Fx(If)39 b Fz(t)g Fx(and)f Fz(t)1872 4983 y Ft(0)1934
+5019 y Fx(ar)l(e)h(gr)l(ound)e(terms)h FA(\(i.e.)h(terms)f(with)g(no)
+166 5139 y(v)n(ariables)d(and)g(hence)h(no)f(suspensions\))f
+Fx(o)o(ver)i(a)f(nominal)f(signatur)l(e)o(,)g(then)i(the)f(r)l(elation)
+166 5259 y Fz(t)d Fu(=)309 5274 y Fq(\013)390 5259 y
+Fz(t)425 5223 y Ft(0)476 5259 y Fx(of)26 b(De\002nition)g(2.15)g(holds)
+g(if)g(and)h(only)g(if)f Fw(;)31 b(`)h Fz(t)g Fw(\031)g
+Fz(t)2375 5223 y Ft(0)2426 5259 y Fx(is)26 b(pr)l(o)o(vable)g(fr)l(om)g
+(the)g(rules)166 5380 y(in)31 b(F)l(igur)l(e)f(2.)h(Furthermor)l(e)o(,)
+f Fw(;)39 b(`)h Fz(a)f Fu(#)h Fz(t)31 b Fx(is)g(pr)l(o)o(vable)f(if)h
+(and)f(only)h(if)g Fz(a)g Fx(is)g(not)f(in)h(the)g(set)1748
+5712 y FA(13)p eop end
+%%Page: 14 14
+TeXDict begin 14 13 bop 166 83 a Fp(F)-10 b(A)p Fu(\()p
+Fz(t)p Fu(\))25 b Fx(of)g(fr)l(ee)g(atoms)f(of)g Fz(t)p
+Fx(,)h(de\002ned)g(by:)1366 325 y Fp(F)-10 b(A)q Fu(\()p
+Fw(hi)p Fu(\))1674 266 y Fg(def)1676 325 y Fu(=)30 b
+Fw(;)1173 498 y Fp(F)-10 b(A)q Fu(\()p Fw(h)p Fz(t)1412
+513 y FG(1)1451 498 y Fz(;)17 b(t)1530 513 y FG(2)1569
+498 y Fw(i)p Fu(\))1674 439 y Fg(def)1676 498 y Fu(=)30
+b Fp(F)-10 b(A)p Fu(\()p Fz(t)1981 513 y FG(1)2021 498
+y Fu(\))22 b Fw([)h Fp(F)-10 b(A)p Fu(\()p Fz(t)2369
+513 y FG(2)2408 498 y Fu(\))1333 672 y Fp(F)g(A)p Fu(\()p
+Fz(f)28 b(t)p Fu(\))1674 613 y Fg(def)1676 672 y Fu(=)i
+Fp(F)-10 b(A)p Fu(\()p Fz(t)p Fu(\))1393 846 y Fp(F)g(A)p
+Fu(\()p Fz(a)p Fu(\))1674 787 y Fg(def)1676 846 y Fu(=)30
+b Fw(f)p Fz(a)p Fw(g)1330 1019 y Fp(F)-10 b(A)q Fu(\()p
+Fz(a:t)p Fu(\))1674 960 y Fg(def)1676 1019 y Fu(=)30
+b Fp(F)-10 b(A)p Fu(\()p Fz(t)p Fu(\))23 b Fw(\000)f(f)p
+Fz(a)p Fw(g)p Fz(:)166 1357 y FB(PR)m(OOF)-11 b(.)49
+b FA(The)24 b(proof)h(is)f(similar)g(to)g(the)h(proof)g(of)g([10,)f
+(Proposition)g(2.2].)p 3382 1357 4 68 v 3386 1293 60
+4 v 3386 1357 V 3445 1357 4 68 v 166 1694 a(F)o(or)33
+b(non-ground)f(terms,)h(the)g(relations)g Fu(=)1722 1709
+y Fq(\013)1804 1694 y FA(and)h Fw(\031)f FA(dif)n(fer!)h(F)o(or)f(e)o
+(xample)f Fz(a:X)52 b Fu(=)3204 1709 y Fq(\013)3296 1694
+y Fz(b:X)166 1815 y FA(al)o(w)o(ays)24 b(holds,)f(whereas)i
+Fw(;)j(`)f Fz(a:X)36 b Fw(\031)28 b Fz(b:X)33 b FA(is)24
+b(not)f(pro)o(v)n(able)g(unless)h Fz(a)k Fu(=)f Fz(b)p
+FA(.)e(This)f(disagree-)166 1935 y(ment)h(is)g(to)g(be)g(e)o(xpected,)g
+(since)h(we)f(noted)g(in)g(the)g(Introduction)g(that)f
+Fu(=)2746 1950 y Fq(\013)2821 1935 y FA(is)h Fx(not)i
+FA(preserv)o(ed)166 2056 y(by)e(substitution,)c(whereas)26
+b(from)e(Lemma)h(2.14)f(we)h(kno)n(w)f(that)g Fw(\031)h
+FA(is.)166 2453 y FB(3)99 b(Uni\002cation)166 2794 y
+FA(Gi)n(v)o(en)37 b(terms)h Fz(t)h FA(and)g Fz(t)996
+2758 y Ft(0)1058 2794 y FA(of)g(the)f(same)h(sort)f(o)o(v)o(er)g(a)h
+(nominal)e(signature,)h(can)h(we)g(decide)166 2914 y(whether)21
+b(or)h(not)e(there)i(is)f(a)h(substitution)c(of)j(terms)g(for)h(the)f
+(v)n(ariables)g(in)f Fz(t)i FA(and)f Fz(t)2984 2878 y
+Ft(0)3029 2914 y FA(that)g(mak)o(es)166 3034 y(them)k(equal)g(in)g(the)
+g(sense)h(of)f(the)g(relation)g Fw(\031)h FA(introduced)f(in)g(the)g
+(pre)n(vious)f(section?)h(Since)166 3155 y(instances)30
+b(of)g Fw(\031)h FA(are)h(established)d(modulo)g(freshness)h
+(constraints,)g(it)f(mak)o(es)i(more)f(sense)166 3275
+y(to)c(ask)g(whether)g(or)g(not)g(there)g(is)g(both)f(a)i(substitution)
+c Fz(\033)30 b FA(and)c(a)h(freshness)f(en)l(vironment)f
+Fw(r)166 3396 y FA(for)k(which)f Fw(r)34 b(`)h Fz(\033)t
+Fu(\()p Fz(t)p Fu(\))f Fw(\031)h Fz(\033)t Fu(\()p Fz(t)1244
+3359 y Ft(0)1268 3396 y Fu(\))28 b FA(holds.)g(As)g(for)h(ordinary)f
+(\002rst-order)h(uni\002cation,)e(solving)166 3516 y(such)h(an)f
+(equational)g(problem)g(may)h(thro)n(w)f(up)g Fx(se)o(ver)o(al)g
+FA(equational)g(subproblems;)e(b)n(ut)j(an)166 3636 y(added)21
+b(complication)f(here)i(is)f(that)g(because)h(of)f(rule)h(\()p
+Fw(\031)p FA(-abstraction-2\))g(in)f(Figure)g(2,)g(equa-)166
+3757 y(tional)30 b(problems)g(may)h(generate)g Fx(fr)l(eshness)g
+FA(problems,)f(i.e.)h(ones)f(in)l(v)n(olving)f(the)i(relation)166
+3877 y Fu(#)p FA(.)c(W)-8 b(e)28 b(are)g(thus)e(led)h(to)f(the)h(follo)
+n(wing)e(de\002nition)h(of)i(uni\002cation)e(problems)g(for)h(nominal)
+166 3998 y(equational)d(logic.)166 4218 y FB(De\002nition)k(3.1.)42
+b FA(A)27 b Fx(uni\002cation)f(pr)l(oblem)f Fz(P)41 b
+FA(o)o(v)o(er)26 b(a)i(nominal)d(signature)i(is)f(a)i(\002nite)f(set)g
+(of)166 4338 y(atomic)32 b(problems,)f(each)i(of)g(which)f(is)g(either)
+g(an)h Fx(equational)e(pr)l(oblem)g Fz(t)42 b Fw(\031)p
+Fu(?)g Fz(t)3086 4302 y Ft(0)3143 4338 y FA(where)32
+b Fz(t)166 4458 y FA(and)27 b Fz(t)372 4422 y Ft(0)422
+4458 y FA(are)g(terms)f(of)h(the)f(same)h(sort)f(o)o(v)o(er)f(the)i
+(signature,)f(or)g(a)h Fx(fr)l(eshness)f(pr)l(oblem)f
+Fz(a)31 b Fu(#?)g Fz(t)166 4579 y FA(where)e Fz(a)g FA(is)f(an)g(atom)g
+(and)g Fz(t)h FA(a)g(term)f(o)o(v)o(er)f(the)i(signature.)e(A)i
+Fx(solution)d FA(for)j Fz(P)42 b FA(consists)27 b(of)h(a)166
+4699 y(pair)d Fu(\()p Fw(r)p Fz(;)17 b(\033)t Fu(\))24
+b FA(where)i Fw(r)f FA(is)f(a)h(freshness)g(en)l(vironment)e(and)i
+Fz(\033)k FA(is)24 b(a)h(substitution)d(satisfying)166
+4919 y Fw(\017)50 b(r)27 b(`)h Fz(a)g Fu(#)g Fz(\033)t
+Fu(\()p Fz(t)p Fu(\))52 b FA(for)26 b(each)53 b Fu(\()p
+Fz(a)27 b Fu(#?)h Fz(t)p Fu(\))g Fw(2)g Fz(P)66 b FA(and)166
+5039 y Fw(\017)50 b(r)27 b(`)h Fz(\033)t Fu(\()p Fz(t)p
+Fu(\))g Fw(\031)g Fz(\033)t Fu(\()p Fz(t)900 5003 y Ft(0)923
+5039 y Fu(\))53 b FA(for)25 b(each)53 b Fu(\()p Fz(t)28
+b Fw(\031)p Fu(?)g Fz(t)1677 5003 y Ft(0)1701 5039 y
+Fu(\))f Fw(2)h Fz(P)14 b FA(.)166 5259 y(W)-8 b(e)24
+b(write)f Ff(U)p Fu(\()p Fz(P)14 b Fu(\))22 b FA(for)i(the)f(set)g(of)g
+(all)g(solutions)e(of)i(a)g(problem)g Fz(P)14 b FA(.)22
+b Fu(\()p Fw(r)p Fz(;)17 b(\033)t Fu(\))28 b Fw(2)g Ff(U)p
+Fu(\()p Fz(P)14 b Fu(\))22 b FA(is)h(a)h Fx(most)166
+5380 y(g)o(ener)o(al)i FA(solution)g(for)i Fz(P)40 b
+FA(if)28 b(gi)n(v)o(en)e(an)o(y)g(other)h(solution)f
+Fu(\()p Fw(r)2296 5344 y Ft(0)2319 5380 y Fz(;)17 b(\033)2422
+5344 y Ft(0)2445 5380 y Fu(\))33 b Fw(2)f Ff(U)p Fu(\()p
+Fz(P)14 b Fu(\))p FA(,)27 b(then)g(there)h(is)e(a)1748
+5712 y(14)p eop end
+%%Page: 15 15
+TeXDict begin 15 14 bop 166 83 a FA(substitution)18 b
+Fz(\033)711 47 y Ft(00)775 83 y FA(satisfying)i Fw(r)1267
+47 y Ft(0)1318 83 y Fw(`)28 b Fz(\033)1466 47 y Ft(00)1508
+83 y Fu(\()p Fw(r)p Fu(\))21 b FA(and)h Fw(r)1937 47
+y Ft(0)1988 83 y Fw(`)27 b Fz(\033)2135 47 y Ft(00)2187
+83 y Fw(\016)9 b Fz(\033)31 b Fw(\031)d Fz(\033)2496
+47 y Ft(0)2520 83 y FA(.)21 b(\(Here)h(we)g(are)g(using)e(the)166
+203 y(notation)25 b(of)h(\(11\);)f(and)h Fz(\033)1076
+167 y Ft(00)1142 203 y Fw(\016)d Fz(\033)30 b FA(denotes)25
+b(the)h Fx(substitution)d(composition)h FA(of)i Fz(\033)k
+FA(follo)n(wed)25 b(by)166 339 y Fz(\033)225 303 y Ft(00)267
+339 y FA(,)j(gi)n(v)o(en)e(by)h Fu(\()p Fz(\033)789 303
+y Ft(00)855 339 y Fw(\016)d Fz(\033)t Fu(\)\()p Fz(X)8
+b Fu(\))1223 287 y Fo(def)1227 339 y Fu(=)37 b Fz(\033)1399
+303 y Ft(00)1441 339 y Fu(\()p Fz(\033)t Fu(\()p Fz(X)8
+b Fu(\)\))p FA(.\))27 b(A)g(solution)f Fu(\()p Fw(r)p
+Fz(;)17 b(\033)t Fu(\))32 b Fw(2)g Ff(U)p Fu(\()p Fz(P)14
+b Fu(\))27 b FA(is)g Fx(idempotent)166 460 y FA(pro)o(vided)d
+Fw(r)j(`)h Fz(\033)e Fw(\016)c Fz(\033)32 b Fw(\031)c
+Fz(\033)t FA(.)166 683 y(W)-8 b(e)31 b(describe)g(an)g(algorithm)e
+(which,)h(gi)n(v)o(en)f(an)o(y)h(nominal)f(uni\002cation)h(problem,)g
+(decides)166 803 y(whether)i(or)h(not)e(it)h(has)g(a)h(solution)d(and)j
+(if)f(it)g(does,)g(returns)g(a)g(most)f(general)i(\(and)f(idem-)166
+923 y(potent\))23 b(solution.)e(The)j(algorithm)e(uses)h(labelled)g
+(transformations,)f(directly)h(generalising)166 1044
+y(the)31 b(presentation)f(of)i(\002rst-order)f(uni\002cation)g(in)g
+([19,)g(Sect.)h(2.6])f(which)f(in)h(turn)g(is)g(based)166
+1164 y(upon)g(the)g(approach)h(in)g([18].)f(\(See)i(also)e([20,)h
+(Sect.)g(4.6])f(for)h(a)g(detailed)f(e)o(xposition,)e(b)n(ut)166
+1284 y(not)d(using)f(labels.\))h(W)-8 b(e)26 b(use)g(tw)o(o)g(types)f
+(of)i(labelled)e(transformation)g(between)i(uni\002cation)166
+1405 y(problems,)d(namely)1248 1547 y Fz(P)1410 1494
+y Fq(\033)1352 1547 y Fu(=)-17 b Fw(\))28 b Fz(P)1616
+1506 y Ft(0)1738 1547 y FA(and)99 b Fz(P)2136 1494 y
+Ft(r)2086 1547 y Fu(=)-17 b Fw(\))27 b Fz(P)2349 1506
+y Ft(0)166 1702 y FA(where)i(the)g(substitution)d Fz(\033)33
+b FA(is)c(either)g(the)f(identity)g Fz(")p FA(,)g(or)h(a)g(single)f
+(replacement)h Fu([)p Fz(X)43 b Fu(:=)36 b Fz(t)p Fu(])p
+FA(;)166 1822 y(and)24 b(where)h(the)f(freshness)g(en)l(vironment)e
+Fw(r)j FA(is)e(either)h(empty)f Fw(;)p FA(,)h(or)h(a)f(singleton)f
+Fw(f)p Fz(a)k Fu(#)h Fz(X)8 b Fw(g)p FA(.)166 1943 y(The)21
+b(le)o(gal)f(transformations)g(are)h(gi)n(v)o(en)f(in)h(Figure)g(3.)g
+(This)f(\002gure)i(uses)f(the)g(notation)e Fz(P)j Fw(])8
+b Fz(P)3430 1906 y Ft(0)166 2063 y FA(to)26 b(indicate)f(the)h(union)g
+(of)g(problems)f Fz(P)40 b FA(and)26 b Fz(P)1869 2027
+y Ft(0)1918 2063 y FA(that)g(are)g(disjoint)f(\()p Fz(P)36
+b Fw(\\)24 b Fz(P)2866 2027 y Ft(0)2919 2063 y Fu(=)30
+b Fw(;)p FA(\);)c(and)g(the)166 2183 y(notation)j Fz(\033)t(P)44
+b FA(to)30 b(indicate)g(the)g(problem)f(resulting)g(from)h(applying)f
+(the)h(substitution)e Fz(\033)34 b FA(to)166 2304 y(all)25
+b(the)f(terms)g(occurring)h(in)g(the)f(problem)g Fz(P)14
+b FA(.)166 2526 y FB(Algorithm.)33 b FA(Gi)n(v)o(en)18
+b(a)h(uni\002cation)f(problem)g Fz(P)c FA(,)19 b(the)g(algorithm)e
+(proceeds)i(in)g(tw)o(o)g(phases.)3398 2490 y FG(2)166
+2647 y FA(In)j(the)g(\002rst)g(phase)g(it)f(applies)g(as)h(man)o(y)1627
+2594 y Fq(\033)1569 2647 y Fu(=)-17 b Fw(\))22 b FA(transformations)e
+(as)i(possible)f(\(non-determin-)166 2767 y(istically\).)30
+b(If)i(this)f(results)g(in)g(a)h(problem)f(containing)f(no)i
+(equational)e(subproblems,)g(then)166 2888 y(it)d(proceeds)h(to)f(the)g
+(second)g(phase;)g(otherwise)g(it)g(halts)g(signalling)f(f)o(ailure.)h
+(In)h(the)f(second)166 3024 y(phase)22 b(it)f(applies)g(as)g(man)o(y)
+1192 2971 y Ft(r)1142 3024 y Fu(=)-17 b Fw(\))21 b FA(transformations)f
+(as)i(possible)e(\(non-deterministically\).)f(If)166
+3145 y(this)24 b(does)h(not)f(result)g(in)h(the)g(empty)f(problem,)g
+(then)g(it)g(halts)h(signalling)e(f)o(ailure;)h(otherwise)166
+3265 y(o)o(v)o(erall)g(it)g(has)h(constructed)f(a)h(transformation)e
+(sequence)i(of)g(the)g(form)1152 3509 y Fz(P)1298 3454
+y Fq(\033)1338 3463 y Fe(1)1256 3509 y Fu(=)-17 b Fw(\))28
+b(\001)17 b(\001)g(\001)1625 3455 y Fq(\033)1665 3463
+y Fd(n)1587 3509 y Fu(=)-17 b Fw(\))27 b Fz(P)1850 3468
+y Ft(0)1933 3454 y(r)1992 3463 y Fe(1)1901 3509 y Fu(=)-17
+b Fw(\))27 b(\001)17 b(\001)g(\001)2252 3455 y Ft(r)2311
+3463 y Fd(m)2231 3509 y Fu(=)-17 b Fw(\))28 b(;)820 b
+FA(\(12\))166 3734 y(\(where)26 b Fz(P)545 3698 y Ft(0)593
+3734 y FA(does)f(not)g(contain)f(an)o(y)h(equational)f(subproblems\))g
+(and)h(the)g(algorithm)f(returns)166 3854 y(the)h(solution)e
+Fu(\()p Fw(r)780 3869 y FG(1)841 3854 y Fw([)g(\001)17
+b(\001)g(\001)j([)j(r)1240 3869 y Fq(m)1306 3854 y Fz(;)34
+b(\033)1422 3869 y Fq(n)1491 3854 y Fw(\016)22 b(\001)17
+b(\001)g(\001)j(\016)i Fz(\033)1828 3869 y FG(1)1868
+3854 y Fu(\))p FA(.)166 4077 y(T)-8 b(o)31 b(sho)n(w)g(the)g
+(correctness)h(of)g(this)f(algorithm,)f(we)i(\002rst)f(establish)g
+(that)g(all)g(sequences)h(of)166 4198 y(uni\002cation)24
+b(transitions)f(must)h(terminate.)166 4420 y FB(Lemma)h(3.2.)41
+b Fx(Ther)l(e)26 b(is)e(no)h(in\002nite)f(series)g(of)g(uni\002cation)f
+(tr)o(ansitions.)166 4783 y FB(PR)m(OOF)-11 b(.)49 b
+FA(Since)28 b(e)n(v)o(ery)f(reduction)f(sequence)i(consists)e(of)i(tw)o
+(o)f(\(possibly)f(empty\))h(subse-)166 4920 y(quences,)j(namely)g(one)h
+(containing)e(only)1757 4867 y Fq(\033)1699 4920 y Fu(=)-17
+b Fw(\))p FA(-steps)30 b(and)h(the)f(other)g(only)2941
+4867 y Ft(r)2891 4920 y Fu(=)-17 b Fw(\))p FA(-steps,)30
+b(we)166 5040 y(can)23 b(sho)n(w)f(termination)f(for)i(both)f
+(subsequences)g(separately)-6 b(.)22 b(F)o(or)h(e)n(v)o(ery)f
+(uni\002cation)g(prob-)166 5161 y(lem)g Fz(P)36 b FA(we)23
+b(de\002ne)h(a)f(measure)f(of)h(the)g(size)f(of)h Fz(P)36
+b FA(to)22 b(be)h(the)g(le)o(xicographically)e(ordered)i(pair)p
+166 5282 299 4 v 166 5346 a FG(2)257 5379 y FD(See)g(Remark)g(3.9)h
+(for)g(discussion)i(of)e(this)g(use)g(of)f(tw)o(o)g(phases.)1748
+5712 y FA(15)p eop end
+%%Page: 16 16
+TeXDict begin 16 15 bop 166 3 3288 4 v 166 3123 4 3121
+v 204 154 a FD(\()p Fl(\031)p Fk(?)p FD(-unit\))798 b
+Fw(fhi)27 b(\031)p Fu(?)h Fw(hig)21 b(])i Fz(P)2063 101
+y Fq(")1999 154 y Fu(=)-17 b Fw(\))34 b Fz(P)204 335
+y FD(\()p Fl(\031)p Fk(?)p FD(-pair\))412 b Fw(fh)p Fz(t)1084
+350 y FG(1)1123 335 y Fz(;)17 b(t)1202 350 y FG(2)1241
+335 y Fw(i)28 b(\031)p Fu(?)g Fw(h)p Fz(t)1533 298 y
+Ft(0)1533 359 y FG(1)1572 335 y Fz(;)17 b(t)1651 298
+y Ft(0)1651 359 y FG(2)1691 335 y Fw(ig)k(])i Fz(P)2063
+281 y Fq(")1999 335 y Fu(=)-17 b Fw(\))34 b(f)p Fz(t)2277
+350 y FG(1)2344 335 y Fw(\031)p Fu(?)28 b Fz(t)2530 298
+y Ft(0)2530 359 y FG(1)2570 335 y Fz(;)17 b(t)2649 350
+y FG(2)2716 335 y Fw(\031)p Fu(?)28 b Fz(t)2902 298 y
+Ft(0)2902 359 y FG(2)2942 335 y Fw(g)21 b([)i Fz(P)204
+515 y FD(\()p Fl(\031)p Fk(?)p FD(-function)j(symbol\))256
+b Fw(f)p Fz(f)27 b(t)h Fw(\031)q Fu(?)f Fz(f)h(t)1706
+479 y Ft(0)1729 515 y Fw(g)22 b(])h Fz(P)2063 462 y Fq(")1999
+515 y Fu(=)-17 b Fw(\))34 b(f)p Fz(t)27 b Fw(\031)q Fu(?)g
+Fz(t)2490 479 y Ft(0)2514 515 y Fw(g)22 b([)h Fz(P)204
+696 y FD(\()p Fl(\031)p Fk(?)p FD(-abstraction-1\))372
+b Fw(f)p Fz(a:t)28 b Fw(\031)q Fu(?)f Fz(a:t)1705 660
+y Ft(0)1729 696 y Fw(g)22 b(])h Fz(P)2063 643 y Fq(")1999
+696 y Fu(=)-17 b Fw(\))34 b(f)p Fz(t)27 b Fw(\031)q Fu(?)g
+Fz(t)2490 660 y Ft(0)2514 696 y Fw(g)22 b([)h Fz(P)204
+876 y FD(\()p Fl(\031)p Fk(?)p FD(-abstraction-2\))349
+b Fw(f)p Fz(a:t)28 b Fw(\031)p Fu(?)g Fz(a)1620 840 y
+Ft(0)1644 876 y Fz(:t)1706 840 y Ft(0)1729 876 y Fw(g)22
+b(])h Fz(P)2063 823 y Fq(")1999 876 y Fu(=)-17 b Fw(\))34
+b(f)p Fz(t)27 b Fw(\031)q Fu(?)g(\()p Fz(a)17 b(a)2612
+840 y Ft(0)2636 876 y Fu(\))p Fr(\001)p Fz(t)2741 840
+y Ft(0)2764 876 y Fz(;)g(a)28 b Fu(#?)g Fz(t)3077 840
+y Ft(0)3100 876 y Fw(g)22 b([)h Fz(P)2781 1010 y FA(pro)o(vided)h
+Fz(a)k Fw(6)p Fu(=)f Fz(a)3392 973 y Ft(0)204 1143 y
+FD(\()p Fl(\031)p Fk(?)p FD(-atom\))809 b Fw(f)p Fz(a)28
+b Fw(\031)p Fu(?)g Fz(a)p Fw(g)22 b(])h Fz(P)2063 1090
+y Fq(")1999 1143 y Fu(=)-17 b Fw(\))34 b Fz(P)204 1324
+y FD(\()p Fl(\031)p Fk(?)p FD(-suspension\))315 b Fw(f)p
+Fz(\031)t Fr(\001)p Fz(X)35 b Fw(\031)q Fu(?)27 b Fz(\031)1585
+1287 y Ft(0)1609 1324 y Fr(\001)p Fz(X)8 b Fw(g)21 b(])i
+Fz(P)2063 1270 y Fq(")1999 1324 y Fu(=)-17 b Fw(\))34
+b(f)p Fz(a)27 b Fu(#?)h Fz(X)36 b Fw(j)27 b Fz(a)h Fw(2)g
+Fp(ds)8 b Fu(\()p Fz(\031)t(;)17 b(\031)3118 1287 y Ft(0)3141
+1324 y Fu(\))p Fw(g)22 b([)g Fz(P)204 1583 y FD(\()p
+Fl(\031)p Fk(?)p FD(-v)n(ariable\))1259 1504 y Fw(f)p
+Fz(t)27 b Fw(\031)q Fu(?)g Fz(\031)t Fr(\001)p Fz(X)8
+b Fw(g)22 b(])h Fz(P)1259 1685 y Fw(f)p Fz(\031)t Fr(\001)o
+Fz(X)36 b Fw(\031)p Fu(?)28 b Fz(t)p Fw(g)22 b(])h Fz(P)1939
+1384 y Fn(9)1939 1459 y(>)1939 1484 y(=)1939 1633 y(>)1939
+1658 y(;)2058 1530 y Fq(\033)1999 1583 y Fu(=)-17 b Fw(\))34
+b Fz(\033)t(P)160 b FA(with)52 b Fz(\033)31 b Fu(=)d([)p
+Fz(X)36 b Fu(:=)27 b Fz(\031)3227 1547 y Ft(\000)p FG(1)3321
+1583 y Fr(\001)p Fz(t)p Fu(])2181 1724 y FA(pro)o(vided)c
+Fz(X)33 b FA(does)25 b(not)f(occur)h(in)f Fz(t)204 1928
+y FD(\()p Fk(#?)p FD(-unit\))815 b Fw(f)p Fz(a)28 b Fu(#?)g
+Fw(hig)21 b(])i Fz(P)2061 1875 y Ft(;)1999 1928 y Fu(=)-17
+b Fw(\))34 b Fz(P)204 2108 y FD(\()p Fk(#?)p FD(-pair\))622
+b Fw(f)p Fz(a)28 b Fu(#?)g Fw(h)p Fz(t)1533 2123 y FG(1)1572
+2108 y Fz(;)17 b(t)1651 2123 y FG(2)1691 2108 y Fw(ig)k(])i
+Fz(P)2061 2055 y Ft(;)1999 2108 y Fu(=)-17 b Fw(\))34
+b(f)p Fz(a)27 b Fu(#?)h Fz(t)2510 2123 y FG(1)2550 2108
+y Fz(;)17 b(a)28 b Fu(#?)g Fz(t)2863 2123 y FG(2)2902
+2108 y Fw(g)22 b([)h Fz(P)204 2289 y FD(\()p Fk(#?)p
+FD(-function)j(symbol\))330 b Fw(f)p Fz(a)28 b Fu(#?)g
+Fz(f)f(t)p Fw(g)22 b(])h Fz(P)2061 2236 y Ft(;)1999 2289
+y Fu(=)-17 b Fw(\))34 b(f)p Fz(a)27 b Fu(#?)h Fz(t)p
+Fw(g)22 b([)h Fz(P)204 2469 y FD(\()p Fk(#?)p FD(-abstraction-1\))449
+b Fw(f)p Fz(a)28 b Fu(#?)g Fz(a:t)p Fw(g)22 b(])h Fz(P)2061
+2416 y Ft(;)1999 2469 y Fu(=)-17 b Fw(\))34 b Fz(P)204
+2650 y FD(\()p Fk(#?)p FD(-abstraction-2\))426 b Fw(f)p
+Fz(a)27 b Fu(#?)h Fz(a)1643 2614 y Ft(0)1667 2650 y Fz(:t)p
+Fw(g)22 b(])h Fz(P)2061 2597 y Ft(;)1999 2650 y Fu(=)-17
+b Fw(\))34 b(f)p Fz(a)27 b Fu(#?)h Fz(t)p Fw(g)22 b([)h
+Fz(P)47 b FA(pro)o(vided)24 b Fz(a)k Fw(6)p Fu(=)f Fz(a)3392
+2614 y Ft(0)204 2831 y FD(\()p Fk(#?)p FD(-atom\))777
+b Fw(f)p Fz(a)28 b Fu(#?)g Fz(a)1706 2794 y Ft(0)1729
+2831 y Fw(g)22 b(])h Fz(P)2061 2777 y Ft(;)1999 2831
+y Fu(=)-17 b Fw(\))34 b Fz(P)526 b FA(pro)o(vided)24
+b Fz(a)k Fw(6)p Fu(=)f Fz(a)3392 2794 y Ft(0)204 3011
+y FD(\()p Fk(#?)p FD(-suspension\))458 b Fw(f)p Fz(a)28
+b Fu(#?)g Fz(\031)t Fr(\001)p Fz(X)8 b Fw(g)21 b(])i
+Fz(P)2049 2958 y Ft(r)1999 3011 y Fu(=)-17 b Fw(\))34
+b Fz(P)113 b FA(with)24 b Fw(r)j Fu(=)h Fw(f)p Fz(\031)2893
+2975 y Ft(\000)p FG(1)2987 3011 y Fr(\001)p Fz(a)g Fu(#)g
+Fz(X)8 b Fw(g)p 3450 3123 V 166 3126 3288 4 v 1214 3263
+a FD(Fig.)23 b(3.)g(Labelled)h(transformations.)166 3422
+y FA(of)29 b(natural)f(numbers)f Fu(\()p Fz(n)1050 3437
+y FG(1)1090 3422 y Fz(;)17 b(n)1192 3437 y FG(2)1231
+3422 y Fu(\))p FA(,)29 b(where)g Fz(n)1653 3437 y FG(1)1721
+3422 y FA(is)f(the)g(number)g(of)g(dif)n(ferent)g(v)n(ariables)g(used)g
+(in)166 3542 y Fz(P)14 b FA(,)24 b(and)h Fz(n)519 3557
+y FG(2)584 3542 y FA(is)f(the)h(size)f(\(see)i(De\002nition)e(2.10\))g
+(of)h(all)g(equational)f(problems)f(in)i Fz(P)14 b FA(,)24
+b(that)g(is)1334 3767 y Fz(n)1392 3782 y FG(2)1459 3714
+y Fo(def)1464 3767 y Fu(=)1671 3684 y Fn(X)1572 3872
+y FG(\()p Fq(t)p Ft(\031)q FG(?)p Fq(t)1738 3853 y Fc(0)1761
+3872 y FG(\))p Ft(2)p Fq(P)1906 3767 y Fw(j)p Fz(t)p
+Fw(j)e Fu(+)g Fw(j)p Fz(t)2180 3726 y Ft(0)2203 3767
+y Fw(j)28 b Fz(:)166 4084 y FA(In)34 b(e)n(v)o(ery)592
+4031 y Fq(\033)533 4084 y Fu(=)-17 b Fw(\))p FA(-step)34
+b(this)e(measure)i(decreases:)g(the)f(\()p Fw(\031)p
+FA(?-v)n(ariable\))h(transition)e(eliminates)166 4205
+y(\(completely\))20 b(one)g(v)n(ariable)h(from)f(the)g(uni\002cation)g
+(problem,)g(and)h(therefore)g Fz(n)2984 4220 y FG(1)3044
+4205 y FA(decreases;)166 4325 y(the)j(\()p Fw(\031)p
+FA(?-suspension\))f(transition)g(may)g(eliminate)g(a)h(v)n(ariable)f
+(and)h(also)g(decreases)g(the)g(size)166 4445 y Fz(n)224
+4460 y FG(2)264 4445 y FA(;)i(all)h(other)f(transitions)f(lea)n(v)o(e)i
+(the)g(number)f(of)h(v)n(ariables)f(unchanged,)g(b)n(ut)g(decrease)i
+Fz(n)3389 4460 y FG(2)3429 4445 y FA(.)166 4582 y(F)o(or)d(the)525
+4529 y Ft(r)474 4582 y Fu(=)-17 b Fw(\))p FA(-steps)25
+b(the)f(size)1696 4619 y Fn(X)1601 4808 y FG(\()p Fq(a)p
+FG(#?)q Fq(t)p FG(\))p Ft(2)p Fq(P)1929 4702 y Fw(j)p
+Fz(t)p Fw(j)166 4927 y FA(decreases)30 b(in)e(e)n(v)o(ery)g(step.)h(T)
+-8 b(aking)28 b(both)g(f)o(acts)h(together)f(means)h(that)f(e)n(v)o
+(ery)g(reduction)h(se-)166 5047 y(quence)c(must)f(terminate.)p
+3382 5047 4 68 v 3386 4983 60 4 v 3386 5047 V 3445 5047
+4 68 v 166 5380 a(The)e(follo)n(wing)f(lemmas)g(help)h(us)g(to)g(sho)n
+(w)f(that)h(the)g(algorithm)f(gi)n(v)o(es)g(correct)i(results)e(upon)
+1748 5712 y(16)p eop end
+%%Page: 17 17
+TeXDict begin 17 16 bop 166 83 a FA(termination.)166
+306 y FB(Lemma)25 b(3.3.)41 b Fx(If)25 b Fw(r)j(`)f Fz(\033)t
+Fu(\()p Fz(\031)t Fr(\001)p Fz(X)8 b Fu(\))27 b Fw(\031)i
+Fz(\033)t Fu(\()p Fz(t)p Fu(\))c Fx(then)f Fw(r)k(`)f
+Fz(\033)g Fw(\016)22 b Fu([)p Fz(X)35 b Fu(:=)28 b Fz(\031)2518
+270 y Ft(\000)p FG(1)2612 306 y Fr(\001)p Fz(t)p Fu(])g
+Fw(\031)g Fz(\033)t Fx(.)166 669 y FB(PR)m(OOF)-11 b(.)49
+b FA(W)-8 b(e)32 b(ha)n(v)o(e)f(to)h(pro)o(v)o(e)e(that)i(both)f
+(substitutions)d(agree)33 b(\(modulo)e Fw(\031)p FA(\))h(on)g(all)f(v)n
+(ari-)166 790 y(ables)f(in)f Fz(dom)p Fu(\()p Fz(\033)t
+Fu(\))d Fw([)g(f)p Fz(X)8 b Fw(g)p FA(.)30 b(The)g(only)f(interesting)g
+(case)h(is)g(for)g(the)g(substitutions)c(applied)166
+910 y(to)35 b Fz(X)8 b FA(,)34 b(when)h(we)g(need)g(to)g(sho)n(w)f
+(that)g Fw(r)46 b(`)h Fz(\033)t Fu(\()p Fz(\031)1985
+874 y Ft(\000)p FG(1)2079 910 y Fr(\001)p Fz(t)p Fu(\))f
+Fw(\031)h Fz(\033)t Fu(\()p Fz(X)8 b Fu(\))p FA(.)35
+b(By)g(Lemma)f(2.14)h(we)166 1030 y(can)i(commute)e(the)h(permutation)f
+(to)h(the)g(outside)f(and)h(mo)o(v)o(e)f(it)h(to)g(the)g(other)g(side)g
+(of)g Fw(\031)166 1151 y FA(by)h(Lemma)g(2.12\227this)e(gi)n(v)o(es)h
+Fw(r)51 b(`)g Fz(\033)t Fu(\()p Fz(t)p Fu(\))g Fw(\031)g
+Fz(\031)t Fr(\001)p Fz(\033)t Fu(\()p Fz(X)8 b Fu(\))p
+FA(.)37 b(The)g(case)h(then)f(follo)n(ws)f(from)166 1271
+y(the)27 b(assumptions)e(by)h(symmetry)g(and)h(commuting)e(the)i
+(permutation)f(inside)g(the)h(substitu-)166 1392 y(tion.)p
+3382 1392 4 68 v 3386 1328 60 4 v 3386 1392 V 3445 1392
+4 68 v 166 1653 a FB(Lemma)g(3.4.)41 b Fx(Given)26 b(a)g(uni\002cation)
+e(pr)l(oblem)h Fz(P)14 b Fx(,)25 b Fu(\()p Fw(r)p Fz(;)17
+b(\033)t Fu(\))29 b Fw(2)h Ff(U)p Fu(\()p Fz(\033)2538
+1616 y Ft(0)2562 1653 y Fz(P)14 b Fu(\))25 b Fx(holds)g(if)g(and)h
+(only)f(if)166 1773 y Fu(\()p Fw(r)p Fz(;)17 b(\033)26
+b Fw(\016)c Fz(\033)543 1737 y Ft(0)566 1773 y Fu(\))28
+b Fw(2)g Ff(U)p Fu(\()p Fz(P)14 b Fu(\))p Fx(.)166 2136
+y FB(PR)m(OOF)-11 b(.)49 b FA(Simple)24 b(calculation)g(using)g(the)g
+(f)o(act)i(that)e Fz(\033)t Fu(\()p Fz(\033)2236 2100
+y Ft(0)2259 2136 y Fu(\()p Fz(t)p Fu(\)\))k(=)f(\()p
+Fz(\033)f Fw(\016)c Fz(\033)2789 2100 y Ft(0)2813 2136
+y Fu(\)\()p Fz(t)p Fu(\))p FA(.)p 3382 2136 V 3386 2073
+60 4 v 3386 2136 V 3445 2136 4 68 v 166 2500 a(The)31
+b(follo)n(wing)e(tw)o(o)i(lemmas)f(sho)n(w)g(that)h(the)g
+(uni\002cation)g(transformations)e(can)j(be)f(used)166
+2620 y(to)d(determine)g(whether)h(or)f(not)g(solutions)f(e)o(xists)g
+(and)h(to)h(describe)f(all)g(of)h(them)f(if)g(the)o(y)g(do)166
+2741 y(e)o(xist.)166 2964 y FB(Lemma)d(3.5.)228 3186
+y Fx(\(i\))50 b(If)25 b Fu(\()p Fw(r)579 3150 y Ft(0)602
+3186 y Fz(;)17 b(\033)705 3150 y Ft(0)728 3186 y Fu(\))37
+b Fw(2)g Ff(U)p Fu(\()p Fz(P)14 b Fu(\))24 b Fx(and)h
+Fz(P)1500 3133 y Fq(\033)1442 3186 y Fu(=)-17 b Fw(\))37
+b Fz(P)1715 3150 y Ft(0)1738 3186 y Fz(;)17 b Fx(then)24
+b Fu(\()p Fw(r)2099 3150 y Ft(0)2122 3186 y Fz(;)17 b(\033)2225
+3150 y Ft(0)2248 3186 y Fu(\))37 b Fw(2)g Ff(U)p Fu(\()p
+Fz(P)2612 3150 y Ft(0)2635 3186 y Fu(\))25 b Fx(and)f
+Fw(r)2955 3150 y Ft(0)3015 3186 y Fw(`)37 b Fz(\033)3172
+3150 y Ft(0)3205 3186 y Fw(\016)25 b Fz(\033)41 b Fw(\031)372
+3307 y Fz(\033)431 3271 y Ft(0)454 3307 y Fz(:)200 3427
+y Fx(\(ii\))50 b(If)25 b Fu(\()p Fw(r)579 3391 y Ft(0)602
+3427 y Fz(;)17 b(\033)705 3391 y Ft(0)728 3427 y Fu(\))28
+b Fw(2)g Ff(U)p Fu(\()p Fz(P)14 b Fu(\))24 b Fx(and)g
+Fz(P)1465 3374 y Ft(r)1415 3427 y Fu(=)-17 b Fw(\))27
+b Fz(P)1678 3391 y Ft(0)1701 3427 y Fz(;)17 b Fx(then)24
+b Fu(\()p Fw(r)2062 3391 y Ft(0)2085 3427 y Fz(;)17 b(\033)2188
+3391 y Ft(0)2211 3427 y Fu(\))28 b Fw(2)g Ff(U)p Fu(\()p
+Fz(P)2557 3391 y Ft(0)2580 3427 y Fu(\))d Fx(and)f Fw(r)2900
+3391 y Ft(0)2951 3427 y Fw(`)j Fz(\033)3098 3391 y Ft(0)3122
+3427 y Fu(\()p Fw(r)p Fu(\))p Fz(:)166 3791 y FB(PR)m(OOF)-11
+b(.)49 b FA(W)-8 b(e)27 b(just)e(gi)n(v)o(e)h(the)g(details)g(for)h(tw)
+o(o)f(uni\002cation)g(transitions:)f(the)i(case)g(for)g(\()p
+Fw(\031)p Fu(?)p FA(-)166 3911 y(suspension\))h(follo)n(ws)f(from)i
+(Lemma)g(2.12\(iii\);)f(and)h(the)g(\()p Fw(\031)p Fu(?)p
+FA(-v)n(ariable\))h(case)f(is)g(a)h(conse-)166 4031 y(quence)25
+b(of)g(Lemmas)f(3.3)g(and)h(3.4.)p 3382 4031 V 3386 3967
+60 4 v 3386 4031 V 3445 4031 4 68 v 166 4292 a FB(Lemma)g(3.6.)228
+4515 y Fx(\(i\))50 b(If)25 b Fu(\()p Fw(r)579 4479 y
+Ft(0)602 4515 y Fz(;)17 b(\033)705 4479 y Ft(0)728 4515
+y Fu(\))28 b Fw(2)g Ff(U)p Fu(\()p Fz(P)1074 4479 y Ft(0)1097
+4515 y Fu(\))c Fx(and)h Fz(P)1496 4462 y Fq(\033)1438
+4515 y Fu(=)-17 b Fw(\))27 b Fz(P)1701 4479 y Ft(0)1724
+4515 y Fz(;)17 b Fx(then)24 b Fu(\()p Fw(r)2085 4479
+y Ft(0)2109 4515 y Fz(;)17 b(\033)2212 4479 y Ft(0)2257
+4515 y Fw(\016)22 b Fz(\033)t Fu(\))27 b Fw(2)h Ff(U)p
+Fu(\()p Fz(P)14 b Fu(\))p Fz(:)200 4655 y Fx(\(ii\))50
+b(If)25 b Fu(\()p Fw(r)579 4619 y Ft(0)602 4655 y Fz(;)17
+b(\033)705 4619 y Ft(0)728 4655 y Fu(\))28 b Fw(2)g Ff(U)p
+Fu(\()p Fz(P)1074 4619 y Ft(0)1097 4655 y Fu(\))p Fz(;)44
+b(P)1360 4602 y Ft(r)1310 4655 y Fu(=)-17 b Fw(\))27
+b Fz(P)1573 4619 y Ft(0)1621 4655 y Fx(and)d Fw(r)1878
+4619 y Ft(00)1949 4655 y Fw(`)j Fz(\033)2096 4619 y Ft(0)2119
+4655 y Fu(\()p Fw(r)p Fu(\))p Fz(;)17 b Fx(then)25 b
+Fu(\()p Fw(r)2640 4619 y Ft(0)2685 4655 y Fw([)d(r)2856
+4619 y Ft(00)2899 4655 y Fz(;)17 b(\033)3002 4619 y Ft(0)3025
+4655 y Fu(\))28 b Fw(2)g Ff(U)p Fu(\()p Fz(P)14 b Fu(\))p
+Fz(:)166 5019 y FB(PR)m(OOF)-11 b(.)49 b FA(Once)33 b(again,)f(we)h
+(just)f(gi)n(v)o(e)g(the)g(details)g(for)i(tw)o(o)e(uni\002cation)g
+(transitions:)f(the)166 5139 y(\()p Fw(\031)p Fu(?)p
+FA(-suspension\))23 b(case)h(follo)n(ws)e(from)h(Lemma)f(2.8;)h(and)g
+(the)h(\()p Fw(\031)p Fu(?)p FA(-v)n(ariable\))f(case)h(follo)n(ws)166
+5259 y(from)34 b(Lemma)f(3.4)h(and)g(the)g(f)o(act)h(that)e
+Fz(t)p Fu([)p Fz(X)53 b Fu(:=)45 b Fz(\031)1994 5223
+y Ft(\000)p FG(1)2088 5259 y Fr(\001)p Fz(t)p Fu(])g(=)g
+Fz(t)p FA(,)35 b(which)e(holds)g(by)h(the)g(side-)166
+5380 y(condition)23 b(on)i(the)g(\()p Fw(\031)p Fu(?)p
+FA(-v)n(ariable\))g(transition)e(about)h(the)h(non-occurrence)h(of)e
+Fz(X)33 b FA(in)25 b Fz(t)p FA(.)p 3382 5380 V 3386 5316
+60 4 v 3386 5380 V 3445 5380 4 68 v 1748 5712 a(17)p
+eop end
+%%Page: 18 18
+TeXDict begin 18 17 bop 166 83 a FA(The)26 b(follo)n(wing)d(theorem)j
+(establishes)e(the)i(correctness)f(of)h(the)g(nominal)e(uni\002cation)h
+(algo-)166 203 y(rithm)f(and)h(is)f(the)h(central)g(result)f(of)h(the)f
+(paper)-5 b(.)166 432 y FB(Theor)n(em)27 b(3.7)d(\(Corr)n(ectness\).)43
+b Fx(Given)25 b(a)f(uni\002cation)g(pr)l(oblem)f Fz(P)228
+660 y Fx(\(i\))50 b(if)24 b(the)h(algorithm)e(fails)g(on)i
+Fz(P)14 b Fx(,)24 b(then)h Fz(P)38 b Fx(has)24 b(no)h(solution;)e(and)
+200 780 y(\(ii\))50 b(if)31 b(the)h(algorithm)e(succeeds)i(on)f
+Fz(P)14 b Fx(,)32 b(then)f(the)h(r)l(esult)f(it)g(pr)l(oduces)g(is)g
+(an)h(idempotent)372 900 y(most)24 b(g)o(ener)o(al)g(solution.)166
+1308 y FB(PR)m(OOF)-11 b(.)49 b FA(When)37 b(f)o(ailure)g(happens)g(it)
+f(is)h(because)h(of)f(certain)g(subproblems)f(that)g(mani-)166
+1428 y(festly)41 b(ha)n(v)o(e)g(no)g(solution)f(\(namely)h(in)g(the)g
+(\002rst)h(phase,)f Fz(a)58 b Fw(\031)q Fu(?)g Fz(a)2641
+1392 y Ft(0)2706 1428 y FA(with)41 b Fz(a)59 b Fw(6)p
+Fu(=)f Fz(a)3220 1392 y Ft(0)3243 1428 y FA(,)42 b(and)166
+1548 y Fz(\031)t Fr(\001)p Fz(X)i Fw(\031)p Fu(?)37 b
+Fz(f)27 b(t)j FA(or)g Fz(f)d(t)37 b Fw(\031)p Fu(?)g
+Fz(\031)t Fr(\001)o Fz(X)h FA(with)28 b Fz(X)38 b FA(occurring)29
+b(in)g Fz(t)p FA(;)h(in)f(the)g(second)h(phase,)f Fz(a)37
+b Fu(#?)f Fz(a)p FA(\).)166 1669 y(Therefore)30 b(part)e(\(i\))h(is)g
+(a)g(consequence)g(of)f(Lemma)h(3.5.)f(F)o(or)h(part)f(\(ii\))h(one)g
+(gets)f(that)g(a)i(se-)166 1789 y(quence)22 b(lik)o(e)g(\(12\))g(e)o
+(xists,)f(and)h(thus)g Fu(\()p Fw(r)p Fz(;)17 b(\033)t
+Fu(\))27 b(=)g(\()p Fw(r)1970 1804 y FG(1)2022 1789 y
+Fw([)12 b(\001)17 b(\001)g(\001)12 b([)g(r)2391 1804
+y Fq(m)2458 1789 y Fz(;)34 b(\033)2574 1804 y Fq(n)2633
+1789 y Fw(\016)12 b(\001)17 b(\001)g(\001)11 b(\016)h
+Fz(\033)2941 1804 y FG(1)2982 1789 y Fu(\))22 b FA(is)f(in)h
+Ff(U)p Fu(\()p Fz(P)14 b Fu(\))166 1910 y FA(by)19 b(Lemma)g(3.6)g(and)
+h(the)f(f)o(act)h(that)f Fu(\()p Fw(;)p Fz(;)e(")p Fu(\))26
+b Fw(2)i Ff(U)p Fu(\()p Fw(;)p Fu(\))p FA(.)20 b(Furthermore)f(from)g
+(Lemma)g(3.5,)g(we)h(get)166 2030 y(that)k(an)o(y)g(other)h(solution)d
+Fu(\()p Fw(r)1203 1994 y Ft(0)1227 2030 y Fz(;)17 b(\033)1330
+1994 y Ft(0)1353 2030 y Fu(\))27 b Fw(2)h Ff(U)p Fu(\()p
+Fz(P)14 b Fu(\))24 b FA(satis\002es)h Fw(r)2184 1994
+y Ft(0)2235 2030 y Fw(`)i Fz(\033)2382 1994 y Ft(0)2405
+2030 y Fu(\()p Fw(r)p Fu(\))e FA(and)f Fw(r)2840 1994
+y Ft(0)2891 2030 y Fw(`)k Fz(\033)3039 1994 y Ft(0)3084
+2030 y Fw(\016)21 b Fz(\033)31 b Fw(\031)d Fz(\033)3405
+1994 y Ft(0)3429 2030 y FA(,)166 2150 y(so)c(that)h Fu(\()p
+Fw(r)p Fz(;)17 b(\033)t Fu(\))24 b FA(is)g(indeed)h(a)g(most)e(general)
+i(solution.)e(Since)i(one)g(of)g(those)f(solutions)f(is)h(the)166
+2271 y(most)31 b(general)h(solution)f Fu(\()p Fw(r)p
+Fz(;)17 b(\033)t Fu(\))p FA(,)31 b(we)i(also)f(kno)n(w)f(that)g
+Fw(r)42 b(`)f Fz(\033)31 b Fw(\016)d Fz(\033)45 b Fw(\031)c
+Fz(\033)c FA(and)32 b(hence)g(that)166 2391 y Fu(\()p
+Fw(r)p Fz(;)17 b(\033)t Fu(\))25 b FA(is)f(idempotent.)p
+3382 2391 4 68 v 3386 2327 60 4 v 3386 2391 V 3445 2391
+4 68 v 166 2798 a(W)-8 b(e)30 b(no)n(w)f(apply)g(the)g(nominal)f
+(uni\002cation)h(algorithm)f(to)h(solv)o(e)g(the)g(quiz)g(questions)f
+(from)166 2919 y(the)d(Introduction.)166 3147 y FB(Example)33
+b(3.8.)45 b FA(Using)32 b(the)g(\002rst)h(three)f(function)g(symbols)f
+(of)h(the)h(nominal)e(signature)h(of)166 3267 y(Example)23
+b(2.2)f(to)h(represent)h Fz(\025)p FA(-terms,)e(the)i(Quiz)f(at)g(the)g
+(end)g(of)g(the)h(Introduction)d(translates)166 3388
+y(into)26 b(the)g(follo)n(wing)f(four)i(uni\002cation)f(problems)f(o)o
+(v)o(er)h(that)g(signature,)g(where)h Fz(a)g FA(and)g
+Fz(b)g FA(are)166 3508 y(distinct)c(atoms)h(of)h(sort)f
+Fp(vid)35 b FA(and)25 b Fz(X)1439 3523 y FG(1)1478 3508
+y Fz(;)17 b(:)g(:)g(:)f(;)h(X)1778 3523 y FG(7)1842 3508
+y FA(are)26 b(distinct)d(v)n(ariables)h(of)h(sort)f Fp(exp)6
+b FA(:)587 3770 y Fz(P)650 3785 y FG(1)717 3717 y Fo(def)722
+3770 y Fu(=)29 b Fw(f)p Fv(fn)17 b Fz(a:)p Fv(fn)h Fz(b:)p
+Fv(app)q Fw(h)p Fz(X)1536 3785 y FG(1)1575 3770 y Fz(;)f
+Fv(vr)g Fz(b)p Fw(i)28 b(\031)p Fu(?)g Fv(fn)17 b Fz(b:)p
+Fv(fn)h Fz(a:)p Fv(app)q Fw(h)p Fv(vr)f Fz(a;)g(X)2870
+3785 y FG(1)2909 3770 y Fw(ig)p FA(,)587 3950 y Fz(P)650
+3965 y FG(2)717 3898 y Fo(def)722 3950 y Fu(=)29 b Fw(f)p
+Fv(fn)17 b Fz(a:)p Fv(fn)h Fz(b:)p Fv(app)q Fw(h)p Fz(X)1536
+3965 y FG(2)1575 3950 y Fz(;)f Fv(vr)g Fz(b)p Fw(i)28
+b(\031)p Fu(?)g Fv(fn)17 b Fz(b:)p Fv(fn)h Fz(a:)p Fv(app)q
+Fw(h)p Fv(vr)f Fz(a;)g(X)2870 3965 y FG(3)2909 3950 y
+Fw(ig)p FA(,)587 4131 y Fz(P)650 4146 y FG(3)717 4079
+y Fo(def)722 4131 y Fu(=)29 b Fw(f)p Fv(fn)17 b Fz(a:)p
+Fv(fn)h Fz(b:)p Fv(app)q Fw(h)p Fv(vr)f Fz(b;)g(X)1740
+4146 y FG(4)1779 4131 y Fw(i)28 b(\031)p Fu(?)g Fv(fn)17
+b Fz(b:)p Fv(fn)h Fz(a:)p Fv(app)q Fw(h)p Fv(vr)f Fz(a;)g(X)2870
+4146 y FG(5)2909 4131 y Fw(ig)p FA(,)587 4312 y Fz(P)650
+4327 y FG(4)717 4259 y Fo(def)722 4312 y Fu(=)29 b Fw(f)p
+Fv(fn)17 b Fz(a:)p Fv(fn)h Fz(b:)p Fv(app)q Fw(h)p Fv(vr)f
+Fz(b;)g(X)1740 4327 y FG(6)1779 4312 y Fw(i)28 b(\031)p
+Fu(?)g Fv(fn)17 b Fz(a:)p Fv(fn)h Fz(a:)p Fv(app)q Fw(h)p
+Fv(vr)f Fz(a;)g(X)2880 4327 y FG(7)2919 4312 y Fw(ig)p
+FA(.)166 4550 y(Applying)23 b(the)i(nominal)e(uni\002cation)i
+(algorithm)e(described)i(abo)o(v)o(e,)e(we)i(\002nd)g(that)166
+4778 y Fw(\017)50 b Fz(P)329 4793 y FG(1)393 4778 y FA(has)25
+b(no)f(solution;)166 4898 y Fw(\017)50 b Fz(P)329 4913
+y FG(2)397 4898 y FA(has)29 b(a)h(most)e(general)h(solution)f(gi)n(v)o
+(en)f(by)i Fw(r)1987 4913 y FG(2)2062 4898 y Fu(=)36
+b Fw(;)29 b FA(and)g Fz(\033)2481 4913 y FG(2)2557 4898
+y Fu(=)35 b([)p Fz(X)2776 4913 y FG(2)2852 4898 y Fu(:=)g
+Fv(vr)17 b Fz(b;)g(X)3275 4913 y FG(3)3351 4898 y Fu(:=)266
+5019 y Fv(vr)g Fz(a)p Fu(])p FA(;)166 5139 y Fw(\017)50
+b Fz(P)329 5154 y FG(3)393 5139 y FA(has)25 b(a)g(most)e(general)j
+(solution)d(gi)n(v)o(en)g(by)h Fw(r)1952 5154 y FG(3)2019
+5139 y Fu(=)k Fw(;)d FA(and)g Fz(\033)2422 5154 y FG(3)2489
+5139 y Fu(=)j([)p Fz(X)2701 5154 y FG(4)2768 5139 y Fu(:=)f(\()p
+Fz(a)17 b(b)p Fu(\))p Fr(\001)p Fz(X)3196 5154 y FG(5)3236
+5139 y Fu(])p FA(;)166 5259 y Fw(\017)50 b Fz(P)329 5274
+y FG(4)403 5259 y FA(has)35 b(a)g(most)f(general)i(solution)d(gi)n(v)o
+(en)g(by)i Fw(r)2034 5274 y FG(4)2120 5259 y Fu(=)47
+b Fw(f)p Fz(b)f Fu(#)h Fz(X)2589 5274 y FG(7)2629 5259
+y Fw(g)35 b FA(and)g Fz(\033)2948 5274 y FG(3)3034 5259
+y Fu(=)47 b([)p Fz(X)3265 5274 y FG(6)3351 5259 y Fu(:=)266
+5380 y(\()p Fz(b)17 b(a)p Fu(\))p Fr(\001)p Fz(X)564
+5395 y FG(7)603 5380 y Fu(])p FA(.)1748 5712 y(18)p eop
+end
+%%Page: 19 19
+TeXDict begin 19 18 bop 166 3 3288 4 v 166 2372 4 2369
+v 283 252 a Fq(P)328 261 y Fe(1)438 210 y Fd(")395 252
+y FG(=)-12 b Ft(\))33 b(f)p Fb(fn)13 b Fq(b:)p Fb(app)p
+Ft(h)p Fq(X)900 261 y Fe(1)936 252 y Fq(;)e Fb(vr)h Fq(b)p
+Ft(i)20 b(\031)p FG(?)g Fb(fn)13 b Fq(b:)p Fb(app)p Ft(h)p
+Fb(vr)g Fq(b;)f FG(\()p Fq(a)g(b)p FG(\))p Fa(\001)p
+Fq(X)1858 261 y Fe(1)1893 252 y Ft(i)p Fq(;)24 b(a)c
+FG(#?)g Fb(fn)12 b Fq(a:)p Fb(app)q Ft(h)p Fb(vr)h Fq(a;)e(X)2614
+261 y Fe(1)2649 252 y Ft(ig)34 b Fo(\()p Ft(\031)p FG(?)p
+Fo(-abstraction-2\))438 329 y Fd(")395 370 y FG(=)-12
+b Ft(\))33 b(f)p Fb(app)q Ft(h)p Fq(X)768 379 y Fe(1)803
+370 y Fq(;)12 b Fb(vr)g Fq(b)p Ft(i)20 b(\031)p FG(?)g
+Fb(app)q Ft(h)p Fb(vr)13 b Fq(b;)e FG(\()p Fq(a)i(b)p
+FG(\))p Fa(\001)p Fq(X)1594 379 y Fe(1)1629 370 y Ft(i)p
+Fq(;)31 b(a)20 b FG(#?)g Fb(fn)12 b Fq(a:)p Fb(app)q
+Ft(h)p Fb(vr)h Fq(a;)f(X)2358 379 y Fe(1)2393 370 y Ft(ig)290
+b Fo(\()p Ft(\031)p FG(?)p Fo(-abstraction-1\))405 489
+y Ft(\001)11 b(\001)h(\001)54 b(\001)12 b(\001)f(\001)2120
+b(\001)11 b(\001)h(\001)438 542 y Fd(")395 583 y FG(=)-12
+b Ft(\))33 b(f)p Fq(X)635 592 y Fe(1)690 583 y Ft(\031)p
+FG(?)20 b Fb(vr)12 b Fq(b;)31 b Fb(vr)12 b Fq(b)20 b
+Ft(\031)p FG(?)g(\()p Fq(a)13 b(b)p FG(\))p Fa(\001)p
+Fq(X)1416 592 y Fe(1)1450 583 y Fq(;)31 b(a)21 b FG(#?)e
+Fb(fn)13 b Fq(a:)p Fb(app)q Ft(h)p Fb(vr)f Fq(a;)g(X)2152
+592 y Fe(1)2187 583 y Ft(ig)496 b Fo(\()p Ft(\031)p FG(?)p
+Fo(-pair\))433 660 y Fd(\033)395 702 y FG(=)-12 b Ft(\))33
+b(f)p Fb(vr)13 b Fq(b)20 b Ft(\031)o FG(?)g Fb(vr)13
+b Fq(a;)31 b(a)20 b FG(#?)g Fb(fn)12 b Fq(a:)p Fb(app)q
+Ft(h)p Fb(vr)h Fq(a;)f Fb(vr)g Fq(b)p Ft(ig)40 b Fo(with)21
+b Fq(\033)h FG(=)e([)p Fq(X)2148 711 y Fe(1)2202 702
+y FG(:=)f Fb(vr)12 b Fq(b)p FG(])317 b Fo(\()p Ft(\031)p
+FG(?)p Fo(-v)n(ariable\))438 778 y Fd(")395 820 y FG(=)-12
+b Ft(\))33 b(f)p Fq(b)20 b Ft(\031)p FG(?)g Fq(a;)32
+b(a)20 b FG(#?)f Fb(fn)13 b Fq(a:)p Fb(app)q Ft(h)p Fb(vr)f
+Fq(a;)g Fb(vr)h Fq(b)p Ft(ig)1154 b Fo(\()p Ft(\031)p
+FG(?)p Fo(-function)19 b(symbol\))542 938 y FG(F)-8 b(AIL)283
+1128 y Fq(P)328 1137 y Fe(4)438 1086 y Fd(")395 1128
+y FG(=)c Ft(\))33 b(f)p Fb(fn)13 b Fq(b:)p Fb(app)p Ft(h)p
+Fb(vr)g Fq(b;)f(X)1045 1137 y Fe(6)1079 1128 y Ft(i)20
+b(\031)p FG(?)g Fb(fn)13 b Fq(a:)p Fb(app)q Ft(h)p Fb(vr)f
+Fq(a;)g(X)1716 1137 y Fe(7)1751 1128 y Ft(ig)932 b Fo(\()p
+Ft(\031)p FG(?)p Fo(-abstraction-1\))438 1204 y Fd(")395
+1246 y FG(=)-12 b Ft(\))33 b(f)p Fb(app)q Ft(h)p Fb(vr)13
+b Fq(b;)f(X)913 1255 y Fe(6)947 1246 y Ft(i)20 b(\031)p
+FG(?)g Fb(app)q Ft(h)p Fb(vr)13 b Fq(b;)e FG(\()p Fq(b)i(a)p
+FG(\))p Fa(\001)p Fq(X)1594 1255 y Fe(7)1629 1246 y Ft(i)p
+Fq(;)31 b(b)20 b FG(#?)g Fb(app)q Ft(h)p Fb(vr)12 b Fq(a;)g(X)2211
+1255 y Fe(7)2246 1246 y Ft(ig)437 b Fo(\()p Ft(\031)p
+FG(?)p Fo(-abstraction-2\))405 1364 y Ft(\001)11 b(\001)h(\001)54
+b(\001)12 b(\001)f(\001)2120 b(\001)11 b(\001)h(\001)438
+1417 y Fd(")395 1459 y FG(=)-12 b Ft(\))33 b(f)p Fq(b)20
+b Ft(\031)p FG(?)g Fq(b;)31 b(X)874 1468 y Fe(6)928 1459
+y Ft(\031)p FG(?)20 b(\()p Fq(b)13 b(a)p FG(\))p Fa(\001)p
+Fq(X)1251 1468 y Fe(7)1286 1459 y Fq(;)31 b(b)20 b FG(#?)f
+Fb(app)q Ft(h)p Fb(vr)13 b Fq(a;)f(X)1841 1468 y Fe(7)1875
+1459 y Ft(ig)808 b Fo(\()p Ft(\031)p FG(?)p Fo(-function)19
+b(symbol\))438 1535 y Fd(")395 1577 y FG(=)-12 b Ft(\))33
+b(f)p Fq(X)635 1586 y Fe(6)690 1577 y Ft(\031)p FG(?)20
+b(\()p Fq(b)12 b(a)p FG(\))p Fa(\001)p Fq(X)1012 1586
+y Fe(7)1047 1577 y Fq(;)31 b(b)20 b FG(#?)g Fb(app)p
+Ft(h)p Fb(vr)13 b Fq(a;)f(X)1602 1586 y Fe(7)1637 1577
+y Ft(ig)1046 b Fo(\()p Ft(\031)p FG(?)p Fo(-atom\))433
+1654 y Fd(\033)395 1695 y FG(=)-12 b Ft(\))33 b(f)p Fq(b)20
+b FG(#?)g Fb(app)q Ft(h)p Fb(vr)13 b Fq(a;)e(X)1081 1704
+y Fe(7)1116 1695 y Ft(ig)67 b Fo(with)21 b Fq(\033)h
+FG(=)e([)p Fq(X)1598 1704 y Fe(6)1652 1695 y FG(:=)f(\()p
+Fq(b)12 b(a)p FG(\))p Fa(\001)q Fq(X)1961 1704 y Fe(7)1995
+1695 y FG(])730 b Fo(\()p Ft(\031)p FG(?)p Fo(-v)n(ariable\))436
+1777 y Fc(;)395 1818 y FG(=)-12 b Ft(\))33 b(f)p Fq(b)20
+b FG(#?)g Ft(h)p Fb(vr)13 b Fq(a;)f(X)976 1827 y Fe(7)1010
+1818 y Ft(ig)1673 b Fo(\()p FG(#?)p Fo(-function)19 b(symbol\))405
+1937 y Ft(\001)11 b(\001)h(\001)54 b(\001)12 b(\001)f(\001)2120
+b(\001)11 b(\001)h(\001)436 1994 y Fc(;)395 2036 y FG(=)-12
+b Ft(\))33 b(f)p Fq(b)20 b FG(#?)g Fq(a;)31 b(b)20 b
+FG(#?)g Fq(X)1047 2045 y Fe(7)1082 2036 y Ft(g)1628 b
+Fo(\()p FG(#?)p Fo(-function)19 b(symbol\))436 2117 y
+Fc(;)395 2159 y FG(=)-12 b Ft(\))33 b(f)p Fq(b)20 b FG(#?)g
+Fq(X)797 2168 y Fe(7)832 2159 y Ft(g)1878 b Fo(\()p FG(#?)p
+Fo(-atom\))427 2237 y Fc(r)395 2278 y FG(=)-12 b Ft(\))33
+b(;)67 b Fo(with)21 b Ft(r)e FG(=)g Ft(f)p Fq(b)i FG(#)e
+Fq(X)1156 2287 y Fe(7)1191 2278 y Ft(g)1519 b Fo(\()p
+FG(#?)p Fo(-suspension\))p 3450 2372 V 166 2375 3288
+4 v 1308 2512 a FD(Fig.)23 b(4.)g(Example)h(deri)n(v)n(ations)166
+2705 y FA(Deri)n(v)n(ations)j(for)j Fz(P)864 2720 y FG(1)933
+2705 y FA(and)f Fz(P)1169 2720 y FG(4)1237 2705 y FA(are)h(sk)o(etched)
+f(in)g(Figure)h(4.)f(Using)f(the)h(Adequac)o(y)g(property)166
+2826 y(of)j(Proposition)e(2.16,)i(one)f(can)i(interpret)e(these)h
+(solutions)e(as)i(the)f(follo)n(wing)f(statements)166
+2946 y(about)24 b(the)h Fz(\025)p FA(-terms)f(mentioned)g(in)g(the)h
+(quiz.)p 166 3085 V 166 4368 4 1284 v 227 3212 a FB(Quiz)g(answers)267
+3332 y FA(\(1\))50 b(There)26 b(is)e(no)h Fz(\025)p FA(-term)f
+Fz(M)1298 3347 y FG(1)1363 3332 y FA(making)g(the)g(\002rst)h(pair)g
+(of)g(terms)f Fz(\013)q FA(-equi)n(v)n(alent.)267 3453
+y(\(2\))50 b(The)25 b(only)f(solution)f(for)i(the)g(second)g(problem)f
+(is)g(to)g(tak)o(e)h Fz(M)2585 3468 y FG(2)2653 3453
+y Fu(=)i Fz(b)f FA(and)e Fz(M)3085 3468 y FG(3)3153 3453
+y Fu(=)j Fz(a)p FA(.)267 3573 y(\(3\))50 b(F)o(or)24
+b(the)g(third)g(problem)f(we)h(can)g(tak)o(e)g Fz(M)1894
+3588 y FG(5)1958 3573 y FA(to)g(be)g(an)o(y)f Fz(\025)p
+FA(-term,)h(so)g(long)f(as)h(we)g(tak)o(e)433 3693 y
+Fz(M)527 3708 y FG(4)592 3693 y FA(to)h(be)g(the)f(result)g(of)h(sw)o
+(apping)f(all)h(occurrences)g(of)g Fz(a)g FA(and)g Fz(b)g
+FA(throughout)f Fz(M)3321 3708 y FG(5)3360 3693 y FA(.)267
+3814 y(\(4\))50 b(F)o(or)21 b(the)f(last)g(problem,)f(we)i(can)g(tak)o
+(e)f Fz(M)1844 3829 y FG(7)1904 3814 y FA(to)g(be)h(an)o(y)f
+Fz(\025)p FA(-term)g(that)g Fx(does)g(not)f(contain)433
+3934 y(fr)l(ee)30 b(occurr)l(ences)f(of)g Fz(b)p FA(,)h(so)e(long)h(as)
+g(we)g(tak)o(e)h Fz(M)2199 3949 y FG(6)2268 3934 y FA(to)e(be)i(the)f
+(result)f(of)h(sw)o(apping)433 4055 y(all)k(occurrences)g(of)g
+Fz(b)g FA(and)f Fz(a)h FA(throughout)e Fz(M)2088 4070
+y FG(7)2128 4055 y FA(,)h(or)h(equi)n(v)n(alently)d(\(since)j
+Fz(b)g FA(is)f(not)433 4175 y(free)c(in)e Fz(M)813 4190
+y FG(7)852 4175 y FA(\),)h(taking)e Fz(M)1306 4190 y
+FG(6)1373 4175 y FA(to)h(be)g(the)g(result)g(of)h(replacing)f(all)g
+(free)h(occurrences)g(of)433 4295 y Fz(a)f FA(in)e Fz(M)706
+4310 y FG(7)771 4295 y FA(with)g Fz(b)p FA(.)p 3450 4368
+V 166 4371 3288 4 v 166 4554 a FB(Remark)34 b(3.9)f(\(Separation)i(of)f
+(the)g(algorithm)f(into)h(tw)o(o)g(phases\).)46 b FA(W)-8
+b(e)34 b(or)n(ganised)f(the)166 4675 y(algorithm)f(into)g(tw)o(o)h
+(phases:)g(equation-solving)e(follo)n(wed)h(by)h(freshness-solving.)e
+(Note)166 4795 y(that)22 b(the)f(second)h(phase)g(is)g(crucial)g(for)g
+(the)g(soundness)f(of)h(the)f(algorithm.)g(Consider)h(for)g(e)o(x-)166
+4915 y(ample)i(the)h(uni\002cation)f(problem)g(consisting)f(of)i(tw)o
+(o)f(terms)h(which)f(are)i Fx(not)g Fz(\013)q FA(-equi)n(v)n(alent:)
+1551 5148 y Fw(f)p Fz(a:b)i Fw(\031)q Fu(?)f Fz(b:a)p
+Fw(g)1220 b FA(\(13\))166 5380 y(After)29 b(applying)d(the)i
+(transformation)f(\()p Fw(\031)q Fu(?)p FA(-abstraction-2\))h(one)g
+(needs)g(to)g(solv)o(e)f(the)h(prob-)1748 5712 y(19)p
+eop end
+%%Page: 20 20
+TeXDict begin 20 19 bop 166 83 a FA(lem)29 b Fw(f)p Fz(a)37
+b Fw(\031)p Fu(?)g Fz(a;)17 b(a)36 b Fu(#?)h Fz(a)p Fw(g)p
+FA(,)30 b(whose)f(\002rst)h(component)e(is)h(solv)o(ed)g(by)g(\()p
+Fw(\031)q Fu(?)p FA(-atom\).)g(F)o(ailure)g(is)166 203
+y(only)g(signalled)g(by)g(the)h(algorithm)e(in)i(the)g(second)f(phase)h
+(when)g(attempting)e(to)h(solv)o(e)g(the)166 324 y(unsolv)n(able)d
+(freshness)h(problem)g Fw(f)p Fz(a)33 b Fu(#?)g Fz(a)p
+Fw(g)p FA(.)28 b(The)g(second)f(phase,)h(i.)f(e.)h(solving)e(all)h
+(fresh-)166 444 y(ness)e(problems,)f(ensures)h(that)f(the)h(uni\002ers)
+g(calculated)g(by)g(the)g(algorithm)f(are)h(sound)g(with)166
+565 y(respect)g(to)g(our)f(notion)g(of)h Fz(\013)q FA(-equi)n(v)n
+(alence.)166 955 y(W)-8 b(e)27 b(used)f(this)g(separation)g(of)h(the)f
+(algorithm)f(into)h(tw)o(o)g(phases)g(in)g(order)h(to)f(mak)o(e)h(the)f
+(cor)n(-)166 1075 y(rectness)c(proof)f(easier)-5 b(.)22
+b(More)f(ef)n(\002cient)h(algorithms)e(w)o(ould)h(seek)g(to)h(minimise)
+d(the)j(amount)166 1196 y(of)29 b(redundant)g(calculations)f(before)i
+(f)o(ailures)f(are)h(signalled,)d(by)i(solving)f(freshness)h(prob-)166
+1316 y(lems)g(more)g(eagerly)-6 b(.)29 b(Ho)n(we)n(v)o(er)l(,)f(care)i
+(needs)f(then)g(to)g(be)g(tak)o(en)g(to)g(not)g(remo)o(v)o(e)f
+(freshness)166 1436 y(constraints)h(from)g(problems)g(too)h(early)-6
+b(.)29 b(F)o(or)h(e)o(xample,)f(consider)h(the)f(follo)n(wing)f
+(uni\002ca-)166 1557 y(tion)c(problem,)g(which)g(has)h(no)f(solution.)
+1420 2117 y Fw(f)p Fz(a)k Fu(#?)g Fz(X)r(;)17 b(a)28
+b Fw(\031)p Fu(?)g Fz(X)8 b Fw(g)1088 b FA(\(14\))166
+2677 y(If)40 b(one)g(applies)f(\002rst)h(\()p Fu(#?)p
+FA(-suspension\))f(follo)n(wed)f(by)i(\()p Fw(\031)p
+Fu(?)q FA(-v)n(ariable\),)f(then)g(one)h(gets)f(a)166
+2798 y Fx(wr)l(ong)31 b FA(result,)g(namely)g Fu(\()p
+Fw(f)p Fz(a)40 b Fu(#)g Fz(X)8 b Fw(g)p Fz(;)17 b Fu([)p
+Fz(X)48 b Fu(:=)40 b Fz(a)p Fu(]\))p FA(.)32 b(The)f(problem)g(is)g
+(that)g(the)h(substitution)166 2918 y Fu([)p Fz(X)46
+b Fu(:=)38 b Fz(a)p Fu(])31 b FA(has)f(not)g(been)g(properly)g
+(propagated)g(to)g(the)g(freshness)g(constraint)g Fz(a)38
+b Fu(#)g Fz(X)8 b FA(.)30 b(If)166 3039 y(freshness)h(problems)g(are)h
+(solv)o(ed)e(more)i(eagerly)-6 b(,)31 b(then)g(proper)h(propagation)e
+(of)i(substitu-)166 3159 y(tions)24 b(into)g(freshness)g(constraints)g
+(needs)h(to)f(be)h(tak)o(en)g(into)f(account.)166 3549
+y FB(Remark)34 b(3.10)f(\(Atoms)g(ar)n(e)h(not)f(v)o(ariables\).)46
+b FA(Nominal)32 b(uni\002cation)h(uni\002es)g(v)n(ariables,)166
+3670 y(b)n(ut)e(it)f(does)h(not)g(unify)f(atoms.)h(Indeed)g(the)g
+(operation)f(of)i(identifying)d(tw)o(o)i(atoms)f(by)h(re-)166
+3790 y(naming)h(one)i(of)f(them)g(to)g(be)h(the)f(other)g(does)g(not)g
+(necessarily)g(preserv)o(e)h(the)f(v)n(alidity)f(of)166
+3910 y(the)k(judgements)f(in)h(Figure)g(2.)g(F)o(or)h(e)o(xample,)e
+Fw(;)49 b(`)g Fz(a:b)g Fw(\031)h Fz(c:b)37 b FA(holds)e(if)h
+Fz(b)50 b Fw(6)p Fu(=)e Fz(a;)17 b(c)p FA(;)36 b(b)n(ut)166
+4031 y(renaming)30 b Fz(b)g FA(to)g(be)g Fz(a)h FA(in)f(this)f
+(judgement)g(we)h(get)g Fw(;)37 b(`)h Fz(a:a)g Fw(\031)g
+Fz(c:a)p FA(,)30 b(which)g(does)g(not)g(hold)166 4151
+y(so)g(long)f(as)i Fz(a)38 b Fw(6)p Fu(=)f Fz(c)p FA(.)31
+b(Referring)g(to)f(De\002nition)f(2.3,)h(you)g(will)f(see)i(that)e(we)i
+(do)f(allo)n(w)f(v)n(ari-)166 4271 y(ables)e(ranging)f(o)o(v)o(er)g
+(sorts)h(of)g(atoms;)f(and)h(such)f(v)n(ariables)h(can)g(be)g
+(uni\002ed)g(lik)o(e)g(an)o(y)f(other)166 4392 y(v)n(ariables.)e(Ho)n
+(we)n(v)o(er)l(,)g(if)h Fz(A)g FA(is)f(such)h(a)g(v)n(ariable,)f(then)h
+(it)f(cannot)h(appear)g(in)g(abstraction)f(po-)166 4512
+y(sition,)19 b(i.e.)h(as)h Fz(A:t)p FA(.)g(This)e(is)h(because)h(we)g
+(speci\002cally)f(restricted)g(abstraction)g(to)g(range)h(o)o(v)o(er)
+166 4633 y(atoms,)27 b(rather)h(than)g(o)o(v)o(er)f(arbitrary)h(terms)f
+(of)h(atom)f(sort.)g(Such)i(a)f(restriction)f(seems)g(nec-)166
+4753 y(essary)35 b(to)g(obtain)f(single,)g(most)g(general,)h(solutions)
+e(to)i(nominal)f(uni\002cation)g(problems.)166 4873 y(F)o(or)i(without)
+f(such)h(a)h(restriction,)e(because)i(of)f(rule)g(\()p
+Fw(\031)p FA(-abstraction-2\))h(in)f(Figure)g(2)h(we)166
+4994 y(w)o(ould)27 b(also)h(ha)n(v)o(e)f(to)h(allo)n(w)f(v)n(ariables)g
+(to)g(appear)i(on)e(the)h(left-hand)g(side)f(of)h(freshness)g(re-)166
+5114 y(lations)c(and)i(in)f(suspended)f(permutations.)g(So)i(then)f(we)
+g(w)o(ould)g(get)g(uni\002cation)g(problems)166 5235
+y(lik)o(e)i Fw(f)p Fu(\()p Fz(A)17 b(B)5 b Fu(\))p Fr(\001)p
+Fz(C)40 b Fw(\031)p Fu(?)33 b Fz(C)7 b Fw(g)p FA(,)28
+b(where)g Fz(A)p FA(,)g Fz(B)33 b FA(and)28 b Fz(C)34
+b FA(are)29 b(v)n(ariables)e(of)h(atom)f(sort;)g(this)g(has)g(tw)o(o)
+166 5355 y(incomparable)d(solutions,)f(namely)h Fu(\()p
+Fw(;)p Fz(;)17 b Fu([)p Fz(A)27 b Fu(:=)h Fz(B)5 b Fu(]\))25
+b FA(and)g Fu(\()p Fw(f)p Fz(A)i Fu(#)h Fz(C)r(;)17 b(B)32
+b Fu(#)c Fz(C)7 b Fw(g)p Fz(;)17 b(")p Fu(\))p FA(.)1748
+5712 y(20)p eop end
+%%Page: 21 21
+TeXDict begin 21 20 bop 166 83 a FB(4)99 b(Related)26
+b(w)o(ork)166 424 y Fx(Higher)n(-or)l(der)e(pattern)g(uni\002cation)166
+765 y FA(Most)32 b(pre)n(vious)h(w)o(ork)g(on)g(uni\002cation)g(for)g
+(languages)g(with)g(binders)g(is)g(based)g(on)g(forms)166
+885 y(of)h(higher)n(-order)g(uni\002cation,)f(i.e.)g(solving)f
+(equations)h(between)h Fz(\025)p FA(-terms)f(modulo)g
+Fz(\013)q(\014)6 b(\021)t FA(-)166 1005 y(equi)n(v)n(alence)40
+b(\()p Fu(=)787 1020 y Fq(\013\014)s(\021)917 1005 y
+FA(\))i(by)e(capture-a)n(v)n(oiding)h(substitution)d(of)j(terms)g(for)g
+(function)f(v)n(ari-)166 1126 y(ables.)33 b(Notable)f(among)g(that)h(w)
+o(ork)g(is)f(Miller')-5 b(s)32 b Fx(higher)n(-or)l(der)g(pattern)g
+(uni\002cation)f FA(used)166 1246 y(in)22 b(his)f Fz(L)470
+1261 y Fq(\025)538 1246 y FA(logic)h(programming)f(language)h([3].)g
+(This)f(kind)h(of)g(uni\002cation)g(retains)f(the)h(good)166
+1367 y(properties)j(of)h(\002rst-order)g(uni\002cation:)e(a)i(linear)n
+(-time)f(decision)g(procedure)h(and)g(e)o(xistence)166
+1487 y(of)34 b(most)f(general)h(uni\002ers.)g(This)f(good)g(beha)n
+(viour)g(of)h(higher)n(-order)g(pattern)g(uni\002cation)166
+1607 y(is)c(the)h(result)f(of)h(equations)f(being)g(solv)o(ed)f(only)h
+(modulo)g Fu(=)2356 1622 y Fq(\013\014)2441 1631 y Fe(0)2475
+1622 y Fq(\021)2548 1607 y FA(\(where)h Fz(\014)2910
+1622 y FG(0)2950 1607 y FA(-equi)n(v)n(alence)166 1728
+y(is)26 b(the)g(restricted)g(form)g(of)g Fz(\014)6 b
+FA(-equi)n(v)n(alence)25 b(that)g(identi\002es)h Fu(\()p
+Fz(\025x:M)10 b Fu(\))p Fz(y)30 b FA(and)d Fz(M)10 b
+Fu([)p Fz(y)t(=x)p Fu(])26 b FA(with)g Fz(y)166 1848
+y FA(being)i(a)i(v)n(ariable\))e(and)h(of)g Fz(\025)p
+FA(-terms)f(being)g(restricted)h(such)g(that)f(function)g(v)n(ariables)
+g(may)166 1969 y(only)f(be)h(applied)f(to)g(distinct)f(bound)h(v)n
+(ariables.)g(An)h(empirical)f(study)f(by)i(Michaylo)o(v)e(and)166
+2089 y(Pfenning)31 b([27])g(suggests)e(that)i(most)f(uni\002cations)g
+(arising)g(dynamically)f(in)i(higher)n(-order)166 2209
+y(logic)k(programming)g(satisfy)g(Miller')-5 b(s)35 b(restrictions,)g
+(b)n(ut)g(that)h(it)f(rules)h(out)g(some)f(useful)166
+2330 y(programming)23 b(idioms.)166 2550 y(The)i(main)g(dif)n(ference)g
+(between)g(higher)n(-order)g(pattern)g(uni\002cation)g(and)g(nominal)e
+(uni\002ca-)166 2671 y(tion)29 b(is)h(that)g(the)g(former)g(solv)o(es)f
+(a)h(set)g(of)g(equations)f(by)h(calculating)g(a)g Fx(captur)l
+(e-avoiding)166 2791 y FA(substitution,)23 b(while)j(the)g(latter)h
+(calculates)f(a)h Fx(possibly-capturing)c FA(substitution)h
+Fx(and)k FA(some)166 2911 y(freshness)33 b(constraints.)f(Moreo)o(v)o
+(er)l(,)g(uni\002ers)h(in)f(higher)n(-order)i(pattern)e(uni\002cation)h
+(solv)o(e)166 3032 y(equations)40 b(with)h(respect)h(to)f
+Fu(=)1327 3047 y Fq(\013\014)1412 3056 y Fe(0)1446 3047
+y Fq(\021)1488 3032 y FA(;)g(whereas)h(in)f(nominal)f(uni\002cation,)h
+(uni\002ers)g(solv)o(e)166 3152 y(equations)f(with)h(respect)g(to)g
+(the)g(equi)n(v)n(alence)f Fw(\031)i FA(de\002ned)f(in)g(Figure)h(2,)f
+(which)f(agrees)166 3272 y(with)29 b Fz(\013)q FA(-equi)n(v)n(alence)f
+(on)i(ground)f(terms)g(\(see)h(Proposition)e(2.16\),)h(b)n(ut)g(dif)n
+(fers)h(from)f(it)g(on)166 3393 y(open)g(terms,)f(since)h(unlik)o(e)g
+Fz(\013)q FA(-equi)n(v)n(alence,)f(it)g(is)h(respected)g(by)g
+(possibly-capturing)e(sub-)166 3513 y(stitutions)c(\(see)j(Lemma)f
+(2.14\).)g(F)o(or)g(us,)g(the)g(main)g(disadv)n(antage)f(of)i(higher)n
+(-order)f(pattern)166 3634 y(uni\002cation)k(is)h(the)f(one)h(common)f
+(to)g(most)g(approaches)h(based)f(on)h(higher)n(-order)g(abstract)166
+3754 y(syntax)23 b(that)h(w)o(as)g(discussed)f(in)g(the)h
+(Introduction:)e(one)i(cannot)g Fx(dir)l(ectly)g FA(e)o(xpress)f(the)h
+(com-)166 3874 y(mon)33 b(idiom)g(of)h(possibly-capturing)d
+(substitution)h(of)i(terms)f(for)h(meta)n(v)n(ariables.)f(Instead)166
+3995 y(one)26 b(has)f(to)g(encode)h(meta)n(v)n(ariables)f
+Fz(X)33 b FA(as)26 b(function)e(v)n(ariables)h(applied)g(to)g(distinct)
+f(lists)h(of)166 4115 y(\(bound\))d(v)n(ariables,)g Fz(X)i(x)1069
+4130 y FG(1)1126 4115 y Fz(:)17 b(:)g(:)f(x)1312 4130
+y Fq(n)1359 4115 y FA(,)23 b(and)f(use)h(capture-a)n(v)n(oiding)e
+(substitution.)f(At)i(\002rst)h(sight,)166 4236 y(there)e(seems)f(to)h
+(be)g(a)g(simple)f(encoding)g(for)h(doing)f(that.)g(Consider)h(for)g(e)
+o(xample)f(the)g(purely)166 4356 y(equational)k(nominal)f
+(uni\002cation)i(problem)1582 4577 y Fz(a:X)36 b Fw(\031)p
+Fu(?)28 b Fz(b:b)1251 b FA(\(15\))166 4798 y(which)29
+b(is)g(solv)o(ed)f(by)i Fu(\()p Fw(;)p Fz(;)17 b Fu([)p
+Fz(X)43 b Fu(:=)37 b Fz(a)p Fu(]\))p FA(.)29 b(The)h(literal)f
+(encoding)g(as)g(the)h(higher)n(-order)f(pattern)166
+4918 y(uni\002cation)23 b(problem)h Fz(\025a:X)35 b Fu(=)1300
+4933 y Fq(\013\014)1385 4942 y Fe(0)1420 4933 y Fq(\021)1462
+4918 y Fu(?)p Fz(\025b:b)25 b FA(does)e(not)h(w)o(ork)g(of)g(course,)g
+(because)g(there)h(is)e(no)166 5039 y(capture-a)n(v)n(oiding)33
+b(substitution)e(that)i(solv)o(es)f(this)h(problem.)f(Ho)n(we)n(v)o(er)
+l(,)h Fz(X)41 b FA(can)34 b(be)g(made)166 5159 y(dependent)24
+b(on)h Fz(a)g FA(yielding)f(the)g(uni\002cation)g(problem)1381
+5380 y Fz(\025a:)p Fu(\()p Fz(X)8 b(a)p Fu(\))28 b(=)1836
+5395 y Fq(\013\014)1921 5404 y Fe(0)1956 5395 y Fq(\021)1998
+5380 y Fu(?)f Fz(\025b:b)1051 b FA(\(16\))1748 5712 y(21)p
+eop end
+%%Page: 22 22
+TeXDict begin 22 21 bop 166 83 a FA(which)30 b(is)g(solv)o(ed)g(by)g
+(the)h(capture-a)n(v)n(oiding)e(substitution)f(of)j Fz(\025c:c)g
+FA(for)g Fz(X)8 b FA(.)30 b(If)h(one)g(further)166 203
+y(applies)d(to)h Fz(\025c:c)g FA(the)g(atom)f Fz(a)h
+FA(used)g(by)g(the)f(encoding,)h(then)f(one)h(can)h(read)f(back)g(the)g
+(orig-)166 324 y(inal)f(solution)f Fu([)p Fz(X)43 b Fu(:=)35
+b Fz(a)p Fu(])29 b FA(by)g(applying)f(some)g Fz(\014)6
+b FA(-reductions.)27 b(There)j(are)f(ho)n(we)n(v)o(er)f(se)n(v)o(eral)
+166 444 y(problems)j(with)h(this)f(encoding.)g(First,)h(the)g(encoding)
+g(in)g(general)g(results)g(in)f(a)i(quadratic)166 565
+y(blo)n(w-up)24 b(in)g(the)h(size)f(of)h(terms.)f(F)o(or)h(e)o(xample)f
+(the)h(nominal)e(uni\002cation)h(problem)1325 861 y Fz(a:b:)p
+Fw(h)p Fz(X)r(;)17 b(Y)22 b Fw(i)28 b(\031)p Fu(?)g Fz(a:b:)p
+Fw(h)p Fz(a;)17 b(b)p Fw(i)994 b FA(\(17\))166 1158 y(solv)o(ed)29
+b(by)g(the)h(uni\002er)g Fu(\()p Fw(;)p Fz(;)17 b Fu([)p
+Fz(X)44 b Fu(:=)37 b Fz(a;)17 b(Y)58 b Fu(:=)36 b Fz(b)p
+Fu(]\))31 b FA(needs)e(to)h(be)g(encoded)f(so)h(that)f
+Fz(X)38 b FA(and)29 b Fz(Y)166 1278 y FA(depend)c(on)f(both)g
+Fz(a)h FA(and)g Fz(b)p FA(.)g(This)f(gi)n(v)o(es)g(the)g(higher)n
+(-order)h(pattern)g(problem)975 1575 y Fz(\025a:\025b:)p
+Fw(h)p Fz(X)g(a)17 b(b;)g(Y)38 b(a)17 b(b)p Fw(i)28 b
+Fu(=)1880 1590 y Fq(\013\014)1965 1599 y Fe(0)1999 1590
+y Fq(\021)2041 1575 y Fu(?)g Fz(\025a:\025b:)p Fw(h)p
+Fz(a;)17 b(b)p Fw(i)28 b Fz(:)644 b FA(\(18\))166 1871
+y(In)35 b(the)g(general)h(case,)g(the)f(encoding)f(needs)i(to)e(mak)o
+(e)i(meta)n(v)n(ariables)e(dependent)h(on)g Fx(all)166
+1992 y FA(atoms)23 b(occurring)g(in)h(a)g(nominal)e(uni\002cation)h
+(problem,)g(re)o(gardless)f(of)i(whether)g(the)o(y)f(actu-)166
+2112 y(ally)f(occur)g(in)g(an)h(indi)n(vidual)d(equational)h(problem.)h
+(F)o(or)g(e)o(xample,)f(if)h Fz(X)31 b FA(occurs)22 b(else)n(where)166
+2232 y(within)31 b(the)i(scope)f(of)h(abstractions)e(of)i
+Fz(c)p FA(,)f Fz(d)p FA(,)g Fz(e)h FA(and)f Fz(f)11 b
+FA(,)33 b(then)f Fz(X)40 b FA(needs)33 b(to)f(be)g(encoded)h(as)166
+2353 y Fu(\()p Fz(X)24 b(a)17 b(b)g(c)g(d)g(e)g(f)11
+b Fu(\))26 b FA(e)n(v)o(en)h(though)f(an)i(indi)n(vidual)d(equational)i
+(problem)g(might)f(contain)h(only)g Fz(a)166 2473 y FA(and)d
+Fz(b)p FA(.)h(Secondly)-6 b(,)23 b(and)h(more)g(importantly)-6
+b(,)22 b(we)i(cannot)g(see)h(ho)n(w)e(to)h(encode)g(our)g(freshness)166
+2594 y(constraints)35 b(using)h(this)g(kind)f(of)i(higher)n(-order)g
+(patterns.)f(\(Note)g(that)g(in)h(nominal)e(uni\002-)166
+2714 y(cation,)28 b(freshness)g(constraints)f(do)h(not)g(necessarily)g
+(come)g(from)g(analysing)f(abstractions,)166 2834 y(rather)e(the)o(y)f
+(can)h(be)g(chosen)g(arbitrarily)-6 b(.\))166 3093 y(A)34
+b(more)g(promising)e(tar)n(get)j(for)f(a)h(reduction)e(of)h(nominal)f
+(uni\002cation)h(to)f(some)h(form)g(of)166 3213 y(higher)n(-order)28
+b(pattern)g(uni\002cation)f(is)h Fz(\025\033)t FA(,)g(a)g
+Fz(\025)p FA(-calculus)g(with)f(de-Bruijn)i(indices)e(and)h(e)o(x-)166
+3333 y(plicit)d(substitutions.)e(Do)n(wek)i Fx(et)h(al)g
+FA([28])g(present)g(a)g(v)o(ersion)f(of)h(higher)n(-order)g(pattern)g
+(uni-)166 3454 y(\002cation)f(for)f Fz(\025\033)29 b
+FA(in)24 b(which)g(uni\002cation)g(problems)f(are)j(solv)o(ed,)d(as)h
+(in)g(nominal)g(uni\002cation,)166 3574 y(by)i(te)o(xtual)e
+(replacements)i(of)g(terms)g(for)g(v)n(ariables;)f(ho)n(we)n(v)o(er)f
+(a)j(\223pre-cooking\224)e(operation)166 3694 y(ensures)30
+b(that)f(the)h(te)o(xtual)f(replacements)h(can)g(be)g(f)o(aithfully)f
+(related)h(to)g(capture-a)n(v)n(oiding)166 3815 y(substitutions.)24
+b(It)k(seems)f(possible)f(that)g(the)i(freshness)f(\(as)h(well)f(as)g
+(the)g(equational\))g(prob-)166 3935 y(lems)34 b(of)g(nominal)f
+(uni\002cation)h(can)h(be)g(encoded)f(into)g(higher)n(-order)g(pattern)
+h(uni\002cation)166 4056 y(problems)21 b(o)o(v)o(er)f
+Fz(\025\033)t FA(,)i(using)f(a)h(non-tri)n(vial)d(translation)i(in)l(v)
+n(olving)e(the)j(use)f(of)h(the)g(shift)e(oper)n(-)166
+4176 y(ator)k(and)f(the)g(introduction)f(of)h(fresh)h(uni\002cation)f
+(v)n(ariables.)f(The)i(details)e(of)i(this)e(encoding)166
+4296 y(still)32 b(remain)i(to)f(be)h(in)l(v)o(estigated.)d
+(Furthermore,)j(it)f(is)h(not)f(clear)h(to)g(us)f(ho)n(w)g(to)g
+(translate)166 4417 y(solutions)k(obtained)i(via)g(the)g(encoding)f
+(back)h(into)g(solutions)e(of)i(the)g(original)f(nominal)166
+4537 y(uni\002cation)33 b(problem.)f(But)h(e)n(v)o(en)f(if)i(it)e
+(turns)h(out)f(that)h(it)g(is)g(possible)f(to)g(reduce)i(nominal)166
+4658 y(uni\002cation)24 b(to)g(the)g(algorithm)f(of)i(Do)n(wek)f
+Fx(et)g(al)p FA(,)g(the)h(calculations)e(in)l(v)n(olv)o(ed)g(in)h
+(translating)166 4778 y(our)h(terms)g(into)f Fz(\025\033)29
+b FA(patterns)c(and)g(then)g(using)f(higher)n(-order)h(pattern)g
+(uni\002cation)f(seem)h(f)o(ar)166 4898 y(more)j(intricate)f(than)g
+(our)h(simple)e(algorithm)h(that)g(solv)o(es)f(nominal)h(uni\002cation)
+g(problems)166 5019 y(directly)-6 b(.)27 b(The)h(conclusion)f(we)h(dra)
+o(w)g(is)g(that)f(an)h(encoding)g(of)g(nominal)f(uni\002cation)g(prob-)
+166 5139 y(lems)34 b(into)g(higher)n(-order)h(pattern)g(uni\002cation)g
+(problems)e(\(using)i(de)g(Bruijn)g(indices)f(and)166
+5259 y(e)o(xplicit)c(substitutions\))f(might)i(be)g(possible,)g(b)n(ut)
+g(such)h(an)f(encoding)h(is)f(no)g(substitute)f(in)166
+5380 y(practice)25 b(for)g(ha)n(ving)f(the)h(simple,)f(direct)g
+(algorithm)g(we)h(presented)f(here.)1748 5712 y(22)p
+eop end
+%%Page: 23 23
+TeXDict begin 23 22 bop 166 83 a Fx(Hamana')l(s)24 b
+Fz(\014)657 98 y FG(0)697 83 y Fx(-uni\002cation)f(of)h
+Fz(\025)p Fx(-terms)h(with)f(\223holes\224)166 433 y
+FA(Hamana)31 b([5,29])f(manages)h(to)f(add)h(possibly-capturing)d
+(substitution)g(to)j(a)g(language)f(lik)o(e)166 554 y(Miller')-5
+b(s)20 b Fz(L)569 569 y Fq(\025)615 554 y FA(.)i(This)f(is)g(achie)n(v)
+o(ed)g(by)h(adding)f(syntax)g(for)h(e)o(xplicit)e(renaming)h
+(operations)g(and)166 674 y(by)31 b(recording)g(implicit)f
+(dependencies)h(of)g(v)n(ariables)g(upon)g(bindable)f(names)h(in)g(a)h
+(typing)166 795 y(conte)o(xt.)20 b(The)h(mathematical)f(foundation)f
+(for)j(Hamana')-5 b(s)20 b(system)g(is)g(the)h(model)f(of)h(binding)166
+915 y(syntax)29 b(of)h(Fiore)g Fx(et)f(al)h FA([24].)f(The)h
+(mathematical)e(foundation)h(for)h(our)f(w)o(ork)h(appeared)g(at)166
+1035 y(the)25 b(same)f(time)g(\(see)i([10]\))f(and)g(is)f(in)g(a)i
+(sense)e(complementary)-6 b(.)23 b(F)o(or)i(in)g(Hamana')-5
+b(s)24 b(system)166 1156 y(the)g(typing)e(conte)o(xt)h(restricts)g
+(which)h(terms)f(may)h(be)g(substituted)d(for)j(a)h(v)n(ariable)e(by)h
+(gi)n(ving)166 1276 y(a)39 b(\002nite)f(set)g(of)g(names)g(that)g
+Fx(must)g(contain)f FA(the)h(free)i(names)e(of)g(such)g(a)h(term;)e
+(whereas)166 1396 y(we)29 b(gi)n(v)o(e)e(a)i(\002nite)f(set)g(of)h
+(names)f(which)g(the)h(term')-5 b(s)27 b(free)j(v)n(ariables)d
+Fx(must)h(avoid)p FA(.)g(Since)h Fz(\013)q FA(-)166 1517
+y(con)l(v)o(ersion)c(is)i(phrased)f(in)g(terms)g(of)h(a)n(v)n(oidance,)
+f(i.e.)h(freshness)f(of)h(names,)f(our)h(approach)166
+1637 y(seems)g(more)h(natural)g(if)g(one)f(w)o(ants)h(to)f(compute)g
+Fz(\013)q FA(-equi)n(v)n(alences)g(concretely)-6 b(.)27
+b(On)g(top)h(of)166 1758 y(that,)d(our)g(use)g(of)h(name)f
+(permutations,)f(rather)i(than)f(arbitrary)g(renaming)g(functions,)f
+(leads)166 1878 y(to)31 b(technical)f(simpli\002cations.)f(In)i(an)o(y)
+f(case,)i(the)f(bottom)e(line)i(is)f(that)h(Hamana')-5
+b(s)30 b(system)166 1998 y(seems)k(more)h(complicated)e(than)h(the)h
+(one)f(presented)h(here)g(and)f(does)g(not)g(possess)g(most)166
+2119 y(general)25 b(uni\002ers.)166 2512 y Fx(Qu-Pr)l(olo)o(g)166
+2863 y FA(The)h(w)o(ork)g([8,9])g(on)g(uni\002cation)g(in)g(Qu-Prolog)g
+(is)f(most)g(closely)h(related)g(to)g(that)g(reported)166
+2983 y(here.)32 b(Qu-Prolog)g(is)g(a)g(mature)g(logic)f(programming)f
+(language)i(addressing)f(man)o(y)g(prob-)166 3103 y(lems)g(we)i(set)f
+(out)f(in)h(the)g(Introduction.)e(T)-8 b(o)32 b(be)o(gin)f(with,)g
+(Qu-Prolog')-5 b(s)32 b(uni\002cation)f(algo-)166 3224
+y(rithm)j(uni\002es)g(terms)g(modulo)g Fz(\013)q FA(-equi)n(v)n(alence)
+f(and)i(may)f(produce)h(solutions)d(that,)i(as)h(in)166
+3344 y(nominal)18 b(uni\002cation,)h(depend)g(on)g(freshness)g
+(constraints)f(\(in)h(Qu-Prolog)g(such)g(constraints)166
+3465 y(are)28 b(represented)g(by)f(a)h(predicate)g(called)f
+Fs(not)p 1845 3465 30 4 v 35 w(free)p 2120 3465 V 35
+w(in)p FA(\).)g(Furthermore,)h(meta)n(v)n(ariables)166
+3585 y(are)k(substituted)d(in)i(a)g(possibly-capturing)e(manner)-5
+b(.)30 b(Ho)n(we)n(v)o(er)l(,)g(there)h(are)h(also)f(a)g(number)166
+3705 y(of)24 b(dif)n(ferences)g(between)f(nominal)g(uni\002cation)g
+(and)h(uni\002cation)f(in)g(Qu-Prolog.)g(The)h(most)166
+3826 y(ob)o(vious)29 b(dif)n(ference)j(is)f(that)g(the)g(term)g
+(language)g(in)g(Qu-Prolog)g(is)g(richer)h(than)f(our)g(term)166
+3946 y(language)22 b(o)o(v)o(er)f(nominal)f(signatures;)h(for)h(e)o
+(xample)f(Qu-Prolog)h(allo)n(ws)f(v)n(ariables)g(in)g(bind-)166
+4066 y(ing)26 b(position)e(and)i(permits)f(e)o(xplicit)g(substitutions)
+e(of)j(terms)g(for)g(v)n(ariables.)g(This)f(richness)166
+4187 y(of)19 b(the)h(term)f(language)g(leads)g(to)g(a)g(number)g(of)g
+(dif)n(\002culties.)f(First,)h(the)g(uni\002cation)g(problems)166
+4307 y(in)31 b(Qu-Prolog)g(are)h(only)f(semi-decidable)f(\(whereas)i
+(the)f(nominal)f(uni\002cation)h(problems)166 4428 y(are)j(decidable\))
+f(and)g(as)g(a)h(result)e(the)h(algorithm)f(emplo)o(yed)g(in)h
+(Qu-Prolog)f(can)i(lea)n(v)o(e)f(as)166 4548 y(unsolv)o(ed)28
+b(some)h(uni\002cation)g(problems)g(that)g(are)i(\223too)e(dif)n
+(\002cult\224.)g(This)g(means)g(the)h(uni\002-)166 4668
+y(cation)g(transformations)g(in)g(Qu-Prolog,)g(while)g(sho)n(wn)g(not)g
+(to)g(delete)h(an)o(y)f(solutions)f(nor)166 4789 y(to)d(introduce)h(an)
+o(y)f(ne)n(w)g(ones,)h(do)f(not)g(al)o(w)o(ays)h(lead)g(to)f(problems)g
+(from)g(which)g(an)h(e)o(xplicit)166 4909 y(solution)e(can)i(be)g
+(obtained.)f(Secondly)-6 b(,)26 b(as)h(we)g(illustrated)e(in)i(Remark)g
+(3.10,)f(the)h(possibil-)166 5029 y(ity)d(of)h(forming)e(terms)h(with)g
+(uni\002cation)g(v)n(ariables)g(in)g(binding)f(position)g(means)h(that)
+g(most)166 5150 y(general)h(solutions)e(may)h(not)h(e)o(xist.)166
+5380 y(Another)j(dif)n(ference)h(arises)f(from)g(the)h(f)o(act)f(that)g
+(in)g(Qu-Prolog)h(binders)e(are)i(renamed)g(via)1748
+5712 y(23)p eop end
+%%Page: 24 24
+TeXDict begin 24 23 bop 166 83 a FA(capture-a)n(v)n(oiding)26
+b(substitutions.)d(This)j(means)g(that)h(fresh)f(names)h(need)g(to)f
+(be)h(introduced)166 203 y(during)c(uni\002cation)g(in)h(order)g(to)g
+(respect)g Fz(\013)q FA(-equi)n(v)n(alence.)f(This)g(is)g(not)h
+(necessary)g(in)g(nomi-)166 324 y(nal)j(uni\002cation,)e(because)j(the)
+e(permutation)f(operation)h(already)h(respects)g Fz(\013)q
+FA(-equi)n(v)n(alence.)166 444 y(In)k(f)o(act)g(the)f(introduction)f
+(of)i(fresh)g(atoms)f(during)g(uni\002cation)g(leads)g(to)g(a)h(more)g
+(compli-)166 565 y(cated)g(notion)f(of)h(most)f(general)h(solution.)e
+(Consider)h(the)h(follo)n(wing)e(v)n(ariant)h(of)h(the)g(\()p
+Fw(\031)p Fu(?)p FA(-)166 685 y(abstraction-2\))24 b(transformation:)
+166 921 y Fu(\()p Fw(\031)p Fu(?)p FA(-abstraction-2)881
+878 y Ft(0)904 921 y Fu(\))k Fw(f)p Fz(a:t)g Fw(\031)p
+Fu(?)g Fz(a)1363 880 y Ft(0)1386 921 y Fz(:t)1448 880
+y Ft(0)1472 921 y Fw(g])p Fz(P)1756 868 y Fq(")1692 921
+y Fu(=)-17 b Fw(\))28 b(f)p Fu(\()p Fz(a)17 b(b)p Fu(\))p
+Fr(\001)p Fz(t)28 b Fw(\031)p Fu(?)g(\()p Fz(a)2449 880
+y Ft(0)2489 921 y Fz(b)p Fu(\))p Fr(\001)p Fz(t)2635
+880 y Ft(0)2659 921 y Fz(;)17 b(b)28 b Fu(#?)g Fz(t;)17
+b(b)28 b Fu(#?)g Fz(t)3265 880 y Ft(0)3288 921 y Fw(g[)p
+Fz(P)166 1157 y FA(which)d(is)g(applicable)g(pro)o(vided)f
+Fz(a)29 b Fw(6)p Fu(=)g Fz(a)1576 1121 y Ft(0)1625 1157
+y FA(and)d Fz(b)g FA(is)f(a)g(fresh)h(atom,)f(not)g(occurring)g(else)n
+(where)166 1277 y(in)39 b(the)g(problem.)g(This)f(rule)i(is)f
+(essentially)f(the)h(re\002nement)h(step)e(that)h(uni\002es)h(tw)o(o)e
+(ab-)166 1398 y(stracted)24 b(terms)f(in)g(Qu-Prolog)h(\(see)g([8,)g(P)
+o(age)g(105]\).)g(If)g(we)g(were)g(to)g(use)f(\()p Fw(\031)p
+Fu(?)p FA(-abstraction-)166 1518 y(2)216 1482 y Ft(0)239
+1518 y FA(\))31 b(instead)g(of)g(\()p Fw(\031)p Fu(?)p
+FA(-abstraction-2\))g(in)f(our)h(nominal)f(uni\002cation)g(algorithm,)f
+(then)i(when)166 1639 y(applied)24 b(to)h(the)f(problem)1486
+1775 y Fw(f)j Fz(a:X)36 b Fw(\031)q Fu(?)27 b Fz(b:Y)50
+b Fw(g)1154 b FA(\(19\))166 1941 y(it)29 b(w)o(ould)g(produce)h(the)g
+(solution)e Fu(\()p Fw(f)p Fz(a)37 b Fu(#)g Fz(Y)5 b(;)17
+b(c)37 b Fu(#)g Fz(Y)21 b Fw(g)p Fz(;)c Fu([)p Fz(X)45
+b Fu(:=)37 b(\()p Fz(a)17 b(c)p Fu(\)\()p Fz(b)g(c)p
+Fu(\))p Fr(\001)o Fz(Y)k Fu(]\))p FA(.)30 b(While)g(this)166
+2061 y(solution)e(solv)o(es)g(the)i(problem,)e(it)h(is)h(not)f(a)h
+(most)e(general)i(solution)e(according)h(to)g(De\002ni-)166
+2181 y(tion)h(3.1\227we)h(lost)g(the)g(information)f(that)g
+Fz(c)i FA(is)f(a)g(completely)g(fresh)g(atom.)g(On)g(the)g(other)166
+2302 y(hand,)21 b(applying)f(transformation)g(\()p Fw(\031)p
+Fu(?)p FA(-abstraction-2\))i(to)e(\(19\))i(leads)f(to)g
+Fu(\()p Fw(f)p Fz(a)27 b Fu(#)h Fz(Y)22 b Fw(g)p Fz(;)17
+b Fu([)p Fz(X)35 b Fu(:=)166 2422 y(\()p Fz(a)17 b(b)p
+Fu(\))p Fr(\001)p Fz(Y)k Fu(]\))p FA(\227a)k(most)f(general)h
+(solution.)166 2650 y(Ov)o(erall,)i(the)h(theory)g(of)g(Qu-Prolog')-5
+b(s)27 b(uni\002cation)h(is)f(more)h(comple)o(x)f(than)h(that)g(of)g
+(nomi-)166 2770 y(nal)22 b(uni\002cation:)g(in)g(nominal)f
+(uni\002cation)h(we)g(do)h(not)e(need)i(to)f(resort)h(to)f(a)g
+(semantic)g(notion)166 2891 y(of)28 b Fz(\013)q FA(-equi)n(v)n(alence)e
+(in)h(order)h(to)g(sho)n(w)e(the)i(correctness)g(of)f(the)h(nominal)e
+(uni\002cation)h(algo-)166 3011 y(rithm;)h(and)i(the)g(use)f(of)h
+(permutations)e(mak)o(es)i(our)f Fw(\031)p FA(-relation)h(much)f
+(simpler)g(compared)166 3132 y(with)f(Qu-Prolog')-5 b(s)28
+b(use)g(of)h(the)f(traditional)f(notion)g(of)i Fz(\013)q
+FA(-equi)n(v)n(alence)e(e)o(xtended)h(to)g(terms)166
+3252 y(with)c(meta)n(v)n(ariables.)166 3707 y FB(5)99
+b(Conclusion)25 b(and)h(Futur)n(e)g(W)-7 b(ork)166 4056
+y FA(In)19 b(this)f(paper)h(we)g(ha)n(v)o(e)g(proposed)f(a)h(ne)n(w)f
+(solution)f(to)i(the)g(problem)f(of)g(computing)g(possibly-)166
+4176 y(capturing)i(substitutions)d(that)j(unify)g(terms)g(in)l(v)n
+(olving)e(binders)i(up)g(to)g Fz(\013)q FA(-con)l(v)o(ersion.)f(T)-8
+b(o)20 b(do)166 4296 y(so)30 b(we)h(considered)f(a)h(man)o(y-sorted)e
+(\002rst-order)i(term)g(language)f(with)g(distinguished)e(col-)166
+4417 y(lections)h(of)h(constants)f(called)h Fx(atoms)f
+FA(and)h(with)f Fx(atom-abstr)o(action)e FA(operations)i(for)h(bind-)
+166 4537 y(ing)35 b(atoms)f(in)i(terms.)e(This)h(pro)o(vides)f(a)i
+(simple,)e(b)n(ut)h(\003e)o(xible,)g(frame)n(w)o(ork)g(for)h(specify-)
+166 4658 y(ing)28 b(binding)g(operations)g(and)h(their)f(scopes,)h(in)f
+(which)h(the)f(bound)g(entities)g(are)i(e)o(xplicitly)166
+4778 y(named.)37 b(By)i(using)d(v)n(ariables)h(pre\002x)o(ed)h(with)f
+(suspended)g(permutations,)f(one)i(can)g(ha)n(v)o(e)166
+4898 y(substitution)27 b(of)j(terms)f(for)h(v)n(ariables)f(both)g(allo)
+n(w)f(capture)i(of)g(atoms)f(by)g(binders)g(and)h(re-)166
+5019 y(spect)i Fz(\013)q FA(-equi)n(v)n(alence)f(\(renaming)h(of)h
+(bound)e(atoms\).)h(The)g(de\002nition)g(of)g Fz(\013)q
+FA(-equi)n(v)n(alence)166 5139 y(for)d(the)g(term)g(language)f(mak)o
+(es)h(use)g(of)g(an)g(auxiliary)f Fx(fr)l(eshness)g FA(relation)h
+(between)g(atoms)166 5259 y(and)i(terms)g(which)g(generalises)h(the)f
+(\223not)g(a)h(free)g(atom)f(of)5 b(\224)32 b(relation)f(from)h(ground)
+e(terms)166 5380 y(to)h(terms)f(with)h(v)n(ariables;)f(furthermore,)h
+(because)g(v)n(ariables)g(stand)f(for)i(unkno)n(wn)d(terms,)1748
+5712 y(24)p eop end
+%%Page: 25 25
+TeXDict begin 25 24 bop 166 3 3288 4 v 166 1030 4 1027
+v 227 127 a Fs(type)59 b(Gamma)g(\(var)g(X\))g(A)h(:-)f(mem)g(\(pair)g
+(X)g(A\))h(Gamma.)227 248 y(type)f(Gamma)g(\(app)g(M)g(N\))h(B)f(:-)h
+(type)f(Gamma)f(M)i(\(arrow)e(A)i(B\),)1782 368 y(type)f(Gamma)f(N)i
+(A.)227 489 y(type)f(Gamma)g(\(lam)g(x.M\))g(\(arrow)f(A)i(B\))f(/)h
+(x#Gamma)e(:-)1782 609 y(type)h(\(pair)f(x)i(A\)::Gamma)d(M)j(B.)227
+850 y(mem)g(A)f(A::Tail.)227 970 y(mem)h(A)f(B::Tail)f(:-)i(mem)f(A)g
+(Tail.)p 3450 1030 V 166 1033 3288 4 v 1134 1170 a FD(Fig.)22
+b(5.)h(An)g(e)o(xample)h FC(\013)p FD(Prolog)h(program)166
+1431 y FA(hence)c(with)g(unkno)n(wn)e(free)j(atoms,)e(it)g(is)g
+(necessary)i(to)e(mak)o(e)h(hypotheses)e(about)i(the)f(fresh-)166
+1551 y(ness)33 b(of)h(atoms)f(for)h(v)n(ariables)e(in)i(judgements)e
+(about)h(term)g(equi)n(v)n(alence)g(and)g(freshness.)166
+1671 y(This)e(reliance)h(on)f(\223freshness\224,)g(coupled)g(with)g
+(name-sw)o(apping)g(rather)g(than)h(renaming,)166 1792
+y(lead)39 b(to)f(a)g(ne)n(w)g(notion)f(of)i(uni\002cation)f(problem)f
+(in)h(which)g(instances)g(of)g(both)g(equi)n(v)n(a-)166
+1912 y(lence)31 b(and)f(freshness)g(ha)n(v)o(e)g(to)g(be)g(solv)o(ed)f
+(by)h(gi)n(ving)f(term-substitutions)e(and)j(\(possibly\))166
+2032 y(freshness)h(conditions)f(on)h(v)n(ariables)g(in)g(the)h
+(solution.)d(W)-8 b(e)32 b(sho)n(wed)f(that)g(this)g(uni\002cation)166
+2153 y(problem)24 b(is)g(decidable)h(and)g(unitary)-6
+b(.)166 2396 y(Chene)o(y)g(,)29 b(Gabbay)h(and)g(Urban)h([30,31])e(are)
+i(in)l(v)o(estigating)c(the)j(e)o(xtent)f(to)h(which)f(nominal)166
+2516 y(uni\002cation)k(can)h(be)g(used)f(in)g(resolution-based)g(proof)
+g(search)h(for)g(a)g(form)g(of)f(\002rst-order)166 2637
+y(logic)38 b(programming)f(for)i(languages)f(with)g(binders)f(\(with)h
+(a)h(vie)n(w)f(to)g(pro)o(viding)f(better)166 2757 y
+(machine-assistance)h(for)i(structural)e(operational)h(semantics\).)f
+(Such)h(a)h(logic)e(program-)166 2877 y(ming)28 b(language)g(should)g
+(permit)g(a)h(concrete,)h(\223nominal\224)e(approach)h(to)f(bound)g
+(entities)g(in)166 2998 y(programs)j(while)f(ensuring)g(that)h
+(computation)e(\(which)i(in)g(this)f(case)i(is)f(the)g(computation)166
+3118 y(of)38 b(answers)f(to)h(queries\))f(respects)h
+Fz(\013)q FA(-equi)n(v)n(alence)e(between)i(terms.)f(This)g(is)g
+(illustrated)166 3238 y(with)25 b(the)h(Prolog-lik)o(e)g(program)f(in)h
+(Figure)g(5,)g(which)g(implements)e(a)i(simple)f(typing)g(algo-)166
+3359 y(rithm)d(for)i Fz(\025)p FA(-terms.)f(The)g(third)f(clause)i(is)e
+(the)i(interesting)e(one.)h(First,)g(note)g(the)g(term)g
+Fs(\(lam)166 3479 y(x.M\))p FA(,)29 b(which)g(uni\002es)g(with)g(an)o
+(y)g Fz(\025)p FA(-abstraction.)g(The)h(binder)f Fs(x)p
+FA(,)g(roughly)g(speaking,)g(has)166 3600 y(in)24 b(the)h
+(\223nominal\224)f(approach)h(a)g(v)n(alue)g(which)f(can)h(be)g(used)g
+(in)f(the)h(body)f(of)h(the)g(clause,)g(for)166 3720
+y(e)o(xample)h(for)h(adding)f Fs(\(pair)59 b(x)h(A\))26
+b FA(to)h(the)f(conte)o(xt)g Fs(Gamma)p FA(.)g(Secondly)-6
+b(,)26 b(the)h(freshness)166 3840 y(constraint)h Fs(x)16
+b(#)h(Gamma)27 b FA(ensures)i(that)f Fs(Gamma)g FA(cannot)g(be)h
+(replaced)g(by)g(a)g(term)f(that)h(con-)166 3961 y(tains)f
+Fs(x)h FA(freely)-6 b(.)28 b(Since)h(this)f(clause)h(is)f(intended)g
+(to)h(implement)e(the)h(usual)g(rule)h(for)g(typing)166
+4081 y Fz(\025)p FA(-abstractions)1361 4211 y Fw(f)p
+Fz(x)f Fu(:)g Fz(A)p Fw(g)22 b([)h Fu(\000)49 b Fz(.)h(M)38
+b Fu(:)28 b Fz(B)p 1361 4256 898 4 v 1399 4342 a Fu(\000)50
+b Fz(.)g(\025x:M)39 b Fu(:)27 b Fz(A)h Fw(\033)g Fz(B)166
+4537 y FA(its)f(operational)h(beha)n(viour)f(is)h(gi)n(v)o(en)f(by:)g
+(choose)h(fresh)g(names)g(for)g Fs(Gamma)p FA(,)f Fs(x)p
+FA(,)h Fs(M)p FA(,)g Fs(A)g FA(and)166 4658 y Fs(B)g
+FA(\(this)f(is)h(standard)f(in)h(Prolog-lik)o(e)f(languages\),)h(unify)
+f(the)h(head)g(of)g(the)g(clause)g(with)f(the)166 4778
+y(goal)19 b(formula,)h(apply)f(the)g(resulting)g(uni\002er)h(to)f(the)h
+(body)f(of)h(the)f(clause)h(and)g(mak)o(e)f(sure)h(that)166
+4898 y Fs(Gamma)f FA(is)h(not)f(replaced)i(by)f(a)g(term)g(that)g
+(contains)f(freely)i(the)e(fresh)i(name)f(we)g(ha)n(v)o(e)g(chosen)166
+5019 y(for)31 b Fs(x)p FA(.)g(Similar)f(f)o(acilities)g(for)h
+Fx(functional)e(pr)l(o)o(gr)o(amming)f FA(already)j(e)o(xist)f(in)g
+(the)h(FreshML)166 5139 y(language,)38 b(b)n(uilt)f(upon)h(the)h(same)f
+(foundations:)f(see)h([13])h(and)f Fv(www)p Fz(:)p Fv(freshml)p
+Fz(:)p Fv(org)t FA(.)g(W)-8 b(e)166 5259 y(are)27 b(also)f(interested)g
+(in)g(the)g(special)g(case)h(of)f(\223nominal)g(matching\224)f(and)h
+(its)g(application)f(to)166 5380 y(term-re)n(writing)e(modulo)h
+Fz(\013)q FA(-equi)n(v)n(alence.)1748 5712 y(25)p eop
+end
+%%Page: 26 26
+TeXDict begin 26 25 bop 166 83 a Fx(A)25 b(note)f(on)h(comple)n(xity)
+166 459 y FA(If)i(these)f(applications)f(sho)n(w)g(that)g(nominal)g
+(uni\002cation)h(is)f(practically)h(useful,)g(then)g(it)f(be-)166
+579 y(comes)33 b(important)e(to)i(study)f(its)h(comple)o(xity)-6
+b(.)30 b(The)j(presentations)f(of)h(the)g(term)g(language)166
+700 y(in)26 b(Section)g(2)g(and)h(of)f(the)g(algorithm)f(in)h(Section)g
+(3)g(were)h(chosen)f(for)h(clarity)f(and)g(to)g(mak)o(e)166
+820 y(the)d(proof)g(of)h(correctness)1121 784 y FG(3)1201
+820 y FA(easier)l(,)g(rather)f(than)g(for)h(ef)n(\002cienc)o(y)-6
+b(.)22 b(One)i(source)f(of)h(increased)166 940 y(ef)n(\002cienc)o(y)j
+(is)g(to)h(delay)f(the)g(application)g(of)g(permutations:)f(instead)g
+(of)i(pushing)e(permuta-)166 1061 y(tion)k(inside)g(terms)g(until)f
+(the)o(y)h(reach)h(suspension)f(as)g(we)h(do)g(here,)g(one)f(should)g
+(just)g(push)166 1181 y(them)38 b(under)h(the)f(\002rst)h(constructor)f
+(\(pairing,)g(function)g(symbol)f(application,)h(or)g(atom-)166
+1301 y(abstraction\))25 b(in)h(order)g(to)f(proceed)i(with)e(the)g(ne)o
+(xt)g(step)h(of)g(decomposition.)d(Ho)n(we)n(v)o(er)l(,)i(the)166
+1422 y(main)30 b(inef)n(\002cienc)o(y)h(of)g(the)f(algorithm)g
+(presented)h(in)f(Section)h(3)g(comes)f(from)h(the)g(lack)g(of)166
+1542 y(sharing)24 b(in)h(terms)f(and)h(substitutions.)c(Thus)j(the)h
+(uni\002cation)f(problem)g(tak)o(en)h(from)f([32])450
+1833 y Fw(f)p Fz(f)11 b Fu(\()p Fz(X)678 1848 y FG(1)717
+1833 y Fz(;)17 b(X)842 1848 y FG(1)881 1833 y Fu(\))27
+b Fw(\031)q Fu(?)h Fz(X)1179 1848 y FG(2)1218 1833 y
+Fz(;)17 b(f)11 b Fu(\()p Fz(X)1440 1848 y FG(2)1479 1833
+y Fz(;)17 b(X)1604 1848 y FG(2)1643 1833 y Fu(\))28 b
+Fw(\031)p Fu(?)g Fz(X)1941 1848 y FG(3)1980 1833 y Fz(;)17
+b(:)g(:)g(:)f(;)h(f)11 b Fu(\()p Fz(X)2377 1848 y Fq(n)p
+Ft(\000)p FG(1)2513 1833 y Fz(;)17 b(X)2638 1848 y Fq(n)p
+Ft(\000)p FG(1)2775 1833 y Fu(\))28 b Fw(\031)p Fu(?)g
+Fz(X)3073 1848 y Fq(n)3120 1833 y Fw(g)166 2124 y FA(which)37
+b(illustrates)f(that)g(the)h(na)m(\250)-30 b(\021v)o(e)37
+b(algorithm)f(for)h(classical)g(\002rst-order)g(uni\002cation)g(has)166
+2244 y(e)o(xponential)27 b(time)g(comple)o(xity)-6 b(,)26
+b(also)i(applies)f(to)h(the)g(algorithm)f(for)i(nominal)e
+(uni\002cation)166 2364 y(gi)n(v)o(en)36 b(here.)i(If)g(one)g(adapts)f
+(a)h(representation)g(for)g(terms)f(using)f(techniques)h(de)n(v)o
+(eloped)166 2485 y(in)26 b([32])h(or)f([33],)h(which)f(are)h(based)f
+(on)g(directed)h(ac)o(yclic)f(graphs,)g(then)g(one)h(easily)f(arri)n(v)
+o(es)166 2605 y(at)c(an)g(algorithm)e(with)h(quadratic)h(time)f(comple)
+o(xity)-6 b(.)19 b(The)j(reason)g(for)g(the)g(quadratic,)f(rather)166
+2725 y(than)j(linear)l(,)h(time-comple)o(xity)d(is)i(that)g
+(permutations)f(need)i(to)f(be)h(applied)f(to)g(some)g(atoms)166
+2846 y(when)30 b(deciding)g(whether)g(the)g(rules)h(\()p
+Fw(\031)p Fu(?)p FA(-abstraction-1\))f(or)h(\()p Fw(\031)p
+Fu(?)p FA(-abstraction-2\))f(are)h(ap-)166 2966 y(plicable,)37
+b(and)g(these)g(permutations)e(\(represented)j(as)f(lists)f(of)h(sw)o
+(appings\))f(might)f(gro)n(w)166 3087 y(linearly)27 b(with)g(the)h
+(number)f(of)h(nodes.)g(Using)f(a)h(representation)f(of)h(permutations)
+e(that)i(al-)166 3207 y(lo)n(ws)36 b(for)i(a)g(more)g(ef)n(\002cient)f
+(calculation)g(of)h(their)f(action)g(on)h(atoms)e(does)i(not)f(impro)o
+(v)o(e)166 3327 y(the)c(quadratic)f(time)g(comple)o(xity)-6
+b(,)30 b(because)k(it)e(mak)o(es)g(the)h(operation)f(of)h(composing)e
+(tw)o(o)166 3448 y(permutations)h(become)h(linear)l(,)g(while)g(this)f
+(can)i(be)f(done)g(in)g(constant)f(time)h(when)g(using)166
+3568 y(the)i(list-of-sw)o(appings)e(representation.)i(F)o(or)g(higher)n
+(-order)g(patterns,)f(Qian)h(managed)g(to)166 3689 y(de)n(v)o(eloped)30
+b(a)i(uni\002cation)g(algorithm)e(with)h(linear)g(time-comple)o(xity)e
+([34].)j(It)g(seems)f(that)166 3809 y(adapting)22 b(Qian')-5
+b(s)23 b(algorithm)e(to)i(nominal)f(uni\002cation)g(via)h(an)g
+(encoding)f(of)i(nominal)d(terms)166 3929 y(into)h(higher)n(-order)h
+(patterns)g(as)g(discussed)f(in)h(Section)g(4)g(will)f(not)h(solv)o(e)f
+(this)g(problem.)g(F)o(or)166 4050 y(the)35 b(encoding)g(mak)o(es)g
+(the)g(resulting)f(higher)n(-order)i(patterns)f(quadratically)f(longer)
+h(than)166 4170 y(the)30 b(original)g(nominal)f(terms,)h(so)h(this)e
+(method)h(w)o(ould)f(only)h(pro)o(vide)g(another)g(algorithm)166
+4290 y(with)24 b(quadratic)h(time)f(comple)o(xity)-6
+b(.)166 4546 y(T)e(o)20 b(sum)g(up,)g(there)h(is)f(a)g(v)o(ersion)g(of)
+g(nominal)f(uni\002cation)h(with)g(quadratic)g(time)f(comple)o(xity)-6
+b(,)166 4666 y(b)n(ut)19 b(is)g(it)f(is)h(still)f(an)h(open)g(question)
+f(whether)i(a)f(v)o(ersion)g(can)g(be)h(de)n(v)o(eloped)d(with)i
+Fx(linear)i FA(time)166 4787 y(comple)o(xity)-6 b(.)p
+166 5169 299 4 v 166 5233 a FG(3)257 5266 y FD(See)47
+b Fi(http://www.cl.ca)o(m.)o(ac)o(.u)o(k/u)o(se)o(rs)o(/c)o(u2)o(00/)o
+(Un)o(if)o(ic)o(at)o(ion)40 b FD(for)49 b(the)f(Is-)166
+5379 y(abelle)25 b(proof)f(scripts.)1748 5712 y FA(26)p
+eop end
+%%Page: 27 27
+TeXDict begin 27 26 bop 166 83 a FB(Ackno)o(wledgements)166
+423 y FA(A)38 b(preliminary)g(v)o(ersion)f(of)h(this)g(paper)g
+(appeared)h(as)g([35].)f(W)-8 b(e)39 b(thank)f(James)f(Chene)o(y)-6
+b(,)166 544 y(Gilles)33 b(Do)n(wek,)h(Ro)o(y)g(Dyckhof)n(f,)f(Dale)i
+(Miller)l(,)e(Frank)i(Pfenning,)f(Francois)g(Pottier)g(and)166
+664 y(Helmut)28 b(Schwichtenber)n(g)i(for)g(comments)e(on)h(this)f(w)o
+(ork.)h(This)g(research)h(w)o(as)g(supported)166 785
+y(by)25 b(UK)f(EPSRC)j(grants)d(GR/R29697)h(\(Urban\))g(and)g
+(GR/R07615)f(\(Pitts)g(and)h(Gabbay\).)166 1168 y FB(Refer)n(ences)166
+1501 y FD([1])71 b(G.)24 b(Do)n(wek,)g(T)-7 b(.)24 b(Hardin,)h(C.)e
+(Kirchner)l(,)k(Higher)n(-order)h(uni\002cation)f(via)e(e)o(xplicit)h
+(substitutions,)342 1614 y(in:)g(10th)h(Annual)g(Symposium)f(on)g
+(Logic)g(in)g(Computer)h(Science,)f(IEEE)e(Computer)j(Society)342
+1727 y(Press,)d(W)-7 b(ashington,)25 b(1995,)g(pp.)e(366\226374.)166
+1911 y([2])71 b(C.)40 b(A.)g(Gunter)l(,)i(Semantics)g(of)f(Programming)
+i(Languages:)g(Structures)h(and)d(T)-6 b(echniques,)342
+2024 y(F)o(oundations)26 b(of)e(Computing,)g(MIT)f(Press,)g(1992.)166
+2209 y([3])71 b(D.)19 b(Miller)l(,)i(A)e(logic)i(programming)h
+(language)g(with)e(lambda-abstraction,)25 b(function)d(v)n(ariables,)
+342 2322 y(and)i(simple)h(uni\002cation,)g(Journal)g(of)f(Logic)g(and)g
+(Computation)h(1)e(\(1991\))i(497\226536.)166 2506 y([4])71
+b(G.)39 b(Do)n(wek,)f(Higher)n(-order)43 b(uni\002cation)e(and)f
+(matching,)h(in:)f(A.)e(Robinson,)j(A.)d(V)-12 b(oronk)o(o)o(v)342
+2619 y(\(Eds.\),)23 b(Handbook)j(of)d(Automated)i(Reasoning,)g(Else)n
+(vier)l(,)g(2001,)f(Ch.)f(16,)g(pp.)g(1009\2261062.)166
+2804 y([5])71 b(M.)61 b(Hamana,)h(A)f(logic)j(programming)g(language)h
+(based)e(on)f(binding)j(algebras,)f(in:)342 2917 y(N.)48
+b(K)m(obayashi,)k(B.)c(C.)g(Pierce)i(\(Eds.\),)f(Theoretical)i(Aspects)
+g(of)e(Computer)i(Softw)o(are,)342 3030 y(4th)d(International)j
+(Symposium,)c(T)-8 b(A)l(CS)45 b(2001,)j(Sendai,)f(Japan,)h(October)g
+(29-31,)g(2001,)342 3143 y(Proceedings,)h(V)-12 b(ol.)46
+b(2215)h(of)f(Lecture)h(Notes)f(in)g(Computer)h(Science,)g(Springer)n
+(-V)-10 b(erlag,)342 3256 y(Berlin,)24 b(2001,)g(pp.)g(243\226262.)166
+3440 y([6])71 b(M.)19 b(Hashimoto,)i(A.)e(Ohori,)h(A)f(typed)i(conte)o
+(xt)g(calculus,)h(Theoretical)h(Computer)d(Science)h(266)342
+3553 y(\(2001\))k(249\226271.)166 3738 y([7])71 b(M.)29
+b(Sato,)h(T)-7 b(.)28 b(Sakurai,)j(Y)-12 b(.)29 b(Kame)o(yama,)g(A)g
+(simply)i(typed)g(conte)o(xt)g(calculus)h(with)e(\002rst-class)342
+3851 y(en)l(vironments,)d(Journal)e(of)f(Functional)i(and)e(Logic)f
+(Programming)i(2002)g(\(4\).)166 4035 y([8])71 b(P)-10
+b(.)28 b(Nick)o(olas,)i(P)-10 b(.)28 b(J.)g(Robinson,)i(The)f
+(Qu-Prolog)h(uni\002cation)h(algorithm:)g(F)o(ormalisation)g(and)342
+4148 y(correctness,)c(Theoretical)e(computer)h(Science)e(169)g
+(\(1996\))h(81\226112.)166 4333 y([9])71 b(R.)19 b(P)o(aterson,)i
+(Uni\002cation)h(of)e(schemes)i(of)e(quanti\002ed)i(terms,)f(in:)f
+(Proc.)g(of)g(UNIF)f(1990,)i(1990,)342 4446 y(unpublished)28
+b(proceedings.)166 4631 y([10])e(M.)h(J.)h(Gabbay)-6
+b(,)28 b(A.)f(M.)g(Pitts,)h(A)f(ne)n(w)g(approach)k(to)d(abstract)i
+(syntax)f(with)f(v)n(ariable)i(binding,)342 4744 y(F)o(ormal)23
+b(Aspects)i(of)e(Computing)i(13)f(\(2002\))h(341\226363.)166
+4928 y([11])h(L.)e(Caires,)h(L.)e(Cardelli,)j(A)e(spatial)i(logic)g
+(for)f(concurrenc)o(y)j(\(part)e(II\),)f(in:)g(L.)e(Brim,)h(P)-10
+b(.)23 b(Jan)5 b(\013)-35 b(car)l(,)342 5041 y(M.)52
+b(K\013)-30 b(ret)m(\264)j(\021nsk)8 b(\264)-38 b(y,)54
+b(A.)d(K)o(u)5 b(\013)-35 b(cera)53 b(\(Eds.\),)g(CONCUR)c(2002)54
+b(\226)f(Concurrenc)o(y)i(Theory)-6 b(,)53 b(13th)342
+5154 y(International)37 b(Conference,)d(Brno,)e(Czech)h(Republic,)h
+(August)f(20-23,)g(2002.)h(Proceedings,)342 5267 y(V)-12
+b(ol.)29 b(2421)g(of)g(Lecture)h(Notes)f(in)f(Computer)i(Science,)f
+(Springer)n(-V)-10 b(erlag,)32 b(Berlin,)d(2002,)g(pp.)342
+5380 y(209\226225.)1748 5712 y FA(27)p eop end
+%%Page: 28 28
+TeXDict begin 28 27 bop 166 83 a FD([12])26 b(L.)32 b(Cardelli,)i(P)-10
+b(.)32 b(Gardner)l(,)i(G.)e(Ghelli,)h(Manipulating)j(trees)e(with)f
+(hidden)i(labels,)f(in:)g(A.)d(D.)342 196 y(Gordon)44
+b(\(Ed.\),)e(F)o(oundations)k(of)d(Softw)o(are)g(Science)h(and)f
+(Computation)i(Structures,)g(6th)342 309 y(International)34
+b(Conference,)e(FOSSA)l(CS)26 b(2003,)31 b(W)-7 b(arsa)o(w)h(,)29
+b(Poland.)h(Proceedings,)i(V)-12 b(ol.)30 b(2620)342
+422 y(of)24 b(Lecture)g(Notes)g(in)f(Computer)i(Science,)f(Springer)n
+(-V)-10 b(erlag,)26 b(Berlin,)e(2003,)g(pp.)g(216\226232.)166
+612 y([13])i(M.)37 b(R.)g(Shinwell,)h(A.)f(M.)g(Pitts,)h(M.)f(J.)g
+(Gabbay)-6 b(,)39 b(FreshML:)f(Programming)h(with)f(binders)342
+725 y(made)j(simple,)h(in:)f(Eighth)g(A)l(CM)e(SIGPLAN)f(International)
+45 b(Conference)e(on)e(Functional)342 838 y(Programming)25
+b(\(ICFP)d(2003\),)j(Uppsala,)f(Sweden,)f(A)l(CM)f(Press,)i(2003,)g
+(pp.)f(263\226274.)166 1028 y([14])j(A.)18 b(Salibra,)h(On)g(the)g
+(algebraic)i(models)f(of)f(lambda)h(calculus,)h(Theoretical)g(Computer)
+f(Science)342 1141 y(249)k(\(2000\))i(197\226240.)166
+1331 y([15])g(F)-7 b(.)34 b(Honsell,)i(M.)f(Miculan,)h(I.)f(Scagnetto,)
+i(An)d(axiomatic)k(approach)g(to)d(metareasoning)k(on)342
+1444 y(nominal)h(algebras)g(in)e(HO)m(AS,)d(in:)j(F)-7
+b(.)36 b(Orejas,)i(P)-10 b(.)37 b(G.)f(Spirakis,)j(J.)e(Leeuwen)i
+(\(Eds.\),)e(28th)342 1557 y(International)c(Colloquium)e(on)e
+(Automata,)g(Languages)i(and)f(Programming,)f(ICALP)e(2001,)342
+1670 y(Crete,)43 b(Greece,)g(July)g(2001.)g(Proceedings,)i(V)-12
+b(ol.)42 b(2076)i(of)e(Lecture)i(Notes)f(in)f(Computer)342
+1783 y(Science,)25 b(Springer)n(-V)-10 b(erlag,)26 b(Heidelber)n(g,)g
+(2001,)e(pp.)f(963\226978.)166 1973 y([16])j(A.)21 b(M.)g(Pitts,)h
+(Nominal)h(logic,)g(a)f(\002rst)g(order)h(theory)h(of)e(names)h(and)g
+(binding,)h(Information)g(and)342 2086 y(Computation)i(186)e(\(2003\))h
+(165\226193.)166 2277 y([17])h(M.)20 b(Sato,)g(T)-7 b(.)20
+b(Sakurai,)i(Y)-12 b(.)19 b(Kame)o(yama,)i(A.)e(Igarashi,)j(Calculi)g
+(of)f(meta-v)n(ariables,)i(in:)e(M.)f(Baaz)342 2390 y(\(Ed.\),)32
+b(Computer)i(Science)g(Logic)f(and)g(8th)g(K)o(urt)f(G)8
+b(\250)-38 b(odel)33 b(Colloquium)h(\(CSL)-8 b('03)32
+b(&)g(KGC\),)342 2502 y(V)-5 b(ienna,)45 b(Austria.)h(Proccedings,)h(V)
+-12 b(ol.)44 b(2803)i(of)e(Lecture)i(Notes)f(in)g(Computer)g(Science,)
+342 2615 y(Springer)n(-V)-10 b(erlag,)27 b(Berlin,)c(2003,)i(pp.)e
+(484\226497.)166 2806 y([18])j(A.)68 b(Martelli,)i(U.)e(Montanari,)j
+(An)d(ef)n(\002cient)i(uni\002cation)i(algorithm,)e(A)l(CM)e(T)m(rans.)
+342 2919 y(Programming)25 b(Languages)h(and)e(Systems)f(4)h(\(2\))f
+(\(1982\))i(258\226282.)166 3109 y([19])h(J.)20 b(W)-8
+b(.)18 b(Klop,)i(T)-6 b(erm)19 b(re)n(writing)i(systems,)g(in:)g(S.)d
+(Abramsk)o(y)-6 b(,)21 b(D.)d(M.)h(Gabbay)-6 b(,)21 b(T)-7
+b(.)19 b(S.)g(E.)f(Maibaum)342 3222 y(\(Eds.\),)27 b(Handbook)i(of)f
+(Logic)f(in)g(Computer)i(Science,)e(V)-12 b(olume)28
+b(2,)f(Oxford)h(Uni)n(v)o(erity)g(Press,)342 3335 y(1992,)d(pp.)e
+(1\226116.)166 3525 y([20])j(F)-7 b(.)36 b(Baader)l(,)i(T)-7
+b(.)36 b(Nipk)o(o)n(w)-6 b(,)38 b(T)-6 b(erm)36 b(Re)n(writing)i(and)g
+(All)f(That,)g(Cambridge)h(Uni)n(v)o(ersity)h(Press,)342
+3638 y(1998.)166 3828 y([21])26 b(H.)33 b(P)-10 b(.)31
+b(Barendre)o(gt,)36 b(The)d(Lambda)g(Calculus:)j(its)d(Syntax)h(and)h
+(Semantics,)f(North-Holland,)342 3941 y(1984.)166 4131
+y([22])26 b(R.)38 b(Milner)l(,)i(J.)f(P)o(arro)n(w)-6
+b(,)38 b(D.)g(W)-7 b(alk)o(er)l(,)40 b(A)e(calculus)j(of)e(mobile)h
+(processes)i(\(parts)f(I)d(and)i(II\),)342 4244 y(Information)26
+b(and)e(Computation)i(100)e(\(1992\))h(1\22677.)166 4435
+y([23])h(G.)k(D.)f(Plotkin,)j(An)e(illati)n(v)o(e)i(theory)h(of)d
+(relations,)j(in:)f(R.)d(Cooper)l(,)j(Mukai,)f(J.)f(Perry)h(\(Eds.\),)
+342 4548 y(Situation)46 b(Theory)f(and)g(its)g(Applications,)i(V)-12
+b(ol.)44 b(22)g(of)g(CSLI)f(Lecture)i(Notes,)g(Stanford)342
+4660 y(Uni)n(v)o(ersity)-6 b(,)25 b(1990,)f(pp.)g(133\226146.)166
+4851 y([24])i(M.)34 b(P)-10 b(.)34 b(Fiore,)g(G.)g(D.)g(Plotkin,)h(D.)f
+(T)l(uri,)g(Abstract)i(syntax)h(and)e(v)n(ariable)i(binding,)g(in:)e
+(14th)342 4964 y(Annual)f(Symposium)f(on)g(Logic)g(in)f(Computer)i
+(Science,)f(IEEE)e(Computer)j(Society)f(Press,)342 5077
+y(W)-7 b(ashington,)26 b(1999,)e(pp.)g(193\226202.)166
+5267 y([25])i(R.)44 b(Milner)l(,)j(M.)d(T)-7 b(ofte,)45
+b(R.)f(Harper)l(,)i(D.)e(MacQueen,)i(The)f(De\002nition)h(of)g
+(Standard)g(ML)342 5380 y(\(Re)n(vised\),)25 b(MIT)d(Press,)i(1997.)
+1748 5712 y FA(28)p eop end
+%%Page: 29 29
+TeXDict begin 29 28 bop 166 83 a FD([26])26 b(F)-7 b(.)49
+b(Pfenning,)i(C.)d(Elliott,)i(Higher)n(-order)j(abstract)f(syntax,)f
+(in:)g(Proc.)e(A)l(CM-SIGPLAN)342 196 y(Conference)d(on)d(Programming)h
+(Language)h(Design)e(and)h(Implementation,)h(A)l(CM)d(Press,)342
+309 y(1988,)25 b(pp.)e(199\226208.)166 495 y([27])j(S.)59
+b(Michaylo)o(v)-6 b(,)62 b(F)-7 b(.)58 b(Pfenning,)k(An)d(empirical)j
+(study)f(of)f(the)h(runtime)g(beha)n(viour)i(of)342 608
+y(higher)n(-order)44 b(logic)e(programs,)f(in:)g(D.)d(Miller)j
+(\(Ed.\),)f(Proc.)f(W)-7 b(orkshop)42 b(on)f(the)f FC(\025)p
+FD(Prolog)342 721 y(Programming)61 b(Language,)g(Uni)n(v)o(ersity)f(of)
+g(Pennsylv)n(ania,)i(1992,)e(pp.)f(257\226271,)i(CIS)342
+834 y(T)-6 b(echnical)25 b(Report)f(MS-CIS-92-86.)166
+1020 y([28])i(G.)i(Do)n(wek,)g(T)-7 b(.)27 b(Hardin,)i(C.)e(Kirchner)l
+(,)j(F)-7 b(.)27 b(Pfenning,)j(Higher)n(-order)i(uni\002cation)e(via)f
+(e)o(xplicit)342 1133 y(substitutions:)36 b(the)31 b(case)h(of)f
+(higher)n(-order)k(patterns,)d(in:)g(Proc.)e(of)h(JICSLP)-10
+b(,)29 b(1996,)j(pp.)f(259\226)342 1246 y(273.)166 1433
+y([29])26 b(M.)33 b(Hamana,)h(Simple)f FC(\014)1179 1447
+y FG(0)1219 1433 y FD(-uni\002cation)j(for)e(terms)g(with)g(conte)o(xt)
+h(holes,)g(in:)f(C.)e(Ringeissen,)342 1545 y(C.)d(T)m(inelli,)i(R.)e(T)
+m(reinen,)h(R.)f(M.)g(V)-10 b(erma)30 b(\(Eds.\),)g(Proc.)f(of)i(UNIF)d
+(2002,)j(2002,)g(unpublished)342 1658 y(proceedings.)166
+1845 y([30])26 b(J.)37 b(Chene)o(y)-6 b(,)37 b(C.)e(Urban,)j
+FC(\013)p FD(Prolog,)f(a)g(fresh)h(approach)h(to)e(logic)h(programming)
+h(modulo)f FC(\013)p FD(-)342 1958 y(equi)n(v)n(alence,)32
+b(in:)d(J.)g(Le)n(vy)-6 b(,)28 b(M.)g(K)m(ohlhase,)j(J.)d(Niehren,)i
+(M.)e(V)-5 b(illaret)29 b(\(Eds.\),)g(Proc.)g(of)g(UNIF)342
+2071 y(2003,)j(no.)e(DSIC-II/12/03)i(in)f(Departamento)i(de)d(Sistemas)
+h(Inform)5 b(\264)-35 b(aticos)34 b(y)c(Computaci)8 b(\264)-38
+b(on)342 2183 y(T)-6 b(echnical)25 b(Report)f(Series,)g(Uni)n(v)o
+(ersidad)h(Polit)5 b(\264)-35 b(ecnica)26 b(de)e(V)-10
+b(alencia,)24 b(2003,)g(pp.)g(15\22619.)166 2370 y([31])i(M.)41
+b(Gabbay)-6 b(,)42 b(J.)f(Chene)o(y)-6 b(,)42 b(A)f(proof)i(theory)g
+(for)f(nominal)h(logic,)f(in:)g(Nineteenth)i(Annual)342
+2483 y(IEEE)38 b(Symposium)j(on)e(Logic)h(in)g(Computer)h(Science,)f
+(IEEE)e(Computer)i(Society)h(Press,)342 2596 y(W)-7 b(ashington,)26
+b(2004.)166 2782 y([32])g(M.)40 b(S.)g(P)o(aterson,)h(M.)f(N.)g(W)-7
+b(e)o(gman,)40 b(Linear)i(uni\002cation,)g(Journal)h(of)e(Computer)h
+(System)342 2895 y(Sciences)25 b(16)f(\(2\))g(\(1978\))h(158\226167.)
+166 3081 y([33])h(A.)j(Martelli,)h(U.)f(Montanari,)i(An)e(ef)n
+(\002cient)i(uni\002cation)h(algorithm,)f(A)l(CM)d(T)m(ransactions)33
+b(on)342 3194 y(Programming)25 b(Languages)h(and)e(Systems)f(4)h
+(\(2\).)166 3381 y([34])i(Z.)20 b(Qian,)g(Uni\002cation)i(of)e(higher)n
+(-order)25 b(patterns)d(in)f(linear)h(time)e(and)h(space,)h(Journal)g
+(of)f(Logic)342 3494 y(and)j(Computation)i(6)d(\(3\))h(\(1996\))h
+(315\226341.)166 3680 y([35])h(C.)18 b(Urban,)i(A.)d(M.)i(Pitts,)f(M.)g
+(J.)h(Gabbay)-6 b(,)20 b(Nominal)g(uni\002cation,)h(in:)e(M.)f(Baaz)i
+(\(Ed.\),)e(Computer)342 3793 y(Science)34 b(Logic)f(and)g(8th)g(K)o
+(urt)g(G)8 b(\250)-38 b(odel)33 b(Colloquium)h(\(CSL)-8
+b('03)32 b(&)g(KGC\),)f(V)-5 b(ienna,)32 b(Austria.)342
+3906 y(Proceedings,)49 b(V)-12 b(ol.)46 b(2803)h(of)f(Lecture)h(Notes)f
+(in)g(Computer)h(Science,)g(Springer)n(-V)-10 b(erlag,)342
+4019 y(Berlin,)24 b(2003,)g(pp.)g(513\226527.)1748 5712
+y FA(29)p eop end
+%%Trailer
+
+userdict /end-hook known{end-hook}if
+%%EOF
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Unification/unification.ml Sun Apr 29 11:29:56 2012 +0100
@@ -0,0 +1,182 @@
+(**
+ * the Nominal Unification algorithm by Urban, Pitts & Gabbay
+ * June, 2003
+*)
+
+open List;;
+exception Fail;; (* exception for unification failure *)
+
+(* terms *)
+type term =
+ Unit (* units *)
+ | Atm of string (* atoms *)
+ | Susp of (string * string) list * string (* suspensions *)
+ | Fun of string * term (* function symbols *)
+ | Abst of string * term (* abstracted terms *)
+ | Pair of term * term;; (* pairs *)
+
+(* swapping operation on atoms *)
+let swap (a,b) c =
+ if a=c then b else if b=c then a else c;;
+let swaps pi c = fold_right swap pi c;;
+
+(* permutation on terms *)
+let rec perm pi t =
+ match t with
+ Unit -> Unit
+ | Atm(a) -> Atm(swaps pi a)
+ | Susp(pi',x) -> Susp(pi@pi',x)
+ | Fun(name,t') -> Fun(name, perm pi t')
+ | Abst(a,t') -> Abst(swaps pi a,perm pi t')
+ | Pair(t1,t2) -> Pair(perm pi t1,perm pi t2);;
+
+(* substitution operation
+ * - implements the simple operation of "hole filling" or "context substitution"
+ *)
+let rec subst t sigma =
+ match t with
+ Unit -> Unit
+ | Atm(a) -> Atm(a)
+ | Susp(pi,y) -> (try (perm pi (assoc y sigma)) with Not_found -> Susp(pi,y))
+ | Fun(name,t') -> Fun(name,subst t' sigma)
+ | Abst(a,t') -> Abst(a,subst t' sigma)
+ | Pair(t1,t2) -> Pair(subst t1 sigma,subst t2 sigma);;
+
+(* substitution composition *)
+let rec subst_compose sigma =
+ match sigma with
+ [] -> []
+ | h::tail -> h::(map (fun (v,t) -> (v,subst t [h])) (subst_compose tail));;
+
+(* occurs check *)
+let rec occurs x t =
+ match t with
+ Unit -> false
+ | Atm(a) -> false
+ | Susp(pi,y) -> if x=y then true else false
+ | Fun(_,t') | Abst(_,t') -> occurs x t'
+ | Pair(t1,t2) -> (occurs x t1) || (occurs x t2);;
+
+(* deletes duplicates from a list
+ * - used for calculating disagreement sets
+ *)
+let rec delete_dups l =
+ match l with
+ [] -> []
+ | h::t -> let t' = delete_dups t in if mem h t' then t' else (h::t');;
+
+(* disagreement set
+ * - takes two permutation (lists of pairs of atoms) as arguments
+ *)
+let rec ds pi pi' =
+ let (l1,l2) = split (pi@pi')
+ in filter (fun a -> (swaps pi a)!=(swaps pi' a)) (delete_dups (l1@l2))
+
+(* eliminates a solved equation *)
+let eliminate (v,t) (eprobs,fprobs) =
+ if occurs v t
+ then raise Fail
+ else (map (fun (t1,t2) -> (subst t1 [(v,t)],subst t2 [(v,t)])) eprobs,
+ map (fun (a,t') -> (a,subst t' [(v,t)])) fprobs);;
+
+(***************)
+(* unification *)
+(***************)
+
+(* checks and solves all freshness problems *)
+let rec check fprobs =
+ match fprobs with
+ [] -> []
+ | (a,Unit)::tail -> check tail
+ | (a,Atm(b))::tail -> if a=b then raise Fail else check tail
+ | (a,Susp(pi,x))::tail -> (swaps (rev pi) a,x)::(check tail)
+ | (a,Fun(_,t))::tail -> check ((a,t)::tail)
+ | (a,Abst(b,t))::tail -> if a=b then (check tail) else (check ((a,t)::tail))
+ | (a,Pair(t1,t2))::tail -> check ((a,t1)::(a,t2)::tail);;
+
+(* solves all equational problems *)
+let rec solve eprobs fprobs = solve_aux eprobs fprobs []
+and solve_aux eprobs fprobs sigma =
+ match eprobs with
+ [] -> (subst_compose sigma, delete_dups (check fprobs))
+ | (Unit,Unit)::tail -> solve_aux tail fprobs sigma
+ | (Atm(a),Atm(b))::tail -> if a=b then (solve_aux tail fprobs sigma) else raise Fail
+ | (Susp(pi,x),Susp(pi',x'))::tail when x=x' ->
+ let new_fps = map (fun a -> (a,Susp([],x))) (ds pi pi')
+ in solve_aux tail (new_fps @ fprobs) sigma
+ | (Susp(pi,x),t)::tail | (t,Susp(pi,x))::tail ->
+ let new_sigma = (x,perm (rev pi) t) in
+ let (new_eprobs,new_fprobs) = eliminate new_sigma (tail,fprobs)
+ in solve_aux new_eprobs new_fprobs (new_sigma::sigma)
+ | (Fun(n1,t1),Fun(n2,t2))::tail ->
+ if n1 = n2
+ then solve_aux ((t1,t2)::tail) fprobs sigma
+ else raise Fail
+ | (Abst(a1,t1),Abst(a2,t2))::tail ->
+ if a1 = a2
+ then solve_aux ((t1,t2)::tail) fprobs sigma
+ else solve_aux ((t1,perm [(a1,a2)] t2)::tail) ((a1,t2)::fprobs) sigma
+ | (Pair(t1,t2),Pair(s1,s2))::tail ->
+ solve_aux ((t1,s1)::(t2,s2)::tail) fprobs sigma
+ | _ -> raise Fail;;
+
+
+(************)
+(* Examples *)
+(************)
+
+(* a few variables*)
+let x = Susp([],"X")
+and y = Susp([],"Y")
+and z = Susp([],"Z");;
+
+(* lam a.(X a) =? lam b.(c b) --> [X:=c] *)
+let t1 = Abst("a",Pair(x,Atm("a")));;
+let t2 = Abst("b",Pair(Atm("c"),Atm("b")));;
+solve [(t1,t2)] [];;
+
+(* lam a.(X a) =? lam b.(a b) --> fails *)
+let t1 = Abst("a",Pair(x,Atm("a")));;
+let t2 = Abst("b",Pair(Atm("a"),Atm("b")));;
+solve [(t1,t2)] [];;
+
+(* lam a.(X a) =? lam b.(X b) --> a#X, b#X *)
+let t1 = Abst("a",Pair(x,Atm("a")));;
+let t2 = Abst("b",Pair(x,Atm("b")));;
+solve [(t1,t2)] [];;
+
+(* lam a.(X a) =? lam b.(Y b) --> [X:=(a b)oY] a#Y *)
+(* --> [Y:=(b a)oX] b#X *)
+let t1 = Abst("a",Pair(x,Atm("a")));;
+let t2 = Abst("b",Pair(y,Atm("b")));;
+solve [(t1,t2)] [];;
+solve [(t2,t1)] [];;
+
+(* quiz-questions from the paper *)
+let m1 = Susp([],"M1")
+and m2 = Susp([],"M2")
+and m3 = Susp([],"M3")
+and m4 = Susp([],"M4")
+and m5 = Susp([],"M5")
+and m6 = Susp([],"M6")
+and m7 = Susp([],"M7");;
+
+(* 1 --> fail *)
+let t1 = Abst("a",Abst("b",Pair(m1,Atm("b"))));;
+let t2 = Abst("b",Abst("a",Pair(Atm("a"),m1)));;
+solve [(t1,t2)] [];;
+
+(* 2 --> [M2:=b ,M3:=a] *)
+let t1 = Abst("a",Abst("b",Pair(m2,Atm("b"))));;
+let t2 = Abst("b",Abst("a",Pair(Atm("a"),m3)));;
+solve [(t1,t2)] [];;
+
+(* 3 --> [M4:=(a b)o M5] *)
+let t1 = Abst("a",Abst("b",Pair(Atm("b"),m4)));;
+let t2 = Abst("b",Abst("a",Pair(Atm("a"),m5)));;
+solve [(t1,t2)] [];;
+
+(* 4 --> [M6:=(b a)oM7] b#M7 *)
+let t1 = Abst("a",Abst("b",Pair(Atm("b"),m6)));;
+let t2 = Abst("a",Abst("a",Pair(Atm("a"),m7)));;
+solve [(t1,t2)] [];;