--- a/IsaMakefile Thu Jan 17 11:51:00 2013 +0000
+++ b/IsaMakefile Fri Jan 18 11:40:01 2013 +0000
@@ -32,6 +32,19 @@
$(ISABELLE_TOOL) document -o pdf generated
cp generated/root.pdf paper.pdf
+## ITP itp
+
+session_itp: Paper/ROOT.ML \
+ Paper/document/root* \
+ Paper/*.thy
+ @$(USEDIR) -D generated -f ROOT.ML HOL Paper
+
+itp: session_itp
+ rm -f Paper/generated/*.aux # otherwise latex will fall over
+ cd Paper/generated ; $(ISABELLE_TOOL) latex -o pdf root.tex
+ cd Paper/generated ; bibtex root
+ cd Paper/generated ; $(ISABELLE_TOOL) latex -o pdf root.tex
+ cp Paper/generated/root.pdf paper.pdf
## clean
Binary file Literature/wang-tiles.pdf has changed
--- a/Paper.thy Thu Jan 17 11:51:00 2013 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,684 +0,0 @@
-(*<*)
-theory Paper
-imports UTM
-begin
-
-hide_const (open) s
-
-abbreviation
- "update p a == new_tape a p"
-
-lemma fetch_def2:
- shows "fetch p 0 b == (Nop, 0)"
- and "fetch p (Suc s) Bk ==
- (case nth_of p (2 * s) of
- Some i \<Rightarrow> i
- | None \<Rightarrow> (Nop, 0))"
- and "fetch p (Suc s) Oc ==
- (case nth_of p ((2 * s) + 1) of
- Some i \<Rightarrow> i
- | None \<Rightarrow> (Nop, 0))"
-apply -
-apply(rule_tac [!] eq_reflection)
-by (auto split: block.splits simp add: fetch.simps)
-
-lemma new_tape_def2:
- shows "new_tape W0 (l, r) == (l, Bk#(tl r))"
- and "new_tape W1 (l, r) == (l, Oc#(tl r))"
- and "new_tape L (l, r) ==
- (if l = [] then ([], Bk#r) else (tl l, (hd l) # r))"
- and "new_tape R (l, r) ==
- (if r = [] then (Bk#l,[]) else ((hd r)#l, tl r))"
- and "new_tape Nop (l, r) == (l, r)"
-apply -
-apply(rule_tac [!] eq_reflection)
-apply(auto split: taction.splits simp add: new_tape.simps)
-done
-
-
-abbreviation
- "read r == if (r = []) then Bk else hd r"
-
-lemma tstep_def2:
- shows "tstep (s, l, r) p == (let (a, s') = fetch p s (read r) in (s', new_tape a (l, r)))"
-apply -
-apply(rule_tac [!] eq_reflection)
-by (auto split: if_splits prod.split list.split simp add: tstep.simps)
-
-abbreviation
- "run p inp out == \<exists>n. steps (1, inp) p n = (0, out)"
-
-lemma haltP_def2:
- "haltP p n = (\<exists>k l m.
- run p ([], exponent Oc n) (exponent Bk k, exponent Oc l @ exponent Bk m))"
-unfolding haltP_def
-apply(auto)
-done
-
-lemma tape_of_nat_list_def2:
- shows "tape_of_nat_list [] = []"
- and "tape_of_nat_list [n] = exponent Oc (n+1)"
- and "ns\<noteq> [] ==> tape_of_nat_list (n#ns) = (exponent Oc (n+1)) @ [Bk] @ (tape_of_nat_list ns)"
-apply(auto simp add: tape_of_nat_list.simps)
-apply(case_tac ns)
-apply(auto simp add: tape_of_nat_list.simps)
-done
-
-lemma tshift_def2:
- fixes p::"tprog"
- shows "tshift p n == (map (\<lambda> (a, s). (a, (if s = 0 then 0 else s + n))) p)"
-apply(rule eq_reflection)
-apply(auto simp add: tshift.simps)
-done
-
-lemma change_termi_state_def2:
- "change_termi_state p ==
- (map (\<lambda> (a, s). (a, if s = 0 then ((length p) div 2) + 1 else s)) p)"
-apply(rule eq_reflection)
-apply(auto simp add: change_termi_state.simps)
-done
-
-
-
-consts DUMMY::'a
-
-notation (latex output)
- Cons ("_::_" [78,77] 73) and
- set ("") and
- W0 ("W\<^bsub>\<^raw:\hspace{-2pt}>Bk\<^esub>") and
- W1 ("W\<^bsub>\<^raw:\hspace{-2pt}>Oc\<^esub>") and
- t_correct ("twf") and
- tstep ("step") and
- steps ("nsteps") and
- abc_lm_v ("lookup") and
- abc_lm_s ("set") and
- haltP ("stdhalt") and
- tshift ("shift") and
- tcopy ("copy") and
- change_termi_state ("adjust") and
- tape_of_nat_list ("\<ulcorner>_\<urcorner>") and
- t_add ("_ \<oplus> _") and
- DUMMY ("\<^raw:\mbox{$\_$}>")
-
-declare [[show_question_marks = false]]
-(*>*)
-
-section {* Introduction *}
-
-text {*
-
-\noindent
-We formalised in earlier work the correctness proofs for two
-algorithms in Isabelle/HOL---one about type-checking in
-LF~\cite{UrbanCheneyBerghofer11} and another about deciding requests
-in access control~\cite{WuZhangUrban12}. The formalisations
-uncovered a gap in the informal correctness proof of the former and
-made us realise that important details were left out in the informal
-model for the latter. However, in both cases we were unable to
-formalise in Isabelle/HOL computability arguments about the
-algorithms. The reason is that both algorithms are formulated in terms
-of inductive predicates. Suppose @{text "P"} stands for one such
-predicate. Decidability of @{text P} usually amounts to showing
-whether \mbox{@{term "P \<or> \<not>P"}} holds. But this does \emph{not} work
-in Isabelle/HOL, since it is a theorem prover based on classical logic
-where the law of excluded middle ensures that \mbox{@{term "P \<or> \<not>P"}}
-is always provable no matter whether @{text P} is constructed by
-computable means. The same problem would arise if we had formulated
-the algorithms as recursive functions, because internally in
-Isabelle/HOL, like in all HOL-based theorem provers, functions are
-represented as inductively defined predicates too.
-
-The only satisfying way out of this problem in a theorem prover based on classical
-logic is to formalise a theory of computability. Norrish provided such
-a formalisation for the HOL4 theorem prover. He choose the
-$\lambda$-calculus as the starting point for his formalisation
-of computability theory,
-because of its ``simplicity'' \cite[Page 297]{Norrish11}. Part of his
-formalisation is a clever infrastructure for reducing
-$\lambda$-terms. He also established the computational equivalence
-between the $\lambda$-calculus and recursive functions. Nevertheless he
-concluded that it would be ``appealing'' to have formalisations for more
-operational models of computations, such as Turing machines or register
-machines. One reason is that many proofs in the literature use
-them. He noted however that in the context of theorem provers
-\cite[Page 310]{Norrish11}:
-
-\begin{quote}
-\it``If register machines are unappealing because of their
-general fiddliness, Turing machines are an even more
-daunting prospect.''
-\end{quote}
-
-\noindent
-In this paper we take on this daunting prospect and provide a
-formalisation of Turing machines, as well as abacus machines (a kind
-of register machines) and recursive functions. To see the difficulties
-involved with this work, one has to understand that interactive
-theorem provers, like Isabelle/HOL, are at their best when the
-data-structures at hand are ``structurally'' defined, like lists,
-natural numbers, regular expressions, etc. Such data-structures come
-with convenient reasoning infrastructures (for example induction
-principles, recursion combinators and so on). But this is \emph{not}
-the case with Turing machines (and also not with register machines):
-underlying their definitions are sets of states together with
-transition functions, all of which are not structurally defined. This
-means we have to implement our own reasoning infrastructure in order
-to prove properties about them. This leads to annoyingly fiddly
-formalisations. We noticed first the difference between both,
-structural and non-structural, ``worlds'' when formalising the
-Myhill-Nerode theorem, where regular expressions fared much better
-than automata \cite{WuZhangUrban11}. However, with Turing machines
-there seems to be no alternative if one wants to formalise the great
-many proofs from the literature that use them. We will analyse one
-example---undecidability of Wang's tiling problem---in Section~\ref{Wang}. The
-standard proof of this property uses the notion of universal
-Turing machines.
-
-We are not the first who formalised Turing machines in a theorem
-prover: we are aware of the preliminary work by Asperti and Ricciotti
-\cite{AspertiRicciotti12}. They describe a complete formalisation of
-Turing machines in the Matita theorem prover, including a universal
-Turing machine. They report that the informal proofs from which they
-started are \emph{not} ``sufficiently accurate to be directly usable as a
-guideline for formalization'' \cite[Page 2]{AspertiRicciotti12}. For
-our formalisation we followed mainly the proofs from the textbook
-\cite{Boolos87} and found that the description there is quite
-detailed. Some details are left out however: for example, it is only
-shown how the universal Turing machine is constructed for Turing
-machines computing unary functions. We had to figure out a way to
-generalise this result to $n$-ary functions. Similarly, when compiling
-recursive functions to abacus machines, the textbook again only shows
-how it can be done for 2- and 3-ary functions, but in the
-formalisation we need arbitrary functions. But the general ideas for
-how to do this are clear enough in \cite{Boolos87}. However, one
-aspect that is completely left out from the informal description in
-\cite{Boolos87}, and similar ones we are aware of, is arguments why certain Turing
-machines are correct. We will introduce Hoare-style proof rules
-which help us with such correctness arguments of Turing machines.
-
-The main difference between our formalisation and the one by Asperti
-and Ricciotti is that their universal Turing machine uses a different
-alphabet than the machines it simulates. They write \cite[Page
-23]{AspertiRicciotti12}:
-
-\begin{quote}\it
-``In particular, the fact that the universal machine operates with a
-different alphabet with respect to the machines it simulates is
-annoying.''
-\end{quote}
-
-\noindent
-In this paper we follow the approach by Boolos et al \cite{Boolos87},
-which goes back to Post \cite{Post36}, where all Turing machines
-operate on tapes that contain only \emph{blank} or \emph{occupied} cells
-(represented by @{term Bk} and @{term Oc}, respectively, in our
-formalisation). Traditionally the content of a cell can be any
-character from a finite alphabet. Although computationally equivalent,
-the more restrictive notion of Turing machines in \cite{Boolos87} makes
-the reasoning more uniform. In addition some proofs \emph{about} Turing
-machines are simpler. The reason is that one often needs to encode
-Turing machines---consequently if the Turing machines are simpler, then the coding
-functions are simpler too. Unfortunately, the restrictiveness also makes
-it harder to design programs for these Turing machines. In order
-to construct a universal Turing machine we therefore do not follow
-\cite{AspertiRicciotti12}, instead follow the proof in
-\cite{Boolos87} by relating abacus machines to Turing machines and in
-turn recursive functions to abacus machines. The universal Turing
-machine can then be constructed as a recursive function.
-
-\smallskip
-\noindent
-{\bf Contributions:} We formalised in Isabelle/HOL Turing machines following the
-description of Boolos et al \cite{Boolos87} where tapes only have blank or
-occupied cells. We mechanise the undecidability of the halting problem and
-prove the correctness of concrete Turing machines that are needed
-in this proof; such correctness proofs are left out in the informal literature.
-We construct the universal Turing machine from \cite{Boolos87} by
-relating recursive functions to abacus machines and abacus machines to
-Turing machines. Since we have set up in Isabelle/HOL a very general computability
-model and undecidability result, we are able to formalise the
-undecidability of Wang's tiling problem. We are not aware of any other
-formalisation of a substantial undecidability problem.
-*}
-
-section {* Turing Machines *}
-
-text {* \noindent
- Turing machines can be thought of as having a read-write-unit, also
- referred to as \emph{head},
- ``gliding'' over a potentially infinite tape. Boolos et
- al~\cite{Boolos87} only consider tapes with cells being either blank
- or occupied, which we represent by a datatype having two
- constructors, namely @{text Bk} and @{text Oc}. One way to
- represent such tapes is to use a pair of lists, written @{term "(l,
- r)"}, where @{term l} stands for the tape on the left-hand side of the
- head and @{term r} for the tape on the right-hand side. We have the
- convention that the head, abbreviated @{term hd}, of the right-list is
- the cell on which the head of the Turing machine currently operates. This can
- be pictured as follows:
-
- \begin{center}
- \begin{tikzpicture}
- \draw[very thick] (-3.0,0) -- ( 3.0,0);
- \draw[very thick] (-3.0,0.5) -- ( 3.0,0.5);
- \draw[very thick] (-0.25,0) -- (-0.25,0.5);
- \draw[very thick] ( 0.25,0) -- ( 0.25,0.5);
- \draw[very thick] (-0.75,0) -- (-0.75,0.5);
- \draw[very thick] ( 0.75,0) -- ( 0.75,0.5);
- \draw[very thick] (-1.25,0) -- (-1.25,0.5);
- \draw[very thick] ( 1.25,0) -- ( 1.25,0.5);
- \draw[very thick] (-1.75,0) -- (-1.75,0.5);
- \draw[very thick] ( 1.75,0) -- ( 1.75,0.5);
- \draw[rounded corners=1mm] (-0.35,-0.1) rectangle (0.35,0.6);
- \draw[fill] (1.35,0.1) rectangle (1.65,0.4);
- \draw[fill] (0.85,0.1) rectangle (1.15,0.4);
- \draw[fill] (-0.35,0.1) rectangle (-0.65,0.4);
- \draw (-0.25,0.8) -- (-0.25,-0.8);
- \draw[<->] (-1.25,-0.7) -- (0.75,-0.7);
- \node [anchor=base] at (-0.8,-0.5) {\small left list};
- \node [anchor=base] at (0.35,-0.5) {\small right list};
- \node [anchor=base] at (0.1,0.7) {\small head};
- \node [anchor=base] at (-2.2,0.2) {\ldots};
- \node [anchor=base] at ( 2.3,0.2) {\ldots};
- \end{tikzpicture}
- \end{center}
-
- \noindent
- Note that by using lists each side of the tape is only finite. The
- potential infinity is achieved by adding an appropriate blank or occupied cell
- whenever the head goes over the ``edge'' of the tape. To
- make this formal we define five possible \emph{actions}
- the Turing machine can perform:
-
- \begin{center}
- \begin{tabular}{rcll}
- @{text "a"} & $::=$ & @{term "W0"} & write blank (@{term Bk})\\
- & $\mid$ & @{term "W1"} & write occupied (@{term Oc})\\
- & $\mid$ & @{term L} & move left\\
- & $\mid$ & @{term R} & move right\\
- & $\mid$ & @{term Nop} & do-nothing operation\\
- \end{tabular}
- \end{center}
-
- \noindent
- We slightly deviate
- from the presentation in \cite{Boolos87} by using the @{term Nop} operation; however its use
- will become important when we formalise halting computations and also universal Turing
- machines. Given a tape and an action, we can define the
- following tape updating function:
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
- @{thm (lhs) new_tape_def2(1)} & @{text "\<equiv>"} & @{thm (rhs) new_tape_def2(1)}\\
- @{thm (lhs) new_tape_def2(2)} & @{text "\<equiv>"} & @{thm (rhs) new_tape_def2(2)}\\
- @{thm (lhs) new_tape_def2(3)} & @{text "\<equiv>"} & \\
- \multicolumn{3}{l}{\hspace{1cm}@{thm (rhs) new_tape_def2(3)}}\\
- @{thm (lhs) new_tape_def2(4)} & @{text "\<equiv>"} & \\
- \multicolumn{3}{l}{\hspace{1cm}@{thm (rhs) new_tape_def2(4)}}\\
- @{thm (lhs) new_tape_def2(5)} & @{text "\<equiv>"} & @{thm (rhs) new_tape_def2(5)}\\
- \end{tabular}
- \end{center}
-
- \noindent
- The first two clauses replace the head of the right-list
- with a new @{term Bk} or @{term Oc}, respectively. To see that
- these two clauses make sense in case where @{text r} is the empty
- list, one has to know that the tail function, @{term tl}, is defined in
- Isabelle/HOL
- such that @{term "tl [] == []"} holds. The third clause
- implements the move of the head one step to the left: we need
- to test if the left-list @{term l} is empty; if yes, then we just prepend a
- blank cell to the right-list; otherwise we have to remove the
- head from the left-list and prepend it to the right-list. Similarly
- in the fourth clause for a right move action. The @{term Nop} operation
- leaves the the tape unchanged (last clause).
-
- Note that our treatment of the tape is rather ``unsymmetric''---we
- have the convention that the head of the right-list is where the
- head is currently positioned. Asperti and Ricciotti
- \cite{AspertiRicciotti12} also considered such a representation, but
- dismiss it as it complicates their definition for \emph{tape
- equality}. The reason is that moving the head one step to
- the left and then back to the right might change the tape (in case
- of going over the ``edge''). Therefore they distinguish four types
- of tapes: one where the tape is empty; another where the head
- is on the left edge, respectively right edge, and in the middle
- of the tape. The reading, writing and moving of the tape is then
- defined in terms of these four cases. In this way they can keep the
- tape in a ``normalised'' form, and thus making a left-move followed
- by a right-move being the identity on tapes. Since we are not using
- the notion of tape equality, we can get away with the unsymmetric
- definition above, and by using the @{term update} function
- cover uniformly all cases including corner cases.
-
- Next we need to define the \emph{states} of a Turing machine. Given
- how little is usually said about how to represent them in informal
- presentations, it might be surprising that in a theorem prover we
- have to select carefully a representation. If we use the naive
- representation where a Turing machine consists of a finite set of
- states, then we will have difficulties composing two Turing
- machines: we would need to combine two finite sets of states,
- possibly renaming states apart whenever both machines share
- states.\footnote{The usual disjoint union operation in Isabelle/HOL
- cannot be used as it does not preserve types.} This renaming can be
- quite cumbersome to reason about. Therefore we made the choice of
- representing a state by a natural number and the states of a Turing
- machine will always consist of the initial segment of natural
- numbers starting from @{text 0} up to the number of states of the
- machine. In doing so we can compose two Turing machine by
- shifting the states of one by an appropriate amount to a higher
- segment and adjusting some ``next states'' in the other.
-
- An \emph{instruction} @{term i} of a Turing machine is a pair consisting of
- an action and a natural number (the next state). A \emph{program} @{term p} of a Turing
- machine is then a list of such pairs. Using as an example the following Turing machine
- program, which consists of four instructions
-
- \begin{equation}
- \begin{tikzpicture}
- \node [anchor=base] at (0,0) {@{thm dither_def}};
- \node [anchor=west] at (-1.5,-0.42) {$\underbrace{\hspace{21mm}}_{\text{1st state}}$};
- \node [anchor=west] at ( 1.1,-0.42) {$\underbrace{\hspace{17mm}}_{\text{2nd state}}$};
- \node [anchor=west] at (-1.5,0.65) {$\overbrace{\hspace{10mm}}^{\text{@{term Bk}-case}}$};
- \node [anchor=west] at (-0.1,0.65) {$\overbrace{\hspace{6mm}}^{\text{@{term Oc}-case}}$};
- \end{tikzpicture}
- \label{dither}
- \end{equation}
-
- \noindent
- the reader can see we have organised our Turing machine programs so
- that segments of two belong to a state. The first component of the
- segment determines what action should be taken and which next state
- should be transitioned to in case the head reads a @{term Bk};
- similarly the second component determines what should be done in
- case of reading @{term Oc}. We have the convention that the first
- state is always the \emph{starting state} of the Turing machine.
- The zeroth state is special in that it will be used as the
- ``halting state''. There are no instructions for the @{text
- 0}-state, but it will always perform a @{term Nop}-operation and
- remain in the @{text 0}-state. Unlike Asperti and Riccioti
- \cite{AspertiRicciotti12}, we have chosen a very concrete
- representation for programs, because when constructing a universal
- Turing machine, we need to define a coding function for programs.
- This can be easily done for our programs-as-lists, but is more
- difficult for the functions used by Asperti and Ricciotti.
-
- Given a program @{term p}, a state
- and the cell being read by the head, we need to fetch
- the corresponding instruction from the program. For this we define
- the function @{term fetch}
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
- \multicolumn{3}{l}{@{thm fetch_def2(1)[where b=DUMMY]}}\\
- @{thm (lhs) fetch_def2(2)} & @{text "\<equiv>"} & \\
- \multicolumn{3}{@ {\hspace{1cm}}l}{@{text "case nth_of p (2 * s) of"}}\\
- \multicolumn{3}{@ {\hspace{1.4cm}}l}{@{text "None \<Rightarrow> (Nop, 0) |"}}\\
- \multicolumn{3}{@ {\hspace{1.4cm}}l}{@{text "Some i \<Rightarrow> i"}}\\
- @{thm (lhs) fetch_def2(3)} & @{text "\<equiv>"} & \\
- \multicolumn{3}{@ {\hspace{1cm}}l}{@{text "case nth_of p (2 * s + 1) of"}}\\
- \multicolumn{3}{@ {\hspace{1.4cm}}l}{@{text "None \<Rightarrow> (Nop, 0) |"}}\\
- \multicolumn{3}{@ {\hspace{1.4cm}}l}{@{text "Some i \<Rightarrow> i"}}
- \end{tabular}
- \end{center}
-
- \noindent
- In this definition the function @{term nth_of} returns the @{text n}th element
- from a list, provided it exists (@{term Some}-case), or if it does not, it
- returns the default action @{term Nop} and the default state @{text 0}
- (@{term None}-case). In doing so we slightly deviate from the description
- in \cite{Boolos87}: if their Turing machines transition to a non-existing
- state, then the computation is halted. We will transition in such cases
- to the @{text 0}-state. However, with introducing the
- notion of \emph{well-formed} Turing machine programs we will later exclude such
- cases and make the @{text 0}-state the only ``halting state''. A program
- @{term p} is said to be well-formed if it satisfies
- the following three properties:
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
- @{term "t_correct p"} & @{text "\<equiv>"} & @{term "2 <= length p"}\\
- & @{text "\<and>"} & @{term "iseven (length p)"}\\
- & @{text "\<and>"} & @{term "\<forall> (a, s) \<in> set p. s <= length p div 2"}
- \end{tabular}
- \end{center}
-
- \noindent
- The first says that @{text p} must have at least an instruction for the starting
- state; the second that @{text p} has a @{term Bk} and @{term Oc} instruction for every
- state, and the third that every next-state is one of the states mentioned in
- the program or being the @{text 0}-state.
-
- A \emph{configuration} @{term c} of a Turing machine is a state together with
- a tape. This is written as @{text "(s, (l, r))"}. If we have a
- configuration and a program, we can calculate
- what the next configuration is by fetching the appropriate action and next state
- from the program, and by updating the state and tape accordingly.
- This single step of execution is defined as the function @{term tstep}
-
- \begin{center}
- \begin{tabular}{l}
- @{text "step (s, (l, r)) p"} @{text "\<equiv>"}\\
- \hspace{10mm}@{text "let (a, s) = fetch p s (read r)"}\\
- \hspace{10mm}@{text "in (s', update (l, r) a)"}
- \end{tabular}
- \end{center}
-
- \noindent
- where @{term "read r"} returns the head of the list @{text r}, or if @{text r} is
- empty it returns @{term Bk}.
- It is impossible in Isabelle/HOL to lift the @{term step}-function realising
- a general evaluation function for Turing machines. The reason is that functions in HOL-based
- provers need to be terminating, and clearly there are Turing machine
- programs that are not. We can however define an evaluation
- function so that it performs exactly @{text n} steps:
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
- @{thm (lhs) steps.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) steps.simps(1)}\\
- @{thm (lhs) steps.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) steps.simps(2)}\\
- \end{tabular}
- \end{center}
-
- \noindent
- Recall our definition of @{term fetch} with the default value for
- the @{text 0}-state. In case a Turing program takes in \cite{Boolos87} less
- then @{text n} steps before it halts, then in our setting the @{term steps}-evaluation
- does not actually halt, but rather transitions to the @{text 0}-state and
- remains there performing @{text Nop}-actions until @{text n} is reached.
-
- Given some input tape @{text "(l\<^isub>i,r\<^isub>i)"}, we can define when a program
- @{term p} generates a specific output tape @{text "(l\<^isub>o,r\<^isub>o)"}
-
- \begin{center}
- \begin{tabular}{l}
- @{term "runs p (l\<^isub>i, r\<^isub>i) (l\<^isub>o,r\<^isub>o)"} @{text "\<equiv>"}\\
- \hspace{6mm}@{text "\<exists>n. nsteps (1, (l\<^isub>i,r\<^isub>i)) p n = (0, (l\<^isub>o,r\<^isub>o))"}
- \end{tabular}
- \end{center}
-
- \noindent
- where @{text 1} stands for the starting state and @{text 0} for our final state.
- A program @{text p} with input tape @{term "(l\<^isub>i, r\<^isub>i)"} \emph{halts} iff
-
- \begin{center}
- @{term "halts p (l\<^isub>i, r\<^isub>i) \<equiv>
- \<exists>l\<^isub>o r\<^isub>o. runs p (l\<^isub>i, r\<^isub>i) (l\<^isub>o,r\<^isub>o)"}
- \end{center}
-
- \noindent
- Later on we need to consider specific Turing machines that
- start with a tape in standard form and halt the computation
- in standard form. To define a tape in standard form, it is
- useful to have an operation @{term "tape_of_nat_list DUMMY"} that
- translates
- lists of natural numbers into tapes.
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
- @{thm (lhs) tape_of_nat_list_def2(1)} & @{text "\<equiv>"} & @{thm (rhs) tape_of_nat_list_def2(1)}\\
- @{thm (lhs) tape_of_nat_list_def2(2)} & @{text "\<equiv>"} & @{thm (rhs) tape_of_nat_list_def2(2)}\\
- @{thm (lhs) tape_of_nat_list_def2(3)} & @{text "\<equiv>"} & @{thm (rhs) tape_of_nat_list_def2(3)}\\
- \end{tabular}
- \end{center}
-
-
-
-
- By this we mean
-
- \begin{center}
- @{thm haltP_def2[where p="p" and n="n", THEN eq_reflection]}
- \end{center}
-
- \noindent
- This means the Turing machine starts with a tape containg @{text n} @{term Oc}s
- and the head pointing to the first one; the Turing machine
- halts with a tape consisting of some @{term Bk}s, followed by a
- ``cluster'' of @{term Oc}s and after that by some @{term Bk}s.
- The head in the output is pointing again at the first @{term Oc}.
- The intuitive meaning of this definition is to start the Turing machine with a
- tape corresponding to a value @{term n} and producing
- a new tape corresponding to the value @{term l} (the number of @{term Oc}s
- clustered on the output tape).
-
- Before we can prove the undecidability of the halting problem for Turing machines,
- we have to define how to compose two Turing machines. Given our setup, this is
- relatively straightforward, if slightly fiddly. We use the following two
- auxiliary functions:
-
- \begin{center}
- \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
- @{thm (lhs) tshift_def2} @{text "\<equiv>"}\\
- \hspace{4mm}@{thm (rhs) tshift_def2}\\
- @{thm (lhs) change_termi_state_def2} @{text "\<equiv>"}\\
- \hspace{4mm}@{text "map (\<lambda> (a, s)."}\\
- \hspace{14mm}@{text "(a, if s = 0 then length p div 2 + 1 else s)) p"}\\
- \end{tabular}
- \end{center}
-
- \noindent
- The first adds @{text n} to all states, exept the @{text 0}-state,
- thus moving all ``regular'' states to the segment starting at @{text
- n}; the second adds @{term "length p div 2 + 1"} to the @{text
- 0}-state, thus ridirecting all references to the ``halting state''
- to the first state after the program @{text p}. With these two
- functions in place, we can define the \emph{sequential composition}
- of two Turing machine programs @{text "p\<^isub>1"} and @{text "p\<^isub>2"}
-
- \begin{center}
- @{thm t_add.simps[where ?t1.0="p\<^isub>1" and ?t2.0="p\<^isub>2", THEN eq_reflection]}
- \end{center}
-
- \noindent
- This means @{text "p\<^isub>1"} is executed first. Whenever it originally
- transitioned to the @{text 0}-state, it will in the composed program transition to the starting
- state of @{text "p\<^isub>2"} instead. All the states of @{text "p\<^isub>2"}
- have been shifted in order to make sure that the states of the composed
- program @{text "p\<^isub>1 \<oplus> p\<^isub>2"} still only ``occupy''
- an initial segment of the natural numbers.
-
- \begin{center}
- \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}p{6.9cm}@ {}}
- @{thm (lhs) tcopy_def} & @{text "\<equiv>"} & @{thm (rhs) tcopy_def}
- \end{tabular}
- \end{center}
-
-
- assertion holds for all tapes
-
- Hoare rule for composition
-
- For showing the undecidability of the halting problem, we need to consider
- two specific Turing machines. copying TM and dithering TM
-
- correctness of the copying TM
-
- measure for the copying TM, which we however omit.
-
- halting problem
-*}
-
-section {* Abacus Machines *}
-
-text {*
- \noindent
- Boolos et al \cite{Boolos87} use abacus machines as a
- stepping stone for making it less laborious to write
- programs for Turing machines. Abacus machines operate
- over an unlimited number of registers $R_0$, $R_1$, \ldots
- each being able to hold an arbitrary large natural number.
- We use natural numbers to refer to registers, but also
- to refer to \emph{opcodes} of abacus
- machines. Obcodes are given by the datatype
-
- \begin{center}
- \begin{tabular}{rcll}
- @{text "o"} & $::=$ & @{term "Inc R\<iota>"} & increment register $R$ by one\\
- & $\mid$ & @{term "Dec R\<iota> o\<iota>"} & if content of $R$ is non-zero,\\
- & & & then decrement it by one\\
- & & & otherwise jump to opcode $o$\\
- & $\mid$ & @{term "Goto o\<iota>"} & jump to opcode $o$
- \end{tabular}
- \end{center}
-
- \noindent
- A \emph{program} of an abacus machine is a list of such
- obcodes. For example the program clearing the register
- $R$ (setting it to 0) can be defined as follows:
-
- \begin{center}
- @{thm clear.simps[where n="R\<iota>" and e="o\<iota>", THEN eq_reflection]}
- \end{center}
-
- \noindent
- The second opcode @{term "Goto 0"} in this programm means we
- jump back to the first opcode, namely @{text "Dec R o"}.
- The \emph{memory} $m$ of an abacus machine holding the values
- of the registers is represented as a list of natural numbers.
- We have a lookup function for this memory, written @{term "abc_lm_v m R\<iota>"},
- which looks up the content of register $R$; if $R$
- is not in this list, then we return 0. Similarly we
- have a setting function, written @{term "abc_lm_s m R\<iota> n"}, which
- sets the value of $R$ to $n$, and if $R$ was not yet in $m$
- it pads it approriately with 0s.
-
-
- Abacus machine halts when it jumps out of range.
-*}
-
-
-section {* Recursive Functions *}
-
-section {* Wang Tiles\label{Wang} *}
-
-text {*
- Used in texture mapings - graphics
-*}
-
-
-section {* Related Work *}
-
-text {*
- The most closely related work is by Norrish \cite{Norrish11}, and Asperti and
- Ricciotti \cite{AspertiRicciotti12}. Norrish bases his approach on
- lambda-terms. For this he introduced a clever rewriting technology
- based on combinators and de-Bruijn indices for
- rewriting modulo $\beta$-equivalence (to keep it manageable)
-*}
-
-
-(*
-Questions:
-
-Can this be done: Ackerman function is not primitive
-recursive (Nora Szasz)
-
-Tape is represented as two lists (finite - usually infinite tape)?
-
-*)
-
-
-(*<*)
-end
-(*>*)
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Paper/Paper.thy Fri Jan 18 11:40:01 2013 +0000
@@ -0,0 +1,610 @@
+(*<*)
+theory Paper
+imports "../thys/uncomputable"
+begin
+
+(*
+hide_const (open) s
+*)
+
+abbreviation
+ "update2 p a \<equiv> update a p"
+
+consts DUMMY::'a
+
+notation (latex output)
+ Cons ("_::_" [78,77] 73) and
+ set ("") and
+ W0 ("W\<^bsub>\<^raw:\hspace{-2pt}>Bk\<^esub>") and
+ W1 ("W\<^bsub>\<^raw:\hspace{-2pt}>Oc\<^esub>") and
+ update2 ("update") and
+(* abc_lm_v ("lookup") and
+ abc_lm_s ("set") and*)
+ haltP ("stdhalt") and
+ tcopy ("copy") and
+ tape_of_nat_list ("\<ulcorner>_\<urcorner>") and
+ tm_comp ("_ \<oplus> _") and
+ DUMMY ("\<^raw:\mbox{$\_$}>")
+
+declare [[show_question_marks = false]]
+(*>*)
+
+section {* Introduction *}
+
+text {*
+
+\noindent
+We formalised in earlier work the correctness proofs for two
+algorithms in Isabelle/HOL---one about type-checking in
+LF~\cite{UrbanCheneyBerghofer11} and another about deciding requests
+in access control~\cite{WuZhangUrban12}. The formalisations
+uncovered a gap in the informal correctness proof of the former and
+made us realise that important details were left out in the informal
+model for the latter. However, in both cases we were unable to
+formalise in Isabelle/HOL computability arguments about the
+algorithms. The reason is that both algorithms are formulated in terms
+of inductive predicates. Suppose @{text "P"} stands for one such
+predicate. Decidability of @{text P} usually amounts to showing
+whether \mbox{@{term "P \<or> \<not>P"}} holds. But this does \emph{not} work
+in Isabelle/HOL, since it is a theorem prover based on classical logic
+where the law of excluded middle ensures that \mbox{@{term "P \<or> \<not>P"}}
+is always provable no matter whether @{text P} is constructed by
+computable means. The same problem would arise if we had formulated
+the algorithms as recursive functions, because internally in
+Isabelle/HOL, like in all HOL-based theorem provers, functions are
+represented as inductively defined predicates too.
+
+The only satisfying way out of this problem in a theorem prover based on classical
+logic is to formalise a theory of computability. Norrish provided such
+a formalisation for the HOL4 theorem prover. He choose the
+$\lambda$-calculus as the starting point for his formalisation
+of computability theory,
+because of its ``simplicity'' \cite[Page 297]{Norrish11}. Part of his
+formalisation is a clever infrastructure for reducing
+$\lambda$-terms. He also established the computational equivalence
+between the $\lambda$-calculus and recursive functions. Nevertheless he
+concluded that it would be ``appealing'' to have formalisations for more
+operational models of computations, such as Turing machines or register
+machines. One reason is that many proofs in the literature use
+them. He noted however that in the context of theorem provers
+\cite[Page 310]{Norrish11}:
+
+\begin{quote}
+\it``If register machines are unappealing because of their
+general fiddliness, Turing machines are an even more
+daunting prospect.''
+\end{quote}
+
+\noindent
+In this paper we take on this daunting prospect and provide a
+formalisation of Turing machines, as well as abacus machines (a kind
+of register machines) and recursive functions. To see the difficulties
+involved with this work, one has to understand that interactive
+theorem provers, like Isabelle/HOL, are at their best when the
+data-structures at hand are ``structurally'' defined, like lists,
+natural numbers, regular expressions, etc. Such data-structures come
+with convenient reasoning infrastructures (for example induction
+principles, recursion combinators and so on). But this is \emph{not}
+the case with Turing machines (and also not with register machines):
+underlying their definitions are sets of states together with
+transition functions, all of which are not structurally defined. This
+means we have to implement our own reasoning infrastructure in order
+to prove properties about them. This leads to annoyingly fiddly
+formalisations. We noticed first the difference between both,
+structural and non-structural, ``worlds'' when formalising the
+Myhill-Nerode theorem, where regular expressions fared much better
+than automata \cite{WuZhangUrban11}. However, with Turing machines
+there seems to be no alternative if one wants to formalise the great
+many proofs from the literature that use them. We will analyse one
+example---undecidability of Wang's tiling problem---in Section~\ref{Wang}. The
+standard proof of this property uses the notion of universal
+Turing machines.
+
+We are not the first who formalised Turing machines in a theorem
+prover: we are aware of the preliminary work by Asperti and Ricciotti
+\cite{AspertiRicciotti12}. They describe a complete formalisation of
+Turing machines in the Matita theorem prover, including a universal
+Turing machine. They report that the informal proofs from which they
+started are \emph{not} ``sufficiently accurate to be directly usable as a
+guideline for formalization'' \cite[Page 2]{AspertiRicciotti12}. For
+our formalisation we followed mainly the proofs from the textbook
+\cite{Boolos87} and found that the description there is quite
+detailed. Some details are left out however: for example, it is only
+shown how the universal Turing machine is constructed for Turing
+machines computing unary functions. We had to figure out a way to
+generalise this result to $n$-ary functions. Similarly, when compiling
+recursive functions to abacus machines, the textbook again only shows
+how it can be done for 2- and 3-ary functions, but in the
+formalisation we need arbitrary functions. But the general ideas for
+how to do this are clear enough in \cite{Boolos87}. However, one
+aspect that is completely left out from the informal description in
+\cite{Boolos87}, and similar ones we are aware of, is arguments why certain Turing
+machines are correct. We will introduce Hoare-style proof rules
+which help us with such correctness arguments of Turing machines.
+
+The main difference between our formalisation and the one by Asperti
+and Ricciotti is that their universal Turing machine uses a different
+alphabet than the machines it simulates. They write \cite[Page
+23]{AspertiRicciotti12}:
+
+\begin{quote}\it
+``In particular, the fact that the universal machine operates with a
+different alphabet with respect to the machines it simulates is
+annoying.''
+\end{quote}
+
+\noindent
+In this paper we follow the approach by Boolos et al \cite{Boolos87},
+which goes back to Post \cite{Post36}, where all Turing machines
+operate on tapes that contain only \emph{blank} or \emph{occupied} cells
+(represented by @{term Bk} and @{term Oc}, respectively, in our
+formalisation). Traditionally the content of a cell can be any
+character from a finite alphabet. Although computationally equivalent,
+the more restrictive notion of Turing machines in \cite{Boolos87} makes
+the reasoning more uniform. In addition some proofs \emph{about} Turing
+machines are simpler. The reason is that one often needs to encode
+Turing machines---consequently if the Turing machines are simpler, then the coding
+functions are simpler too. Unfortunately, the restrictiveness also makes
+it harder to design programs for these Turing machines. In order
+to construct a universal Turing machine we therefore do not follow
+\cite{AspertiRicciotti12}, instead follow the proof in
+\cite{Boolos87} by relating abacus machines to Turing machines and in
+turn recursive functions to abacus machines. The universal Turing
+machine can then be constructed as a recursive function.
+
+\smallskip
+\noindent
+{\bf Contributions:} We formalised in Isabelle/HOL Turing machines following the
+description of Boolos et al \cite{Boolos87} where tapes only have blank or
+occupied cells. We mechanise the undecidability of the halting problem and
+prove the correctness of concrete Turing machines that are needed
+in this proof; such correctness proofs are left out in the informal literature.
+We construct the universal Turing machine from \cite{Boolos87} by
+relating recursive functions to abacus machines and abacus machines to
+Turing machines. Since we have set up in Isabelle/HOL a very general computability
+model and undecidability result, we are able to formalise the
+undecidability of Wang's tiling problem. We are not aware of any other
+formalisation of a substantial undecidability problem.
+*}
+
+section {* Turing Machines *}
+
+text {* \noindent
+ Turing machines can be thought of as having a read-write-unit, also
+ referred to as \emph{head},
+ ``gliding'' over a potentially infinite tape. Boolos et
+ al~\cite{Boolos87} only consider tapes with cells being either blank
+ or occupied, which we represent by a datatype having two
+ constructors, namely @{text Bk} and @{text Oc}. One way to
+ represent such tapes is to use a pair of lists, written @{term "(l,
+ r)"}, where @{term l} stands for the tape on the left-hand side of the
+ head and @{term r} for the tape on the right-hand side. We have the
+ convention that the head, abbreviated @{term hd}, of the right-list is
+ the cell on which the head of the Turing machine currently operates. This can
+ be pictured as follows:
+
+ \begin{center}
+ \begin{tikzpicture}
+ \draw[very thick] (-3.0,0) -- ( 3.0,0);
+ \draw[very thick] (-3.0,0.5) -- ( 3.0,0.5);
+ \draw[very thick] (-0.25,0) -- (-0.25,0.5);
+ \draw[very thick] ( 0.25,0) -- ( 0.25,0.5);
+ \draw[very thick] (-0.75,0) -- (-0.75,0.5);
+ \draw[very thick] ( 0.75,0) -- ( 0.75,0.5);
+ \draw[very thick] (-1.25,0) -- (-1.25,0.5);
+ \draw[very thick] ( 1.25,0) -- ( 1.25,0.5);
+ \draw[very thick] (-1.75,0) -- (-1.75,0.5);
+ \draw[very thick] ( 1.75,0) -- ( 1.75,0.5);
+ \draw[rounded corners=1mm] (-0.35,-0.1) rectangle (0.35,0.6);
+ \draw[fill] (1.35,0.1) rectangle (1.65,0.4);
+ \draw[fill] (0.85,0.1) rectangle (1.15,0.4);
+ \draw[fill] (-0.35,0.1) rectangle (-0.65,0.4);
+ \draw (-0.25,0.8) -- (-0.25,-0.8);
+ \draw[<->] (-1.25,-0.7) -- (0.75,-0.7);
+ \node [anchor=base] at (-0.8,-0.5) {\small left list};
+ \node [anchor=base] at (0.35,-0.5) {\small right list};
+ \node [anchor=base] at (0.1,0.7) {\small head};
+ \node [anchor=base] at (-2.2,0.2) {\ldots};
+ \node [anchor=base] at ( 2.3,0.2) {\ldots};
+ \end{tikzpicture}
+ \end{center}
+
+ \noindent
+ Note that by using lists each side of the tape is only finite. The
+ potential infinity is achieved by adding an appropriate blank or occupied cell
+ whenever the head goes over the ``edge'' of the tape. To
+ make this formal we define five possible \emph{actions}
+ the Turing machine can perform:
+
+ \begin{center}
+ \begin{tabular}{rcll}
+ @{text "a"} & $::=$ & @{term "W0"} & write blank (@{term Bk})\\
+ & $\mid$ & @{term "W1"} & write occupied (@{term Oc})\\
+ & $\mid$ & @{term L} & move left\\
+ & $\mid$ & @{term R} & move right\\
+ & $\mid$ & @{term Nop} & do-nothing operation\\
+ \end{tabular}
+ \end{center}
+
+ \noindent
+ We slightly deviate
+ from the presentation in \cite{Boolos87} by using the @{term Nop} operation; however its use
+ will become important when we formalise halting computations and also universal Turing
+ machines. Given a tape and an action, we can define the
+ following tape updating function:
+
+ \begin{center}
+ \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
+ @{thm (lhs) update.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) update.simps(1)}\\
+ @{thm (lhs) update.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) update.simps(2)}\\
+ @{thm (lhs) update.simps(3)} & @{text "\<equiv>"} & \\
+ \multicolumn{3}{l}{\hspace{1cm}@{thm (rhs) update.simps(3)}}\\
+ @{thm (lhs) update.simps(4)} & @{text "\<equiv>"} & \\
+ \multicolumn{3}{l}{\hspace{1cm}@{thm (rhs) update.simps(4)}}\\
+ @{thm (lhs) update.simps(5)} & @{text "\<equiv>"} & @{thm (rhs) update.simps(5)}\\
+ \end{tabular}
+ \end{center}
+
+ \noindent
+ The first two clauses replace the head of the right-list
+ with a new @{term Bk} or @{term Oc}, respectively. To see that
+ these two clauses make sense in case where @{text r} is the empty
+ list, one has to know that the tail function, @{term tl}, is defined in
+ Isabelle/HOL
+ such that @{term "tl [] == []"} holds. The third clause
+ implements the move of the head one step to the left: we need
+ to test if the left-list @{term l} is empty; if yes, then we just prepend a
+ blank cell to the right-list; otherwise we have to remove the
+ head from the left-list and prepend it to the right-list. Similarly
+ in the fourth clause for a right move action. The @{term Nop} operation
+ leaves the the tape unchanged (last clause).
+
+ Note that our treatment of the tape is rather ``unsymmetric''---we
+ have the convention that the head of the right-list is where the
+ head is currently positioned. Asperti and Ricciotti
+ \cite{AspertiRicciotti12} also considered such a representation, but
+ dismiss it as it complicates their definition for \emph{tape
+ equality}. The reason is that moving the head one step to
+ the left and then back to the right might change the tape (in case
+ of going over the ``edge''). Therefore they distinguish four types
+ of tapes: one where the tape is empty; another where the head
+ is on the left edge, respectively right edge, and in the middle
+ of the tape. The reading, writing and moving of the tape is then
+ defined in terms of these four cases. In this way they can keep the
+ tape in a ``normalised'' form, and thus making a left-move followed
+ by a right-move being the identity on tapes. Since we are not using
+ the notion of tape equality, we can get away with the unsymmetric
+ definition above, and by using the @{term update} function
+ cover uniformly all cases including corner cases.
+
+ Next we need to define the \emph{states} of a Turing machine. Given
+ how little is usually said about how to represent them in informal
+ presentations, it might be surprising that in a theorem prover we
+ have to select carefully a representation. If we use the naive
+ representation where a Turing machine consists of a finite set of
+ states, then we will have difficulties composing two Turing
+ machines: we would need to combine two finite sets of states,
+ possibly renaming states apart whenever both machines share
+ states.\footnote{The usual disjoint union operation in Isabelle/HOL
+ cannot be used as it does not preserve types.} This renaming can be
+ quite cumbersome to reason about. Therefore we made the choice of
+ representing a state by a natural number and the states of a Turing
+ machine will always consist of the initial segment of natural
+ numbers starting from @{text 0} up to the number of states of the
+ machine. In doing so we can compose two Turing machine by
+ shifting the states of one by an appropriate amount to a higher
+ segment and adjusting some ``next states'' in the other.
+
+ An \emph{instruction} @{term i} of a Turing machine is a pair consisting of
+ an action and a natural number (the next state). A \emph{program} @{term p} of a Turing
+ machine is then a list of such pairs. Using as an example the following Turing machine
+ program, which consists of four instructions
+
+ \begin{equation}
+ \begin{tikzpicture}
+ \node [anchor=base] at (0,0) {@{thm dither_def}};
+ \node [anchor=west] at (-1.5,-0.42) {$\underbrace{\hspace{21mm}}_{\text{1st state}}$};
+ \node [anchor=west] at ( 1.1,-0.42) {$\underbrace{\hspace{17mm}}_{\text{2nd state}}$};
+ \node [anchor=west] at (-1.5,0.65) {$\overbrace{\hspace{10mm}}^{\text{@{term Bk}-case}}$};
+ \node [anchor=west] at (-0.1,0.65) {$\overbrace{\hspace{6mm}}^{\text{@{term Oc}-case}}$};
+ \end{tikzpicture}
+ \label{dither}
+ \end{equation}
+
+ \noindent
+ the reader can see we have organised our Turing machine programs so
+ that segments of two belong to a state. The first component of the
+ segment determines what action should be taken and which next state
+ should be transitioned to in case the head reads a @{term Bk};
+ similarly the second component determines what should be done in
+ case of reading @{term Oc}. We have the convention that the first
+ state is always the \emph{starting state} of the Turing machine.
+ The zeroth state is special in that it will be used as the
+ ``halting state''. There are no instructions for the @{text
+ 0}-state, but it will always perform a @{term Nop}-operation and
+ remain in the @{text 0}-state. Unlike Asperti and Riccioti
+ \cite{AspertiRicciotti12}, we have chosen a very concrete
+ representation for programs, because when constructing a universal
+ Turing machine, we need to define a coding function for programs.
+ This can be easily done for our programs-as-lists, but is more
+ difficult for the functions used by Asperti and Ricciotti.
+
+ Given a program @{term p}, a state
+ and the cell being read by the head, we need to fetch
+ the corresponding instruction from the program. For this we define
+ the function @{term fetch}
+
+ \begin{center}
+ \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
+ \multicolumn{3}{l}{@{thm fetch.simps(1)[where b=DUMMY]}}\\
+ @{thm (lhs) fetch.simps(2)} & @{text "\<equiv>"} & \\
+ \multicolumn{3}{@ {\hspace{1cm}}l}{@{text "case nth_of p (2 * s) of"}}\\
+ \multicolumn{3}{@ {\hspace{1.4cm}}l}{@{text "None \<Rightarrow> (Nop, 0) |"}}\\
+ \multicolumn{3}{@ {\hspace{1.4cm}}l}{@{text "Some i \<Rightarrow> i"}}\\
+ @{thm (lhs) fetch.simps(3)} & @{text "\<equiv>"} & \\
+ \multicolumn{3}{@ {\hspace{1cm}}l}{@{text "case nth_of p (2 * s + 1) of"}}\\
+ \multicolumn{3}{@ {\hspace{1.4cm}}l}{@{text "None \<Rightarrow> (Nop, 0) |"}}\\
+ \multicolumn{3}{@ {\hspace{1.4cm}}l}{@{text "Some i \<Rightarrow> i"}}
+ \end{tabular}
+ \end{center}
+
+ \noindent
+ In this definition the function @{term nth_of} returns the @{text n}th element
+ from a list, provided it exists (@{term Some}-case), or if it does not, it
+ returns the default action @{term Nop} and the default state @{text 0}
+ (@{term None}-case). In doing so we slightly deviate from the description
+ in \cite{Boolos87}: if their Turing machines transition to a non-existing
+ state, then the computation is halted. We will transition in such cases
+ to the @{text 0}-state. However, with introducing the
+ notion of \emph{well-formed} Turing machine programs we will later exclude such
+ cases and make the @{text 0}-state the only ``halting state''. A program
+ @{term p} is said to be well-formed if it satisfies
+ the following three properties:
+
+ \begin{center}
+ \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
+ @{term "t_correct p"} & @{text "\<equiv>"} & @{term "2 <= length p"}\\
+ & @{text "\<and>"} & @{term "iseven (length p)"}\\
+ & @{text "\<and>"} & @{term "\<forall> (a, s) \<in> set p. s <= length p div 2"}
+ \end{tabular}
+ \end{center}
+
+ \noindent
+ The first says that @{text p} must have at least an instruction for the starting
+ state; the second that @{text p} has a @{term Bk} and @{term Oc} instruction for every
+ state, and the third that every next-state is one of the states mentioned in
+ the program or being the @{text 0}-state.
+
+ A \emph{configuration} @{term c} of a Turing machine is a state together with
+ a tape. This is written as @{text "(s, (l, r))"}. If we have a
+ configuration and a program, we can calculate
+ what the next configuration is by fetching the appropriate action and next state
+ from the program, and by updating the state and tape accordingly.
+ This single step of execution is defined as the function @{term tstep}
+
+ \begin{center}
+ \begin{tabular}{l}
+ @{text "step (s, (l, r)) p"} @{text "\<equiv>"}\\
+ \hspace{10mm}@{text "let (a, s) = fetch p s (read r)"}\\
+ \hspace{10mm}@{text "in (s', update (l, r) a)"}
+ \end{tabular}
+ \end{center}
+
+ \noindent
+ where @{term "read r"} returns the head of the list @{text r}, or if @{text r} is
+ empty it returns @{term Bk}.
+ It is impossible in Isabelle/HOL to lift the @{term step}-function realising
+ a general evaluation function for Turing machines. The reason is that functions in HOL-based
+ provers need to be terminating, and clearly there are Turing machine
+ programs that are not. We can however define an evaluation
+ function so that it performs exactly @{text n} steps:
+
+ \begin{center}
+ \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
+ @{thm (lhs) steps.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) steps.simps(1)}\\
+ @{thm (lhs) steps.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) steps.simps(2)}\\
+ \end{tabular}
+ \end{center}
+
+ \noindent
+ Recall our definition of @{term fetch} with the default value for
+ the @{text 0}-state. In case a Turing program takes in \cite{Boolos87} less
+ then @{text n} steps before it halts, then in our setting the @{term steps}-evaluation
+ does not actually halt, but rather transitions to the @{text 0}-state and
+ remains there performing @{text Nop}-actions until @{text n} is reached.
+
+ Given some input tape @{text "(l\<^isub>i,r\<^isub>i)"}, we can define when a program
+ @{term p} generates a specific output tape @{text "(l\<^isub>o,r\<^isub>o)"}
+
+ \begin{center}
+ \begin{tabular}{l}
+ @{term "runs p (l\<^isub>i, r\<^isub>i) (l\<^isub>o,r\<^isub>o)"} @{text "\<equiv>"}\\
+ \hspace{6mm}@{text "\<exists>n. nsteps (1, (l\<^isub>i,r\<^isub>i)) p n = (0, (l\<^isub>o,r\<^isub>o))"}
+ \end{tabular}
+ \end{center}
+
+ \noindent
+ where @{text 1} stands for the starting state and @{text 0} for our final state.
+ A program @{text p} with input tape @{term "(l\<^isub>i, r\<^isub>i)"} \emph{halts} iff
+
+ \begin{center}
+ @{term "halts p (l\<^isub>i, r\<^isub>i) \<equiv>
+ \<exists>l\<^isub>o r\<^isub>o. runs p (l\<^isub>i, r\<^isub>i) (l\<^isub>o,r\<^isub>o)"}
+ \end{center}
+
+ \noindent
+ Later on we need to consider specific Turing machines that
+ start with a tape in standard form and halt the computation
+ in standard form. To define a tape in standard form, it is
+ useful to have an operation %@{ term "tape_of_nat_list DUMMY"}
+ that translates lists of natural numbers into tapes.
+
+
+ \begin{center}
+ \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
+ %@ { thm (lhs) tape_of_nat_list_def2(1)} & @{text "\<equiv>"} & @ { thm (rhs) tape_of_nat_list_def2(1)}\\
+ %@ { thm (lhs) tape_of_nat_list_def2(2)} & @{text "\<equiv>"} & @ { thm (rhs) tape_of_nat_list_def2(2)}\\
+ %@ { thm (lhs) tape_of_nat_list_def2(3)} & @{text "\<equiv>"} & @ { thm (rhs) tape_of_nat_list_def2(3)}\\
+ \end{tabular}
+ \end{center}
+
+
+
+
+ By this we mean
+
+ \begin{center}
+ %@ {thm haltP_def2[where p="p" and n="n", THEN eq_reflection]}
+ \end{center}
+
+ \noindent
+ This means the Turing machine starts with a tape containg @{text n} @{term Oc}s
+ and the head pointing to the first one; the Turing machine
+ halts with a tape consisting of some @{term Bk}s, followed by a
+ ``cluster'' of @{term Oc}s and after that by some @{term Bk}s.
+ The head in the output is pointing again at the first @{term Oc}.
+ The intuitive meaning of this definition is to start the Turing machine with a
+ tape corresponding to a value @{term n} and producing
+ a new tape corresponding to the value @{term l} (the number of @{term Oc}s
+ clustered on the output tape).
+
+ Before we can prove the undecidability of the halting problem for Turing machines,
+ we have to define how to compose two Turing machines. Given our setup, this is
+ relatively straightforward, if slightly fiddly. We use the following two
+ auxiliary functions:
+
+ \begin{center}
+ \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
+ @{thm (lhs) shift.simps} @{text "\<equiv>"}\\
+ \hspace{4mm}@{thm (rhs) shift.simps}\\
+ @{thm (lhs) adjust.simps} @{text "\<equiv>"}\\
+ \hspace{4mm}@{text "map (\<lambda> (a, s)."}\\
+ \hspace{14mm}@{text "(a, if s = 0 then length p div 2 + 1 else s)) p"}\\
+ \end{tabular}
+ \end{center}
+
+ \noindent
+ The first adds @{text n} to all states, exept the @{text 0}-state,
+ thus moving all ``regular'' states to the segment starting at @{text
+ n}; the second adds @{term "length p div 2 + 1"} to the @{text
+ 0}-state, thus ridirecting all references to the ``halting state''
+ to the first state after the program @{text p}. With these two
+ functions in place, we can define the \emph{sequential composition}
+ of two Turing machine programs @{text "p\<^isub>1"} and @{text "p\<^isub>2"}
+
+ \begin{center}
+ @{thm tm_comp.simps[THEN eq_reflection]}
+ \end{center}
+
+ \noindent
+ This means @{text "p\<^isub>1"} is executed first. Whenever it originally
+ transitioned to the @{text 0}-state, it will in the composed program transition to the starting
+ state of @{text "p\<^isub>2"} instead. All the states of @{text "p\<^isub>2"}
+ have been shifted in order to make sure that the states of the composed
+ program @{text "p\<^isub>1 \<oplus> p\<^isub>2"} still only ``occupy''
+ an initial segment of the natural numbers.
+
+ \begin{center}
+ \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}p{6.9cm}@ {}}
+ @{thm (lhs) tcopy_def} & @{text "\<equiv>"} & @{thm (rhs) tcopy_def}
+ \end{tabular}
+ \end{center}
+
+
+ assertion holds for all tapes
+
+ Hoare rule for composition
+
+ For showing the undecidability of the halting problem, we need to consider
+ two specific Turing machines. copying TM and dithering TM
+
+ correctness of the copying TM
+
+ measure for the copying TM, which we however omit.
+
+ halting problem
+*}
+
+section {* Abacus Machines *}
+
+text {*
+ \noindent
+ Boolos et al \cite{Boolos87} use abacus machines as a
+ stepping stone for making it less laborious to write
+ programs for Turing machines. Abacus machines operate
+ over an unlimited number of registers $R_0$, $R_1$, \ldots
+ each being able to hold an arbitrary large natural number.
+ We use natural numbers to refer to registers, but also
+ to refer to \emph{opcodes} of abacus
+ machines. Obcodes are given by the datatype
+
+ \begin{center}
+ \begin{tabular}{rcll}
+ @{text "o"} & $::=$ & @{term "Inc R\<iota>"} & increment register $R$ by one\\
+ & $\mid$ & @{term "Dec R\<iota> o\<iota>"} & if content of $R$ is non-zero,\\
+ & & & then decrement it by one\\
+ & & & otherwise jump to opcode $o$\\
+ & $\mid$ & @{term "Goto o\<iota>"} & jump to opcode $o$
+ \end{tabular}
+ \end{center}
+
+ \noindent
+ A \emph{program} of an abacus machine is a list of such
+ obcodes. For example the program clearing the register
+ $R$ (setting it to 0) can be defined as follows:
+
+ \begin{center}
+ %@ {thm clear.simps[where n="R\<iota>" and e="o\<iota>", THEN eq_reflection]}
+ \end{center}
+
+ \noindent
+ The second opcode @{term "Goto 0"} in this programm means we
+ jump back to the first opcode, namely @{text "Dec R o"}.
+ The \emph{memory} $m$ of an abacus machine holding the values
+ of the registers is represented as a list of natural numbers.
+ We have a lookup function for this memory, written @{term "abc_lm_v m R\<iota>"},
+ which looks up the content of register $R$; if $R$
+ is not in this list, then we return 0. Similarly we
+ have a setting function, written @{term "abc_lm_s m R\<iota> n"}, which
+ sets the value of $R$ to $n$, and if $R$ was not yet in $m$
+ it pads it approriately with 0s.
+
+
+ Abacus machine halts when it jumps out of range.
+*}
+
+
+section {* Recursive Functions *}
+
+section {* Wang Tiles\label{Wang} *}
+
+text {*
+ Used in texture mapings - graphics
+*}
+
+
+section {* Related Work *}
+
+text {*
+ The most closely related work is by Norrish \cite{Norrish11}, and Asperti and
+ Ricciotti \cite{AspertiRicciotti12}. Norrish bases his approach on
+ lambda-terms. For this he introduced a clever rewriting technology
+ based on combinators and de-Bruijn indices for
+ rewriting modulo $\beta$-equivalence (to keep it manageable)
+*}
+
+
+(*
+Questions:
+
+Can this be done: Ackerman function is not primitive
+recursive (Nora Szasz)
+
+Tape is represented as two lists (finite - usually infinite tape)?
+
+*)
+
+
+(*<*)
+end
+(*>*)
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Paper/ROOT.ML Fri Jan 18 11:40:01 2013 +0000
@@ -0,0 +1,6 @@
+no_document
+use_thys ["../thys/turing_basic",
+ "../thys/uncomputable"(*,
+ "../thys/abacus"*)];
+
+use_thys ["Paper"]
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Paper/document/IEEEtran.cls Fri Jan 18 11:40:01 2013 +0000
@@ -0,0 +1,4722 @@
+%%
+%% IEEEtran.cls 2007/03/05 version V1.7a
+%%
+%%
+%% This is the official IEEE LaTeX class for authors of the Institute of
+%% Electrical and Electronics Engineers (IEEE) Transactions journals and
+%% conferences.
+%%
+%% Support sites:
+%% http://www.michaelshell.org/tex/ieeetran/
+%% http://www.ctan.org/tex-archive/macros/latex/contrib/IEEEtran/
+%% and
+%% http://www.ieee.org/
+%%
+%% Based on the original 1993 IEEEtran.cls, but with many bug fixes
+%% and enhancements (from both JVH and MDS) over the 1996/7 version.
+%%
+%%
+%% Contributors:
+%% Gerry Murray (1993), Silvano Balemi (1993),
+%% Jon Dixon (1996), Peter N"uchter (1996),
+%% Juergen von Hagen (2000), and Michael Shell (2001-2007)
+%%
+%%
+%% Copyright (c) 1993-2000 by Gerry Murray, Silvano Balemi,
+%% Jon Dixon, Peter N"uchter,
+%% Juergen von Hagen
+%% and
+%% Copyright (c) 2001-2007 by Michael Shell
+%%
+%% Current maintainer (V1.3 to V1.7): Michael Shell
+%% See:
+%% http://www.michaelshell.org/
+%% for current contact information.
+%%
+%% Special thanks to Peter Wilson (CUA) and Donald Arseneau
+%% for allowing the inclusion of the \@ifmtarg command
+%% from their ifmtarg LaTeX package.
+%%
+%%*************************************************************************
+%% Legal Notice:
+%% This code is offered as-is without any warranty either expressed or
+%% implied; without even the implied warranty of MERCHANTABILITY or
+%% FITNESS FOR A PARTICULAR PURPOSE!
+%% User assumes all risk.
+%% In no event shall IEEE or any contributor to this code be liable for
+%% any damages or losses, including, but not limited to, incidental,
+%% consequential, or any other damages, resulting from the use or misuse
+%% of any information contained here.
+%%
+%% All comments are the opinions of their respective authors and are not
+%% necessarily endorsed by the IEEE.
+%%
+%% This work is distributed under the LaTeX Project Public License (LPPL)
+%% ( http://www.latex-project.org/ ) version 1.3, and may be freely used,
+%% distributed and modified. A copy of the LPPL, version 1.3, is included
+%% in the base LaTeX documentation of all distributions of LaTeX released
+%% 2003/12/01 or later.
+%% Retain all contribution notices and credits.
+%% ** Modified files should be clearly indicated as such, including **
+%% ** renaming them and changing author support contact information. **
+%%
+%% File list of work: IEEEtran.cls, IEEEtran_HOWTO.pdf, bare_adv.tex,
+%% bare_conf.tex, bare_jrnl.tex, bare_jrnl_compsoc.tex
+%%
+%% Major changes to the user interface should be indicated by an
+%% increase in the version numbers. If a version is a beta, it will
+%% be indicated with a BETA suffix, i.e., 1.4 BETA.
+%% Small changes can be indicated by appending letters to the version
+%% such as "IEEEtran_v14a.cls".
+%% In all cases, \Providesclass, any \typeout messages to the user,
+%% \IEEEtransversionmajor and \IEEEtransversionminor must reflect the
+%% correct version information.
+%% The changes should also be documented via source comments.
+%%*************************************************************************
+%%
+%
+% Available class options
+% e.g., \documentclass[10pt,conference]{IEEEtran}
+%
+% *** choose only one from each category ***
+%
+% 9pt, 10pt, 11pt, 12pt
+% Sets normal font size. The default is 10pt.
+%
+% conference, journal, technote, peerreview, peerreviewca
+% determines format mode - conference papers, journal papers,
+% correspondence papers (technotes), or peer review papers. The user
+% should also select 9pt when using technote. peerreview is like
+% journal mode, but provides for a single-column "cover" title page for
+% anonymous peer review. The paper title (without the author names) is
+% repeated at the top of the page after the cover page. For peer review
+% papers, the \IEEEpeerreviewmaketitle command must be executed (will
+% automatically be ignored for non-peerreview modes) at the place the
+% cover page is to end, usually just after the abstract (keywords are
+% not normally used with peer review papers). peerreviewca is like
+% peerreview, but allows the author names to be entered and formatted
+% as with conference mode so that author affiliation and contact
+% information can be easily seen on the cover page.
+% The default is journal.
+%
+% draft, draftcls, draftclsnofoot, final
+% determines if paper is formatted as a widely spaced draft (for
+% handwritten editor comments) or as a properly typeset final version.
+% draftcls restricts draft mode to the class file while all other LaTeX
+% packages (i.e., \usepackage{graphicx}) will behave as final - allows
+% for a draft paper with visible figures, etc. draftclsnofoot is like
+% draftcls, but does not display the date and the word "DRAFT" at the foot
+% of the pages. If using one of the draft modes, the user will probably
+% also want to select onecolumn.
+% The default is final.
+%
+% letterpaper, a4paper
+% determines paper size: 8.5in X 11in or 210mm X 297mm. CHANGING THE PAPER
+% SIZE WILL NOT ALTER THE TYPESETTING OF THE DOCUMENT - ONLY THE MARGINS
+% WILL BE AFFECTED. In particular, documents using the a4paper option will
+% have reduced side margins (A4 is narrower than US letter) and a longer
+% bottom margin (A4 is longer than US letter). For both cases, the top
+% margins will be the same and the text will be horizontally centered.
+% For final submission to IEEE, authors should use US letter (8.5 X 11in)
+% paper. Note that authors should ensure that all post-processing
+% (ps, pdf, etc.) uses the same paper specificiation as the .tex document.
+% Problems here are by far the number one reason for incorrect margins.
+% IEEEtran will automatically set the default paper size under pdflatex
+% (without requiring a change to pdftex.cfg), so this issue is more
+% important to dvips users. Fix config.ps, config.pdf, or ~/.dvipsrc for
+% dvips, or use the dvips -t papersize option instead as needed. See the
+% testflow documentation
+% http://www.ctan.org/tex-archive/macros/latex/contrib/IEEEtran/testflow
+% for more details on dvips paper size configuration.
+% The default is letterpaper.
+%
+% oneside, twoside
+% determines if layout follows single sided or two sided (duplex)
+% printing. The only notable change is with the headings at the top of
+% the pages.
+% The default is oneside.
+%
+% onecolumn, twocolumn
+% determines if text is organized into one or two columns per page. One
+% column mode is usually used only with draft papers.
+% The default is twocolumn.
+%
+% compsoc
+% Use the format of the IEEE Computer Society.
+%
+% compsocconf
+% Use the format of IEEE Computer Society conferencs (CPS)
+%
+% romanappendices
+% Use the "Appendix I" convention when numbering appendices. IEEEtran.cls
+% now defaults to Alpha "Appendix A" convention - the opposite of what
+% v1.6b and earlier did.
+%
+% captionsoff
+% disables the display of the figure/table captions. Some IEEE journals
+% request that captions be removed and figures/tables be put on pages
+% of their own at the end of an initial paper submission. The endfloat
+% package can be used with this class option to achieve this format.
+%
+% nofonttune
+% turns off tuning of the font interword spacing. Maybe useful to those
+% not using the standard Times fonts or for those who have already "tuned"
+% their fonts.
+% The default is to enable IEEEtran to tune font parameters.
+%
+%
+%----------
+% Available CLASSINPUTs provided (all are macros unless otherwise noted):
+% \CLASSINPUTbaselinestretch
+% \CLASSINPUTinnersidemargin
+% \CLASSINPUToutersidemargin
+% \CLASSINPUTtoptextmargin
+% \CLASSINPUTbottomtextmargin
+%
+% Available CLASSINFOs provided:
+% \ifCLASSINFOpdf (TeX if conditional)
+% \CLASSINFOpaperwidth (macro)
+% \CLASSINFOpaperheight (macro)
+% \CLASSINFOnormalsizebaselineskip (length)
+% \CLASSINFOnormalsizeunitybaselineskip (length)
+%
+% Available CLASSOPTIONs provided:
+% all class option flags (TeX if conditionals) unless otherwise noted,
+% e.g., \ifCLASSOPTIONcaptionsoff
+% point size options provided as a single macro:
+% \CLASSOPTIONpt
+% which will be defined as 9, 10, 11, or 12 depending on the document's
+% normalsize point size.
+% also, class option peerreviewca implies the use of class option peerreview
+% and classoption draft implies the use of class option draftcls
+
+
+
+
+
+\ProvidesClass{IEEEtran}[2007/03/05 V1.7a by Michael Shell]
+\typeout{-- See the "IEEEtran_HOWTO" manual for usage information.}
+\typeout{-- http://www.michaelshell.org/tex/ieeetran/}
+\NeedsTeXFormat{LaTeX2e}
+
+% IEEEtran.cls version numbers, provided as of V1.3
+% These values serve as a way a .tex file can
+% determine if the new features are provided.
+% The version number of this IEEEtrans.cls can be obtained from
+% these values. i.e., V1.4
+% KEEP THESE AS INTEGERS! i.e., NO {4a} or anything like that-
+% (no need to enumerate "a" minor changes here)
+\def\IEEEtransversionmajor{1}
+\def\IEEEtransversionminor{7}
+
+% These do nothing, but provide them like in article.cls
+\newif\if@restonecol
+\newif\if@titlepage
+
+
+% class option conditionals
+\newif\ifCLASSOPTIONonecolumn \CLASSOPTIONonecolumnfalse
+\newif\ifCLASSOPTIONtwocolumn \CLASSOPTIONtwocolumntrue
+
+\newif\ifCLASSOPTIONoneside \CLASSOPTIONonesidetrue
+\newif\ifCLASSOPTIONtwoside \CLASSOPTIONtwosidefalse
+
+\newif\ifCLASSOPTIONfinal \CLASSOPTIONfinaltrue
+\newif\ifCLASSOPTIONdraft \CLASSOPTIONdraftfalse
+\newif\ifCLASSOPTIONdraftcls \CLASSOPTIONdraftclsfalse
+\newif\ifCLASSOPTIONdraftclsnofoot \CLASSOPTIONdraftclsnofootfalse
+
+\newif\ifCLASSOPTIONpeerreview \CLASSOPTIONpeerreviewfalse
+\newif\ifCLASSOPTIONpeerreviewca \CLASSOPTIONpeerreviewcafalse
+
+\newif\ifCLASSOPTIONjournal \CLASSOPTIONjournaltrue
+\newif\ifCLASSOPTIONconference \CLASSOPTIONconferencefalse
+\newif\ifCLASSOPTIONtechnote \CLASSOPTIONtechnotefalse
+
+\newif\ifCLASSOPTIONnofonttune \CLASSOPTIONnofonttunefalse
+
+\newif\ifCLASSOPTIONcaptionsoff \CLASSOPTIONcaptionsofffalse
+
+\newif\ifCLASSOPTIONcompsoc \CLASSOPTIONcompsocfalse
+
+\newif\ifCLASSOPTIONcompsocconf \CLASSOPTIONcompsocconffalse
+
+\newif\ifCLASSOPTIONromanappendices \CLASSOPTIONromanappendicesfalse
+
+
+% class info conditionals
+
+% indicates if pdf (via pdflatex) output
+\newif\ifCLASSINFOpdf \CLASSINFOpdffalse
+
+
+% V1.6b internal flag to show if using a4paper
+\newif\if@IEEEusingAfourpaper \@IEEEusingAfourpaperfalse
+
+
+
+% IEEEtran class scratch pad registers
+% dimen
+\newdimen\@IEEEtrantmpdimenA
+\newdimen\@IEEEtrantmpdimenB
+% count
+\newcount\@IEEEtrantmpcountA
+\newcount\@IEEEtrantmpcountB
+% token list
+\newtoks\@IEEEtrantmptoksA
+
+% we use \CLASSOPTIONpt so that we can ID the point size (even for 9pt docs)
+% as well as LaTeX's \@ptsize to retain some compatability with some
+% external packages
+\def\@ptsize{0}
+% LaTeX does not support 9pt, so we set \@ptsize to 0 - same as that of 10pt
+\DeclareOption{9pt}{\def\CLASSOPTIONpt{9}\def\@ptsize{0}}
+\DeclareOption{10pt}{\def\CLASSOPTIONpt{10}\def\@ptsize{0}}
+\DeclareOption{11pt}{\def\CLASSOPTIONpt{11}\def\@ptsize{1}}
+\DeclareOption{12pt}{\def\CLASSOPTIONpt{12}\def\@ptsize{2}}
+
+
+
+\DeclareOption{letterpaper}{\setlength{\paperheight}{11in}%
+ \setlength{\paperwidth}{8.5in}%
+ \@IEEEusingAfourpaperfalse
+ \def\CLASSOPTIONpaper{letter}%
+ \def\CLASSINFOpaperwidth{8.5in}%
+ \def\CLASSINFOpaperheight{11in}}
+
+
+\DeclareOption{a4paper}{\setlength{\paperheight}{297mm}%
+ \setlength{\paperwidth}{210mm}%
+ \@IEEEusingAfourpapertrue
+ \def\CLASSOPTIONpaper{a4}%
+ \def\CLASSINFOpaperwidth{210mm}%
+ \def\CLASSINFOpaperheight{297mm}}
+
+\DeclareOption{oneside}{\@twosidefalse\@mparswitchfalse
+ \CLASSOPTIONonesidetrue\CLASSOPTIONtwosidefalse}
+\DeclareOption{twoside}{\@twosidetrue\@mparswitchtrue
+ \CLASSOPTIONtwosidetrue\CLASSOPTIONonesidefalse}
+
+\DeclareOption{onecolumn}{\CLASSOPTIONonecolumntrue\CLASSOPTIONtwocolumnfalse}
+\DeclareOption{twocolumn}{\CLASSOPTIONtwocolumntrue\CLASSOPTIONonecolumnfalse}
+
+% If the user selects draft, then this class AND any packages
+% will go into draft mode.
+\DeclareOption{draft}{\CLASSOPTIONdrafttrue\CLASSOPTIONdraftclstrue
+ \CLASSOPTIONdraftclsnofootfalse}
+% draftcls is for a draft mode which will not affect any packages
+% used by the document.
+\DeclareOption{draftcls}{\CLASSOPTIONdraftfalse\CLASSOPTIONdraftclstrue
+ \CLASSOPTIONdraftclsnofootfalse}
+% draftclsnofoot is like draftcls, but without the footer.
+\DeclareOption{draftclsnofoot}{\CLASSOPTIONdraftfalse\CLASSOPTIONdraftclstrue
+ \CLASSOPTIONdraftclsnofoottrue}
+\DeclareOption{final}{\CLASSOPTIONdraftfalse\CLASSOPTIONdraftclsfalse
+ \CLASSOPTIONdraftclsnofootfalse}
+
+\DeclareOption{journal}{\CLASSOPTIONpeerreviewfalse\CLASSOPTIONpeerreviewcafalse
+ \CLASSOPTIONjournaltrue\CLASSOPTIONconferencefalse\CLASSOPTIONtechnotefalse}
+
+\DeclareOption{conference}{\CLASSOPTIONpeerreviewfalse\CLASSOPTIONpeerreviewcafalse
+ \CLASSOPTIONjournalfalse\CLASSOPTIONconferencetrue\CLASSOPTIONtechnotefalse}
+
+\DeclareOption{technote}{\CLASSOPTIONpeerreviewfalse\CLASSOPTIONpeerreviewcafalse
+ \CLASSOPTIONjournalfalse\CLASSOPTIONconferencefalse\CLASSOPTIONtechnotetrue}
+
+\DeclareOption{peerreview}{\CLASSOPTIONpeerreviewtrue\CLASSOPTIONpeerreviewcafalse
+ \CLASSOPTIONjournalfalse\CLASSOPTIONconferencefalse\CLASSOPTIONtechnotefalse}
+
+\DeclareOption{peerreviewca}{\CLASSOPTIONpeerreviewtrue\CLASSOPTIONpeerreviewcatrue
+ \CLASSOPTIONjournalfalse\CLASSOPTIONconferencefalse\CLASSOPTIONtechnotefalse}
+
+\DeclareOption{nofonttune}{\CLASSOPTIONnofonttunetrue}
+
+\DeclareOption{captionsoff}{\CLASSOPTIONcaptionsofftrue}
+
+\DeclareOption{compsoc}{\CLASSOPTIONcompsoctrue}
+
+\DeclareOption{compsocconf}{\CLASSOPTIONcompsocconftrue}
+
+\DeclareOption{romanappendices}{\CLASSOPTIONromanappendicestrue}
+
+
+% default to US letter paper, 10pt, twocolumn, one sided, final, journal
+\ExecuteOptions{letterpaper,10pt,twocolumn,oneside,final,journal}
+% overrride these defaults per user requests
+\ProcessOptions
+
+
+
+% Computer Society conditional execution command
+\long\def\@IEEEcompsoconly#1{\relax\ifCLASSOPTIONcompsoc\relax#1\relax\fi\relax}
+% inverse
+\long\def\@IEEEnotcompsoconly#1{\relax\ifCLASSOPTIONcompsoc\else\relax#1\relax\fi\relax}
+% compsoc conference
+\long\def\@IEEEcompsocconfonly#1{\relax\ifCLASSOPTIONcompsocconf\ifCLASSOPTIONconference\relax#1\relax\fi\fi\relax}
+% compsoc not conference
+\long\def\@IEEEcompsocnotconfonly#1{\relax\ifCLASSOPTIONcompsoc\ifCLASSOPTIONconference\else\relax#1\relax\fi\fi\relax}
+
+
+% IEEE uses Times Roman font, so we'll default to Times.
+% These three commands make up the entire times.sty package.
+\renewcommand{\sfdefault}{phv}
+\renewcommand{\rmdefault}{ptm}
+\renewcommand{\ttdefault}{pcr}
+
+\@IEEEcompsoconly{\typeout{-- Using IEEE Computer Society mode.}}
+
+% V1.7 compsoc nonconference papers, use Palatino/Palladio as the main text font,
+% not Times Roman.
+\@IEEEcompsocnotconfonly{\renewcommand{\rmdefault}{ppl}}
+
+% enable Times/Palatino main text font
+\normalfont\selectfont
+
+
+
+
+
+% V1.7 conference notice message hook
+\def\@IEEEconsolenoticeconference{\typeout{}%
+\typeout{** Conference Paper **}%
+\typeout{Before submitting the final camera ready copy, remember to:}%
+\typeout{}%
+\typeout{ 1. Manually equalize the lengths of two columns on the last page}%
+\typeout{ of your paper;}%
+\typeout{}%
+\typeout{ 2. Ensure that any PostScript and/or PDF output post-processing}%
+\typeout{ uses only Type 1 fonts and that every step in the generation}%
+\typeout{ process uses the appropriate paper size.}%
+\typeout{}}
+
+
+% we can send console reminder messages to the user here
+\AtEndDocument{\ifCLASSOPTIONconference\@IEEEconsolenoticeconference\fi}
+
+
+% warn about the use of single column other than for draft mode
+\ifCLASSOPTIONtwocolumn\else%
+ \ifCLASSOPTIONdraftcls\else%
+ \typeout{** ATTENTION: Single column mode is not typically used with IEEE publications.}%
+ \fi%
+\fi
+
+
+% V1.7 improved paper size setting code.
+% Set pdfpage and dvips paper sizes. Conditional tests are similar to that
+% of ifpdf.sty. Retain within {} to ensure tested macros are never altered,
+% even if only effect is to set them to \relax.
+% if \pdfoutput is undefined or equal to relax, output a dvips special
+{\@ifundefined{pdfoutput}{\AtBeginDvi{\special{papersize=\CLASSINFOpaperwidth,\CLASSINFOpaperheight}}}{%
+% pdfoutput is defined and not equal to \relax
+% check for pdfpageheight existence just in case someone sets pdfoutput
+% under non-pdflatex. If exists, set them regardless of value of \pdfoutput.
+\@ifundefined{pdfpageheight}{\relax}{\global\pdfpagewidth\paperwidth
+\global\pdfpageheight\paperheight}%
+% if using \pdfoutput=0 under pdflatex, send dvips papersize special
+\ifcase\pdfoutput
+\AtBeginDvi{\special{papersize=\CLASSINFOpaperwidth,\CLASSINFOpaperheight}}%
+\else
+% we are using pdf output, set CLASSINFOpdf flag
+\global\CLASSINFOpdftrue
+\fi}}
+
+% let the user know the selected papersize
+\typeout{-- Using \CLASSINFOpaperwidth\space x \CLASSINFOpaperheight\space
+(\CLASSOPTIONpaper)\space paper.}
+
+\ifCLASSINFOpdf
+\typeout{-- Using PDF output.}
+\else
+\typeout{-- Using DVI output.}
+\fi
+
+
+% The idea hinted here is for LaTeX to generate markleft{} and markright{}
+% automatically for you after you enter \author{}, \journal{},
+% \journaldate{}, journalvol{}, \journalnum{}, etc.
+% However, there may be some backward compatibility issues here as
+% well as some special applications for IEEEtran.cls and special issues
+% that may require the flexible \markleft{}, \markright{} and/or \markboth{}.
+% We'll leave this as an open future suggestion.
+%\newcommand{\journal}[1]{\def\@journal{#1}}
+%\def\@journal{}
+
+
+
+% pointsize values
+% used with ifx to determine the document's normal size
+\def\@IEEEptsizenine{9}
+\def\@IEEEptsizeten{10}
+\def\@IEEEptsizeeleven{11}
+\def\@IEEEptsizetwelve{12}
+
+
+
+% FONT DEFINITIONS (No sizexx.clo file needed)
+% V1.6 revised font sizes, displayskip values and
+% revised normalsize baselineskip to reduce underfull vbox problems
+% on the 58pc = 696pt = 9.5in text height we want
+% normalsize #lines/column baselineskip (aka leading)
+% 9pt 63 11.0476pt (truncated down)
+% 10pt 58 12pt (exact)
+% 11pt 52 13.3846pt (truncated down)
+% 12pt 50 13.92pt (exact)
+%
+
+% we need to store the nominal baselineskip for the given font size
+% in case baselinestretch ever changes.
+% this is a dimen, so it will not hold stretch or shrink
+\newdimen\@IEEEnormalsizeunitybaselineskip
+\@IEEEnormalsizeunitybaselineskip\baselineskip
+
+\ifx\CLASSOPTIONpt\@IEEEptsizenine
+\typeout{-- This is a 9 point document.}
+\def\normalsize{\@setfontsize{\normalsize}{9}{11.0476pt}}%
+\setlength{\@IEEEnormalsizeunitybaselineskip}{11.0476pt}%
+\normalsize
+\abovedisplayskip 1.5ex plus3pt minus1pt%
+\belowdisplayskip \abovedisplayskip%
+\abovedisplayshortskip 0pt plus3pt%
+\belowdisplayshortskip 1.5ex plus3pt minus1pt
+\def\small{\@setfontsize{\small}{8.5}{10pt}}
+\def\footnotesize{\@setfontsize{\footnotesize}{8}{9pt}}
+\def\scriptsize{\@setfontsize{\scriptsize}{7}{8pt}}
+\def\tiny{\@setfontsize{\tiny}{5}{6pt}}
+% sublargesize is the same as large - 10pt
+\def\sublargesize{\@setfontsize{\sublargesize}{10}{12pt}}
+\def\large{\@setfontsize{\large}{10}{12pt}}
+\def\Large{\@setfontsize{\Large}{12}{14pt}}
+\def\LARGE{\@setfontsize{\LARGE}{14}{17pt}}
+\def\huge{\@setfontsize{\huge}{17}{20pt}}
+\def\Huge{\@setfontsize{\Huge}{20}{24pt}}
+\fi
+
+
+% Check if we have selected 10 points
+\ifx\CLASSOPTIONpt\@IEEEptsizeten
+\typeout{-- This is a 10 point document.}
+\def\normalsize{\@setfontsize{\normalsize}{10}{12.00pt}}%
+\setlength{\@IEEEnormalsizeunitybaselineskip}{12pt}%
+\normalsize
+\abovedisplayskip 1.5ex plus4pt minus2pt%
+\belowdisplayskip \abovedisplayskip%
+\abovedisplayshortskip 0pt plus4pt%
+\belowdisplayshortskip 1.5ex plus4pt minus2pt
+\def\small{\@setfontsize{\small}{9}{10pt}}
+\def\footnotesize{\@setfontsize{\footnotesize}{8}{9pt}}
+\def\scriptsize{\@setfontsize{\scriptsize}{7}{8pt}}
+\def\tiny{\@setfontsize{\tiny}{5}{6pt}}
+% sublargesize is a tad smaller than large - 11pt
+\def\sublargesize{\@setfontsize{\sublargesize}{11}{13.4pt}}
+\def\large{\@setfontsize{\large}{12}{14pt}}
+\def\Large{\@setfontsize{\Large}{14}{17pt}}
+\def\LARGE{\@setfontsize{\LARGE}{17}{20pt}}
+\def\huge{\@setfontsize{\huge}{20}{24pt}}
+\def\Huge{\@setfontsize{\Huge}{24}{28pt}}
+\fi
+
+
+% Check if we have selected 11 points
+\ifx\CLASSOPTIONpt\@IEEEptsizeeleven
+\typeout{-- This is an 11 point document.}
+\def\normalsize{\@setfontsize{\normalsize}{11}{13.3846pt}}%
+\setlength{\@IEEEnormalsizeunitybaselineskip}{13.3846pt}%
+\normalsize
+\abovedisplayskip 1.5ex plus5pt minus3pt%
+\belowdisplayskip \abovedisplayskip%
+\abovedisplayshortskip 0pt plus5pt%
+\belowdisplayshortskip 1.5ex plus5pt minus3pt
+\def\small{\@setfontsize{\small}{10}{12pt}}
+\def\footnotesize{\@setfontsize{\footnotesize}{9}{10.5pt}}
+\def\scriptsize{\@setfontsize{\scriptsize}{8}{9pt}}
+\def\tiny{\@setfontsize{\tiny}{6}{7pt}}
+% sublargesize is the same as large - 12pt
+\def\sublargesize{\@setfontsize{\sublargesize}{12}{14pt}}
+\def\large{\@setfontsize{\large}{12}{14pt}}
+\def\Large{\@setfontsize{\Large}{14}{17pt}}
+\def\LARGE{\@setfontsize{\LARGE}{17}{20pt}}
+\def\huge{\@setfontsize{\huge}{20}{24pt}}
+\def\Huge{\@setfontsize{\Huge}{24}{28pt}}
+\fi
+
+
+% Check if we have selected 12 points
+\ifx\CLASSOPTIONpt\@IEEEptsizetwelve
+\typeout{-- This is a 12 point document.}
+\def\normalsize{\@setfontsize{\normalsize}{12}{13.92pt}}%
+\setlength{\@IEEEnormalsizeunitybaselineskip}{13.92pt}%
+\normalsize
+\abovedisplayskip 1.5ex plus6pt minus4pt%
+\belowdisplayskip \abovedisplayskip%
+\abovedisplayshortskip 0pt plus6pt%
+\belowdisplayshortskip 1.5ex plus6pt minus4pt
+\def\small{\@setfontsize{\small}{10}{12pt}}
+\def\footnotesize{\@setfontsize{\footnotesize}{9}{10.5pt}}
+\def\scriptsize{\@setfontsize{\scriptsize}{8}{9pt}}
+\def\tiny{\@setfontsize{\tiny}{6}{7pt}}
+% sublargesize is the same as large - 14pt
+\def\sublargesize{\@setfontsize{\sublargesize}{14}{17pt}}
+\def\large{\@setfontsize{\large}{14}{17pt}}
+\def\Large{\@setfontsize{\Large}{17}{20pt}}
+\def\LARGE{\@setfontsize{\LARGE}{20}{24pt}}
+\def\huge{\@setfontsize{\huge}{22}{26pt}}
+\def\Huge{\@setfontsize{\Huge}{24}{28pt}}
+\fi
+
+
+% V1.6 The Computer Modern Fonts will issue a substitution warning for
+% 24pt titles (24.88pt is used instead) increase the substitution
+% tolerance to turn off this warning
+\def\fontsubfuzz{.9pt}
+% However, the default (and correct) Times font will scale exactly as needed.
+
+
+% warn the user in case they forget to use the 9pt option with
+% technote
+\ifCLASSOPTIONtechnote%
+ \ifx\CLASSOPTIONpt\@IEEEptsizenine\else%
+ \typeout{** ATTENTION: Technotes are normally 9pt documents.}%
+ \fi%
+\fi
+
+
+% V1.7
+% Improved \textunderscore to provide a much better fake _ when used with
+% OT1 encoding. Under OT1, detect use of pcr or cmtt \ttfamily and use
+% available true _ glyph for those two typewriter fonts.
+\def\@IEEEstringptm{ptm} % Times Roman family
+\def\@IEEEstringppl{ppl} % Palatino Roman family
+\def\@IEEEstringphv{phv} % Helvetica Sans Serif family
+\def\@IEEEstringpcr{pcr} % Courier typewriter family
+\def\@IEEEstringcmtt{cmtt} % Computer Modern typewriter family
+\DeclareTextCommandDefault{\textunderscore}{\leavevmode
+\ifx\f@family\@IEEEstringpcr\string_\else
+\ifx\f@family\@IEEEstringcmtt\string_\else
+\ifx\f@family\@IEEEstringptm\kern 0em\vbox{\hrule\@width 0.5em\@height 0.5pt\kern -0.3ex}\else
+\ifx\f@family\@IEEEstringppl\kern 0em\vbox{\hrule\@width 0.5em\@height 0.5pt\kern -0.3ex}\else
+\ifx\f@family\@IEEEstringphv\kern -0.03em\vbox{\hrule\@width 0.62em\@height 0.52pt\kern -0.33ex}\kern -0.03em\else
+\kern 0.09em\vbox{\hrule\@width 0.6em\@height 0.44pt\kern -0.63pt\kern -0.42ex}\kern 0.09em\fi\fi\fi\fi\fi\relax}
+
+
+
+
+% set the default \baselinestretch
+\def\baselinestretch{1}
+\ifCLASSOPTIONdraftcls
+ \def\baselinestretch{1.5}% default baselinestretch for draft modes
+\fi
+
+
+% process CLASSINPUT baselinestretch
+\ifx\CLASSINPUTbaselinestretch\@IEEEundefined
+\else
+ \edef\baselinestretch{\CLASSINPUTbaselinestretch} % user CLASSINPUT override
+ \typeout{** ATTENTION: Overriding \string\baselinestretch\space to
+ \baselinestretch\space via \string\CLASSINPUT.}
+\fi
+
+\normalsize % make \baselinestretch take affect
+
+
+
+
+% store the normalsize baselineskip
+\newdimen\CLASSINFOnormalsizebaselineskip
+\CLASSINFOnormalsizebaselineskip=\baselineskip\relax
+% and the normalsize unity (baselinestretch=1) baselineskip
+% we could save a register by giving the user access to
+% \@IEEEnormalsizeunitybaselineskip. However, let's protect
+% its read only internal status
+\newdimen\CLASSINFOnormalsizeunitybaselineskip
+\CLASSINFOnormalsizeunitybaselineskip=\@IEEEnormalsizeunitybaselineskip\relax
+% store the nominal value of jot
+\newdimen\IEEEnormaljot
+\IEEEnormaljot=0.25\baselineskip\relax
+
+% set \jot
+\jot=\IEEEnormaljot\relax
+
+
+
+
+% V1.6, we are now going to fine tune the interword spacing
+% The default interword glue for Times under TeX appears to use a
+% nominal interword spacing of 25% (relative to the font size, i.e., 1em)
+% a maximum of 40% and a minimum of 19%.
+% For example, 10pt text uses an interword glue of:
+%
+% 2.5pt plus 1.49998pt minus 0.59998pt
+%
+% However, IEEE allows for a more generous range which reduces the need
+% for hyphenation, especially for two column text. Furthermore, IEEE
+% tends to use a little bit more nominal space between the words.
+% IEEE's interword spacing percentages appear to be:
+% 35% nominal
+% 23% minimum
+% 50% maximum
+% (They may even be using a tad more for the largest fonts such as 24pt.)
+%
+% for bold text, IEEE increases the spacing a little more:
+% 37.5% nominal
+% 23% minimum
+% 55% maximum
+
+% here are the interword spacing ratios we'll use
+% for medium (normal weight)
+\def\@IEEEinterspaceratioM{0.35}
+\def\@IEEEinterspaceMINratioM{0.23}
+\def\@IEEEinterspaceMAXratioM{0.50}
+
+% for bold
+\def\@IEEEinterspaceratioB{0.375}
+\def\@IEEEinterspaceMINratioB{0.23}
+\def\@IEEEinterspaceMAXratioB{0.55}
+
+
+% command to revise the interword spacing for the current font under TeX:
+% \fontdimen2 = nominal interword space
+% \fontdimen3 = interword stretch
+% \fontdimen4 = interword shrink
+% since all changes to the \fontdimen are global, we can enclose these commands
+% in braces to confine any font attribute or length changes
+\def\@@@IEEEsetfontdimens#1#2#3{{%
+\setlength{\@IEEEtrantmpdimenB}{\f@size pt}% grab the font size in pt, could use 1em instead.
+\setlength{\@IEEEtrantmpdimenA}{#1\@IEEEtrantmpdimenB}%
+\fontdimen2\font=\@IEEEtrantmpdimenA\relax
+\addtolength{\@IEEEtrantmpdimenA}{-#2\@IEEEtrantmpdimenB}%
+\fontdimen3\font=-\@IEEEtrantmpdimenA\relax
+\setlength{\@IEEEtrantmpdimenA}{#1\@IEEEtrantmpdimenB}%
+\addtolength{\@IEEEtrantmpdimenA}{-#3\@IEEEtrantmpdimenB}%
+\fontdimen4\font=\@IEEEtrantmpdimenA\relax}}
+
+% revise the interword spacing for each font weight
+\def\@@IEEEsetfontdimens{{%
+\mdseries
+\@@@IEEEsetfontdimens{\@IEEEinterspaceratioM}{\@IEEEinterspaceMAXratioM}{\@IEEEinterspaceMINratioM}%
+\bfseries
+\@@@IEEEsetfontdimens{\@IEEEinterspaceratioB}{\@IEEEinterspaceMAXratioB}{\@IEEEinterspaceMINratioB}%
+}}
+
+% revise the interword spacing for each font shape
+% \slshape is not often used for IEEE work and is not altered here. The \scshape caps are
+% already a tad too large in the free LaTeX fonts (as compared to what IEEE uses) so we
+% won't alter these either.
+\def\@IEEEsetfontdimens{{%
+\normalfont
+\@@IEEEsetfontdimens
+\normalfont\itshape
+\@@IEEEsetfontdimens
+}}
+
+% command to revise the interword spacing for each font size (and shape
+% and weight). Only the \rmfamily is done here as \ttfamily uses a
+% fixed spacing and \sffamily is not used as the main text of IEEE papers.
+\def\@IEEEtunefonts{{\selectfont\rmfamily
+\tiny\@IEEEsetfontdimens
+\scriptsize\@IEEEsetfontdimens
+\footnotesize\@IEEEsetfontdimens
+\small\@IEEEsetfontdimens
+\normalsize\@IEEEsetfontdimens
+\sublargesize\@IEEEsetfontdimens
+\large\@IEEEsetfontdimens
+\LARGE\@IEEEsetfontdimens
+\huge\@IEEEsetfontdimens
+\Huge\@IEEEsetfontdimens}}
+
+% if the nofonttune class option is not given, revise the interword spacing
+% now - in case IEEEtran makes any default length measurements, and make
+% sure all the default fonts are loaded
+\ifCLASSOPTIONnofonttune\else
+\@IEEEtunefonts
+\fi
+
+% and again at the start of the document in case the user loaded different fonts
+\AtBeginDocument{\ifCLASSOPTIONnofonttune\else\@IEEEtunefonts\fi}
+
+
+
+% V1.6
+% LaTeX is a little to quick to use hyphenations
+% So, we increase the penalty for their use and raise
+% the badness level that triggers an underfull hbox
+% warning. The author may still have to tweak things,
+% but the appearance will be much better "right out
+% of the box" than that under V1.5 and prior.
+% TeX default is 50
+\hyphenpenalty=750
+% If we didn't adjust the interword spacing, 2200 might be better.
+% The TeX default is 1000
+\hbadness=1350
+% IEEE does not use extra spacing after punctuation
+\frenchspacing
+
+% V1.7 increase this a tad to discourage equation breaks
+\binoppenalty=1000 % default 700
+\relpenalty=800 % default 500
+
+
+% margin note stuff
+\marginparsep 10pt
+\marginparwidth 20pt
+\marginparpush 25pt
+
+
+% if things get too close, go ahead and let them touch
+\lineskip 0pt
+\normallineskip 0pt
+\lineskiplimit 0pt
+\normallineskiplimit 0pt
+
+% The distance from the lower edge of the text body to the
+% footline
+\footskip 0.4in
+
+% normally zero, should be relative to font height.
+% put in a little rubber to help stop some bad breaks (underfull vboxes)
+\parskip 0ex plus 0.2ex minus 0.1ex
+
+\parindent 1.0em
+
+\topmargin -49.0pt
+\headheight 12pt
+\headsep 0.25in
+
+% use the normal font baselineskip
+% so that \topskip is unaffected by changes in \baselinestretch
+\topskip=\@IEEEnormalsizeunitybaselineskip
+\textheight 58pc % 9.63in, 696pt
+% Tweak textheight to a perfect integer number of lines/page.
+% The normal baselineskip for each document point size is used
+% to determine these values.
+\ifx\CLASSOPTIONpt\@IEEEptsizenine\textheight=63\@IEEEnormalsizeunitybaselineskip\fi % 63 lines/page
+\ifx\CLASSOPTIONpt\@IEEEptsizeten\textheight=58\@IEEEnormalsizeunitybaselineskip\fi % 58 lines/page
+\ifx\CLASSOPTIONpt\@IEEEptsizeeleven\textheight=52\@IEEEnormalsizeunitybaselineskip\fi % 52 lines/page
+\ifx\CLASSOPTIONpt\@IEEEptsizetwelve\textheight=50\@IEEEnormalsizeunitybaselineskip\fi % 50 lines/page
+
+
+\columnsep 1pc
+\textwidth 43pc % 2 x 21pc + 1pc = 43pc
+
+
+% the default side margins are equal
+\if@IEEEusingAfourpaper
+\oddsidemargin 19.05mm
+\evensidemargin 19.05mm
+\else
+\oddsidemargin 0.680in
+\evensidemargin 0.680in
+\fi
+% compensate for LaTeX's 1in offset
+\addtolength{\oddsidemargin}{-1in}
+\addtolength{\evensidemargin}{-1in}
+
+
+
+% adjust margins for conference mode
+\ifCLASSOPTIONconference
+ \topmargin -0.25in
+ % we retain the reserved, but unused space for headers
+ \addtolength{\topmargin}{-\headheight}
+ \addtolength{\topmargin}{-\headsep}
+ \textheight 9.25in % The standard for conferences (668.4975pt)
+ % Tweak textheight to a perfect integer number of lines/page.
+ \ifx\CLASSOPTIONpt\@IEEEptsizenine\textheight=61\@IEEEnormalsizeunitybaselineskip\fi % 61 lines/page
+ \ifx\CLASSOPTIONpt\@IEEEptsizeten\textheight=56\@IEEEnormalsizeunitybaselineskip\fi % 56 lines/page
+ \ifx\CLASSOPTIONpt\@IEEEptsizeeleven\textheight=50\@IEEEnormalsizeunitybaselineskip\fi % 50 lines/page
+ \ifx\CLASSOPTIONpt\@IEEEptsizetwelve\textheight=48\@IEEEnormalsizeunitybaselineskip\fi % 48 lines/page
+\fi
+
+
+% compsoc conference
+\ifCLASSOPTIONcompsocconf
+\ifCLASSOPTIONconference
+ % compsoc conference use a larger value for columnsep
+ \columnsep 0.25in
+ % compsoc conferences want 1in top margin, 1.125in bottom margin
+ \topmargin 0in
+ %\addtolength{\topmargin}{-6pt}% we tweak this a tad to better comply with top of line stuff
+ % we retain the reserved, but unused space for headers
+ \addtolength{\topmargin}{-\headheight}
+ \addtolength{\topmargin}{-\headsep}
+ \textheight 9.0in % (641.39625pt)
+
+ % Tweak textheight to a perfect integer number of lines/page.
+ \ifx\CLASSOPTIONpt\@IEEEptsizenine\textheight=58\@IEEEnormalsizeunitybaselineskip\fi % 58 lines/page
+ \ifx\CLASSOPTIONpt\@IEEEptsizeten\textheight=54\@IEEEnormalsizeunitybaselineskip\fi % 54 lines/page
+ \ifx\CLASSOPTIONpt\@IEEEptsizeeleven\textheight=48\@IEEEnormalsizeunitybaselineskip\fi % 48 lines/page
+ \ifx\CLASSOPTIONpt\@IEEEptsizetwelve\textheight=46\@IEEEnormalsizeunitybaselineskip\fi % 46 lines/page
+ \textwidth 7in
+
+
+ %adjust text h/w for A4 paper
+ \if@IEEEusingAfourpaper
+ \textheight 9.69in
+ \textwidth 6.77in
+ \fi
+
+ % the default side margins are equal
+ \if@IEEEusingAfourpaper
+ \oddsidemargin 19.05mm
+ \evensidemargin 19.05mm
+ \else
+ \oddsidemargin 0.75in
+ \evensidemargin 0.75in
+ \fi
+ % compensate for LaTeX's 1in offset
+ \addtolength{\oddsidemargin}{-1in}
+ \addtolength{\evensidemargin}{-1in}
+\fi\fi
+
+
+
+% draft mode settings override that of all other modes
+% provides a nice 1in margin all around the paper and extra
+% space between the lines for editor's comments
+\ifCLASSOPTIONdraftcls
+ % want 1in from top of paper to text
+ \setlength{\topmargin}{-\headsep}%
+ \addtolength{\topmargin}{-\headheight}%
+ % we want 1in side margins regardless of paper type
+ \oddsidemargin 0in
+ \evensidemargin 0in
+ % set the text width
+ \setlength{\textwidth}{\paperwidth}%
+ \addtolength{\textwidth}{-2.0in}%
+ \setlength{\textheight}{\paperheight}%
+ \addtolength{\textheight}{-2.0in}%
+ % digitize textheight to be an integer number of lines.
+ % this may cause the bottom margin to be off a tad
+ \addtolength{\textheight}{-1\topskip}%
+ \divide\textheight by \baselineskip%
+ \multiply\textheight by \baselineskip%
+ \addtolength{\textheight}{\topskip}%
+\fi
+
+
+
+% process CLASSINPUT inner/outer margin
+% if inner margin defined, but outer margin not, set outer to inner.
+\ifx\CLASSINPUTinnersidemargin\@IEEEundefined
+\else
+ \ifx\CLASSINPUToutersidemargin\@IEEEundefined
+ \edef\CLASSINPUToutersidemargin{\CLASSINPUTinnersidemargin}
+ \fi
+\fi
+
+\ifx\CLASSINPUToutersidemargin\@IEEEundefined
+\else
+ % if outer margin defined, but inner margin not, set inner to outer.
+ \ifx\CLASSINPUTinnersidemargin\@IEEEundefined
+ \edef\CLASSINPUTinnersidemargin{\CLASSINPUToutersidemargin}
+ \fi
+ \setlength{\oddsidemargin}{\CLASSINPUTinnersidemargin}
+ \ifCLASSOPTIONtwoside
+ \setlength{\evensidemargin}{\CLASSINPUToutersidemargin}
+ \else
+ \setlength{\evensidemargin}{\CLASSINPUTinnersidemargin}
+ \fi
+ \addtolength{\oddsidemargin}{-1in}
+ \addtolength{\evensidemargin}{-1in}
+ \setlength{\textwidth}{\paperwidth}
+ \addtolength{\textwidth}{-\CLASSINPUTinnersidemargin}
+ \addtolength{\textwidth}{-\CLASSINPUToutersidemargin}
+ \typeout{** ATTENTION: Overriding inner side margin to \CLASSINPUTinnersidemargin\space and
+ outer side margin to \CLASSINPUToutersidemargin\space via \string\CLASSINPUT.}
+\fi
+
+
+
+% process CLASSINPUT top/bottom text margin
+% if toptext margin defined, but bottomtext margin not, set bottomtext to toptext margin
+\ifx\CLASSINPUTtoptextmargin\@IEEEundefined
+\else
+ \ifx\CLASSINPUTbottomtextmargin\@IEEEundefined
+ \edef\CLASSINPUTbottomtextmargin{\CLASSINPUTtoptextmargin}
+ \fi
+\fi
+
+\ifx\CLASSINPUTbottomtextmargin\@IEEEundefined
+\else
+ % if bottomtext margin defined, but toptext margin not, set toptext to bottomtext margin
+ \ifx\CLASSINPUTtoptextmargin\@IEEEundefined
+ \edef\CLASSINPUTtoptextmargin{\CLASSINPUTbottomtextmargin}
+ \fi
+ \setlength{\topmargin}{\CLASSINPUTtoptextmargin}
+ \addtolength{\topmargin}{-1in}
+ \addtolength{\topmargin}{-\headheight}
+ \addtolength{\topmargin}{-\headsep}
+ \setlength{\textheight}{\paperheight}
+ \addtolength{\textheight}{-\CLASSINPUTtoptextmargin}
+ \addtolength{\textheight}{-\CLASSINPUTbottomtextmargin}
+ % in the default format we use the normal baselineskip as topskip
+ % we only need 0.7 of this to clear typical top text and we need
+ % an extra 0.3 spacing at the bottom for descenders. This will
+ % correct for both.
+ \addtolength{\topmargin}{-0.3\@IEEEnormalsizeunitybaselineskip}
+ \typeout{** ATTENTION: Overriding top text margin to \CLASSINPUTtoptextmargin\space and
+ bottom text margin to \CLASSINPUTbottomtextmargin\space via \string\CLASSINPUT.}
+\fi
+
+
+
+
+
+
+
+% LIST SPACING CONTROLS
+
+% Controls the amount of EXTRA spacing
+% above and below \trivlist
+% Both \list and IED lists override this.
+% However, \trivlist will use this as will most
+% things built from \trivlist like the \center
+% environment.
+\topsep 0.5\baselineskip
+
+% Controls the additional spacing around lists preceded
+% or followed by blank lines. IEEE does not increase
+% spacing before or after paragraphs so it is set to zero.
+% \z@ is the same as zero, but faster.
+\partopsep \z@
+
+% Controls the spacing between paragraphs in lists.
+% IEEE does not increase spacing before or after paragraphs
+% so this is also zero.
+% With IEEEtran.cls, global changes to
+% this value DO affect lists (but not IED lists).
+\parsep \z@
+
+% Controls the extra spacing between list items.
+% IEEE does not put extra spacing between items.
+% With IEEEtran.cls, global changes to this value DO affect
+% lists (but not IED lists).
+\itemsep \z@
+
+% \itemindent is the amount to indent the FIRST line of a list
+% item. It is auto set to zero within the \list environment. To alter
+% it, you have to do so when you call the \list.
+% However, IEEE uses this for the theorem environment
+% There is an alternative value for this near \leftmargini below
+\itemindent -1em
+
+% \leftmargin, the spacing from the left margin of the main text to
+% the left of the main body of a list item is set by \list.
+% Hence this statement does nothing for lists.
+% But, quote and verse do use it for indention.
+\leftmargin 2em
+
+% we retain this stuff from the older IEEEtran.cls so that \list
+% will work the same way as before. However, itemize, enumerate and
+% description (IED) could care less about what these are as they
+% all are overridden.
+\leftmargini 2em
+%\itemindent 2em % Alternative values: sometimes used.
+%\leftmargini 0em
+\leftmarginii 1em
+\leftmarginiii 1.5em
+\leftmarginiv 1.5em
+\leftmarginv 1.0em
+\leftmarginvi 1.0em
+\labelsep 0.5em
+\labelwidth \z@
+
+
+% The old IEEEtran.cls behavior of \list is retained.
+% However, the new V1.3 IED list environments override all the
+% @list stuff (\@listX is called within \list for the
+% appropriate level just before the user's list_decl is called).
+% \topsep is now 2pt as IEEE puts a little extra space around
+% lists - used by those non-IED macros that depend on \list.
+% Note that \parsep and \itemsep are not redefined as in
+% the sizexx.clo \@listX (which article.cls uses) so global changes
+% of these values DO affect \list
+%
+\def\@listi{\leftmargin\leftmargini \topsep 2pt plus 1pt minus 1pt}
+\let\@listI\@listi
+\def\@listii{\leftmargin\leftmarginii\labelwidth\leftmarginii%
+ \advance\labelwidth-\labelsep \topsep 2pt}
+\def\@listiii{\leftmargin\leftmarginiii\labelwidth\leftmarginiii%
+ \advance\labelwidth-\labelsep \topsep 2pt}
+\def\@listiv{\leftmargin\leftmarginiv\labelwidth\leftmarginiv%
+ \advance\labelwidth-\labelsep \topsep 2pt}
+\def\@listv{\leftmargin\leftmarginv\labelwidth\leftmarginv%
+ \advance\labelwidth-\labelsep \topsep 2pt}
+\def\@listvi{\leftmargin\leftmarginvi\labelwidth\leftmarginvi%
+ \advance\labelwidth-\labelsep \topsep 2pt}
+
+
+% IEEE uses 5) not 5.
+\def\labelenumi{\theenumi)} \def\theenumi{\arabic{enumi}}
+
+% IEEE uses a) not (a)
+\def\labelenumii{\theenumii)} \def\theenumii{\alph{enumii}}
+
+% IEEE uses iii) not iii.
+\def\labelenumiii{\theenumiii)} \def\theenumiii{\roman{enumiii}}
+
+% IEEE uses A) not A.
+\def\labelenumiv{\theenumiv)} \def\theenumiv{\Alph{enumiv}}
+
+% exactly the same as in article.cls
+\def\p@enumii{\theenumi}
+\def\p@enumiii{\theenumi(\theenumii)}
+\def\p@enumiv{\p@enumiii\theenumiii}
+
+% itemized list label styles
+\def\labelitemi{$\scriptstyle\bullet$}
+\def\labelitemii{\textbf{--}}
+\def\labelitemiii{$\ast$}
+\def\labelitemiv{$\cdot$}
+
+
+
+% **** V1.3 ENHANCEMENTS ****
+% Itemize, Enumerate and Description (IED) List Controls
+% ***************************
+%
+%
+% IEEE seems to use at least two different values by
+% which ITEMIZED list labels are indented to the right
+% For The Journal of Lightwave Technology (JLT) and The Journal
+% on Selected Areas in Communications (JSAC), they tend to use
+% an indention equal to \parindent. For Transactions on Communications
+% they tend to indent ITEMIZED lists a little more--- 1.3\parindent.
+% We'll provide both values here for you so that you can choose
+% which one you like in your document using a command such as:
+% setlength{\IEEEilabelindent}{\IEEEilabelindentB}
+\newdimen\IEEEilabelindentA
+\IEEEilabelindentA \parindent
+
+\newdimen\IEEEilabelindentB
+\IEEEilabelindentB 1.3\parindent
+% However, we'll default to using \parindent
+% which makes more sense to me
+\newdimen\IEEEilabelindent
+\IEEEilabelindent \IEEEilabelindentA
+
+
+% This controls the default amount the enumerated list labels
+% are indented to the right.
+% Normally, this is the same as the paragraph indention
+\newdimen\IEEEelabelindent
+\IEEEelabelindent \parindent
+
+% This controls the default amount the description list labels
+% are indented to the right.
+% Normally, this is the same as the paragraph indention
+\newdimen\IEEEdlabelindent
+\IEEEdlabelindent \parindent
+
+% This is the value actually used within the IED lists.
+% The IED environments automatically set its value to
+% one of the three values above, so global changes do
+% not have any effect
+\newdimen\IEEElabelindent
+\IEEElabelindent \parindent
+
+% The actual amount labels will be indented is
+% \IEEElabelindent multiplied by the factor below
+% corresponding to the level of nesting depth
+% This provides a means by which the user can
+% alter the effective \IEEElabelindent for deeper
+% levels
+% There may not be such a thing as correct "standard IEEE"
+% values. What IEEE actually does may depend on the specific
+% circumstances.
+% The first list level almost always has full indention.
+% The second levels I've seen have only 75% of the normal indentation
+% Three level or greater nestings are very rare. I am guessing
+% that they don't use any indentation.
+\def\IEEElabelindentfactori{1.0} % almost always one
+\def\IEEElabelindentfactorii{0.75} % 0.0 or 1.0 may be used in some cases
+\def\IEEElabelindentfactoriii{0.0} % 0.75? 0.5? 0.0?
+\def\IEEElabelindentfactoriv{0.0}
+\def\IEEElabelindentfactorv{0.0}
+\def\IEEElabelindentfactorvi{0.0}
+
+% value actually used within IED lists, it is auto
+% set to one of the 6 values above
+% global changes here have no effect
+\def\IEEElabelindentfactor{1.0}
+
+% This controls the default spacing between the end of the IED
+% list labels and the list text, when normal text is used for
+% the labels.
+\newdimen\IEEEiednormlabelsep
+\IEEEiednormlabelsep 0.6em
+
+% This controls the default spacing between the end of the IED
+% list labels and the list text, when math symbols are used for
+% the labels (nomenclature lists). IEEE usually increases the
+% spacing in these cases
+\newdimen\IEEEiedmathlabelsep
+\IEEEiedmathlabelsep 1.2em
+
+% This controls the extra vertical separation put above and
+% below each IED list. IEEE usually puts a little extra spacing
+% around each list. However, this spacing is barely noticeable.
+\newskip\IEEEiedtopsep
+\IEEEiedtopsep 2pt plus 1pt minus 1pt
+
+
+% This command is executed within each IED list environment
+% at the beginning of the list. You can use this to set the
+% parameters for some/all your IED list(s) without disturbing
+% global parameters that affect things other than lists.
+% i.e., renewcommand{\IEEEiedlistdecl}{\setlength{\labelsep}{5em}}
+% will alter the \labelsep for the next list(s) until
+% \IEEEiedlistdecl is redefined.
+\def\IEEEiedlistdecl{\relax}
+
+% This command provides an easy way to set \leftmargin based
+% on the \labelwidth, \labelsep and the argument \IEEElabelindent
+% Usage: \IEEEcalcleftmargin{width-to-indent-the-label}
+% output is in the \leftmargin variable, i.e., effectively:
+% \leftmargin = argument + \labelwidth + \labelsep
+% Note controlled spacing here, shield end of lines with %
+\def\IEEEcalcleftmargin#1{\setlength{\leftmargin}{#1}%
+\addtolength{\leftmargin}{\labelwidth}%
+\addtolength{\leftmargin}{\labelsep}}
+
+% This command provides an easy way to set \labelwidth to the
+% width of the given text. It is the same as
+% \settowidth{\labelwidth}{label-text}
+% and useful as a shorter alternative.
+% Typically used to set \labelwidth to be the width
+% of the longest label in the list
+\def\IEEEsetlabelwidth#1{\settowidth{\labelwidth}{#1}}
+
+% When this command is executed, IED lists will use the
+% IEEEiedmathlabelsep label separation rather than the normal
+% spacing. To have an effect, this command must be executed via
+% the \IEEEiedlistdecl or within the option of the IED list
+% environments.
+\def\IEEEusemathlabelsep{\setlength{\labelsep}{\IEEEiedmathlabelsep}}
+
+% A flag which controls whether the IED lists automatically
+% calculate \leftmargin from \IEEElabelindent, \labelwidth and \labelsep
+% Useful if you want to specify your own \leftmargin
+% This flag must be set (\IEEEnocalcleftmargintrue or \IEEEnocalcleftmarginfalse)
+% via the \IEEEiedlistdecl or within the option of the IED list
+% environments to have an effect.
+\newif\ifIEEEnocalcleftmargin
+\IEEEnocalcleftmarginfalse
+
+% A flag which controls whether \IEEElabelindent is multiplied by
+% the \IEEElabelindentfactor for each list level.
+% This flag must be set via the \IEEEiedlistdecl or within the option
+% of the IED list environments to have an effect.
+\newif\ifIEEEnolabelindentfactor
+\IEEEnolabelindentfactorfalse
+
+
+% internal variable to indicate type of IED label
+% justification
+% 0 - left; 1 - center; 2 - right
+\def\@IEEEiedjustify{0}
+
+
+% commands to allow the user to control IED
+% label justifications. Use these commands within
+% the IED environment option or in the \IEEEiedlistdecl
+% Note that changing the normal list justifications
+% is nonstandard and IEEE may not like it if you do so!
+% I include these commands as they may be helpful to
+% those who are using these enhanced list controls for
+% other non-IEEE related LaTeX work.
+% itemize and enumerate automatically default to right
+% justification, description defaults to left.
+\def\IEEEiedlabeljustifyl{\def\@IEEEiedjustify{0}}%left
+\def\IEEEiedlabeljustifyc{\def\@IEEEiedjustify{1}}%center
+\def\IEEEiedlabeljustifyr{\def\@IEEEiedjustify{2}}%right
+
+
+
+
+% commands to save to and restore from the list parameter copies
+% this allows us to set all the list parameters within
+% the list_decl and prevent \list (and its \@list)
+% from overriding any of our parameters
+% V1.6 use \edefs instead of dimen's to conserve dimen registers
+% Note controlled spacing here, shield end of lines with %
+\def\@IEEEsavelistparams{\edef\@IEEEiedtopsep{\the\topsep}%
+\edef\@IEEEiedlabelwidth{\the\labelwidth}%
+\edef\@IEEEiedlabelsep{\the\labelsep}%
+\edef\@IEEEiedleftmargin{\the\leftmargin}%
+\edef\@IEEEiedpartopsep{\the\partopsep}%
+\edef\@IEEEiedparsep{\the\parsep}%
+\edef\@IEEEieditemsep{\the\itemsep}%
+\edef\@IEEEiedrightmargin{\the\rightmargin}%
+\edef\@IEEEiedlistparindent{\the\listparindent}%
+\edef\@IEEEieditemindent{\the\itemindent}}
+
+% Note controlled spacing here
+\def\@IEEErestorelistparams{\topsep\@IEEEiedtopsep\relax%
+\labelwidth\@IEEEiedlabelwidth\relax%
+\labelsep\@IEEEiedlabelsep\relax%
+\leftmargin\@IEEEiedleftmargin\relax%
+\partopsep\@IEEEiedpartopsep\relax%
+\parsep\@IEEEiedparsep\relax%
+\itemsep\@IEEEieditemsep\relax%
+\rightmargin\@IEEEiedrightmargin\relax%
+\listparindent\@IEEEiedlistparindent\relax%
+\itemindent\@IEEEieditemindent\relax}
+
+
+% v1.6b provide original LaTeX IED list environments
+% note that latex.ltx defines \itemize and \enumerate, but not \description
+% which must be created by the base classes
+% save original LaTeX itemize and enumerate
+\let\LaTeXitemize\itemize
+\let\endLaTeXitemize\enditemize
+\let\LaTeXenumerate\enumerate
+\let\endLaTeXenumerate\endenumerate
+
+% provide original LaTeX description environment from article.cls
+\newenvironment{LaTeXdescription}
+ {\list{}{\labelwidth\z@ \itemindent-\leftmargin
+ \let\makelabel\descriptionlabel}}
+ {\endlist}
+\newcommand*\descriptionlabel[1]{\hspace\labelsep
+ \normalfont\bfseries #1}
+
+
+% override LaTeX's default IED lists
+\def\itemize{\@IEEEitemize}
+\def\enditemize{\@endIEEEitemize}
+\def\enumerate{\@IEEEenumerate}
+\def\endenumerate{\@endIEEEenumerate}
+\def\description{\@IEEEdescription}
+\def\enddescription{\@endIEEEdescription}
+
+% provide the user with aliases - may help those using packages that
+% override itemize, enumerate, or description
+\def\IEEEitemize{\@IEEEitemize}
+\def\endIEEEitemize{\@endIEEEitemize}
+\def\IEEEenumerate{\@IEEEenumerate}
+\def\endIEEEenumerate{\@endIEEEenumerate}
+\def\IEEEdescription{\@IEEEdescription}
+\def\endIEEEdescription{\@endIEEEdescription}
+
+
+% V1.6 we want to keep the IEEEtran IED list definitions as our own internal
+% commands so they are protected against redefinition
+\def\@IEEEitemize{\@ifnextchar[{\@@IEEEitemize}{\@@IEEEitemize[\relax]}}
+\def\@IEEEenumerate{\@ifnextchar[{\@@IEEEenumerate}{\@@IEEEenumerate[\relax]}}
+\def\@IEEEdescription{\@ifnextchar[{\@@IEEEdescription}{\@@IEEEdescription[\relax]}}
+\def\@endIEEEitemize{\endlist}
+\def\@endIEEEenumerate{\endlist}
+\def\@endIEEEdescription{\endlist}
+
+
+% DO NOT ALLOW BLANK LINES TO BE IN THESE IED ENVIRONMENTS
+% AS THIS WILL FORCE NEW PARAGRAPHS AFTER THE IED LISTS
+% IEEEtran itemized list MDS 1/2001
+% Note controlled spacing here, shield end of lines with %
+\def\@@IEEEitemize[#1]{%
+ \ifnum\@itemdepth>3\relax\@toodeep\else%
+ \ifnum\@listdepth>5\relax\@toodeep\else%
+ \advance\@itemdepth\@ne%
+ \edef\@itemitem{labelitem\romannumeral\the\@itemdepth}%
+ % get the labelindentfactor for this level
+ \advance\@listdepth\@ne% we need to know what the level WILL be
+ \edef\IEEElabelindentfactor{\csname IEEElabelindentfactor\romannumeral\the\@listdepth\endcsname}%
+ \advance\@listdepth-\@ne% undo our increment
+ \def\@IEEEiedjustify{2}% right justified labels are default
+ % set other defaults
+ \IEEEnocalcleftmarginfalse%
+ \IEEEnolabelindentfactorfalse%
+ \topsep\IEEEiedtopsep%
+ \IEEElabelindent\IEEEilabelindent%
+ \labelsep\IEEEiednormlabelsep%
+ \partopsep 0ex%
+ \parsep 0ex%
+ \itemsep 0ex%
+ \rightmargin 0em%
+ \listparindent 0em%
+ \itemindent 0em%
+ % calculate the label width
+ % the user can override this later if
+ % they specified a \labelwidth
+ \settowidth{\labelwidth}{\csname labelitem\romannumeral\the\@itemdepth\endcsname}%
+ \@IEEEsavelistparams% save our list parameters
+ \list{\csname\@itemitem\endcsname}{%
+ \@IEEErestorelistparams% override any list{} changes
+ % to our globals
+ \let\makelabel\@IEEEiedmakelabel% v1.6b setup \makelabel
+ \IEEEiedlistdecl% let user alter parameters
+ #1\relax%
+ % If the user has requested not to use the
+ % labelindent factor, don't revise \labelindent
+ \ifIEEEnolabelindentfactor\relax%
+ \else\IEEElabelindent=\IEEElabelindentfactor\labelindent%
+ \fi%
+ % Unless the user has requested otherwise,
+ % calculate our left margin based
+ % on \IEEElabelindent, \labelwidth and
+ % \labelsep
+ \ifIEEEnocalcleftmargin\relax%
+ \else\IEEEcalcleftmargin{\IEEElabelindent}%
+ \fi}\fi\fi}%
+
+
+% DO NOT ALLOW BLANK LINES TO BE IN THESE IED ENVIRONMENTS
+% AS THIS WILL FORCE NEW PARAGRAPHS AFTER THE IED LISTS
+% IEEEtran enumerate list MDS 1/2001
+% Note controlled spacing here, shield end of lines with %
+\def\@@IEEEenumerate[#1]{%
+ \ifnum\@enumdepth>3\relax\@toodeep\else%
+ \ifnum\@listdepth>5\relax\@toodeep\else%
+ \advance\@enumdepth\@ne%
+ \edef\@enumctr{enum\romannumeral\the\@enumdepth}%
+ % get the labelindentfactor for this level
+ \advance\@listdepth\@ne% we need to know what the level WILL be
+ \edef\IEEElabelindentfactor{\csname IEEElabelindentfactor\romannumeral\the\@listdepth\endcsname}%
+ \advance\@listdepth-\@ne% undo our increment
+ \def\@IEEEiedjustify{2}% right justified labels are default
+ % set other defaults
+ \IEEEnocalcleftmarginfalse%
+ \IEEEnolabelindentfactorfalse%
+ \topsep\IEEEiedtopsep%
+ \IEEElabelindent\IEEEelabelindent%
+ \labelsep\IEEEiednormlabelsep%
+ \partopsep 0ex%
+ \parsep 0ex%
+ \itemsep 0ex%
+ \rightmargin 0em%
+ \listparindent 0em%
+ \itemindent 0em%
+ % calculate the label width
+ % We'll set it to the width suitable for all labels using
+ % normalfont 1) to 9)
+ % The user can override this later
+ \settowidth{\labelwidth}{9)}%
+ \@IEEEsavelistparams% save our list parameters
+ \list{\csname label\@enumctr\endcsname}{\usecounter{\@enumctr}%
+ \@IEEErestorelistparams% override any list{} changes
+ % to our globals
+ \let\makelabel\@IEEEiedmakelabel% v1.6b setup \makelabel
+ \IEEEiedlistdecl% let user alter parameters
+ #1\relax%
+ % If the user has requested not to use the
+ % IEEElabelindent factor, don't revise \IEEElabelindent
+ \ifIEEEnolabelindentfactor\relax%
+ \else\IEEElabelindent=\IEEElabelindentfactor\IEEElabelindent%
+ \fi%
+ % Unless the user has requested otherwise,
+ % calculate our left margin based
+ % on \IEEElabelindent, \labelwidth and
+ % \labelsep
+ \ifIEEEnocalcleftmargin\relax%
+ \else\IEEEcalcleftmargin{\IEEElabelindent}%
+ \fi}\fi\fi}%
+
+
+% DO NOT ALLOW BLANK LINES TO BE IN THESE IED ENVIRONMENTS
+% AS THIS WILL FORCE NEW PARAGRAPHS AFTER THE IED LISTS
+% IEEEtran description list MDS 1/2001
+% Note controlled spacing here, shield end of lines with %
+\def\@@IEEEdescription[#1]{%
+ \ifnum\@listdepth>5\relax\@toodeep\else%
+ % get the labelindentfactor for this level
+ \advance\@listdepth\@ne% we need to know what the level WILL be
+ \edef\IEEElabelindentfactor{\csname IEEElabelindentfactor\romannumeral\the\@listdepth\endcsname}%
+ \advance\@listdepth-\@ne% undo our increment
+ \def\@IEEEiedjustify{0}% left justified labels are default
+ % set other defaults
+ \IEEEnocalcleftmarginfalse%
+ \IEEEnolabelindentfactorfalse%
+ \topsep\IEEEiedtopsep%
+ \IEEElabelindent\IEEEdlabelindent%
+ % assume normal labelsep
+ \labelsep\IEEEiednormlabelsep%
+ \partopsep 0ex%
+ \parsep 0ex%
+ \itemsep 0ex%
+ \rightmargin 0em%
+ \listparindent 0em%
+ \itemindent 0em%
+ % Bogus label width in case the user forgets
+ % to set it.
+ % TIP: If you want to see what a variable's width is you
+ % can use the TeX command \showthe\width-variable to
+ % display it on the screen during compilation
+ % (This might be helpful to know when you need to find out
+ % which label is the widest)
+ \settowidth{\labelwidth}{Hello}%
+ \@IEEEsavelistparams% save our list parameters
+ \list{}{\@IEEErestorelistparams% override any list{} changes
+ % to our globals
+ \let\makelabel\@IEEEiedmakelabel% v1.6b setup \makelabel
+ \IEEEiedlistdecl% let user alter parameters
+ #1\relax%
+ % If the user has requested not to use the
+ % labelindent factor, don't revise \IEEElabelindent
+ \ifIEEEnolabelindentfactor\relax%
+ \else\IEEElabelindent=\IEEElabelindentfactor\IEEElabelindent%
+ \fi%
+ % Unless the user has requested otherwise,
+ % calculate our left margin based
+ % on \IEEElabelindent, \labelwidth and
+ % \labelsep
+ \ifIEEEnocalcleftmargin\relax%
+ \else\IEEEcalcleftmargin{\IEEElabelindent}\relax%
+ \fi}\fi}
+
+% v1.6b we use one makelabel that does justification as needed.
+\def\@IEEEiedmakelabel#1{\relax\if\@IEEEiedjustify 0\relax
+\makebox[\labelwidth][l]{\normalfont #1}\else
+\if\@IEEEiedjustify 1\relax
+\makebox[\labelwidth][c]{\normalfont #1}\else
+\makebox[\labelwidth][r]{\normalfont #1}\fi\fi}
+
+
+% VERSE and QUOTE
+% V1.7 define environments with newenvironment
+\newenvironment{verse}{\let\\=\@centercr
+ \list{}{\itemsep\z@ \itemindent -1.5em \listparindent \itemindent
+ \rightmargin\leftmargin\advance\leftmargin 1.5em}\item\relax}
+ {\endlist}
+\newenvironment{quotation}{\list{}{\listparindent 1.5em \itemindent\listparindent
+ \rightmargin\leftmargin \parsep 0pt plus 1pt}\item\relax}
+ {\endlist}
+\newenvironment{quote}{\list{}{\rightmargin\leftmargin}\item\relax}
+ {\endlist}
+
+
+% \titlepage
+% provided only for backward compatibility. \maketitle is the correct
+% way to create the title page.
+\newif\if@restonecol
+\def\titlepage{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
+ \else \newpage \fi \thispagestyle{empty}\c@page\z@}
+\def\endtitlepage{\if@restonecol\twocolumn \else \newpage \fi}
+
+% standard values from article.cls
+\arraycolsep 5pt
+\arrayrulewidth .4pt
+\doublerulesep 2pt
+
+\tabcolsep 6pt
+\tabbingsep 0.5em
+
+
+%% FOOTNOTES
+%
+%\skip\footins 10pt plus 4pt minus 2pt
+% V1.6 respond to changes in font size
+% space added above the footnotes (if present)
+\skip\footins 0.9\baselineskip plus 0.4\baselineskip minus 0.2\baselineskip
+
+% V1.6, we need to make \footnotesep responsive to changes
+% in \baselineskip or strange spacings will result when in
+% draft mode. Here is a little LaTeX secret - \footnotesep
+% determines the height of an invisible strut that is placed
+% *above* the baseline of footnotes after the first. Since
+% LaTeX considers the space for characters to be 0.7/baselineskip
+% above the baseline and 0.3/baselineskip below it, we need to
+% use 0.7/baselineskip as a \footnotesep to maintain equal spacing
+% between all the lines of the footnotes. IEEE often uses a tad
+% more, so use 0.8\baselineskip. This slightly larger value also helps
+% the text to clear the footnote marks. Note that \thanks in IEEEtran
+% uses its own value of \footnotesep which is set in \maketitle.
+{\footnotesize
+\global\footnotesep 0.8\baselineskip}
+
+
+\skip\@mpfootins = \skip\footins
+\fboxsep = 3pt
+\fboxrule = .4pt
+% V1.6 use 1em, then use LaTeX2e's \@makefnmark
+% Note that IEEE normally *left* aligns the footnote marks, so we don't need
+% box resizing tricks here.
+\long\def\@makefntext#1{\parindent 1em\indent\hbox{\@makefnmark}#1}% V1.6 use 1em
+% V1.7 compsoc does not use superscipts for footnote marks
+\ifCLASSOPTIONcompsoc
+\def\@IEEEcompsocmakefnmark{\hbox{\normalfont\@thefnmark.\ }}
+\long\def\@makefntext#1{\parindent 1em\indent\hbox{\@IEEEcompsocmakefnmark}#1}
+\fi
+
+% IEEE does not use footnote rules
+\def\footnoterule{}
+
+% V1.7 for compsoc, IEEE uses a footnote rule only for \thanks. We devise a "one-shot"
+% system to implement this.
+\newif\if@IEEEenableoneshotfootnoterule
+\@IEEEenableoneshotfootnoterulefalse
+\ifCLASSOPTIONcompsoc
+\def\footnoterule{\relax\if@IEEEenableoneshotfootnoterule
+\kern-5pt
+\hbox to \columnwidth{\hfill\vrule width 0.5\columnwidth height 0.4pt\hfill}
+\kern4.6pt
+\global\@IEEEenableoneshotfootnoterulefalse
+\else
+\relax
+\fi}
+\fi
+
+% V1.6 do not allow LaTeX to break a footnote across multiple pages
+\interfootnotelinepenalty=10000
+
+% V1.6 discourage breaks within equations
+% Note that amsmath normally sets this to 10000,
+% but LaTeX2e normally uses 100.
+\interdisplaylinepenalty=2500
+
+% default allows section depth up to /paragraph
+\setcounter{secnumdepth}{4}
+
+% technotes do not allow /paragraph
+\ifCLASSOPTIONtechnote
+ \setcounter{secnumdepth}{3}
+\fi
+% neither do compsoc conferences
+\@IEEEcompsocconfonly{\setcounter{secnumdepth}{3}}
+
+
+\newcounter{section}
+\newcounter{subsection}[section]
+\newcounter{subsubsection}[subsection]
+\newcounter{paragraph}[subsubsection]
+
+% used only by IEEEtran's IEEEeqnarray as other packages may
+% have their own, different, implementations
+\newcounter{IEEEsubequation}[equation]
+
+% as shown when called by user from \ref, \label and in table of contents
+\def\theequation{\arabic{equation}} % 1
+\def\theIEEEsubequation{\theequation\alph{IEEEsubequation}} % 1a (used only by IEEEtran's IEEEeqnarray)
+\ifCLASSOPTIONcompsoc
+% compsoc is all arabic
+\def\thesection{\arabic{section}}
+\def\thesubsection{\thesection.\arabic{subsection}}
+\def\thesubsubsection{\thesubsection.\arabic{subsubsection}}
+\def\theparagraph{\thesubsubsection.\arabic{paragraph}}
+\else
+\def\thesection{\Roman{section}} % I
+% V1.7, \mbox prevents breaks around -
+\def\thesubsection{\mbox{\thesection-\Alph{subsection}}} % I-A
+% V1.7 use I-A1 format used by IEEE rather than I-A.1
+\def\thesubsubsection{\thesubsection\arabic{subsubsection}} % I-A1
+\def\theparagraph{\thesubsubsection\alph{paragraph}} % I-A1a
+\fi
+
+% From Heiko Oberdiek. Because of the \mbox in \thesubsection, we need to
+% tell hyperref to disable the \mbox command when making PDF bookmarks.
+% This done already with hyperref.sty version 6.74o and later, but
+% it will not hurt to do it here again for users of older versions.
+\@ifundefined{pdfstringdefPreHook}{\let\pdfstringdefPreHook\@empty}{}%
+\g@addto@macro\pdfstringdefPreHook{\let\mbox\relax}
+
+
+% Main text forms (how shown in main text headings)
+% V1.6, using \thesection in \thesectiondis allows changes
+% in the former to automatically appear in the latter
+\ifCLASSOPTIONcompsoc
+ \ifCLASSOPTIONconference% compsoc conference
+ \def\thesectiondis{\thesection.}
+ \def\thesubsectiondis{\thesectiondis\arabic{subsection}.}
+ \def\thesubsubsectiondis{\thesubsectiondis\arabic{subsubsection}.}
+ \def\theparagraphdis{\thesubsubsectiondis\arabic{paragraph}.}
+ \else% compsoc not conferencs
+ \def\thesectiondis{\thesection}
+ \def\thesubsectiondis{\thesectiondis.\arabic{subsection}}
+ \def\thesubsubsectiondis{\thesubsectiondis.\arabic{subsubsection}}
+ \def\theparagraphdis{\thesubsubsectiondis.\arabic{paragraph}}
+ \fi
+\else% not compsoc
+ \def\thesectiondis{\thesection.} % I.
+ \def\thesubsectiondis{\Alph{subsection}.} % B.
+ \def\thesubsubsectiondis{\arabic{subsubsection})} % 3)
+ \def\theparagraphdis{\alph{paragraph})} % d)
+\fi
+
+% just like LaTeX2e's \@eqnnum
+\def\theequationdis{{\normalfont \normalcolor (\theequation)}}% (1)
+% IEEEsubequation used only by IEEEtran's IEEEeqnarray
+\def\theIEEEsubequationdis{{\normalfont \normalcolor (\theIEEEsubequation)}}% (1a)
+% redirect LaTeX2e's equation number display and all that depend on
+% it, through IEEEtran's \theequationdis
+\def\@eqnnum{\theequationdis}
+
+
+
+% V1.7 provide string macros as article.cls does
+\def\contentsname{Contents}
+\def\listfigurename{List of Figures}
+\def\listtablename{List of Tables}
+\def\refname{References}
+\def\indexname{Index}
+\def\figurename{Fig.}
+\def\tablename{TABLE}
+\@IEEEcompsocconfonly{\def\figurename{Figure}\def\tablename{Table}}
+\def\partname{Part}
+\def\appendixname{Appendix}
+\def\abstractname{Abstract}
+% IEEE specific names
+\def\IEEEkeywordsname{Keywords}
+\def\IEEEproofname{Proof}
+
+
+% LIST OF FIGURES AND TABLES AND TABLE OF CONTENTS
+%
+\def\@pnumwidth{1.55em}
+\def\@tocrmarg{2.55em}
+\def\@dotsep{4.5}
+\setcounter{tocdepth}{3}
+
+% adjusted some spacings here so that section numbers will not easily
+% collide with the section titles.
+% VIII; VIII-A; and VIII-A.1 are usually the worst offenders.
+% MDS 1/2001
+\def\tableofcontents{\section*{\contentsname}\@starttoc{toc}}
+\def\l@section#1#2{\addpenalty{\@secpenalty}\addvspace{1.0em plus 1pt}%
+ \@tempdima 2.75em \begingroup \parindent \z@ \rightskip \@pnumwidth%
+ \parfillskip-\@pnumwidth {\bfseries\leavevmode #1}\hfil\hbox to\@pnumwidth{\hss #2}\par%
+ \endgroup}
+% argument format #1:level, #2:labelindent,#3:labelsep
+\def\l@subsection{\@dottedtocline{2}{2.75em}{3.75em}}
+\def\l@subsubsection{\@dottedtocline{3}{6.5em}{4.5em}}
+% must provide \l@ defs for ALL sublevels EVEN if tocdepth
+% is such as they will not appear in the table of contents
+% these defs are how TOC knows what level these things are!
+\def\l@paragraph{\@dottedtocline{4}{6.5em}{5.5em}}
+\def\l@subparagraph{\@dottedtocline{5}{6.5em}{6.5em}}
+\def\listoffigures{\section*{\listfigurename}\@starttoc{lof}}
+\def\l@figure{\@dottedtocline{1}{0em}{2.75em}}
+\def\listoftables{\section*{\listtablename}\@starttoc{lot}}
+\let\l@table\l@figure
+
+
+%% Definitions for floats
+%%
+%% Normal Floats
+\floatsep 1\baselineskip plus 0.2\baselineskip minus 0.2\baselineskip
+\textfloatsep 1.7\baselineskip plus 0.2\baselineskip minus 0.4\baselineskip
+\@fptop 0pt plus 1fil
+\@fpsep 0.75\baselineskip plus 2fil
+\@fpbot 0pt plus 1fil
+\def\topfraction{0.9}
+\def\bottomfraction{0.4}
+\def\floatpagefraction{0.8}
+% V1.7, let top floats approach 90% of page
+\def\textfraction{0.1}
+
+%% Double Column Floats
+\dblfloatsep 1\baselineskip plus 0.2\baselineskip minus 0.2\baselineskip
+
+\dbltextfloatsep 1.7\baselineskip plus 0.2\baselineskip minus 0.4\baselineskip
+% Note that it would be nice if the rubber here actually worked in LaTeX2e.
+% There is a long standing limitation in LaTeX, first discovered (to the best
+% of my knowledge) by Alan Jeffrey in 1992. LaTeX ignores the stretchable
+% portion of \dbltextfloatsep, and as a result, double column figures can and
+% do result in an non-integer number of lines in the main text columns with
+% underfull vbox errors as a consequence. A post to comp.text.tex
+% by Donald Arseneau confirms that this had not yet been fixed in 1998.
+% IEEEtran V1.6 will fix this problem for you in the titles, but it doesn't
+% protect you from other double floats. Happy vspace'ing.
+
+\@dblfptop 0pt plus 1fil
+\@dblfpsep 0.75\baselineskip plus 2fil
+\@dblfpbot 0pt plus 1fil
+\def\dbltopfraction{0.8}
+\def\dblfloatpagefraction{0.8}
+\setcounter{dbltopnumber}{4}
+
+\intextsep 1\baselineskip plus 0.2\baselineskip minus 0.2\baselineskip
+\setcounter{topnumber}{2}
+\setcounter{bottomnumber}{2}
+\setcounter{totalnumber}{4}
+
+
+
+% article class provides these, we should too.
+\newlength\abovecaptionskip
+\newlength\belowcaptionskip
+% but only \abovecaptionskip is used above figure captions and *below* table
+% captions
+\setlength\abovecaptionskip{0.5\baselineskip}
+\setlength\belowcaptionskip{0pt}
+% V1.6 create hooks in case the caption spacing ever needs to be
+% overridden by a user
+\def\@IEEEfigurecaptionsepspace{\vskip\abovecaptionskip\relax}%
+\def\@IEEEtablecaptionsepspace{\vskip\abovecaptionskip\relax}%
+
+
+% 1.6b revise caption system so that \@makecaption uses two arguments
+% as with LaTeX2e. Otherwise, there will be problems when using hyperref.
+\def\@IEEEtablestring{table}
+
+\ifCLASSOPTIONcompsoc
+% V1.7 compsoc \@makecaption
+\ifCLASSOPTIONconference% compsoc conference
+\long\def\@makecaption#1#2{%
+% test if is a for a figure or table
+\ifx\@captype\@IEEEtablestring%
+% if a table, do table caption
+\normalsize\begin{center}{\normalfont\sffamily\normalsize {#1.}~ #2}\end{center}%
+\@IEEEtablecaptionsepspace
+% if not a table, format it as a figure
+\else
+\@IEEEfigurecaptionsepspace
+\setbox\@tempboxa\hbox{\normalfont\sffamily\normalsize {#1.}~ #2}%
+\ifdim \wd\@tempboxa >\hsize%
+% if caption is longer than a line, let it wrap around
+\setbox\@tempboxa\hbox{\normalfont\sffamily\normalsize {#1.}~ }%
+\parbox[t]{\hsize}{\normalfont\sffamily\normalsize \noindent\unhbox\@tempboxa#2}%
+% if caption is shorter than a line, center
+\else%
+\hbox to\hsize{\normalfont\sffamily\normalsize\hfil\box\@tempboxa\hfil}%
+\fi\fi}
+\else% nonconference compsoc
+\long\def\@makecaption#1#2{%
+% test if is a for a figure or table
+\ifx\@captype\@IEEEtablestring%
+% if a table, do table caption
+\normalsize\begin{center}{\normalfont\sffamily\normalsize #1}\\{\normalfont\sffamily\normalsize #2}\end{center}%
+\@IEEEtablecaptionsepspace
+% if not a table, format it as a figure
+\else
+\@IEEEfigurecaptionsepspace
+\setbox\@tempboxa\hbox{\normalfont\sffamily\normalsize {#1.}~ #2}%
+\ifdim \wd\@tempboxa >\hsize%
+% if caption is longer than a line, let it wrap around
+\setbox\@tempboxa\hbox{\normalfont\sffamily\normalsize {#1.}~ }%
+\parbox[t]{\hsize}{\normalfont\sffamily\normalsize \noindent\unhbox\@tempboxa#2}%
+% if caption is shorter than a line, left justify
+\else%
+\hbox to\hsize{\normalfont\sffamily\normalsize\box\@tempboxa\hfil}%
+\fi\fi}
+\fi
+
+\else% traditional noncompsoc \@makecaption
+\long\def\@makecaption#1#2{%
+% test if is a for a figure or table
+\ifx\@captype\@IEEEtablestring%
+% if a table, do table caption
+\footnotesize\begin{center}{\normalfont\footnotesize #1}\\{\normalfont\footnotesize\scshape #2}\end{center}%
+\@IEEEtablecaptionsepspace
+% if not a table, format it as a figure
+\else
+\@IEEEfigurecaptionsepspace
+% 3/2001 use footnotesize, not small; use two nonbreaking spaces, not one
+\setbox\@tempboxa\hbox{\normalfont\footnotesize {#1.}~~ #2}%
+\ifdim \wd\@tempboxa >\hsize%
+% if caption is longer than a line, let it wrap around
+\setbox\@tempboxa\hbox{\normalfont\footnotesize {#1.}~~ }%
+\parbox[t]{\hsize}{\normalfont\footnotesize\noindent\unhbox\@tempboxa#2}%
+% if caption is shorter than a line, center if conference, left justify otherwise
+\else%
+\ifCLASSOPTIONconference \hbox to\hsize{\normalfont\footnotesize\hfil\box\@tempboxa\hfil}%
+\else \hbox to\hsize{\normalfont\footnotesize\box\@tempboxa\hfil}%
+\fi\fi\fi}
+\fi
+
+
+
+% V1.7 disable captions class option, do so in a way that retains operation of \label
+% within \caption
+\ifCLASSOPTIONcaptionsoff
+\long\def\@makecaption#1#2{\vspace*{2em}\footnotesize\begin{center}{\footnotesize #1}\end{center}%
+\let\@IEEEtemporiglabeldefsave\label
+\let\@IEEEtemplabelargsave\relax
+\def\label##1{\gdef\@IEEEtemplabelargsave{##1}}%
+\setbox\@tempboxa\hbox{#2}%
+\let\label\@IEEEtemporiglabeldefsave
+\ifx\@IEEEtemplabelargsave\relax\else\label{\@IEEEtemplabelargsave}\fi}
+\fi
+
+
+% V1.7 define end environments with \def not \let so as to work OK with
+% preview-latex
+\newcounter{figure}
+\def\thefigure{\@arabic\c@figure}
+\def\fps@figure{tbp}
+\def\ftype@figure{1}
+\def\ext@figure{lof}
+\def\fnum@figure{\figurename~\thefigure}
+\def\figure{\@float{figure}}
+\def\endfigure{\end@float}
+\@namedef{figure*}{\@dblfloat{figure}}
+\@namedef{endfigure*}{\end@dblfloat}
+\newcounter{table}
+\ifCLASSOPTIONcompsoc
+\def\thetable{\arabic{table}}
+\else
+\def\thetable{\@Roman\c@table}
+\fi
+\def\fps@table{tbp}
+\def\ftype@table{2}
+\def\ext@table{lot}
+\def\fnum@table{\tablename~\thetable}
+% V1.6 IEEE uses 8pt text for tables
+% to default to footnotesize, we hack into LaTeX2e's \@floatboxreset and pray
+\def\table{\def\@floatboxreset{\reset@font\footnotesize\@setminipage}\@float{table}}
+\def\endtable{\end@float}
+% v1.6b double column tables need to default to footnotesize as well.
+\@namedef{table*}{\def\@floatboxreset{\reset@font\footnotesize\@setminipage}\@dblfloat{table}}
+\@namedef{endtable*}{\end@dblfloat}
+
+
+
+
+%%
+%% START OF IEEEeqnarry DEFINITIONS
+%%
+%% Inspired by the concepts, examples, and previous works of LaTeX
+%% coders and developers such as Donald Arseneau, Fred Bartlett,
+%% David Carlisle, Tony Liu, Frank Mittelbach, Piet van Oostrum,
+%% Roland Winkler and Mark Wooding.
+%% I don't make the claim that my work here is even near their calibre. ;)
+
+
+% hook to allow easy changeover to IEEEtran.cls/tools.sty error reporting
+\def\@IEEEclspkgerror{\ClassError{IEEEtran}}
+
+\newif\if@IEEEeqnarraystarform% flag to indicate if the environment was called as the star form
+\@IEEEeqnarraystarformfalse
+
+\newif\if@advanceIEEEeqncolcnt% tracks if the environment should advance the col counter
+% allows a way to make an \IEEEeqnarraybox that can be used within an \IEEEeqnarray
+% used by IEEEeqnarraymulticol so that it can work properly in both
+\@advanceIEEEeqncolcnttrue
+
+\newcount\@IEEEeqnnumcols % tracks how many IEEEeqnarray cols are defined
+\newcount\@IEEEeqncolcnt % tracks how many IEEEeqnarray cols the user actually used
+
+
+% The default math style used by the columns
+\def\IEEEeqnarraymathstyle{\displaystyle}
+% The default text style used by the columns
+% default to using the current font
+\def\IEEEeqnarraytextstyle{\relax}
+
+% like the iedlistdecl but for \IEEEeqnarray
+\def\IEEEeqnarraydecl{\relax}
+\def\IEEEeqnarrayboxdecl{\relax}
+
+% \yesnumber is the opposite of \nonumber
+% a novel concept with the same def as the equationarray package
+% However, we give IEEE versions too since some LaTeX packages such as
+% the MDWtools mathenv.sty redefine \nonumber to something else.
+\providecommand{\yesnumber}{\global\@eqnswtrue}
+\def\IEEEyesnumber{\global\@eqnswtrue}
+\def\IEEEnonumber{\global\@eqnswfalse}
+
+
+\def\IEEEyessubnumber{\global\@IEEEissubequationtrue\global\@eqnswtrue%
+\if@IEEEeqnarrayISinner% only do something inside an IEEEeqnarray
+\if@IEEElastlinewassubequation\addtocounter{equation}{-1}\else\setcounter{IEEEsubequation}{1}\fi%
+\def\@currentlabel{\p@IEEEsubequation\theIEEEsubequation}\fi}
+
+% flag to indicate that an equation is a sub equation
+\newif\if@IEEEissubequation%
+\@IEEEissubequationfalse
+
+% allows users to "push away" equations that get too close to the equation numbers
+\def\IEEEeqnarraynumspace{\hphantom{\if@IEEEissubequation\theIEEEsubequationdis\else\theequationdis\fi}}
+
+% provides a way to span multiple columns within IEEEeqnarray environments
+% will consider \if@advanceIEEEeqncolcnt before globally advancing the
+% column counter - so as to work within \IEEEeqnarraybox
+% usage: \IEEEeqnarraymulticol{number cols. to span}{col type}{cell text}
+\long\def\IEEEeqnarraymulticol#1#2#3{\multispan{#1}%
+% check if column is defined
+\relax\expandafter\ifx\csname @IEEEeqnarraycolDEF#2\endcsname\@IEEEeqnarraycolisdefined%
+\csname @IEEEeqnarraycolPRE#2\endcsname#3\relax\relax\relax\relax\relax%
+\relax\relax\relax\relax\relax\csname @IEEEeqnarraycolPOST#2\endcsname%
+\else% if not, error and use default type
+\@IEEEclspkgerror{Invalid column type "#2" in \string\IEEEeqnarraymulticol.\MessageBreak
+Using a default centering column instead}%
+{You must define IEEEeqnarray column types before use.}%
+\csname @IEEEeqnarraycolPRE@IEEEdefault\endcsname#3\relax\relax\relax\relax\relax%
+\relax\relax\relax\relax\relax\csname @IEEEeqnarraycolPOST@IEEEdefault\endcsname%
+\fi%
+% advance column counter only if the IEEEeqnarray environment wants it
+\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by #1\relax\fi}
+
+% like \omit, but maintains track of the column counter for \IEEEeqnarray
+\def\IEEEeqnarrayomit{\omit\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by 1\relax\fi}
+
+
+% provides a way to define a letter referenced column type
+% usage: \IEEEeqnarraydefcol{col. type letter/name}{pre insertion text}{post insertion text}
+\def\IEEEeqnarraydefcol#1#2#3{\expandafter\def\csname @IEEEeqnarraycolPRE#1\endcsname{#2}%
+\expandafter\def\csname @IEEEeqnarraycolPOST#1\endcsname{#3}%
+\expandafter\def\csname @IEEEeqnarraycolDEF#1\endcsname{1}}
+
+
+% provides a way to define a numerically referenced inter-column glue types
+% usage: \IEEEeqnarraydefcolsep{col. glue number}{glue definition}
+\def\IEEEeqnarraydefcolsep#1#2{\expandafter\def\csname @IEEEeqnarraycolSEP\romannumeral #1\endcsname{#2}%
+\expandafter\def\csname @IEEEeqnarraycolSEPDEF\romannumeral #1\endcsname{1}}
+
+
+\def\@IEEEeqnarraycolisdefined{1}% just a macro for 1, used for checking undefined column types
+
+
+% expands and appends the given argument to the \@IEEEtrantmptoksA token list
+% used to build up the \halign preamble
+\def\@IEEEappendtoksA#1{\edef\@@IEEEappendtoksA{\@IEEEtrantmptoksA={\the\@IEEEtrantmptoksA #1}}%
+\@@IEEEappendtoksA}
+
+% also appends to \@IEEEtrantmptoksA, but does not expand the argument
+% uses \toks8 as a scratchpad register
+\def\@IEEEappendNOEXPANDtoksA#1{\toks8={#1}%
+\edef\@@IEEEappendNOEXPANDtoksA{\@IEEEtrantmptoksA={\the\@IEEEtrantmptoksA\the\toks8}}%
+\@@IEEEappendNOEXPANDtoksA}
+
+% define some common column types for the user
+% math
+\IEEEeqnarraydefcol{l}{$\IEEEeqnarraymathstyle}{$\hfil}
+\IEEEeqnarraydefcol{c}{\hfil$\IEEEeqnarraymathstyle}{$\hfil}
+\IEEEeqnarraydefcol{r}{\hfil$\IEEEeqnarraymathstyle}{$}
+\IEEEeqnarraydefcol{L}{$\IEEEeqnarraymathstyle{}}{{}$\hfil}
+\IEEEeqnarraydefcol{C}{\hfil$\IEEEeqnarraymathstyle{}}{{}$\hfil}
+\IEEEeqnarraydefcol{R}{\hfil$\IEEEeqnarraymathstyle{}}{{}$}
+% text
+\IEEEeqnarraydefcol{s}{\IEEEeqnarraytextstyle}{\hfil}
+\IEEEeqnarraydefcol{t}{\hfil\IEEEeqnarraytextstyle}{\hfil}
+\IEEEeqnarraydefcol{u}{\hfil\IEEEeqnarraytextstyle}{}
+
+% vertical rules
+\IEEEeqnarraydefcol{v}{}{\vrule width\arrayrulewidth}
+\IEEEeqnarraydefcol{vv}{\vrule width\arrayrulewidth\hfil}{\hfil\vrule width\arrayrulewidth}
+\IEEEeqnarraydefcol{V}{}{\vrule width\arrayrulewidth\hskip\doublerulesep\vrule width\arrayrulewidth}
+\IEEEeqnarraydefcol{VV}{\vrule width\arrayrulewidth\hskip\doublerulesep\vrule width\arrayrulewidth\hfil}%
+{\hfil\vrule width\arrayrulewidth\hskip\doublerulesep\vrule width\arrayrulewidth}
+
+% horizontal rules
+\IEEEeqnarraydefcol{h}{}{\leaders\hrule height\arrayrulewidth\hfil}
+\IEEEeqnarraydefcol{H}{}{\leaders\vbox{\hrule width\arrayrulewidth\vskip\doublerulesep\hrule width\arrayrulewidth}\hfil}
+
+% plain
+\IEEEeqnarraydefcol{x}{}{}
+\IEEEeqnarraydefcol{X}{$}{$}
+
+% the default column type to use in the event a column type is not defined
+\IEEEeqnarraydefcol{@IEEEdefault}{\hfil$\IEEEeqnarraymathstyle}{$\hfil}
+
+
+% a zero tabskip (used for "-" col types)
+\def\@IEEEeqnarraycolSEPzero{0pt plus 0pt minus 0pt}
+% a centering tabskip (used for "+" col types)
+\def\@IEEEeqnarraycolSEPcenter{1000pt plus 0pt minus 1000pt}
+
+% top level default tabskip glues for the start, end, and inter-column
+% may be reset within environments not always at the top level, e.g., \IEEEeqnarraybox
+\edef\@IEEEeqnarraycolSEPdefaultstart{\@IEEEeqnarraycolSEPcenter}% default start glue
+\edef\@IEEEeqnarraycolSEPdefaultend{\@IEEEeqnarraycolSEPcenter}% default end glue
+\edef\@IEEEeqnarraycolSEPdefaultmid{\@IEEEeqnarraycolSEPzero}% default inter-column glue
+
+
+
+% creates a vertical rule that extends from the bottom to the top a a cell
+% Provided in case other packages redefine \vline some other way.
+% usage: \IEEEeqnarrayvrule[rule thickness]
+% If no argument is provided, \arrayrulewidth will be used for the rule thickness.
+\newcommand\IEEEeqnarrayvrule[1][\arrayrulewidth]{\vrule\@width#1\relax}
+
+% creates a blank separator row
+% usage: \IEEEeqnarrayseprow[separation length][font size commands]
+% default is \IEEEeqnarrayseprow[0.25\normalbaselineskip][\relax]
+% blank arguments inherit the default values
+% uses \skip5 as a scratch register - calls \@IEEEeqnarraystrutsize which uses more scratch registers
+\def\IEEEeqnarrayseprow{\relax\@ifnextchar[{\@IEEEeqnarrayseprow}{\@IEEEeqnarrayseprow[0.25\normalbaselineskip]}}
+\def\@IEEEeqnarrayseprow[#1]{\relax\@ifnextchar[{\@@IEEEeqnarrayseprow[#1]}{\@@IEEEeqnarrayseprow[#1][\relax]}}
+\def\@@IEEEeqnarrayseprow[#1][#2]{\def\@IEEEeqnarrayseprowARGONE{#1}%
+\ifx\@IEEEeqnarrayseprowARGONE\@empty%
+% get the skip value, based on the font commands
+% use skip5 because \IEEEeqnarraystrutsize uses \skip0, \skip2, \skip3
+% assign within a bogus box to confine the font changes
+{\setbox0=\hbox{#2\relax\global\skip5=0.25\normalbaselineskip}}%
+\else%
+{\setbox0=\hbox{#2\relax\global\skip5=#1}}%
+\fi%
+\@IEEEeqnarrayhoptolastcolumn\IEEEeqnarraystrutsize{\skip5}{0pt}[\relax]\relax}
+
+% creates a blank separator row, but omits all the column templates
+% usage: \IEEEeqnarrayseprowcut[separation length][font size commands]
+% default is \IEEEeqnarrayseprowcut[0.25\normalbaselineskip][\relax]
+% blank arguments inherit the default values
+% uses \skip5 as a scratch register - calls \@IEEEeqnarraystrutsize which uses more scratch registers
+\def\IEEEeqnarrayseprowcut{\multispan{\@IEEEeqnnumcols}\relax% span all the cols
+% advance column counter only if the IEEEeqnarray environment wants it
+\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by \@IEEEeqnnumcols\relax\fi%
+\@ifnextchar[{\@IEEEeqnarrayseprowcut}{\@IEEEeqnarrayseprowcut[0.25\normalbaselineskip]}}
+\def\@IEEEeqnarrayseprowcut[#1]{\relax\@ifnextchar[{\@@IEEEeqnarrayseprowcut[#1]}{\@@IEEEeqnarrayseprowcut[#1][\relax]}}
+\def\@@IEEEeqnarrayseprowcut[#1][#2]{\def\@IEEEeqnarrayseprowARGONE{#1}%
+\ifx\@IEEEeqnarrayseprowARGONE\@empty%
+% get the skip value, based on the font commands
+% use skip5 because \IEEEeqnarraystrutsize uses \skip0, \skip2, \skip3
+% assign within a bogus box to confine the font changes
+{\setbox0=\hbox{#2\relax\global\skip5=0.25\normalbaselineskip}}%
+\else%
+{\setbox0=\hbox{#2\relax\global\skip5=#1}}%
+\fi%
+\IEEEeqnarraystrutsize{\skip5}{0pt}[\relax]\relax}
+
+
+
+% draws a single rule across all the columns optional
+% argument determines the rule width, \arrayrulewidth is the default
+% updates column counter as needed and turns off struts
+% usage: \IEEEeqnarrayrulerow[rule line thickness]
+\def\IEEEeqnarrayrulerow{\multispan{\@IEEEeqnnumcols}\relax% span all the cols
+% advance column counter only if the IEEEeqnarray environment wants it
+\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by \@IEEEeqnnumcols\relax\fi%
+\@ifnextchar[{\@IEEEeqnarrayrulerow}{\@IEEEeqnarrayrulerow[\arrayrulewidth]}}
+\def\@IEEEeqnarrayrulerow[#1]{\leaders\hrule height#1\hfil\relax% put in our rule
+% turn off any struts
+\IEEEeqnarraystrutsize{0pt}{0pt}[\relax]\relax}
+
+
+% draws a double rule by using a single rule row, a separator row, and then
+% another single rule row
+% first optional argument determines the rule thicknesses, \arrayrulewidth is the default
+% second optional argument determines the rule spacing, \doublerulesep is the default
+% usage: \IEEEeqnarraydblrulerow[rule line thickness][rule spacing]
+\def\IEEEeqnarraydblrulerow{\multispan{\@IEEEeqnnumcols}\relax% span all the cols
+% advance column counter only if the IEEEeqnarray environment wants it
+\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by \@IEEEeqnnumcols\relax\fi%
+\@ifnextchar[{\@IEEEeqnarraydblrulerow}{\@IEEEeqnarraydblrulerow[\arrayrulewidth]}}
+\def\@IEEEeqnarraydblrulerow[#1]{\relax\@ifnextchar[{\@@IEEEeqnarraydblrulerow[#1]}%
+{\@@IEEEeqnarraydblrulerow[#1][\doublerulesep]}}
+\def\@@IEEEeqnarraydblrulerow[#1][#2]{\def\@IEEEeqnarraydblrulerowARG{#1}%
+% we allow the user to say \IEEEeqnarraydblrulerow[][]
+\ifx\@IEEEeqnarraydblrulerowARG\@empty%
+\@IEEEeqnarrayrulerow[\arrayrulewidth]%
+\else%
+\@IEEEeqnarrayrulerow[#1]\relax%
+\fi%
+\def\@IEEEeqnarraydblrulerowARG{#2}%
+\ifx\@IEEEeqnarraydblrulerowARG\@empty%
+\\\IEEEeqnarrayseprow[\doublerulesep][\relax]%
+\else%
+\\\IEEEeqnarrayseprow[#2][\relax]%
+\fi%
+\\\multispan{\@IEEEeqnnumcols}%
+% advance column counter only if the IEEEeqnarray environment wants it
+\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by \@IEEEeqnnumcols\relax\fi%
+\def\@IEEEeqnarraydblrulerowARG{#1}%
+\ifx\@IEEEeqnarraydblrulerowARG\@empty%
+\@IEEEeqnarrayrulerow[\arrayrulewidth]%
+\else%
+\@IEEEeqnarrayrulerow[#1]%
+\fi%
+}
+
+% draws a double rule by using a single rule row, a separator (cutting) row, and then
+% another single rule row
+% first optional argument determines the rule thicknesses, \arrayrulewidth is the default
+% second optional argument determines the rule spacing, \doublerulesep is the default
+% usage: \IEEEeqnarraydblrulerow[rule line thickness][rule spacing]
+\def\IEEEeqnarraydblrulerowcut{\multispan{\@IEEEeqnnumcols}\relax% span all the cols
+% advance column counter only if the IEEEeqnarray environment wants it
+\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by \@IEEEeqnnumcols\relax\fi%
+\@ifnextchar[{\@IEEEeqnarraydblrulerowcut}{\@IEEEeqnarraydblrulerowcut[\arrayrulewidth]}}
+\def\@IEEEeqnarraydblrulerowcut[#1]{\relax\@ifnextchar[{\@@IEEEeqnarraydblrulerowcut[#1]}%
+{\@@IEEEeqnarraydblrulerowcut[#1][\doublerulesep]}}
+\def\@@IEEEeqnarraydblrulerowcut[#1][#2]{\def\@IEEEeqnarraydblrulerowARG{#1}%
+% we allow the user to say \IEEEeqnarraydblrulerow[][]
+\ifx\@IEEEeqnarraydblrulerowARG\@empty%
+\@IEEEeqnarrayrulerow[\arrayrulewidth]%
+\else%
+\@IEEEeqnarrayrulerow[#1]%
+\fi%
+\def\@IEEEeqnarraydblrulerowARG{#2}%
+\ifx\@IEEEeqnarraydblrulerowARG\@empty%
+\\\IEEEeqnarrayseprowcut[\doublerulesep][\relax]%
+\else%
+\\\IEEEeqnarrayseprowcut[#2][\relax]%
+\fi%
+\\\multispan{\@IEEEeqnnumcols}%
+% advance column counter only if the IEEEeqnarray environment wants it
+\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by \@IEEEeqnnumcols\relax\fi%
+\def\@IEEEeqnarraydblrulerowARG{#1}%
+\ifx\@IEEEeqnarraydblrulerowARG\@empty%
+\@IEEEeqnarrayrulerow[\arrayrulewidth]%
+\else%
+\@IEEEeqnarrayrulerow[#1]%
+\fi%
+}
+
+
+
+% inserts a full row's worth of &'s
+% relies on \@IEEEeqnnumcols to provide the correct number of columns
+% uses \@IEEEtrantmptoksA, \count0 as scratch registers
+\def\@IEEEeqnarrayhoptolastcolumn{\@IEEEtrantmptoksA={}\count0=1\relax%
+\loop% add cols if the user did not use them all
+\ifnum\count0<\@IEEEeqnnumcols\relax%
+\@IEEEappendtoksA{&}%
+\advance\count0 by 1\relax% update the col count
+\repeat%
+\the\@IEEEtrantmptoksA%execute the &'s
+}
+
+
+
+\newif\if@IEEEeqnarrayISinner % flag to indicate if we are within the lines
+\@IEEEeqnarrayISinnerfalse % of an IEEEeqnarray - after the IEEEeqnarraydecl
+
+\edef\@IEEEeqnarrayTHEstrutheight{0pt} % height and depth of IEEEeqnarray struts
+\edef\@IEEEeqnarrayTHEstrutdepth{0pt}
+
+\edef\@IEEEeqnarrayTHEmasterstrutheight{0pt} % default height and depth of
+\edef\@IEEEeqnarrayTHEmasterstrutdepth{0pt} % struts within an IEEEeqnarray
+
+\edef\@IEEEeqnarrayTHEmasterstrutHSAVE{0pt} % saved master strut height
+\edef\@IEEEeqnarrayTHEmasterstrutDSAVE{0pt} % and depth
+
+\newif\if@IEEEeqnarrayusemasterstrut % flag to indicate that the master strut value
+\@IEEEeqnarrayusemasterstruttrue % is to be used
+
+
+
+% saves the strut height and depth of the master strut
+\def\@IEEEeqnarraymasterstrutsave{\relax%
+\expandafter\skip0=\@IEEEeqnarrayTHEmasterstrutheight\relax%
+\expandafter\skip2=\@IEEEeqnarrayTHEmasterstrutdepth\relax%
+% remove stretchability
+\dimen0\skip0\relax%
+\dimen2\skip2\relax%
+% save values
+\edef\@IEEEeqnarrayTHEmasterstrutHSAVE{\the\dimen0}%
+\edef\@IEEEeqnarrayTHEmasterstrutDSAVE{\the\dimen2}}
+
+% restores the strut height and depth of the master strut
+\def\@IEEEeqnarraymasterstrutrestore{\relax%
+\expandafter\skip0=\@IEEEeqnarrayTHEmasterstrutHSAVE\relax%
+\expandafter\skip2=\@IEEEeqnarrayTHEmasterstrutDSAVE\relax%
+% remove stretchability
+\dimen0\skip0\relax%
+\dimen2\skip2\relax%
+% restore values
+\edef\@IEEEeqnarrayTHEmasterstrutheight{\the\dimen0}%
+\edef\@IEEEeqnarrayTHEmasterstrutdepth{\the\dimen2}}
+
+
+% globally restores the strut height and depth to the
+% master values and sets the master strut flag to true
+\def\@IEEEeqnarraystrutreset{\relax%
+\expandafter\skip0=\@IEEEeqnarrayTHEmasterstrutheight\relax%
+\expandafter\skip2=\@IEEEeqnarrayTHEmasterstrutdepth\relax%
+% remove stretchability
+\dimen0\skip0\relax%
+\dimen2\skip2\relax%
+% restore values
+\xdef\@IEEEeqnarrayTHEstrutheight{\the\dimen0}%
+\xdef\@IEEEeqnarrayTHEstrutdepth{\the\dimen2}%
+\global\@IEEEeqnarrayusemasterstruttrue}
+
+
+% if the master strut is not to be used, make the current
+% values of \@IEEEeqnarrayTHEstrutheight, \@IEEEeqnarrayTHEstrutdepth
+% and the use master strut flag, global
+% this allows user strut commands issued in the last column to be carried
+% into the isolation/strut column
+\def\@IEEEeqnarrayglobalizestrutstatus{\relax%
+\if@IEEEeqnarrayusemasterstrut\else%
+\xdef\@IEEEeqnarrayTHEstrutheight{\@IEEEeqnarrayTHEstrutheight}%
+\xdef\@IEEEeqnarrayTHEstrutdepth{\@IEEEeqnarrayTHEstrutdepth}%
+\global\@IEEEeqnarrayusemasterstrutfalse%
+\fi}
+
+
+
+% usage: \IEEEeqnarraystrutsize{height}{depth}[font size commands]
+% If called outside the lines of an IEEEeqnarray, sets the height
+% and depth of both the master and local struts. If called inside
+% an IEEEeqnarray line, sets the height and depth of the local strut
+% only and sets the flag to indicate the use of the local strut
+% values. If the height or depth is left blank, 0.7\normalbaselineskip
+% and 0.3\normalbaselineskip will be used, respectively.
+% The optional argument can be used to evaluate the lengths under
+% a different font size and styles. If none is specified, the current
+% font is used.
+% uses scratch registers \skip0, \skip2, \skip3, \dimen0, \dimen2
+\def\IEEEeqnarraystrutsize#1#2{\relax\@ifnextchar[{\@IEEEeqnarraystrutsize{#1}{#2}}{\@IEEEeqnarraystrutsize{#1}{#2}[\relax]}}
+\def\@IEEEeqnarraystrutsize#1#2[#3]{\def\@IEEEeqnarraystrutsizeARG{#1}%
+\ifx\@IEEEeqnarraystrutsizeARG\@empty%
+{\setbox0=\hbox{#3\relax\global\skip3=0.7\normalbaselineskip}}%
+\skip0=\skip3\relax%
+\else% arg one present
+{\setbox0=\hbox{#3\relax\global\skip3=#1\relax}}%
+\skip0=\skip3\relax%
+\fi% if null arg
+\def\@IEEEeqnarraystrutsizeARG{#2}%
+\ifx\@IEEEeqnarraystrutsizeARG\@empty%
+{\setbox0=\hbox{#3\relax\global\skip3=0.3\normalbaselineskip}}%
+\skip2=\skip3\relax%
+\else% arg two present
+{\setbox0=\hbox{#3\relax\global\skip3=#2\relax}}%
+\skip2=\skip3\relax%
+\fi% if null arg
+% remove stretchability, just to be safe
+\dimen0\skip0\relax%
+\dimen2\skip2\relax%
+% dimen0 = height, dimen2 = depth
+\if@IEEEeqnarrayISinner% inner does not touch master strut size
+\edef\@IEEEeqnarrayTHEstrutheight{\the\dimen0}%
+\edef\@IEEEeqnarrayTHEstrutdepth{\the\dimen2}%
+\@IEEEeqnarrayusemasterstrutfalse% do not use master
+\else% outer, have to set master strut too
+\edef\@IEEEeqnarrayTHEmasterstrutheight{\the\dimen0}%
+\edef\@IEEEeqnarrayTHEmasterstrutdepth{\the\dimen2}%
+\edef\@IEEEeqnarrayTHEstrutheight{\the\dimen0}%
+\edef\@IEEEeqnarrayTHEstrutdepth{\the\dimen2}%
+\@IEEEeqnarrayusemasterstruttrue% use master strut
+\fi}
+
+
+% usage: \IEEEeqnarraystrutsizeadd{added height}{added depth}[font size commands]
+% If called outside the lines of an IEEEeqnarray, adds the given height
+% and depth to both the master and local struts.
+% If called inside an IEEEeqnarray line, adds the given height and depth
+% to the local strut only and sets the flag to indicate the use
+% of the local strut values.
+% In both cases, if a height or depth is left blank, 0pt is used instead.
+% The optional argument can be used to evaluate the lengths under
+% a different font size and styles. If none is specified, the current
+% font is used.
+% uses scratch registers \skip0, \skip2, \skip3, \dimen0, \dimen2
+\def\IEEEeqnarraystrutsizeadd#1#2{\relax\@ifnextchar[{\@IEEEeqnarraystrutsizeadd{#1}{#2}}{\@IEEEeqnarraystrutsizeadd{#1}{#2}[\relax]}}
+\def\@IEEEeqnarraystrutsizeadd#1#2[#3]{\def\@IEEEeqnarraystrutsizearg{#1}%
+\ifx\@IEEEeqnarraystrutsizearg\@empty%
+\skip0=0pt\relax%
+\else% arg one present
+{\setbox0=\hbox{#3\relax\global\skip3=#1}}%
+\skip0=\skip3\relax%
+\fi% if null arg
+\def\@IEEEeqnarraystrutsizearg{#2}%
+\ifx\@IEEEeqnarraystrutsizearg\@empty%
+\skip2=0pt\relax%
+\else% arg two present
+{\setbox0=\hbox{#3\relax\global\skip3=#2}}%
+\skip2=\skip3\relax%
+\fi% if null arg
+% remove stretchability, just to be safe
+\dimen0\skip0\relax%
+\dimen2\skip2\relax%
+% dimen0 = height, dimen2 = depth
+\if@IEEEeqnarrayISinner% inner does not touch master strut size
+% get local strut size
+\expandafter\skip0=\@IEEEeqnarrayTHEstrutheight\relax%
+\expandafter\skip2=\@IEEEeqnarrayTHEstrutdepth\relax%
+% add it to the user supplied values
+\advance\dimen0 by \skip0\relax%
+\advance\dimen2 by \skip2\relax%
+% update the local strut size
+\edef\@IEEEeqnarrayTHEstrutheight{\the\dimen0}%
+\edef\@IEEEeqnarrayTHEstrutdepth{\the\dimen2}%
+\@IEEEeqnarrayusemasterstrutfalse% do not use master
+\else% outer, have to set master strut too
+% get master strut size
+\expandafter\skip0=\@IEEEeqnarrayTHEmasterstrutheight\relax%
+\expandafter\skip2=\@IEEEeqnarrayTHEmasterstrutdepth\relax%
+% add it to the user supplied values
+\advance\dimen0 by \skip0\relax%
+\advance\dimen2 by \skip2\relax%
+% update the local and master strut sizes
+\edef\@IEEEeqnarrayTHEmasterstrutheight{\the\dimen0}%
+\edef\@IEEEeqnarrayTHEmasterstrutdepth{\the\dimen2}%
+\edef\@IEEEeqnarrayTHEstrutheight{\the\dimen0}%
+\edef\@IEEEeqnarrayTHEstrutdepth{\the\dimen2}%
+\@IEEEeqnarrayusemasterstruttrue% use master strut
+\fi}
+
+
+% allow user a way to see the struts
+\newif\ifIEEEvisiblestruts
+\IEEEvisiblestrutsfalse
+
+% inserts an invisible strut using the master or local strut values
+% uses scratch registers \skip0, \skip2, \dimen0, \dimen2
+\def\@IEEEeqnarrayinsertstrut{\relax%
+\if@IEEEeqnarrayusemasterstrut
+% get master strut size
+\expandafter\skip0=\@IEEEeqnarrayTHEmasterstrutheight\relax%
+\expandafter\skip2=\@IEEEeqnarrayTHEmasterstrutdepth\relax%
+\else%
+% get local strut size
+\expandafter\skip0=\@IEEEeqnarrayTHEstrutheight\relax%
+\expandafter\skip2=\@IEEEeqnarrayTHEstrutdepth\relax%
+\fi%
+% remove stretchability, probably not needed
+\dimen0\skip0\relax%
+\dimen2\skip2\relax%
+% dimen0 = height, dimen2 = depth
+% allow user to see struts if desired
+\ifIEEEvisiblestruts%
+\vrule width0.2pt height\dimen0 depth\dimen2\relax%
+\else%
+\vrule width0pt height\dimen0 depth\dimen2\relax\fi}
+
+
+% creates an invisible strut, useable even outside \IEEEeqnarray
+% if \IEEEvisiblestrutstrue, the strut will be visible and 0.2pt wide.
+% usage: \IEEEstrut[height][depth][font size commands]
+% default is \IEEEstrut[0.7\normalbaselineskip][0.3\normalbaselineskip][\relax]
+% blank arguments inherit the default values
+% uses \dimen0, \dimen2, \skip0, \skip2
+\def\IEEEstrut{\relax\@ifnextchar[{\@IEEEstrut}{\@IEEEstrut[0.7\normalbaselineskip]}}
+\def\@IEEEstrut[#1]{\relax\@ifnextchar[{\@@IEEEstrut[#1]}{\@@IEEEstrut[#1][0.3\normalbaselineskip]}}
+\def\@@IEEEstrut[#1][#2]{\relax\@ifnextchar[{\@@@IEEEstrut[#1][#2]}{\@@@IEEEstrut[#1][#2][\relax]}}
+\def\@@@IEEEstrut[#1][#2][#3]{\mbox{#3\relax%
+\def\@IEEEstrutARG{#1}%
+\ifx\@IEEEstrutARG\@empty%
+\skip0=0.7\normalbaselineskip\relax%
+\else%
+\skip0=#1\relax%
+\fi%
+\def\@IEEEstrutARG{#2}%
+\ifx\@IEEEstrutARG\@empty%
+\skip2=0.3\normalbaselineskip\relax%
+\else%
+\skip2=#2\relax%
+\fi%
+% remove stretchability, probably not needed
+\dimen0\skip0\relax%
+\dimen2\skip2\relax%
+\ifIEEEvisiblestruts%
+\vrule width0.2pt height\dimen0 depth\dimen2\relax%
+\else%
+\vrule width0.0pt height\dimen0 depth\dimen2\relax\fi}}
+
+
+% enables strut mode by setting a default strut size and then zeroing the
+% \baselineskip, \lineskip, \lineskiplimit and \jot
+\def\IEEEeqnarraystrutmode{\IEEEeqnarraystrutsize{0.7\normalbaselineskip}{0.3\normalbaselineskip}[\relax]%
+\baselineskip=0pt\lineskip=0pt\lineskiplimit=0pt\jot=0pt}
+
+
+
+\def\IEEEeqnarray{\@IEEEeqnarraystarformfalse\@IEEEeqnarray}
+\def\endIEEEeqnarray{\end@IEEEeqnarray}
+
+\@namedef{IEEEeqnarray*}{\@IEEEeqnarraystarformtrue\@IEEEeqnarray}
+\@namedef{endIEEEeqnarray*}{\end@IEEEeqnarray}
+
+
+% \IEEEeqnarray is an enhanced \eqnarray.
+% The star form defaults to not putting equation numbers at the end of each row.
+% usage: \IEEEeqnarray[decl]{cols}
+\def\@IEEEeqnarray{\relax\@ifnextchar[{\@@IEEEeqnarray}{\@@IEEEeqnarray[\relax]}}
+\def\@@IEEEeqnarray[#1]#2{%
+ % default to showing the equation number or not based on whether or not
+ % the star form was involked
+ \if@IEEEeqnarraystarform\global\@eqnswfalse
+ \else% not the star form
+ \global\@eqnswtrue
+ \fi% if star form
+ \@IEEEissubequationfalse% default to no subequations
+ \@IEEElastlinewassubequationfalse% assume last line is not a sub equation
+ \@IEEEeqnarrayISinnerfalse% not yet within the lines of the halign
+ \@IEEEeqnarraystrutsize{0pt}{0pt}[\relax]% turn off struts by default
+ \@IEEEeqnarrayusemasterstruttrue% use master strut till user asks otherwise
+ \IEEEvisiblestrutsfalse% diagnostic mode defaults to off
+ % no extra space unless the user specifically requests it
+ \lineskip=0pt\relax
+ \lineskiplimit=0pt\relax
+ \baselineskip=\normalbaselineskip\relax%
+ \jot=\IEEEnormaljot\relax%
+ \mathsurround\z@\relax% no extra spacing around math
+ \@advanceIEEEeqncolcnttrue% advance the col counter for each col the user uses,
+ % used in \IEEEeqnarraymulticol and in the preamble build
+ \stepcounter{equation}% advance equation counter before first line
+ \setcounter{IEEEsubequation}{0}% no subequation yet
+ \def\@currentlabel{\p@equation\theequation}% redefine the ref label
+ \IEEEeqnarraydecl\relax% allow a way for the user to make global overrides
+ #1\relax% allow user to override defaults
+ \let\\\@IEEEeqnarraycr% replace newline with one that can put in eqn. numbers
+ \global\@IEEEeqncolcnt\z@% col. count = 0 for first line
+ \@IEEEbuildpreamble #2\end\relax% build the preamble and put it into \@IEEEtrantmptoksA
+ % put in the column for the equation number
+ \ifnum\@IEEEeqnnumcols>0\relax\@IEEEappendtoksA{&}\fi% col separator for those after the first
+ \toks0={##}%
+ % advance the \@IEEEeqncolcnt for the isolation col, this helps with error checking
+ \@IEEEappendtoksA{\global\advance\@IEEEeqncolcnt by 1\relax}%
+ % add the isolation column
+ \@IEEEappendtoksA{\tabskip\z@skip\bgroup\the\toks0\egroup}%
+ % advance the \@IEEEeqncolcnt for the equation number col, this helps with error checking
+ \@IEEEappendtoksA{&\global\advance\@IEEEeqncolcnt by 1\relax}%
+ % add the equation number col to the preamble
+ \@IEEEappendtoksA{\tabskip\z@skip\hb@xt@\z@\bgroup\hss\the\toks0\egroup}%
+ % note \@IEEEeqnnumcols does not count the equation col or isolation col
+ % set the starting tabskip glue as determined by the preamble build
+ \tabskip=\@IEEEBPstartglue\relax
+ % begin the display alignment
+ \@IEEEeqnarrayISinnertrue% commands are now within the lines
+ $$\everycr{}\halign to\displaywidth\bgroup
+ % "exspand" the preamble
+ \span\the\@IEEEtrantmptoksA\cr}
+
+% enter isolation/strut column (or the next column if the user did not use
+% every column), record the strut status, complete the columns, do the strut if needed,
+% restore counters to correct values and exit
+\def\end@IEEEeqnarray{\@IEEEeqnarrayglobalizestrutstatus&\@@IEEEeqnarraycr\egroup%
+\if@IEEElastlinewassubequation\global\advance\c@IEEEsubequation\m@ne\fi%
+\global\advance\c@equation\m@ne%
+$$\@ignoretrue}
+
+% need a way to remember if last line is a subequation
+\newif\if@IEEElastlinewassubequation%
+\@IEEElastlinewassubequationfalse
+
+% IEEEeqnarray uses a modifed \\ instead of the plain \cr to
+% end rows. This allows for things like \\*[vskip amount]
+% This "cr" macros are modified versions those for LaTeX2e's eqnarray
+% the {\ifnum0=`} braces must be kept away from the last column to avoid
+% altering spacing of its math, so we use & to advance to the next column
+% as there is an isolation/strut column after the user's columns
+\def\@IEEEeqnarraycr{\@IEEEeqnarrayglobalizestrutstatus&% save strut status and advance to next column
+ {\ifnum0=`}\fi
+ \@ifstar{%
+ \global\@eqpen\@M\@IEEEeqnarrayYCR
+ }{%
+ \global\@eqpen\interdisplaylinepenalty \@IEEEeqnarrayYCR
+ }%
+}
+
+\def\@IEEEeqnarrayYCR{\@testopt\@IEEEeqnarrayXCR\z@skip}
+
+\def\@IEEEeqnarrayXCR[#1]{%
+ \ifnum0=`{\fi}%
+ \@@IEEEeqnarraycr
+ \noalign{\penalty\@eqpen\vskip\jot\vskip #1\relax}}%
+
+\def\@@IEEEeqnarraycr{\@IEEEtrantmptoksA={}% clear token register
+ \advance\@IEEEeqncolcnt by -1\relax% adjust col count because of the isolation column
+ \ifnum\@IEEEeqncolcnt>\@IEEEeqnnumcols\relax
+ \@IEEEclspkgerror{Too many columns within the IEEEeqnarray\MessageBreak
+ environment}%
+ {Use fewer \string &'s or put more columns in the IEEEeqnarry column\MessageBreak
+ specifications.}\relax%
+ \else
+ \loop% add cols if the user did not use them all
+ \ifnum\@IEEEeqncolcnt<\@IEEEeqnnumcols\relax
+ \@IEEEappendtoksA{&}%
+ \advance\@IEEEeqncolcnt by 1\relax% update the col count
+ \repeat
+ % this number of &'s will take us the the isolation column
+ \fi
+ % execute the &'s
+ \the\@IEEEtrantmptoksA%
+ % handle the strut/isolation column
+ \@IEEEeqnarrayinsertstrut% do the strut if needed
+ \@IEEEeqnarraystrutreset% reset the strut system for next line or IEEEeqnarray
+ &% and enter the equation number column
+ % is this line needs an equation number, display it and advance the
+ % (sub)equation counters, record what type this line was
+ \if@eqnsw%
+ \if@IEEEissubequation\theIEEEsubequationdis\addtocounter{equation}{1}\stepcounter{IEEEsubequation}%
+ \global\@IEEElastlinewassubequationtrue%
+ \else% display a standard equation number, initialize the IEEEsubequation counter
+ \theequationdis\stepcounter{equation}\setcounter{IEEEsubequation}{0}%
+ \global\@IEEElastlinewassubequationfalse\fi%
+ \fi%
+ % reset the eqnsw flag to indicate default preference of the display of equation numbers
+ \if@IEEEeqnarraystarform\global\@eqnswfalse\else\global\@eqnswtrue\fi
+ \global\@IEEEissubequationfalse% reset the subequation flag
+ % reset the number of columns the user actually used
+ \global\@IEEEeqncolcnt\z@\relax
+ % the real end of the line
+ \cr}
+
+
+
+
+
+% \IEEEeqnarraybox is like \IEEEeqnarray except the box form puts everything
+% inside a vtop, vbox, or vcenter box depending on the letter in the second
+% optional argument (t,b,c). Vbox is the default. Unlike \IEEEeqnarray,
+% equation numbers are not displayed and \IEEEeqnarraybox can be nested.
+% \IEEEeqnarrayboxm is for math mode (like \array) and does not put the vbox
+% within an hbox.
+% \IEEEeqnarrayboxt is for text mode (like \tabular) and puts the vbox within
+% a \hbox{$ $} construct.
+% \IEEEeqnarraybox will auto detect whether to use \IEEEeqnarrayboxm or
+% \IEEEeqnarrayboxt depending on the math mode.
+% The third optional argument specifies the width this box is to be set to -
+% natural width is the default.
+% The * forms do not add \jot line spacing
+% usage: \IEEEeqnarraybox[decl][pos][width]{cols}
+\def\IEEEeqnarrayboxm{\@IEEEeqnarraystarformfalse\@IEEEeqnarrayboxHBOXSWfalse\@IEEEeqnarraybox}
+\def\endIEEEeqnarrayboxm{\end@IEEEeqnarraybox}
+\@namedef{IEEEeqnarrayboxm*}{\@IEEEeqnarraystarformtrue\@IEEEeqnarrayboxHBOXSWfalse\@IEEEeqnarraybox}
+\@namedef{endIEEEeqnarrayboxm*}{\end@IEEEeqnarraybox}
+
+\def\IEEEeqnarrayboxt{\@IEEEeqnarraystarformfalse\@IEEEeqnarrayboxHBOXSWtrue\@IEEEeqnarraybox}
+\def\endIEEEeqnarrayboxt{\end@IEEEeqnarraybox}
+\@namedef{IEEEeqnarrayboxt*}{\@IEEEeqnarraystarformtrue\@IEEEeqnarrayboxHBOXSWtrue\@IEEEeqnarraybox}
+\@namedef{endIEEEeqnarrayboxt*}{\end@IEEEeqnarraybox}
+
+\def\IEEEeqnarraybox{\@IEEEeqnarraystarformfalse\ifmmode\@IEEEeqnarrayboxHBOXSWfalse\else\@IEEEeqnarrayboxHBOXSWtrue\fi%
+\@IEEEeqnarraybox}
+\def\endIEEEeqnarraybox{\end@IEEEeqnarraybox}
+
+\@namedef{IEEEeqnarraybox*}{\@IEEEeqnarraystarformtrue\ifmmode\@IEEEeqnarrayboxHBOXSWfalse\else\@IEEEeqnarrayboxHBOXSWtrue\fi%
+\@IEEEeqnarraybox}
+\@namedef{endIEEEeqnarraybox*}{\end@IEEEeqnarraybox}
+
+% flag to indicate if the \IEEEeqnarraybox needs to put things into an hbox{$ $}
+% for \vcenter in non-math mode
+\newif\if@IEEEeqnarrayboxHBOXSW%
+\@IEEEeqnarrayboxHBOXSWfalse
+
+\def\@IEEEeqnarraybox{\relax\@ifnextchar[{\@@IEEEeqnarraybox}{\@@IEEEeqnarraybox[\relax]}}
+\def\@@IEEEeqnarraybox[#1]{\relax\@ifnextchar[{\@@@IEEEeqnarraybox[#1]}{\@@@IEEEeqnarraybox[#1][b]}}
+\def\@@@IEEEeqnarraybox[#1][#2]{\relax\@ifnextchar[{\@@@@IEEEeqnarraybox[#1][#2]}{\@@@@IEEEeqnarraybox[#1][#2][\relax]}}
+
+% #1 = decl; #2 = t,b,c; #3 = width, #4 = col specs
+\def\@@@@IEEEeqnarraybox[#1][#2][#3]#4{\@IEEEeqnarrayISinnerfalse % not yet within the lines of the halign
+ \@IEEEeqnarraymasterstrutsave% save current master strut values
+ \@IEEEeqnarraystrutsize{0pt}{0pt}[\relax]% turn off struts by default
+ \@IEEEeqnarrayusemasterstruttrue% use master strut till user asks otherwise
+ \IEEEvisiblestrutsfalse% diagnostic mode defaults to off
+ % no extra space unless the user specifically requests it
+ \lineskip=0pt\relax%
+ \lineskiplimit=0pt\relax%
+ \baselineskip=\normalbaselineskip\relax%
+ \jot=\IEEEnormaljot\relax%
+ \mathsurround\z@\relax% no extra spacing around math
+ % the default end glues are zero for an \IEEEeqnarraybox
+ \edef\@IEEEeqnarraycolSEPdefaultstart{\@IEEEeqnarraycolSEPzero}% default start glue
+ \edef\@IEEEeqnarraycolSEPdefaultend{\@IEEEeqnarraycolSEPzero}% default end glue
+ \edef\@IEEEeqnarraycolSEPdefaultmid{\@IEEEeqnarraycolSEPzero}% default inter-column glue
+ \@advanceIEEEeqncolcntfalse% do not advance the col counter for each col the user uses,
+ % used in \IEEEeqnarraymulticol and in the preamble build
+ \IEEEeqnarrayboxdecl\relax% allow a way for the user to make global overrides
+ #1\relax% allow user to override defaults
+ \let\\\@IEEEeqnarrayboxcr% replace newline with one that allows optional spacing
+ \@IEEEbuildpreamble #4\end\relax% build the preamble and put it into \@IEEEtrantmptoksA
+ % add an isolation column to the preamble to stop \\'s {} from getting into the last col
+ \ifnum\@IEEEeqnnumcols>0\relax\@IEEEappendtoksA{&}\fi% col separator for those after the first
+ \toks0={##}%
+ % add the isolation column to the preamble
+ \@IEEEappendtoksA{\tabskip\z@skip\bgroup\the\toks0\egroup}%
+ % set the starting tabskip glue as determined by the preamble build
+ \tabskip=\@IEEEBPstartglue\relax
+ % begin the alignment
+ \everycr{}%
+ % use only the very first token to determine the positioning
+ % this stops some problems when the user uses more than one letter,
+ % but is probably not worth the effort
+ % \noindent is used as a delimiter
+ \def\@IEEEgrabfirstoken##1##2\noindent{\let\@IEEEgrabbedfirstoken=##1}%
+ \@IEEEgrabfirstoken#2\relax\relax\noindent
+ % \@IEEEgrabbedfirstoken has the first token, the rest are discarded
+ % if we need to put things into and hbox and go into math mode, do so now
+ \if@IEEEeqnarrayboxHBOXSW \leavevmode \hbox \bgroup $\fi%
+ % use the appropriate vbox type
+ \if\@IEEEgrabbedfirstoken t\relax\vtop\else\if\@IEEEgrabbedfirstoken c\relax%
+ \vcenter\else\vbox\fi\fi\bgroup%
+ \@IEEEeqnarrayISinnertrue% commands are now within the lines
+ \ifx#3\relax\halign\else\halign to #3\relax\fi%
+ \bgroup
+ % "exspand" the preamble
+ \span\the\@IEEEtrantmptoksA\cr}
+
+% carry strut status and enter the isolation/strut column,
+% exit from math mode if needed, and exit
+\def\end@IEEEeqnarraybox{\@IEEEeqnarrayglobalizestrutstatus% carry strut status
+&% enter isolation/strut column
+\@IEEEeqnarrayinsertstrut% do strut if needed
+\@IEEEeqnarraymasterstrutrestore% restore the previous master strut values
+% reset the strut system for next IEEEeqnarray
+% (sets local strut values back to previous master strut values)
+\@IEEEeqnarraystrutreset%
+% ensure last line, exit from halign, close vbox
+\crcr\egroup\egroup%
+% exit from math mode and close hbox if needed
+\if@IEEEeqnarrayboxHBOXSW $\egroup\fi}
+
+
+
+% IEEEeqnarraybox uses a modifed \\ instead of the plain \cr to
+% end rows. This allows for things like \\[vskip amount]
+% This "cr" macros are modified versions those for LaTeX2e's eqnarray
+% For IEEEeqnarraybox, \\* is the same as \\
+% the {\ifnum0=`} braces must be kept away from the last column to avoid
+% altering spacing of its math, so we use & to advance to the isolation/strut column
+% carry strut status into isolation/strut column
+\def\@IEEEeqnarrayboxcr{\@IEEEeqnarrayglobalizestrutstatus% carry strut status
+&% enter isolation/strut column
+\@IEEEeqnarrayinsertstrut% do strut if needed
+% reset the strut system for next line or IEEEeqnarray
+\@IEEEeqnarraystrutreset%
+{\ifnum0=`}\fi%
+\@ifstar{\@IEEEeqnarrayboxYCR}{\@IEEEeqnarrayboxYCR}}
+
+% test and setup the optional argument to \\[]
+\def\@IEEEeqnarrayboxYCR{\@testopt\@IEEEeqnarrayboxXCR\z@skip}
+
+% IEEEeqnarraybox does not automatically increase line spacing by \jot
+\def\@IEEEeqnarrayboxXCR[#1]{\ifnum0=`{\fi}%
+\cr\noalign{\if@IEEEeqnarraystarform\else\vskip\jot\fi\vskip#1\relax}}
+
+
+
+% starts the halign preamble build
+\def\@IEEEbuildpreamble{\@IEEEtrantmptoksA={}% clear token register
+\let\@IEEEBPcurtype=u%current column type is not yet known
+\let\@IEEEBPprevtype=s%the previous column type was the start
+\let\@IEEEBPnexttype=u%next column type is not yet known
+% ensure these are valid
+\def\@IEEEBPcurglue={0pt plus 0pt minus 0pt}%
+\def\@IEEEBPcurcolname{@IEEEdefault}% name of current column definition
+% currently acquired numerically referenced glue
+% use a name that is easier to remember
+\let\@IEEEBPcurnum=\@IEEEtrantmpcountA%
+\@IEEEBPcurnum=0%
+% tracks number of columns in the preamble
+\@IEEEeqnnumcols=0%
+% record the default end glues
+\edef\@IEEEBPstartglue{\@IEEEeqnarraycolSEPdefaultstart}%
+\edef\@IEEEBPendglue{\@IEEEeqnarraycolSEPdefaultend}%
+% now parse the user's column specifications
+\@@IEEEbuildpreamble}
+
+
+% parses and builds the halign preamble
+\def\@@IEEEbuildpreamble#1#2{\let\@@nextIEEEbuildpreamble=\@@IEEEbuildpreamble%
+% use only the very first token to check the end
+% \noindent is used as a delimiter as \end can be present here
+\def\@IEEEgrabfirstoken##1##2\noindent{\let\@IEEEgrabbedfirstoken=##1}%
+\@IEEEgrabfirstoken#1\relax\relax\noindent
+\ifx\@IEEEgrabbedfirstoken\end\let\@@nextIEEEbuildpreamble=\@@IEEEfinishpreamble\else%
+% identify current and next token type
+\@IEEEgetcoltype{#1}{\@IEEEBPcurtype}{1}% current, error on invalid
+\@IEEEgetcoltype{#2}{\@IEEEBPnexttype}{0}% next, no error on invalid next
+% if curtype is a glue, get the glue def
+\if\@IEEEBPcurtype g\@IEEEgetcurglue{#1}{\@IEEEBPcurglue}\fi%
+% if curtype is a column, get the column def and set the current column name
+\if\@IEEEBPcurtype c\@IEEEgetcurcol{#1}\fi%
+% if curtype is a numeral, acquire the user defined glue
+\if\@IEEEBPcurtype n\@IEEEprocessNcol{#1}\fi%
+% process the acquired glue
+\if\@IEEEBPcurtype g\@IEEEprocessGcol\fi%
+% process the acquired col
+\if\@IEEEBPcurtype c\@IEEEprocessCcol\fi%
+% ready prevtype for next col spec.
+\let\@IEEEBPprevtype=\@IEEEBPcurtype%
+% be sure and put back the future token(s) as a group
+\fi\@@nextIEEEbuildpreamble{#2}}
+
+
+% executed just after preamble build is completed
+% warn about zero cols, and if prevtype type = u, put in end tabskip glue
+\def\@@IEEEfinishpreamble#1{\ifnum\@IEEEeqnnumcols<1\relax
+\@IEEEclspkgerror{No column specifiers declared for IEEEeqnarray}%
+{At least one column type must be declared for each IEEEeqnarray.}%
+\fi%num cols less than 1
+%if last type undefined, set default end tabskip glue
+\if\@IEEEBPprevtype u\@IEEEappendtoksA{\tabskip=\@IEEEBPendglue}\fi}
+
+
+% Identify and return the column specifier's type code
+\def\@IEEEgetcoltype#1#2#3{%
+% use only the very first token to determine the type
+% \noindent is used as a delimiter as \end can be present here
+\def\@IEEEgrabfirstoken##1##2\noindent{\let\@IEEEgrabbedfirstoken=##1}%
+\@IEEEgrabfirstoken#1\relax\relax\noindent
+% \@IEEEgrabfirstoken has the first token, the rest are discarded
+% n = number
+% g = glue (any other char in catagory 12)
+% c = letter
+% e = \end
+% u = undefined
+% third argument: 0 = no error message, 1 = error on invalid char
+\let#2=u\relax% assume invalid until know otherwise
+\ifx\@IEEEgrabbedfirstoken\end\let#2=e\else
+\ifcat\@IEEEgrabbedfirstoken\relax\else% screen out control sequences
+\if0\@IEEEgrabbedfirstoken\let#2=n\else
+\if1\@IEEEgrabbedfirstoken\let#2=n\else
+\if2\@IEEEgrabbedfirstoken\let#2=n\else
+\if3\@IEEEgrabbedfirstoken\let#2=n\else
+\if4\@IEEEgrabbedfirstoken\let#2=n\else
+\if5\@IEEEgrabbedfirstoken\let#2=n\else
+\if6\@IEEEgrabbedfirstoken\let#2=n\else
+\if7\@IEEEgrabbedfirstoken\let#2=n\else
+\if8\@IEEEgrabbedfirstoken\let#2=n\else
+\if9\@IEEEgrabbedfirstoken\let#2=n\else
+\ifcat,\@IEEEgrabbedfirstoken\let#2=g\relax
+\else\ifcat a\@IEEEgrabbedfirstoken\let#2=c\relax\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi
+\if#2u\relax
+\if0\noexpand#3\relax\else\@IEEEclspkgerror{Invalid character in column specifications}%
+{Only letters, numerals and certain other symbols are allowed \MessageBreak
+as IEEEeqnarray column specifiers.}\fi\fi}
+
+
+% identify the current letter referenced column
+% if invalid, use a default column
+\def\@IEEEgetcurcol#1{\expandafter\ifx\csname @IEEEeqnarraycolDEF#1\endcsname\@IEEEeqnarraycolisdefined%
+\def\@IEEEBPcurcolname{#1}\else% invalid column name
+\@IEEEclspkgerror{Invalid column type "#1" in column specifications.\MessageBreak
+Using a default centering column instead}%
+{You must define IEEEeqnarray column types before use.}%
+\def\@IEEEBPcurcolname{@IEEEdefault}\fi}
+
+
+% identify and return the predefined (punctuation) glue value
+\def\@IEEEgetcurglue#1#2{%
+% ! = \! (neg small) -0.16667em (-3/18 em)
+% , = \, (small) 0.16667em ( 3/18 em)
+% : = \: (med) 0.22222em ( 4/18 em)
+% ; = \; (large) 0.27778em ( 5/18 em)
+% ' = \quad 1em
+% " = \qquad 2em
+% . = 0.5\arraycolsep
+% / = \arraycolsep
+% ? = 2\arraycolsep
+% * = 1fil
+% + = \@IEEEeqnarraycolSEPcenter
+% - = \@IEEEeqnarraycolSEPzero
+% Note that all em values are referenced to the math font (textfont2) fontdimen6
+% value for 1em.
+%
+% use only the very first token to determine the type
+% this prevents errant tokens from getting in the main text
+% \noindent is used as a delimiter here
+\def\@IEEEgrabfirstoken##1##2\noindent{\let\@IEEEgrabbedfirstoken=##1}%
+\@IEEEgrabfirstoken#1\relax\relax\noindent
+% get the math font 1em value
+% LaTeX2e's NFSS2 does not preload the fonts, but \IEEEeqnarray needs
+% to gain access to the math (\textfont2) font's spacing parameters.
+% So we create a bogus box here that uses the math font to ensure
+% that \textfont2 is loaded and ready. If this is not done,
+% the \textfont2 stuff here may not work.
+% Thanks to Bernd Raichle for his 1997 post on this topic.
+{\setbox0=\hbox{$\displaystyle\relax$}}%
+% fontdimen6 has the width of 1em (a quad).
+\@IEEEtrantmpdimenA=\fontdimen6\textfont2\relax%
+% identify the glue value based on the first token
+% we discard anything after the first
+\if!\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=-0.16667\@IEEEtrantmpdimenA\edef#2{\the\@IEEEtrantmpdimenA}\else
+\if,\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=0.16667\@IEEEtrantmpdimenA\edef#2{\the\@IEEEtrantmpdimenA}\else
+\if:\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=0.22222\@IEEEtrantmpdimenA\edef#2{\the\@IEEEtrantmpdimenA}\else
+\if;\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=0.27778\@IEEEtrantmpdimenA\edef#2{\the\@IEEEtrantmpdimenA}\else
+\if'\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=1\@IEEEtrantmpdimenA\edef#2{\the\@IEEEtrantmpdimenA}\else
+\if"\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=2\@IEEEtrantmpdimenA\edef#2{\the\@IEEEtrantmpdimenA}\else
+\if.\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=0.5\arraycolsep\edef#2{\the\@IEEEtrantmpdimenA}\else
+\if/\@IEEEgrabbedfirstoken\edef#2{\the\arraycolsep}\else
+\if?\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=2\arraycolsep\edef#2{\the\@IEEEtrantmpdimenA}\else
+\if *\@IEEEgrabbedfirstoken\edef#2{0pt plus 1fil minus 0pt}\else
+\if+\@IEEEgrabbedfirstoken\edef#2{\@IEEEeqnarraycolSEPcenter}\else
+\if-\@IEEEgrabbedfirstoken\edef#2{\@IEEEeqnarraycolSEPzero}\else
+\edef#2{\@IEEEeqnarraycolSEPzero}%
+\@IEEEclspkgerror{Invalid predefined inter-column glue type "#1" in\MessageBreak
+column specifications. Using a default value of\MessageBreak
+0pt instead}%
+{Only !,:;'"./?*+ and - are valid predefined glue types in the\MessageBreak
+IEEEeqnarray column specifications.}\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi}
+
+
+
+% process a numerical digit from the column specification
+% and look up the corresponding user defined glue value
+% can transform current type from n to g or a as the user defined glue is acquired
+\def\@IEEEprocessNcol#1{\if\@IEEEBPprevtype g%
+\@IEEEclspkgerror{Back-to-back inter-column glue specifiers in column\MessageBreak
+specifications. Ignoring consecutive glue specifiers\MessageBreak
+after the first}%
+{You cannot have two or more glue types next to each other\MessageBreak
+in the IEEEeqnarray column specifications.}%
+\let\@IEEEBPcurtype=a% abort this glue, future digits will be discarded
+\@IEEEBPcurnum=0\relax%
+\else% if we previously aborted a glue
+\if\@IEEEBPprevtype a\@IEEEBPcurnum=0\let\@IEEEBPcurtype=a%maintain digit abortion
+\else%acquire this number
+% save the previous type before the numerical digits started
+\if\@IEEEBPprevtype n\else\let\@IEEEBPprevsavedtype=\@IEEEBPprevtype\fi%
+\multiply\@IEEEBPcurnum by 10\relax%
+\advance\@IEEEBPcurnum by #1\relax% add in number, \relax is needed to stop TeX's number scan
+\if\@IEEEBPnexttype n\else%close acquisition
+\expandafter\ifx\csname @IEEEeqnarraycolSEPDEF\expandafter\romannumeral\number\@IEEEBPcurnum\endcsname\@IEEEeqnarraycolisdefined%
+\edef\@IEEEBPcurglue{\csname @IEEEeqnarraycolSEP\expandafter\romannumeral\number\@IEEEBPcurnum\endcsname}%
+\else%user glue not defined
+\@IEEEclspkgerror{Invalid user defined inter-column glue type "\number\@IEEEBPcurnum" in\MessageBreak
+column specifications. Using a default value of\MessageBreak
+0pt instead}%
+{You must define all IEEEeqnarray numerical inter-column glue types via\MessageBreak
+\string\IEEEeqnarraydefcolsep \space before they are used in column specifications.}%
+\edef\@IEEEBPcurglue{\@IEEEeqnarraycolSEPzero}%
+\fi% glue defined or not
+\let\@IEEEBPcurtype=g% change the type to reflect the acquired glue
+\let\@IEEEBPprevtype=\@IEEEBPprevsavedtype% restore the prev type before this number glue
+\@IEEEBPcurnum=0\relax%ready for next acquisition
+\fi%close acquisition, get glue
+\fi%discard or acquire number
+\fi%prevtype glue or not
+}
+
+
+% process an acquired glue
+% add any acquired column/glue pair to the preamble
+\def\@IEEEprocessGcol{\if\@IEEEBPprevtype a\let\@IEEEBPcurtype=a%maintain previous glue abortions
+\else
+% if this is the start glue, save it, but do nothing else
+% as this is not used in the preamble, but before
+\if\@IEEEBPprevtype s\edef\@IEEEBPstartglue{\@IEEEBPcurglue}%
+\else%not the start glue
+\if\@IEEEBPprevtype g%ignore if back to back glues
+\@IEEEclspkgerror{Back-to-back inter-column glue specifiers in column\MessageBreak
+specifications. Ignoring consecutive glue specifiers\MessageBreak
+after the first}%
+{You cannot have two or more glue types next to each other\MessageBreak
+in the IEEEeqnarray column specifications.}%
+\let\@IEEEBPcurtype=a% abort this glue
+\else% not a back to back glue
+\if\@IEEEBPprevtype c\relax% if the previoustype was a col, add column/glue pair to preamble
+\ifnum\@IEEEeqnnumcols>0\relax\@IEEEappendtoksA{&}\fi
+\toks0={##}%
+% make preamble advance col counter if this environment needs this
+\if@advanceIEEEeqncolcnt\@IEEEappendtoksA{\global\advance\@IEEEeqncolcnt by 1\relax}\fi
+% insert the column defintion into the preamble, being careful not to expand
+% the column definition
+\@IEEEappendtoksA{\tabskip=\@IEEEBPcurglue}%
+\@IEEEappendNOEXPANDtoksA{\begingroup\csname @IEEEeqnarraycolPRE}%
+\@IEEEappendtoksA{\@IEEEBPcurcolname}%
+\@IEEEappendNOEXPANDtoksA{\endcsname}%
+\@IEEEappendtoksA{\the\toks0}%
+\@IEEEappendNOEXPANDtoksA{\relax\relax\relax\relax\relax%
+\relax\relax\relax\relax\relax\csname @IEEEeqnarraycolPOST}%
+\@IEEEappendtoksA{\@IEEEBPcurcolname}%
+\@IEEEappendNOEXPANDtoksA{\endcsname\relax\relax\relax\relax\relax%
+\relax\relax\relax\relax\relax\endgroup}%
+\advance\@IEEEeqnnumcols by 1\relax%one more column in the preamble
+\else% error: non-start glue with no pending column
+\@IEEEclspkgerror{Inter-column glue specifier without a prior column\MessageBreak
+type in the column specifications. Ignoring this glue\MessageBreak
+specifier}%
+{Except for the first and last positions, glue can be placed only\MessageBreak
+between column types.}%
+\let\@IEEEBPcurtype=a% abort this glue
+\fi% previous was a column
+\fi% back-to-back glues
+\fi% is start column glue
+\fi% prev type not a
+}
+
+
+% process an acquired letter referenced column and, if necessary, add it to the preamble
+\def\@IEEEprocessCcol{\if\@IEEEBPnexttype g\else
+\if\@IEEEBPnexttype n\else
+% we have a column followed by something other than a glue (or numeral glue)
+% so we must add this column to the preamble now
+\ifnum\@IEEEeqnnumcols>0\relax\@IEEEappendtoksA{&}\fi%col separator for those after the first
+\if\@IEEEBPnexttype e\@IEEEappendtoksA{\tabskip=\@IEEEBPendglue\relax}\else%put in end glue
+\@IEEEappendtoksA{\tabskip=\@IEEEeqnarraycolSEPdefaultmid\relax}\fi% or default mid glue
+\toks0={##}%
+% make preamble advance col counter if this environment needs this
+\if@advanceIEEEeqncolcnt\@IEEEappendtoksA{\global\advance\@IEEEeqncolcnt by 1\relax}\fi
+% insert the column definition into the preamble, being careful not to expand
+% the column definition
+\@IEEEappendNOEXPANDtoksA{\begingroup\csname @IEEEeqnarraycolPRE}%
+\@IEEEappendtoksA{\@IEEEBPcurcolname}%
+\@IEEEappendNOEXPANDtoksA{\endcsname}%
+\@IEEEappendtoksA{\the\toks0}%
+\@IEEEappendNOEXPANDtoksA{\relax\relax\relax\relax\relax%
+\relax\relax\relax\relax\relax\csname @IEEEeqnarraycolPOST}%
+\@IEEEappendtoksA{\@IEEEBPcurcolname}%
+\@IEEEappendNOEXPANDtoksA{\endcsname\relax\relax\relax\relax\relax%
+\relax\relax\relax\relax\relax\endgroup}%
+\advance\@IEEEeqnnumcols by 1\relax%one more column in the preamble
+\fi%next type not numeral
+\fi%next type not glue
+}
+
+
+%%
+%% END OF IEEEeqnarry DEFINITIONS
+%%
+
+
+
+
+% set up the running headings, this complex because of all the different
+% modes IEEEtran supports
+\if@twoside
+ \ifCLASSOPTIONtechnote
+ \def\ps@headings{%
+ \def\@oddhead{\hbox{}\scriptsize\leftmark \hfil \thepage}
+ \def\@evenhead{\scriptsize\thepage \hfil \leftmark\hbox{}}
+ \ifCLASSOPTIONdraftcls
+ \ifCLASSOPTIONdraftclsnofoot
+ \def\@oddfoot{}\def\@evenfoot{}%
+ \else
+ \def\@oddfoot{\scriptsize\@date\hfil DRAFT}
+ \def\@evenfoot{\scriptsize DRAFT\hfil\@date}
+ \fi
+ \else
+ \def\@oddfoot{}\def\@evenfoot{}
+ \fi}
+ \else % not a technote
+ \def\ps@headings{%
+ \ifCLASSOPTIONconference
+ \def\@oddhead{}
+ \def\@evenhead{}
+ \else
+ \def\@oddhead{\hbox{}\scriptsize\rightmark \hfil \thepage}
+ \def\@evenhead{\scriptsize\thepage \hfil \leftmark\hbox{}}
+ \fi
+ \ifCLASSOPTIONdraftcls
+ \def\@oddhead{\hbox{}\scriptsize\rightmark \hfil \thepage}
+ \def\@evenhead{\scriptsize\thepage \hfil \leftmark\hbox{}}
+ \ifCLASSOPTIONdraftclsnofoot
+ \def\@oddfoot{}\def\@evenfoot{}%
+ \else
+ \def\@oddfoot{\scriptsize\@date\hfil DRAFT}
+ \def\@evenfoot{\scriptsize DRAFT\hfil\@date}
+ \fi
+ \else
+ \def\@oddfoot{}\def\@evenfoot{}%
+ \fi}
+ \fi
+\else % single side
+\def\ps@headings{%
+ \ifCLASSOPTIONconference
+ \def\@oddhead{}
+ \def\@evenhead{}
+ \else
+ \def\@oddhead{\hbox{}\scriptsize\leftmark \hfil \thepage}
+ \def\@evenhead{}
+ \fi
+ \ifCLASSOPTIONdraftcls
+ \def\@oddhead{\hbox{}\scriptsize\leftmark \hfil \thepage}
+ \def\@evenhead{}
+ \ifCLASSOPTIONdraftclsnofoot
+ \def\@oddfoot{}
+ \else
+ \def\@oddfoot{\scriptsize \@date \hfil DRAFT}
+ \fi
+ \else
+ \def\@oddfoot{}
+ \fi
+ \def\@evenfoot{}}
+\fi
+
+
+% title page style
+\def\ps@IEEEtitlepagestyle{\def\@oddfoot{}\def\@evenfoot{}%
+\ifCLASSOPTIONconference
+ \def\@oddhead{}%
+ \def\@evenhead{}%
+\else
+ \def\@oddhead{\hbox{}\scriptsize\leftmark \hfil \thepage}%
+ \def\@evenhead{\scriptsize\thepage \hfil \leftmark\hbox{}}%
+\fi
+\ifCLASSOPTIONdraftcls
+ \def\@oddhead{\hbox{}\scriptsize\leftmark \hfil \thepage}%
+ \def\@evenhead{\scriptsize\thepage \hfil \leftmark\hbox{}}%
+ \ifCLASSOPTIONdraftclsnofoot\else
+ \def\@oddfoot{\scriptsize \@date\hfil DRAFT}%
+ \def\@evenfoot{\scriptsize DRAFT\hfil \@date}%
+ \fi
+\else
+ % all non-draft mode footers
+ \if@IEEEusingpubid
+ % for title pages that are using a pubid
+ % do not repeat pubid if using peer review option
+ \ifCLASSOPTIONpeerreview
+ \else
+ \footskip 0pt%
+ \ifCLASSOPTIONcompsocconf
+ \def\@oddfoot{\hss\normalfont\scriptsize\raisebox{-1.5\@IEEEnormalsizeunitybaselineskip}[0ex][0ex]{\@IEEEpubid}\hss}%
+ \def\@evenfoot{\hss\normalfont\scriptsize\raisebox{-1.5\@IEEEnormalsizeunitybaselineskip}[0ex][0ex]{\@IEEEpubid}\hss}%
+ \else
+ \def\@oddfoot{\hss\normalfont\footnotesize\raisebox{1.5ex}[1.5ex]{\@IEEEpubid}\hss}%
+ \def\@evenfoot{\hss\normalfont\footnotesize\raisebox{1.5ex}[1.5ex]{\@IEEEpubid}\hss}%
+ \fi
+ \fi
+ \fi
+\fi}
+
+
+% peer review cover page style
+\def\ps@IEEEpeerreviewcoverpagestyle{%
+\def\@oddhead{}\def\@evenhead{}%
+\def\@oddfoot{}\def\@evenfoot{}%
+\ifCLASSOPTIONdraftcls
+ \ifCLASSOPTIONdraftclsnofoot\else
+ \def\@oddfoot{\scriptsize \@date\hfil DRAFT}%
+ \def\@evenfoot{\scriptsize DRAFT\hfil \@date}%
+ \fi
+\else
+ % non-draft mode footers
+ \if@IEEEusingpubid
+ \footskip 0pt%
+ \ifCLASSOPTIONcompsoc
+ \def\@oddfoot{\hss\normalfont\scriptsize\raisebox{-1.5\@IEEEnormalsizeunitybaselineskip}[0ex][0ex]{\@IEEEpubid}\hss}%
+ \def\@evenfoot{\hss\normalfont\scriptsize\raisebox{-1.5\@IEEEnormalsizeunitybaselineskip}[0ex][0ex]{\@IEEEpubid}\hss}%
+ \else
+ \def\@oddfoot{\hss\normalfont\footnotesize\raisebox{1.5ex}[1.5ex]{\@IEEEpubid}\hss}%
+ \def\@evenfoot{\hss\normalfont\footnotesize\raisebox{1.5ex}[1.5ex]{\@IEEEpubid}\hss}%
+ \fi
+ \fi
+\fi}
+
+
+% start with empty headings
+\def\rightmark{}\def\leftmark{}
+
+
+%% Defines the command for putting the header. \footernote{TEXT} is the same
+%% as \markboth{TEXT}{TEXT}.
+%% Note that all the text is forced into uppercase, if you have some text
+%% that needs to be in lower case, for instance et. al., then either manually
+%% set \leftmark and \rightmark or use \MakeLowercase{et. al.} within the
+%% arguments to \markboth.
+\def\markboth#1#2{\def\leftmark{\@IEEEcompsoconly{\sffamily}\MakeUppercase{#1}}%
+\def\rightmark{\@IEEEcompsoconly{\sffamily}\MakeUppercase{#2}}}
+\def\footernote#1{\markboth{#1}{#1}}
+
+\def\today{\ifcase\month\or
+ January\or February\or March\or April\or May\or June\or
+ July\or August\or September\or October\or November\or December\fi
+ \space\number\day, \number\year}
+
+
+
+
+%% CITATION AND BIBLIOGRAPHY COMMANDS
+%%
+%% V1.6 no longer supports the older, nonstandard \shortcite and \citename setup stuff
+%
+%
+% Modify Latex2e \@citex to separate citations with "], ["
+\def\@citex[#1]#2{%
+ \let\@citea\@empty
+ \@cite{\@for\@citeb:=#2\do
+ {\@citea\def\@citea{], [}%
+ \edef\@citeb{\expandafter\@firstofone\@citeb\@empty}%
+ \if@filesw\immediate\write\@auxout{\string\citation{\@citeb}}\fi
+ \@ifundefined{b@\@citeb}{\mbox{\reset@font\bfseries ?}%
+ \G@refundefinedtrue
+ \@latex@warning
+ {Citation `\@citeb' on page \thepage \space undefined}}%
+ {\hbox{\csname b@\@citeb\endcsname}}}}{#1}}
+
+% V1.6 we create hooks for the optional use of Donald Arseneau's
+% cite.sty package. cite.sty is "smart" and will notice that the
+% following format controls are already defined and will not
+% redefine them. The result will be the proper sorting of the
+% citation numbers and auto detection of 3 or more entry "ranges" -
+% all in IEEE style: [1], [2], [5]--[7], [12]
+% This also allows for an optional note, i.e., \cite[mynote]{..}.
+% If the \cite with note has more than one reference, the note will
+% be applied to the last of the listed references. It is generally
+% desired that if a note is given, only one reference is listed in
+% that \cite.
+% Thanks to Mr. Arseneau for providing the required format arguments
+% to produce the IEEE style.
+\def\citepunct{], [}
+\def\citedash{]--[}
+
+% V1.7 default to using same font for urls made by url.sty
+\AtBeginDocument{\csname url@samestyle\endcsname}
+
+% V1.6 class files should always provide these
+\def\newblock{\hskip .11em\@plus.33em\@minus.07em}
+\let\@openbib@code\@empty
+
+
+% Provide support for the control entries of IEEEtran.bst V1.00 and later.
+% V1.7 optional argument allows for a different aux file to be specified in
+% order to handle multiple bibliographies. For example, with multibib.sty:
+% \newcites{sec}{Secondary Literature}
+% \bstctlcite[@auxoutsec]{BSTcontrolhak}
+\def\bstctlcite{\@ifnextchar[{\@bstctlcite}{\@bstctlcite[@auxout]}}
+\def\@bstctlcite[#1]#2{\@bsphack
+ \@for\@citeb:=#2\do{%
+ \edef\@citeb{\expandafter\@firstofone\@citeb}%
+ \if@filesw\immediate\write\csname #1\endcsname{\string\citation{\@citeb}}\fi}%
+ \@esphack}
+
+% V1.6 provide a way for a user to execute a command just before
+% a given reference number - used to insert a \newpage to balance
+% the columns on the last page
+\edef\@IEEEtriggerrefnum{0} % the default of zero means that
+ % the command is not executed
+\def\@IEEEtriggercmd{\newpage}
+
+% allow the user to alter the triggered command
+\long\def\IEEEtriggercmd#1{\long\def\@IEEEtriggercmd{#1}}
+
+% allow user a way to specify the reference number just before the
+% command is executed
+\def\IEEEtriggeratref#1{\@IEEEtrantmpcountA=#1%
+\edef\@IEEEtriggerrefnum{\the\@IEEEtrantmpcountA}}%
+
+% trigger command at the given reference
+\def\@IEEEbibitemprefix{\@IEEEtrantmpcountA=\@IEEEtriggerrefnum\relax%
+\advance\@IEEEtrantmpcountA by -1\relax%
+\ifnum\c@enumiv=\@IEEEtrantmpcountA\relax\@IEEEtriggercmd\relax\fi}
+
+
+\def\@biblabel#1{[#1]}
+
+% compsoc journals left align the reference numbers
+\@IEEEcompsocnotconfonly{\def\@biblabel#1{[#1]\hfill}}
+
+% controls bib item spacing
+\def\IEEEbibitemsep{0pt plus .5pt}
+
+\@IEEEcompsocconfonly{\def\IEEEbibitemsep{1\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip}}
+
+
+\def\thebibliography#1{\section*{\refname}%
+ \addcontentsline{toc}{section}{\refname}%
+ % V1.6 add some rubber space here and provide a command trigger
+ \footnotesize\@IEEEcompsocconfonly{\small}\vskip 0.3\baselineskip plus 0.1\baselineskip minus 0.1\baselineskip%
+ \list{\@biblabel{\@arabic\c@enumiv}}%
+ {\settowidth\labelwidth{\@biblabel{#1}}%
+ \leftmargin\labelwidth
+ \advance\leftmargin\labelsep\relax
+ \itemsep \IEEEbibitemsep\relax
+ \usecounter{enumiv}%
+ \let\p@enumiv\@empty
+ \renewcommand\theenumiv{\@arabic\c@enumiv}}%
+ \let\@IEEElatexbibitem\bibitem%
+ \def\bibitem{\@IEEEbibitemprefix\@IEEElatexbibitem}%
+\def\newblock{\hskip .11em plus .33em minus .07em}%
+% originally:
+% \sloppy\clubpenalty4000\widowpenalty4000%
+% by adding the \interlinepenalty here, we make it more
+% difficult, but not impossible, for LaTeX to break within a reference.
+% IEEE almost never breaks a reference (but they do it more often with
+% technotes). You may get an underfull vbox warning around the bibliography,
+% but the final result will be much more like what IEEE will publish.
+% MDS 11/2000
+\ifCLASSOPTIONtechnote\sloppy\clubpenalty4000\widowpenalty4000\interlinepenalty100%
+\else\sloppy\clubpenalty4000\widowpenalty4000\interlinepenalty500\fi%
+ \sfcode`\.=1000\relax}
+\let\endthebibliography=\endlist
+
+
+
+
+% TITLE PAGE COMMANDS
+%
+%
+% \IEEEmembership is used to produce the sublargesize italic font used to indicate author
+% IEEE membership. compsoc uses a large size sans slant font
+\def\IEEEmembership#1{{\@IEEEnotcompsoconly{\sublargesize}\normalfont\@IEEEcompsoconly{\sffamily}\textit{#1}}}
+
+
+% \IEEEauthorrefmark{} produces a footnote type symbol to indicate author affiliation.
+% When given an argument of 1 to 9, \IEEEauthorrefmark{} follows the standard LaTeX footnote
+% symbol sequence convention. However, for arguments 10 and above, \IEEEauthorrefmark{}
+% reverts to using lower case roman numerals, so it cannot overflow. Do note that you
+% cannot use \footnotemark[] in place of \IEEEauthorrefmark{} within \author as the footnote
+% symbols will have been turned off to prevent \thanks from creating footnote marks.
+% \IEEEauthorrefmark{} produces a symbol that appears to LaTeX as having zero vertical
+% height - this allows for a more compact line packing, but the user must ensure that
+% the interline spacing is large enough to prevent \IEEEauthorrefmark{} from colliding
+% with the text above.
+% V1.7 make this a robust command
+\DeclareRobustCommand*{\IEEEauthorrefmark}[1]{\raisebox{0pt}[0pt][0pt]{\textsuperscript{\footnotesize\ensuremath{\ifcase#1\or *\or \dagger\or \ddagger\or%
+ \mathsection\or \mathparagraph\or \|\or **\or \dagger\dagger%
+ \or \ddagger\ddagger \else\textsuperscript{\expandafter\romannumeral#1}\fi}}}}
+
+
+% FONT CONTROLS AND SPACINGS FOR CONFERENCE MODE AUTHOR NAME AND AFFILIATION BLOCKS
+%
+% The default font styles for the author name and affiliation blocks (confmode)
+%\def\@IEEEauthorblockNstyle{\normalfont\@IEEEcompsocnotconfonly{\sffamily}\sublargesize\@IEEEcompsocconfonly{\large}}
+%\def\@IEEEauthorblockAstyle{\normalfont\@IEEEcompsocnotconfonly{\sffamily}\@IEEEcompsocconfonly{\itshape}\normalsize\@IEEEcompsocconfonly{\large}}
+\def\@IEEEauthorblockNstyle{\normalfont\normalsize}
+\def\@IEEEauthorblockAstyle{\normalfont\@IEEEcompsocnotconfonly{\sffamily}\@IEEEcompsocconfonly{\itshape}\normalsize}
+
+% The default if the user does not use an author block
+\def\@IEEEauthordefaulttextstyle{\normalfont\@IEEEcompsocnotconfonly{\sffamily}\sublargesize}
+
+% spacing from title (or special paper notice) to author name blocks (confmode)
+% can be negative
+\def\@IEEEauthorblockconfadjspace{-0.25em}
+% compsoc conferences need more space here
+\@IEEEcompsocconfonly{\def\@IEEEauthorblockconfadjspace{0.75\@IEEEnormalsizeunitybaselineskip}}
+
+% spacing between name and affiliation blocks (confmode)
+% This can be negative.
+% IEEE doesn't want any added spacing here, but I will leave these
+% controls in place in case they ever change their mind.
+% Personally, I like 0.75ex.
+%\def\@IEEEauthorblockNtopspace{0.75ex}
+%\def\@IEEEauthorblockAtopspace{0.75ex}
+\def\@IEEEauthorblockNtopspace{0.0ex}
+\def\@IEEEauthorblockAtopspace{0.0ex}
+% baseline spacing within name and affiliation blocks (confmode)
+% must be positive, spacings below certain values will make
+% the position of line of text sensitive to the contents of the
+% line above it i.e., whether or not the prior line has descenders,
+% subscripts, etc. For this reason it is a good idea to keep
+% these above 2.6ex
+\def\@IEEEauthorblockNinterlinespace{2.6ex}
+\def\@IEEEauthorblockAinterlinespace{2.75ex}
+
+% This tracks the required strut size.
+% See the \@IEEEauthorhalign command for the actual default value used.
+\def\@IEEEauthorblockXinterlinespace{2.7ex}
+
+% variables to retain font size and style across groups
+% values given here have no effect as they will be overwritten later
+\gdef\@IEEESAVESTATEfontsize{10}
+\gdef\@IEEESAVESTATEfontbaselineskip{12}
+\gdef\@IEEESAVESTATEfontencoding{OT1}
+\gdef\@IEEESAVESTATEfontfamily{ptm}
+\gdef\@IEEESAVESTATEfontseries{m}
+\gdef\@IEEESAVESTATEfontshape{n}
+
+% saves the current font attributes
+\def\@IEEEcurfontSAVE{\global\let\@IEEESAVESTATEfontsize\f@size%
+\global\let\@IEEESAVESTATEfontbaselineskip\f@baselineskip%
+\global\let\@IEEESAVESTATEfontencoding\f@encoding%
+\global\let\@IEEESAVESTATEfontfamily\f@family%
+\global\let\@IEEESAVESTATEfontseries\f@series%
+\global\let\@IEEESAVESTATEfontshape\f@shape}
+
+% restores the saved font attributes
+\def\@IEEEcurfontRESTORE{\fontsize{\@IEEESAVESTATEfontsize}{\@IEEESAVESTATEfontbaselineskip}%
+\fontencoding{\@IEEESAVESTATEfontencoding}%
+\fontfamily{\@IEEESAVESTATEfontfamily}%
+\fontseries{\@IEEESAVESTATEfontseries}%
+\fontshape{\@IEEESAVESTATEfontshape}%
+\selectfont}
+
+
+% variable to indicate if the current block is the first block in the column
+\newif\if@IEEEprevauthorblockincol \@IEEEprevauthorblockincolfalse
+
+
+% the command places a strut with height and depth = \@IEEEauthorblockXinterlinespace
+% we use this technique to have complete manual control over the spacing of the lines
+% within the halign environment.
+% We set the below baseline portion at 30%, the above
+% baseline portion at 70% of the total length.
+% Responds to changes in the document's \baselinestretch
+\def\@IEEEauthorstrutrule{\@IEEEtrantmpdimenA\@IEEEauthorblockXinterlinespace%
+\@IEEEtrantmpdimenA=\baselinestretch\@IEEEtrantmpdimenA%
+\rule[-0.3\@IEEEtrantmpdimenA]{0pt}{\@IEEEtrantmpdimenA}}
+
+
+% blocks to hold the authors' names and affilations.
+% Makes formatting easy for conferences
+%
+% use real definitions in conference mode
+% name block
+\def\IEEEauthorblockN#1{\relax\@IEEEauthorblockNstyle% set the default text style
+\gdef\@IEEEauthorblockXinterlinespace{0pt}% disable strut for spacer row
+% the \expandafter hides the \cr in conditional tex, see the array.sty docs
+% for details, probably not needed here as the \cr is in a macro
+% do a spacer row if needed
+\if@IEEEprevauthorblockincol\expandafter\@IEEEauthorblockNtopspaceline\fi
+\global\@IEEEprevauthorblockincoltrue% we now have a block in this column
+%restore the correct strut value
+\gdef\@IEEEauthorblockXinterlinespace{\@IEEEauthorblockNinterlinespace}%
+% input the author names
+#1%
+% end the row if the user did not already
+\crcr}
+% spacer row for names
+\def\@IEEEauthorblockNtopspaceline{\cr\noalign{\vskip\@IEEEauthorblockNtopspace}}
+%
+% affiliation block
+\def\IEEEauthorblockA#1{\relax\@IEEEauthorblockAstyle% set the default text style
+\gdef\@IEEEauthorblockXinterlinespace{0pt}%disable strut for spacer row
+% the \expandafter hides the \cr in conditional tex, see the array.sty docs
+% for details, probably not needed here as the \cr is in a macro
+% do a spacer row if needed
+\if@IEEEprevauthorblockincol\expandafter\@IEEEauthorblockAtopspaceline\fi
+\global\@IEEEprevauthorblockincoltrue% we now have a block in this column
+%restore the correct strut value
+\gdef\@IEEEauthorblockXinterlinespace{\@IEEEauthorblockAinterlinespace}%
+% input the author affiliations
+#1%
+% end the row if the user did not already
+\crcr}
+% spacer row for affiliations
+\def\@IEEEauthorblockAtopspaceline{\cr\noalign{\vskip\@IEEEauthorblockAtopspace}}
+
+
+% allow papers to compile even if author blocks are used in modes other
+% than conference or peerreviewca. For such cases, we provide dummy blocks.
+\ifCLASSOPTIONconference
+\else
+ \ifCLASSOPTIONpeerreviewca\else
+ % not conference or peerreviewca mode
+ \def\IEEEauthorblockN#1{#1}%
+ \def\IEEEauthorblockA#1{#1}%
+ \fi
+\fi
+
+
+
+% we provide our own halign so as not to have to depend on tabular
+\def\@IEEEauthorhalign{\@IEEEauthordefaulttextstyle% default text style
+ \lineskip=0pt\relax% disable line spacing
+ \lineskiplimit=0pt\relax%
+ \baselineskip=0pt\relax%
+ \@IEEEcurfontSAVE% save the current font
+ \mathsurround\z@\relax% no extra spacing around math
+ \let\\\@IEEEauthorhaligncr% replace newline with halign friendly one
+ \tabskip=0pt\relax% no column spacing
+ \everycr{}% ensure no problems here
+ \@IEEEprevauthorblockincolfalse% no author blocks yet
+ \def\@IEEEauthorblockXinterlinespace{2.7ex}% default interline space
+ \vtop\bgroup%vtop box
+ \halign\bgroup&\relax\hfil\@IEEEcurfontRESTORE\relax ##\relax
+ \hfil\@IEEEcurfontSAVE\@IEEEauthorstrutrule\cr}
+
+% ensure last line, exit from halign, close vbox
+\def\end@IEEEauthorhalign{\crcr\egroup\egroup}
+
+% handle bogus star form
+\def\@IEEEauthorhaligncr{{\ifnum0=`}\fi\@ifstar{\@@IEEEauthorhaligncr}{\@@IEEEauthorhaligncr}}
+
+% test and setup the optional argument to \\[]
+\def\@@IEEEauthorhaligncr{\@testopt\@@@IEEEauthorhaligncr\z@skip}
+
+% end the line and do the optional spacer
+\def\@@@IEEEauthorhaligncr[#1]{\ifnum0=`{\fi}\cr\noalign{\vskip#1\relax}}
+
+
+
+% flag to prevent multiple \and warning messages
+\newif\if@IEEEWARNand
+\@IEEEWARNandtrue
+
+% if in conference or peerreviewca modes, we support the use of \and as \author is a
+% tabular environment, otherwise we warn the user that \and is invalid
+% outside of conference or peerreviewca modes.
+\def\and{\relax} % provide a bogus \and that we will then override
+
+\renewcommand{\and}[1][\relax]{\if@IEEEWARNand\typeout{** WARNING: \noexpand\and is valid only
+ when in conference or peerreviewca}\typeout{modes (line \the\inputlineno).}\fi\global\@IEEEWARNandfalse}
+
+\ifCLASSOPTIONconference%
+\renewcommand{\and}[1][\hfill]{\end{@IEEEauthorhalign}#1\begin{@IEEEauthorhalign}}%
+\fi
+\ifCLASSOPTIONpeerreviewca
+\renewcommand{\and}[1][\hfill]{\end{@IEEEauthorhalign}#1\begin{@IEEEauthorhalign}}%
+\fi
+
+
+% page clearing command
+% based on LaTeX2e's \cleardoublepage, but allows different page styles
+% for the inserted blank pages
+\def\@IEEEcleardoublepage#1{\clearpage\if@twoside\ifodd\c@page\else
+\hbox{}\thispagestyle{#1}\newpage\if@twocolumn\hbox{}\thispagestyle{#1}\newpage\fi\fi\fi}
+
+
+% user command to invoke the title page
+\def\maketitle{\par%
+ \begingroup%
+ \normalfont%
+ \def\thefootnote{}% the \thanks{} mark type is empty
+ \def\footnotemark{}% and kill space from \thanks within author
+ \let\@makefnmark\relax% V1.7, must *really* kill footnotemark to remove all \textsuperscript spacing as well.
+ \footnotesize% equal spacing between thanks lines
+ \footnotesep 0.7\baselineskip%see global setting of \footnotesep for more info
+ % V1.7 disable \thanks note indention for compsoc
+ \@IEEEcompsoconly{\long\def\@makefntext##1{\parindent 1em\noindent\hbox{\@makefnmark}##1}}%
+ \normalsize%
+ \ifCLASSOPTIONpeerreview
+ \newpage\global\@topnum\z@ \@maketitle\@IEEEstatictitlevskip\@IEEEaftertitletext%
+ \thispagestyle{IEEEpeerreviewcoverpagestyle}\@thanks%
+ \else
+ \if@twocolumn%
+ \ifCLASSOPTIONtechnote%
+ \newpage\global\@topnum\z@ \@maketitle\@IEEEstatictitlevskip\@IEEEaftertitletext%
+ \else
+ \twocolumn[\@maketitle\@IEEEdynamictitlevspace\@IEEEaftertitletext]%
+ \fi
+ \else
+ \newpage\global\@topnum\z@ \@maketitle\@IEEEstatictitlevskip\@IEEEaftertitletext%
+ \fi
+ \thispagestyle{IEEEtitlepagestyle}\@thanks%
+ \fi
+ % pullup page for pubid if used.
+ \if@IEEEusingpubid
+ \enlargethispage{-\@IEEEpubidpullup}%
+ \fi
+ \endgroup
+ \setcounter{footnote}{0}\let\maketitle\relax\let\@maketitle\relax
+ \gdef\@thanks{}%
+ % v1.6b do not clear these as we will need the title again for peer review papers
+ % \gdef\@author{}\gdef\@title{}%
+ \let\thanks\relax}
+
+
+
+% V1.7 parbox to format \@IEEEcompsoctitleabstractindextext
+\long\def\@IEEEcompsoctitleabstractindextextbox#1{\parbox{0.915\textwidth}{#1}}
+
+
+% formats the Title, authors names, affiliations and special paper notice
+% THIS IS A CONTROLLED SPACING COMMAND! Do not allow blank lines or unintentional
+% spaces to enter the definition - use % at the end of each line
+\def\@maketitle{\newpage
+\begin{center}%
+\ifCLASSOPTIONtechnote% technotes
+ {\bfseries\large\@IEEEcompsoconly{\sffamily}\@title\par}\vskip 1.3em{\lineskip .5em\@IEEEcompsoconly{\sffamily}\@author
+ \@IEEEspecialpapernotice\par{\@IEEEcompsoconly{\vskip 1.5em\relax
+ \@IEEEcompsoctitleabstractindextextbox{\@IEEEcompsoctitleabstractindextext}\par
+ \hfill\@IEEEcompsocdiamondline\hfill\hbox{}\par}}}\relax
+\else% not a technote
+ \vskip0.2em{\Huge\@IEEEcompsoconly{\sffamily}\@IEEEcompsocconfonly{\normalfont\normalsize\vskip 0\@IEEEnormalsizeunitybaselineskip
+ \bfseries\Large}\@title\par}\vskip 1.0em\par%
+ % V1.6 handle \author differently if in conference mode
+ \ifCLASSOPTIONconference%
+ {\@IEEEspecialpapernotice\mbox{}\vskip\@IEEEauthorblockconfadjspace%
+ \mbox{}\hfill\begin{@IEEEauthorhalign}\@author\end{@IEEEauthorhalign}\hfill\mbox{}\par}\relax
+ \else% peerreviewca, peerreview or journal
+ \ifCLASSOPTIONpeerreviewca
+ % peerreviewca handles author names just like conference mode
+ {\@IEEEcompsoconly{\sffamily}\@IEEEspecialpapernotice\mbox{}\vskip\@IEEEauthorblockconfadjspace%
+ \mbox{}\hfill\begin{@IEEEauthorhalign}\@author\end{@IEEEauthorhalign}\hfill\mbox{}\par
+ {\@IEEEcompsoconly{\vskip 1.5em\relax
+ \@IEEEcompsoctitleabstractindextextbox{\@IEEEcompsoctitleabstractindextext}\par\hfill
+ \@IEEEcompsocdiamondline\hfill\hbox{}\par}}}\relax
+ \else% journal or peerreview
+ {\lineskip.5em\@IEEEcompsoconly{\sffamily}\sublargesize\@author\@IEEEspecialpapernotice\par
+ {\@IEEEcompsoconly{\vskip 1.5em\relax
+ \@IEEEcompsoctitleabstractindextextbox{\@IEEEcompsoctitleabstractindextext}\par\hfill
+ \@IEEEcompsocdiamondline\hfill\hbox{}\par}}}\relax
+ \fi
+ \fi
+\fi\end{center}}
+
+
+
+% V1.7 Computer Society "diamond line" which follows index terms for nonconference papers
+\def\@IEEEcompsocdiamondline{\vrule depth 0pt height 0.5pt width 4cm\hspace{7.5pt}%
+\raisebox{-3.5pt}{\fontfamily{pzd}\fontencoding{U}\fontseries{m}\fontshape{n}\fontsize{11}{12}\selectfont\char70}%
+\hspace{7.5pt}\vrule depth 0pt height 0.5pt width 4cm\relax}
+
+% V1.7 standard LateX2e \thanks, but with \itshape under compsoc. Also make it a \long\def
+% We also need to trigger the one-shot footnote rule
+\def\@IEEEtriggeroneshotfootnoterule{\global\@IEEEenableoneshotfootnoteruletrue}
+
+
+\long\def\thanks#1{\footnotemark
+ \protected@xdef\@thanks{\@thanks
+ \protect\footnotetext[\the\c@footnote]{\@IEEEcompsoconly{\itshape
+ \protect\@IEEEtriggeroneshotfootnoterule\relax}\ignorespaces#1}}}
+\let\@thanks\@empty
+
+% V1.7 allow \author to contain \par's. This is needed to allow \thanks to contain \par.
+\long\def\author#1{\gdef\@author{#1}}
+
+
+% in addition to setting up IEEEitemize, we need to remove a baselineskip space above and
+% below it because \list's \pars introduce blank lines because of the footnote struts.
+\def\@IEEEsetupcompsocitemizelist{\def\labelitemi{$\bullet$}%
+\setlength{\IEEElabelindent}{0pt}\setlength{\parskip}{0pt}%
+\setlength{\partopsep}{0pt}\setlength{\topsep}{0.5\baselineskip}\vspace{-1\baselineskip}\relax}
+
+
+% flag for fake non-compsoc \IEEEcompsocthanksitem - prevents line break on very first item
+\newif\if@IEEEbreakcompsocthanksitem \@IEEEbreakcompsocthanksitemfalse
+
+\ifCLASSOPTIONcompsoc
+% V1.7 compsoc bullet item \thanks
+% also, we need to redefine this to destroy the argument in \@IEEEdynamictitlevspace
+\long\def\IEEEcompsocitemizethanks#1{\relax\@IEEEbreakcompsocthanksitemfalse\footnotemark
+ \protected@xdef\@thanks{\@thanks
+ \protect\footnotetext[\the\c@footnote]{\itshape\protect\@IEEEtriggeroneshotfootnoterule
+ {\let\IEEEiedlistdecl\relax\protect\begin{IEEEitemize}[\protect\@IEEEsetupcompsocitemizelist]\ignorespaces#1\relax
+ \protect\end{IEEEitemize}}\protect\vspace{-1\baselineskip}}}}
+\DeclareRobustCommand*{\IEEEcompsocthanksitem}{\item}
+\else
+% non-compsoc, allow for dual compilation via rerouting to normal \thanks
+\long\def\IEEEcompsocitemizethanks#1{\thanks{#1}}
+% redirect to "pseudo-par" \hfil\break\indent after swallowing [] from \IEEEcompsocthanksitem[]
+\DeclareRobustCommand{\IEEEcompsocthanksitem}{\@ifnextchar [{\@IEEEthanksswallowoptionalarg}%
+{\@IEEEthanksswallowoptionalarg[\relax]}}
+% be sure and break only after first item, be sure and ignore spaces after optional argument
+\def\@IEEEthanksswallowoptionalarg[#1]{\relax\if@IEEEbreakcompsocthanksitem\hfil\break
+\indent\fi\@IEEEbreakcompsocthanksitemtrue\ignorespaces}
+\fi
+
+
+% V1.6b define the \IEEEpeerreviewmaketitle as needed
+\ifCLASSOPTIONpeerreview
+\def\IEEEpeerreviewmaketitle{\@IEEEcleardoublepage{empty}%
+\ifCLASSOPTIONtwocolumn
+\twocolumn[\@IEEEpeerreviewmaketitle\@IEEEdynamictitlevspace]
+\else
+\newpage\@IEEEpeerreviewmaketitle\@IEEEstatictitlevskip
+\fi
+\thispagestyle{IEEEtitlepagestyle}}
+\else
+% \IEEEpeerreviewmaketitle does nothing if peer review option has not been selected
+\def\IEEEpeerreviewmaketitle{\relax}
+\fi
+
+% peerreview formats the repeated title like the title in journal papers.
+\def\@IEEEpeerreviewmaketitle{\begin{center}\@IEEEcompsoconly{\sffamily}%
+\normalfont\normalsize\vskip0.2em{\Huge\@title\par}\vskip1.0em\par
+\end{center}}
+
+
+
+% V1.6
+% this is a static rubber spacer between the title/authors and the main text
+% used for single column text, or when the title appears in the first column
+% of two column text (technotes).
+\def\@IEEEstatictitlevskip{{\normalfont\normalsize
+% adjust spacing to next text
+% v1.6b handle peer review papers
+\ifCLASSOPTIONpeerreview
+% for peer review papers, the same value is used for both title pages
+% regardless of the other paper modes
+ \vskip 1\baselineskip plus 0.375\baselineskip minus 0.1875\baselineskip
+\else
+ \ifCLASSOPTIONconference% conference
+ \vskip 1\baselineskip plus 0.375\baselineskip minus 0.1875\baselineskip%
+ \else%
+ \ifCLASSOPTIONtechnote% technote
+ \vskip 1\baselineskip plus 0.375\baselineskip minus 0.1875\baselineskip%
+ \else% journal uses more space
+ \vskip 2.5\baselineskip plus 0.75\baselineskip minus 0.375\baselineskip%
+ \fi
+ \fi
+\fi}}
+
+
+% V1.6
+% This is a dynamically determined rigid spacer between the title/authors
+% and the main text. This is used only for single column titles over two
+% column text (most common)
+% This is bit tricky because we have to ensure that the textheight of the
+% main text is an integer multiple of \baselineskip
+% otherwise underfull vbox problems may develop in the second column of the
+% text on the titlepage
+% The possible use of \IEEEpubid must also be taken into account.
+\def\@IEEEdynamictitlevspace{{%
+ % we run within a group so that all the macros can be forgotten when we are done
+ \long\def\thanks##1{\relax}%don't allow \thanks to run when we evaluate the vbox height
+ \long\def\IEEEcompsocitemizethanks##1{\relax}%don't allow \IEEEcompsocitemizethanks to run when we evaluate the vbox height
+ \normalfont\normalsize% we declare more descriptive variable names
+ \let\@IEEEmaintextheight=\@IEEEtrantmpdimenA%height of the main text columns
+ \let\@IEEEINTmaintextheight=\@IEEEtrantmpdimenB%height of the main text columns with integer # lines
+ % set the nominal and minimum values for the title spacer
+ % the dynamic algorithm will not allow the spacer size to
+ % become less than \@IEEEMINtitlevspace - instead it will be
+ % lengthened
+ % default to journal values
+ \def\@IEEENORMtitlevspace{2.5\baselineskip}%
+ \def\@IEEEMINtitlevspace{2\baselineskip}%
+ % conferences and technotes need tighter spacing
+ \ifCLASSOPTIONconference%conference
+ \def\@IEEENORMtitlevspace{1\baselineskip}%
+ \def\@IEEEMINtitlevspace{0.75\baselineskip}%
+ \fi
+ \ifCLASSOPTIONtechnote%technote
+ \def\@IEEENORMtitlevspace{1\baselineskip}%
+ \def\@IEEEMINtitlevspace{0.75\baselineskip}%
+ \fi%
+ % get the height that the title will take up
+ \ifCLASSOPTIONpeerreview
+ \settoheight{\@IEEEmaintextheight}{\vbox{\hsize\textwidth \@IEEEpeerreviewmaketitle}}%
+ \else
+ \settoheight{\@IEEEmaintextheight}{\vbox{\hsize\textwidth \@maketitle}}%
+ \fi
+ \@IEEEmaintextheight=-\@IEEEmaintextheight% title takes away from maintext, so reverse sign
+ % add the height of the page textheight
+ \advance\@IEEEmaintextheight by \textheight%
+ % correct for title pages using pubid
+ \ifCLASSOPTIONpeerreview\else
+ % peerreview papers use the pubid on the cover page only.
+ % And the cover page uses a static spacer.
+ \if@IEEEusingpubid\advance\@IEEEmaintextheight by -\@IEEEpubidpullup\fi
+ \fi%
+ % subtract off the nominal value of the title bottom spacer
+ \advance\@IEEEmaintextheight by -\@IEEENORMtitlevspace%
+ % \topskip takes away some too
+ \advance\@IEEEmaintextheight by -\topskip%
+ % calculate the column height of the main text for lines
+ % now we calculate the main text height as if holding
+ % an integer number of \normalsize lines after the first
+ % and discard any excess fractional remainder
+ % we subtracted the first line, because the first line
+ % is placed \topskip into the maintext, not \baselineskip like the
+ % rest of the lines.
+ \@IEEEINTmaintextheight=\@IEEEmaintextheight%
+ \divide\@IEEEINTmaintextheight by \baselineskip%
+ \multiply\@IEEEINTmaintextheight by \baselineskip%
+ % now we calculate how much the title spacer height will
+ % have to be reduced from nominal (\@IEEEREDUCEmaintextheight is always
+ % a positive value) so that the maintext area will contain an integer
+ % number of normal size lines
+ % we change variable names here (to avoid confusion) as we no longer
+ % need \@IEEEINTmaintextheight and can reuse its dimen register
+ \let\@IEEEREDUCEmaintextheight=\@IEEEINTmaintextheight%
+ \advance\@IEEEREDUCEmaintextheight by -\@IEEEmaintextheight%
+ \advance\@IEEEREDUCEmaintextheight by \baselineskip%
+ % this is the calculated height of the spacer
+ % we change variable names here (to avoid confusion) as we no longer
+ % need \@IEEEmaintextheight and can reuse its dimen register
+ \let\@IEEECOMPENSATElen=\@IEEEmaintextheight%
+ \@IEEECOMPENSATElen=\@IEEENORMtitlevspace% set the nominal value
+ % we go with the reduced length if it is smaller than an increase
+ \ifdim\@IEEEREDUCEmaintextheight < 0.5\baselineskip\relax%
+ \advance\@IEEECOMPENSATElen by -\@IEEEREDUCEmaintextheight%
+ % if the resulting spacer is too small back out and go with an increase instead
+ \ifdim\@IEEECOMPENSATElen<\@IEEEMINtitlevspace\relax%
+ \advance\@IEEECOMPENSATElen by \baselineskip%
+ \fi%
+ \else%
+ % go with an increase because it is closer to the nominal than a decrease
+ \advance\@IEEECOMPENSATElen by -\@IEEEREDUCEmaintextheight%
+ \advance\@IEEECOMPENSATElen by \baselineskip%
+ \fi%
+ % set the calculated rigid spacer
+ \vspace{\@IEEECOMPENSATElen}}}
+
+
+
+% V1.6
+% we allow the user access to the last part of the title area
+% useful in emergencies such as when a different spacing is needed
+% This text is NOT compensated for in the dynamic sizer.
+\let\@IEEEaftertitletext=\relax
+\long\def\IEEEaftertitletext#1{\def\@IEEEaftertitletext{#1}}
+
+% V1.7 provide a way for users to enter abstract and keywords
+% into the onecolumn title are. This text is compensated for
+% in the dynamic sizer.
+\let\@IEEEcompsoctitleabstractindextext=\relax
+\long\def\IEEEcompsoctitleabstractindextext#1{\def\@IEEEcompsoctitleabstractindextext{#1}}
+% V1.7 provide a way for users to get the \@IEEEcompsoctitleabstractindextext if
+% not in compsoc journal mode - this way abstract and keywords can be placed
+% in their conventional position if not in compsoc mode.
+\def\IEEEdisplaynotcompsoctitleabstractindextext{%
+\ifCLASSOPTIONcompsoc% display if compsoc conf
+\ifCLASSOPTIONconference\@IEEEcompsoctitleabstractindextext\fi
+\else% or if not compsoc
+\@IEEEcompsoctitleabstractindextext\fi}
+
+
+% command to allow alteration of baselinestretch, but only if the current
+% baselineskip is unity. Used to tweak the compsoc abstract and keywords line spacing.
+\def\@IEEEtweakunitybaselinestretch#1{{\def\baselinestretch{1}\selectfont
+\global\@tempskipa\baselineskip}\ifnum\@tempskipa=\baselineskip%
+\def\baselinestretch{#1}\selectfont\fi\relax}
+
+
+% abstract and keywords are in \small, except
+% for 9pt docs in which they are in \footnotesize
+% Because 9pt docs use an 8pt footnotesize, \small
+% becomes a rather awkward 8.5pt
+\def\@IEEEabskeysecsize{\small}
+\ifx\CLASSOPTIONpt\@IEEEptsizenine
+ \def\@IEEEabskeysecsize{\footnotesize}
+\fi
+
+% compsoc journals use \footnotesize, compsoc conferences use normalsize
+\@IEEEcompsoconly{\def\@IEEEabskeysecsize{\footnotesize}}
+%\@IEEEcompsocconfonly{\def\@IEEEabskeysecsize{\normalsize}}
+
+
+
+
+% V1.6 have abstract and keywords strip leading spaces, pars and newlines
+% so that spacing is more tightly controlled.
+\def\abstract{\normalfont
+ \if@twocolumn
+ \@IEEEabskeysecsize\bfseries\textit{\abstractname}---\relax
+ \else
+ \begin{center}\vspace{-1.78ex}\@IEEEabskeysecsize\textbf{\abstractname}\end{center}\quotation\@IEEEabskeysecsize
+ \fi\@IEEEgobbleleadPARNLSP}
+% V1.6 IEEE wants only 1 pica from end of abstract to introduction heading when in
+% conference mode (the heading already has this much above it)
+\def\endabstract{\relax\ifCLASSOPTIONconference\vspace{1.34ex}\else\vspace{1.34ex}\fi\par\if@twocolumn\else\endquotation\fi
+ \normalfont\normalsize}
+
+\def\IEEEkeywords{\normalfont
+ \if@twocolumn
+ \@IEEEabskeysecsize\bfseries\textit{\IEEEkeywordsname}-\relax
+ \else
+ \begin{center}\@IEEEabskeysecsize\textbf{\IEEEkeywordsname}\end{center}\quotation\@IEEEabskeysecsize
+ \fi\@IEEEgobbleleadPARNLSP}
+\def\endIEEEkeywords{\relax\ifCLASSOPTIONtechnote\vspace{1.34ex}\else\vspace{0.67ex}\fi
+ \par\if@twocolumn\else\endquotation\fi%
+ \normalfont\normalsize}
+
+% V1.7 compsoc keywords index terms
+\ifCLASSOPTIONcompsoc
+ \ifCLASSOPTIONconference% compsoc conference
+\def\abstract{\normalfont
+ \begin{center}\@IEEEabskeysecsize\textbf{\large\abstractname}\end{center}\vskip 0.5\baselineskip plus 0.1\baselineskip minus 0.1\baselineskip
+ \if@twocolumn\else\quotation\fi\itshape\@IEEEabskeysecsize%
+ \par\@IEEEgobbleleadPARNLSP}
+\def\IEEEkeywords{\normalfont\vskip 1.5\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip
+ \begin{center}\@IEEEabskeysecsize\textbf{\large\IEEEkeywordsname}\end{center}\vskip 0.5\baselineskip plus 0.1\baselineskip minus 0.1\baselineskip
+ \if@twocolumn\else\quotation\fi\itshape\@IEEEabskeysecsize%
+ \par\@IEEEgobbleleadPARNLSP}
+ \else% compsoc not conference
+\def\abstract{\normalfont\@IEEEtweakunitybaselinestretch{1.15}\sffamily
+ \if@twocolumn
+ \@IEEEabskeysecsize\noindent\textbf{\abstractname}---\relax
+ \else
+ \begin{center}\vspace{-1.78ex}\@IEEEabskeysecsize\textbf{\abstractname}\end{center}\quotation\@IEEEabskeysecsize%
+ \fi\@IEEEgobbleleadPARNLSP}
+\def\IEEEkeywords{\normalfont\@IEEEtweakunitybaselinestretch{1.15}\sffamily
+ \if@twocolumn
+ \@IEEEabskeysecsize\vskip 0.5\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip\noindent
+ \textbf{\IEEEkeywordsname}---\relax
+ \else
+ \begin{center}\@IEEEabskeysecsize\textbf{\IEEEkeywordsname}\end{center}\quotation\@IEEEabskeysecsize%
+ \fi\@IEEEgobbleleadPARNLSP}
+ \fi
+\fi
+
+
+
+% gobbles all leading \, \\ and \par, upon finding first token that
+% is not a \ , \\ or a \par, it ceases and returns that token
+%
+% used to strip leading \, \\ and \par from the input
+% so that such things in the beginning of an environment will not
+% affect the formatting of the text
+\long\def\@IEEEgobbleleadPARNLSP#1{\let\@IEEEswallowthistoken=0%
+\let\@IEEEgobbleleadPARNLSPtoken#1%
+\let\@IEEEgobbleleadPARtoken=\par%
+\let\@IEEEgobbleleadNLtoken=\\%
+\let\@IEEEgobbleleadSPtoken=\ %
+\def\@IEEEgobbleleadSPMACRO{\ }%
+\ifx\@IEEEgobbleleadPARNLSPtoken\@IEEEgobbleleadPARtoken%
+\let\@IEEEswallowthistoken=1%
+\fi%
+\ifx\@IEEEgobbleleadPARNLSPtoken\@IEEEgobbleleadNLtoken%
+\let\@IEEEswallowthistoken=1%
+\fi%
+\ifx\@IEEEgobbleleadPARNLSPtoken\@IEEEgobbleleadSPtoken%
+\let\@IEEEswallowthistoken=1%
+\fi%
+% a control space will come in as a macro
+% when it is the last one on a line
+\ifx\@IEEEgobbleleadPARNLSPtoken\@IEEEgobbleleadSPMACRO%
+\let\@IEEEswallowthistoken=1%
+\fi%
+% if we have to swallow this token, do so and taste the next one
+% else spit it out and stop gobbling
+\ifx\@IEEEswallowthistoken 1\let\@IEEEnextgobbleleadPARNLSP=\@IEEEgobbleleadPARNLSP\else%
+\let\@IEEEnextgobbleleadPARNLSP=#1\fi%
+\@IEEEnextgobbleleadPARNLSP}%
+
+
+
+
+% TITLING OF SECTIONS
+\def\@IEEEsectpunct{:\ \,} % Punctuation after run-in section heading (headings which are
+ % part of the paragraphs), need little bit more than a single space
+ % spacing from section number to title
+% compsoc conferences use regular period/space punctuation
+\ifCLASSOPTIONcompsoc
+\ifCLASSOPTIONconference
+\def\@IEEEsectpunct{.\ }
+\fi\fi
+
+
+\def\@seccntformat#1{\csname the#1dis\endcsname\hskip 0.5em\relax}
+
+\ifCLASSOPTIONcompsoc
+% compsoc journals need extra spacing
+\ifCLASSOPTIONconference\else
+\def\@seccntformat#1{\csname the#1dis\endcsname\hskip 1em\relax}
+\fi\fi
+
+%v1.7 put {} after #6 to allow for some types of user font control
+%and use \@@par rather than \par
+\def\@sect#1#2#3#4#5#6[#7]#8{%
+ \ifnum #2>\c@secnumdepth
+ \let\@svsec\@empty
+ \else
+ \refstepcounter{#1}%
+ % load section label and spacer into \@svsec
+ \protected@edef\@svsec{\@seccntformat{#1}\relax}%
+ \fi%
+ \@tempskipa #5\relax
+ \ifdim \@tempskipa>\z@% tempskipa determines whether is treated as a high
+ \begingroup #6{\relax% or low level heading
+ \noindent % subsections are NOT indented
+ % print top level headings. \@svsec is label, #8 is heading title
+ % IEEE does not block indent the section title text, it flows like normal
+ {\hskip #3\relax\@svsec}{\interlinepenalty \@M #8\@@par}}%
+ \endgroup
+ \addcontentsline{toc}{#1}{\ifnum #2>\c@secnumdepth\relax\else
+ \protect\numberline{\csname the#1\endcsname}\fi#7}%
+ \else % printout low level headings
+ % svsechd seems to swallow the trailing space, protect it with \mbox{}
+ % got rid of sectionmark stuff
+ \def\@svsechd{#6{\hskip #3\relax\@svsec #8\@IEEEsectpunct\mbox{}}%
+ \addcontentsline{toc}{#1}{\ifnum #2>\c@secnumdepth\relax\else
+ \protect\numberline{\csname the#1\endcsname}\fi#7}}%
+ \fi%skip down
+ \@xsect{#5}}
+
+
+% section* handler
+%v1.7 put {} after #4 to allow for some types of user font control
+%and use \@@par rather than \par
+\def\@ssect#1#2#3#4#5{\@tempskipa #3\relax
+ \ifdim \@tempskipa>\z@
+ %\begingroup #4\@hangfrom{\hskip #1}{\interlinepenalty \@M #5\par}\endgroup
+ % IEEE does not block indent the section title text, it flows like normal
+ \begingroup \noindent #4{\relax{\hskip #1}{\interlinepenalty \@M #5\@@par}}\endgroup
+ % svsechd swallows the trailing space, protect it with \mbox{}
+ \else \def\@svsechd{#4{\hskip #1\relax #5\@IEEEsectpunct\mbox{}}}\fi
+ \@xsect{#3}}
+
+
+%% SECTION heading spacing and font
+%%
+% arguments are: #1 - sectiontype name
+% (for \@sect) #2 - section level
+% #3 - section heading indent
+% #4 - top separation (absolute value used, neg indicates not to indent main text)
+% If negative, make stretch parts negative too!
+% #5 - (absolute value used) positive: bottom separation after heading,
+% negative: amount to indent main text after heading
+% Both #4 and #5 negative means to indent main text and use negative top separation
+% #6 - font control
+% You've got to have \normalfont\normalsize in the font specs below to prevent
+% trouble when you do something like:
+% \section{Note}{\ttfamily TT-TEXT} is known to ...
+% IEEE sometimes REALLY stretches the area before a section
+% heading by up to about 0.5in. However, it may not be a good
+% idea to let LaTeX have quite this much rubber.
+\ifCLASSOPTIONconference%
+% IEEE wants section heading spacing to decrease for conference mode
+\def\section{\@startsection{section}{1}{\z@}{1.5ex plus 1.5ex minus 0.5ex}%
+{0.7ex plus 1ex minus 0ex}{\normalfont\normalsize\centering\scshape}}%
+\def\subsection{\@startsection{subsection}{2}{\z@}{1.5ex plus 1.5ex minus 0.5ex}%
+{0.7ex plus .5ex minus 0ex}{\normalfont\normalsize\itshape}}%
+\else % for journals
+\def\section{\@startsection{section}{1}{\z@}{3.0ex plus 1.5ex minus 1.5ex}% V1.6 3.0ex from 3.5ex
+{0.7ex plus 1ex minus 0ex}{\normalfont\normalsize\centering\scshape}}%
+\def\subsection{\@startsection{subsection}{2}{\z@}{3.5ex plus 1.5ex minus 1.5ex}%
+{0.7ex plus .5ex minus 0ex}{\normalfont\normalsize\itshape}}%
+\fi
+
+% for both journals and conferences
+% decided to put in a little rubber above the section, might help somebody
+\def\subsubsection{\@startsection{subsubsection}{3}{\parindent}{0ex plus 0.1ex minus 0.1ex}%
+{0ex}{\normalfont\normalsize\itshape}}%
+\def\paragraph{\@startsection{paragraph}{4}{2\parindent}{0ex plus 0.1ex minus 0.1ex}%
+{0ex}{\normalfont\normalsize\itshape}}%
+
+
+% compsoc
+\ifCLASSOPTIONcompsoc
+\ifCLASSOPTIONconference
+% compsoc conference
+\def\section{\@startsection{section}{1}{\z@}{1\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip}%
+{1\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip}{\normalfont\large\bfseries}}%
+\def\subsection{\@startsection{subsection}{2}{\z@}{1\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip}%
+{1\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip}{\normalfont\sublargesize\bfseries}}%
+\def\subsubsection{\@startsection{subsubsection}{3}{\z@}{1\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip}%
+{0ex}{\normalfont\normalsize\bfseries}}%
+\def\paragraph{\@startsection{paragraph}{4}{2\parindent}{0ex plus 0.1ex minus 0.1ex}%
+{0ex}{\normalfont\normalsize}}%
+\else% compsoc journals
+% use negative top separation as compsoc journals do not indent paragraphs after section titles
+\def\section{\@startsection{section}{1}{\z@}{-3ex plus -2ex minus -1.5ex}%
+{0.7ex plus 1ex minus 0ex}{\normalfont\large\sffamily\bfseries\scshape}}%
+% Note that subsection and smaller may not be correct for the Computer Society,
+% I have to look up an example.
+\def\subsection{\@startsection{subsection}{2}{\z@}{-3.5ex plus -1.5ex minus -1.5ex}%
+{0.7ex plus .5ex minus 0ex}{\normalfont\normalsize\sffamily\bfseries}}%
+\def\subsubsection{\@startsection{subsubsection}{3}{\z@}{-2.5ex plus -1ex minus -1ex}%
+{0.5ex plus 0.5ex minus 0ex}{\normalfont\normalsize\sffamily\itshape}}%
+\def\paragraph{\@startsection{paragraph}{4}{2\parindent}{-0ex plus -0.1ex minus -0.1ex}%
+{0ex}{\normalfont\normalsize}}%
+\fi\fi
+
+
+
+
+%% ENVIRONMENTS
+% "box" symbols at end of proofs
+\def\IEEEQEDclosed{\mbox{\rule[0pt]{1.3ex}{1.3ex}}} % for a filled box
+% V1.6 some journals use an open box instead that will just fit around a closed one
+\def\IEEEQEDopen{{\setlength{\fboxsep}{0pt}\setlength{\fboxrule}{0.2pt}\fbox{\rule[0pt]{0pt}{1.3ex}\rule[0pt]{1.3ex}{0pt}}}}
+\ifCLASSOPTIONcompsoc
+\def\IEEEQED{\IEEEQEDopen} % default to open for compsoc
+\else
+\def\IEEEQED{\IEEEQEDclosed} % otherwise default to closed
+\fi
+
+% v1.7 name change to avoid namespace collision with amsthm. Also add support
+% for an optional argument.
+\def\IEEEproof{\@ifnextchar[{\@IEEEproof}{\@IEEEproof[\IEEEproofname]}}
+\def\@IEEEproof[#1]{\par\noindent\hspace{2em}{\itshape #1: }}
+\def\endIEEEproof{\hspace*{\fill}~\IEEEQED\par}
+
+
+%\itemindent is set to \z@ by list, so define new temporary variable
+\newdimen\@IEEEtmpitemindent
+\def\@begintheorem#1#2{\@IEEEtmpitemindent\itemindent\topsep 0pt\rmfamily\trivlist%
+ \item[\hskip \labelsep{\indent\itshape #1\ #2:}]\itemindent\@IEEEtmpitemindent}
+\def\@opargbegintheorem#1#2#3{\@IEEEtmpitemindent\itemindent\topsep 0pt\rmfamily \trivlist%
+% V1.6 IEEE is back to using () around theorem names which are also in italics
+% Thanks to Christian Peel for reporting this.
+ \item[\hskip\labelsep{\indent\itshape #1\ #2\ (#3):}]\itemindent\@IEEEtmpitemindent}
+% V1.7 remove bogus \unskip that caused equations in theorems to collide with
+% lines below.
+\def\@endtheorem{\endtrivlist}
+
+% V1.6
+% display command for the section the theorem is in - so that \thesection
+% is not used as this will be in Roman numerals when we want arabic.
+% LaTeX2e uses \def\@thmcounter#1{\noexpand\arabic{#1}} for the theorem number
+% (second part) display and \def\@thmcountersep{.} as a separator.
+% V1.7 intercept calls to the section counter and reroute to \@IEEEthmcounterinsection
+% to allow \appendix(ices} to override as needed.
+%
+% special handler for sections, allows appendix(ices) to override
+\gdef\@IEEEthmcounterinsection#1{\arabic{#1}}
+% string macro
+\edef\@IEEEstringsection{section}
+
+% redefine the #1#2[#3] form of newtheorem to use a hook to \@IEEEthmcounterinsection
+% if section in_counter is used
+\def\@xnthm#1#2[#3]{%
+ \expandafter\@ifdefinable\csname #1\endcsname
+ {\@definecounter{#1}\@newctr{#1}[#3]%
+ \edef\@IEEEstringtmp{#3}
+ \ifx\@IEEEstringtmp\@IEEEstringsection
+ \expandafter\xdef\csname the#1\endcsname{%
+ \noexpand\@IEEEthmcounterinsection{#3}\@thmcountersep
+ \@thmcounter{#1}}%
+ \else
+ \expandafter\xdef\csname the#1\endcsname{%
+ \expandafter\noexpand\csname the#3\endcsname \@thmcountersep
+ \@thmcounter{#1}}%
+ \fi
+ \global\@namedef{#1}{\@thm{#1}{#2}}%
+ \global\@namedef{end#1}{\@endtheorem}}}
+
+
+
+%% SET UP THE DEFAULT PAGESTYLE
+\ps@headings
+\pagenumbering{arabic}
+
+% normally the page counter starts at 1
+\setcounter{page}{1}
+% however, for peerreview the cover sheet is page 0 or page -1
+% (for duplex printing)
+\ifCLASSOPTIONpeerreview
+ \if@twoside
+ \setcounter{page}{-1}
+ \else
+ \setcounter{page}{0}
+ \fi
+\fi
+
+% standard book class behavior - let bottom line float up and down as
+% needed when single sided
+\ifCLASSOPTIONtwoside\else\raggedbottom\fi
+% if two column - turn on twocolumn, allow word spacings to stretch more and
+% enforce a rigid position for the last lines
+\ifCLASSOPTIONtwocolumn
+% the peer review option delays invoking twocolumn
+ \ifCLASSOPTIONpeerreview\else
+ \twocolumn
+ \fi
+\sloppy
+\flushbottom
+\fi
+
+
+
+
+% \APPENDIX and \APPENDICES definitions
+
+% This is the \@ifmtarg command from the LaTeX ifmtarg package
+% by Peter Wilson (CUA) and Donald Arseneau
+% \@ifmtarg is used to determine if an argument to a command
+% is present or not.
+% For instance:
+% \@ifmtarg{#1}{\typeout{empty}}{\typeout{has something}}
+% \@ifmtarg is used with our redefined \section command if
+% \appendices is invoked.
+% The command \section will behave slightly differently depending
+% on whether the user specifies a title:
+% \section{My appendix title}
+% or not:
+% \section{}
+% This way, we can eliminate the blank lines where the title
+% would be, and the unneeded : after Appendix in the table of
+% contents
+\begingroup
+\catcode`\Q=3
+\long\gdef\@ifmtarg#1{\@xifmtarg#1QQ\@secondoftwo\@firstoftwo\@nil}
+\long\gdef\@xifmtarg#1#2Q#3#4#5\@nil{#4}
+\endgroup
+% end of \@ifmtarg defs
+
+
+% V1.7
+% command that allows the one time saving of the original definition
+% of section to \@IEEEappendixsavesection for \appendix or \appendices
+% we don't save \section here as it may be redefined later by other
+% packages (hyperref.sty, etc.)
+\def\@IEEEsaveoriginalsectiononce{\let\@IEEEappendixsavesection\section
+\let\@IEEEsaveoriginalsectiononce\relax}
+
+% neat trick to grab and process the argument from \section{argument}
+% we process differently if the user invoked \section{} with no
+% argument (title)
+% note we reroute the call to the old \section*
+\def\@IEEEprocessthesectionargument#1{%
+\@ifmtarg{#1}{%
+\@IEEEappendixsavesection*{\appendixname~\thesectiondis}%
+\addcontentsline{toc}{section}{\appendixname~\thesection}}{%
+\@IEEEappendixsavesection*{\appendixname~\thesectiondis \\* #1}%
+\addcontentsline{toc}{section}{\appendixname~\thesection: #1}}}
+
+% we use this if the user calls \section{} after
+% \appendix-- which has no meaning. So, we ignore the
+% command and its argument. Then, warn the user.
+\def\@IEEEdestroythesectionargument#1{\typeout{** WARNING: Ignoring useless
+\protect\section\space in Appendix (line \the\inputlineno).}}
+
+
+% remember \thesection forms will be displayed in \ref calls
+% and in the Table of Contents.
+% The \sectiondis form is used in the actual heading itself
+
+% appendix command for one single appendix
+% normally has no heading. However, if you want a
+% heading, you can do so via the optional argument:
+% \appendix[Optional Heading]
+\def\appendix{\relax}
+\renewcommand{\appendix}[1][]{\@IEEEsaveoriginalsectiononce\par
+ % v1.6 keep hyperref's identifiers unique
+ \gdef\theHsection{Appendix.A}%
+ % v1.6 adjust hyperref's string name for the section
+ \xdef\Hy@chapapp{appendix}%
+ \setcounter{section}{0}%
+ \setcounter{subsection}{0}%
+ \setcounter{subsubsection}{0}%
+ \setcounter{paragraph}{0}%
+ \gdef\thesection{A}%
+ \gdef\thesectiondis{}%
+ \gdef\thesubsection{\Alph{subsection}}%
+ \gdef\@IEEEthmcounterinsection##1{A}
+ \refstepcounter{section}% update the \ref counter
+ \@ifmtarg{#1}{\@IEEEappendixsavesection*{\appendixname}%
+ \addcontentsline{toc}{section}{\appendixname}}{%
+ \@IEEEappendixsavesection*{\appendixname~\\* #1}%
+ \addcontentsline{toc}{section}{\appendixname: #1}}%
+ % redefine \section command for appendix
+ % leave \section* as is
+ \def\section{\@ifstar{\@IEEEappendixsavesection*}{%
+ \@IEEEdestroythesectionargument}}% throw out the argument
+ % of the normal form
+}
+
+
+
+% appendices command for multiple appendices
+% user then calls \section with an argument (possibly empty) to
+% declare the individual appendices
+\def\appendices{\@IEEEsaveoriginalsectiononce\par
+ % v1.6 keep hyperref's identifiers unique
+ \gdef\theHsection{Appendix.\Alph{section}}%
+ % v1.6 adjust hyperref's string name for the section
+ \xdef\Hy@chapapp{appendix}%
+ \setcounter{section}{-1}% we want \refstepcounter to use section 0
+ \setcounter{subsection}{0}%
+ \setcounter{subsubsection}{0}%
+ \setcounter{paragraph}{0}%
+ \ifCLASSOPTIONromanappendices%
+ \gdef\thesection{\Roman{section}}%
+ \gdef\thesectiondis{\Roman{section}}%
+ \@IEEEcompsocconfonly{\gdef\thesectiondis{\Roman{section}.}}%
+ \gdef\@IEEEthmcounterinsection##1{A\arabic{##1}}
+ \else%
+ \gdef\thesection{\Alph{section}}%
+ \gdef\thesectiondis{\Alph{section}}%
+ \@IEEEcompsocconfonly{\gdef\thesectiondis{\Alph{section}.}}%
+ \gdef\@IEEEthmcounterinsection##1{\Alph{##1}}
+ \fi%
+ \refstepcounter{section}% update the \ref counter
+ \setcounter{section}{0}% NEXT \section will be the FIRST appendix
+ % redefine \section command for appendices
+ % leave \section* as is
+ \def\section{\@ifstar{\@IEEEappendixsavesection*}{% process the *-form
+ \refstepcounter{section}% or is a new section so,
+ \@IEEEprocessthesectionargument}}% process the argument
+ % of the normal form
+}
+
+
+
+% \IEEEPARstart
+% Definition for the big two line drop cap letter at the beginning of the
+% first paragraph of journal papers. The first argument is the first letter
+% of the first word, the second argument is the remaining letters of the
+% first word which will be rendered in upper case.
+% In V1.6 this has been completely rewritten to:
+%
+% 1. no longer have problems when the user begins an environment
+% within the paragraph that uses \IEEEPARstart.
+% 2. auto-detect and use the current font family
+% 3. revise handling of the space at the end of the first word so that
+% interword glue will now work as normal.
+% 4. produce correctly aligned edges for the (two) indented lines.
+%
+% We generalize things via control macros - playing with these is fun too.
+%
+% V1.7 added more control macros to make it easy for IEEEtrantools.sty users
+% to change the font style.
+%
+% the number of lines that are indented to clear it
+% may need to increase if using decenders
+\def\@IEEEPARstartDROPLINES{2}
+% minimum number of lines left on a page to allow a \@IEEEPARstart
+% Does not take into consideration rubber shrink, so it tends to
+% be overly cautious
+\def\@IEEEPARstartMINPAGELINES{2}
+% V1.7 the height of the drop cap is adjusted to match the height of this text
+% in the current font (when \IEEEPARstart is called).
+\def\@IEEEPARstartHEIGHTTEXT{T}
+% the depth the letter is lowered below the baseline
+% the height (and size) of the letter is determined by the sum
+% of this value and the height of the \@IEEEPARstartHEIGHTTEXT in the current
+% font. It is a good idea to set this value in terms of the baselineskip
+% so that it can respond to changes therein.
+\def\@IEEEPARstartDROPDEPTH{1.1\baselineskip}
+% V1.7 the font the drop cap will be rendered in,
+% can take zero or one argument.
+\def\@IEEEPARstartFONTSTYLE{\bfseries}
+% V1.7 any additional, non-font related commands needed to modify
+% the drop cap letter, can take zero or one argument.
+\def\@IEEEPARstartCAPSTYLE{\MakeUppercase}
+% V1.7 the font that will be used to render the rest of the word,
+% can take zero or one argument.
+\def\@IEEEPARstartWORDFONTSTYLE{\relax}
+% V1.7 any additional, non-font related commands needed to modify
+% the rest of the word, can take zero or one argument.
+\def\@IEEEPARstartWORDCAPSTYLE{\MakeUppercase}
+% This is the horizontal separation distance from the drop letter to the main text.
+% Lengths that depend on the font (e.g., ex, em, etc.) will be referenced
+% to the font that is active when \IEEEPARstart is called.
+\def\@IEEEPARstartSEP{0.15em}
+% V1.7 horizontal offset applied to the left of the drop cap.
+\def\@IEEEPARstartHOFFSET{0em}
+% V1.7 Italic correction command applied at the end of the drop cap.
+\def\@IEEEPARstartITLCORRECT{\/}
+
+% V1.7 compoc uses nonbold drop cap and small caps word style
+\ifCLASSOPTIONcompsoc
+\def\@IEEEPARstartFONTSTYLE{\mdseries}
+\def\@IEEEPARstartWORDFONTSTYLE{\scshape}
+\def\@IEEEPARstartWORDCAPSTYLE{\relax}
+\fi
+
+% definition of \IEEEPARstart
+% THIS IS A CONTROLLED SPACING AREA, DO NOT ALLOW SPACES WITHIN THESE LINES
+%
+% The token \@IEEEPARstartfont will be globally defined after the first use
+% of \IEEEPARstart and will be a font command which creates the big letter
+% The first argument is the first letter of the first word and the second
+% argument is the rest of the first word(s).
+\def\IEEEPARstart#1#2{\par{%
+% if this page does not have enough space, break it and lets start
+% on a new one
+\@IEEEtranneedspace{\@IEEEPARstartMINPAGELINES\baselineskip}{\relax}%
+% V1.7 move this up here in case user uses \textbf for \@IEEEPARstartFONTSTYLE
+% which uses command \leavevmode which causes an unwanted \indent to be issued
+\noindent
+% calculate the desired height of the big letter
+% it extends from the top of \@IEEEPARstartHEIGHTTEXT in the current font
+% down to \@IEEEPARstartDROPDEPTH below the current baseline
+\settoheight{\@IEEEtrantmpdimenA}{\@IEEEPARstartHEIGHTTEXT}%
+\addtolength{\@IEEEtrantmpdimenA}{\@IEEEPARstartDROPDEPTH}%
+% extract the name of the current font in bold
+% and place it in \@IEEEPARstartFONTNAME
+\def\@IEEEPARstartGETFIRSTWORD##1 ##2\relax{##1}%
+{\@IEEEPARstartFONTSTYLE{\selectfont\edef\@IEEEPARstartFONTNAMESPACE{\fontname\font\space}%
+\xdef\@IEEEPARstartFONTNAME{\expandafter\@IEEEPARstartGETFIRSTWORD\@IEEEPARstartFONTNAMESPACE\relax}}}%
+% define a font based on this name with a point size equal to the desired
+% height of the drop letter
+\font\@IEEEPARstartsubfont\@IEEEPARstartFONTNAME\space at \@IEEEtrantmpdimenA\relax%
+% save this value as a counter (integer) value (sp points)
+\@IEEEtrantmpcountA=\@IEEEtrantmpdimenA%
+% now get the height of the actual letter produced by this font size
+\settoheight{\@IEEEtrantmpdimenB}{\@IEEEPARstartsubfont\@IEEEPARstartCAPSTYLE{#1}}%
+% If something bogus happens like the first argument is empty or the
+% current font is strange, do not allow a zero height.
+\ifdim\@IEEEtrantmpdimenB=0pt\relax%
+\typeout{** WARNING: IEEEPARstart drop letter has zero height! (line \the\inputlineno)}%
+\typeout{ Forcing the drop letter font size to 10pt.}%
+\@IEEEtrantmpdimenB=10pt%
+\fi%
+% and store it as a counter
+\@IEEEtrantmpcountB=\@IEEEtrantmpdimenB%
+% Since a font size doesn't exactly correspond to the height of the capital
+% letters in that font, the actual height of the letter, \@IEEEtrantmpcountB,
+% will be less than that desired, \@IEEEtrantmpcountA
+% we need to raise the font size, \@IEEEtrantmpdimenA
+% by \@IEEEtrantmpcountA / \@IEEEtrantmpcountB
+% But, TeX doesn't have floating point division, so we have to use integer
+% division. Hence the use of the counters.
+% We need to reduce the denominator so that the loss of the remainder will
+% have minimal affect on the accuracy of the result
+\divide\@IEEEtrantmpcountB by 200%
+\divide\@IEEEtrantmpcountA by \@IEEEtrantmpcountB%
+% Then reequalize things when we use TeX's ability to multiply by
+% floating point values
+\@IEEEtrantmpdimenB=0.005\@IEEEtrantmpdimenA%
+\multiply\@IEEEtrantmpdimenB by \@IEEEtrantmpcountA%
+% \@IEEEPARstartfont is globaly set to the calculated font of the big letter
+% We need to carry this out of the local calculation area to to create the
+% big letter.
+\global\font\@IEEEPARstartfont\@IEEEPARstartFONTNAME\space at \@IEEEtrantmpdimenB%
+% Now set \@IEEEtrantmpdimenA to the width of the big letter
+% We need to carry this out of the local calculation area to set the
+% hanging indent
+\settowidth{\global\@IEEEtrantmpdimenA}{\@IEEEPARstartfont
+\@IEEEPARstartCAPSTYLE{#1\@IEEEPARstartITLCORRECT}}}%
+% end of the isolated calculation environment
+% add in the extra clearance we want
+\advance\@IEEEtrantmpdimenA by \@IEEEPARstartSEP\relax%
+% add in the optional offset
+\advance\@IEEEtrantmpdimenA by \@IEEEPARstartHOFFSET\relax%
+% V1.7 don't allow negative offsets to produce negative hanging indents
+\@IEEEtrantmpdimenB\@IEEEtrantmpdimenA
+\ifnum\@IEEEtrantmpdimenB < 0 \@IEEEtrantmpdimenB 0pt\fi
+% \@IEEEtrantmpdimenA has the width of the big letter plus the
+% separation space and \@IEEEPARstartfont is the font we need to use
+% Now, we make the letter and issue the hanging indent command
+% The letter is placed in a box of zero width and height so that other
+% text won't be displaced by it.
+\hangindent\@IEEEtrantmpdimenB\hangafter=-\@IEEEPARstartDROPLINES%
+\makebox[0pt][l]{\hspace{-\@IEEEtrantmpdimenA}%
+\raisebox{-\@IEEEPARstartDROPDEPTH}[0pt][0pt]{\hspace{\@IEEEPARstartHOFFSET}%
+\@IEEEPARstartfont\@IEEEPARstartCAPSTYLE{#1\@IEEEPARstartITLCORRECT}%
+\hspace{\@IEEEPARstartSEP}}}%
+{\@IEEEPARstartWORDFONTSTYLE{\@IEEEPARstartWORDCAPSTYLE{\selectfont#2}}}}
+
+
+
+
+
+
+% determines if the space remaining on a given page is equal to or greater
+% than the specified space of argument one
+% if not, execute argument two (only if the remaining space is greater than zero)
+% and issue a \newpage
+%
+% example: \@IEEEtranneedspace{2in}{\vfill}
+%
+% Does not take into consideration rubber shrinkage, so it tends to
+% be overly cautious
+% Based on an example posted by Donald Arseneau
+% Note this macro uses \@IEEEtrantmpdimenB internally for calculations,
+% so DO NOT PASS \@IEEEtrantmpdimenB to this routine
+% if you need a dimen register, import with \@IEEEtrantmpdimenA instead
+\def\@IEEEtranneedspace#1#2{\penalty-100\begingroup%shield temp variable
+\@IEEEtrantmpdimenB\pagegoal\advance\@IEEEtrantmpdimenB-\pagetotal% space left
+\ifdim #1>\@IEEEtrantmpdimenB\relax% not enough space left
+\ifdim\@IEEEtrantmpdimenB>\z@\relax #2\fi%
+\newpage%
+\fi\endgroup}
+
+
+
+% IEEEbiography ENVIRONMENT
+% Allows user to enter biography leaving place for picture (adapts to font size)
+% As of V1.5, a new optional argument allows you to have a real graphic!
+% V1.5 and later also fixes the "colliding biographies" which could happen when a
+% biography's text was shorter than the space for the photo.
+% MDS 7/2001
+% V1.6 prevent multiple biographies from making multiple TOC entries
+\newif\if@IEEEbiographyTOCentrynotmade
+\global\@IEEEbiographyTOCentrynotmadetrue
+
+% biography counter so hyperref can jump directly to the biographies
+% and not just the previous section
+\newcounter{IEEEbiography}
+\setcounter{IEEEbiography}{0}
+
+% photo area size
+\def\@IEEEBIOphotowidth{1.0in} % width of the biography photo area
+\def\@IEEEBIOphotodepth{1.25in} % depth (height) of the biography photo area
+% area cleared for photo
+\def\@IEEEBIOhangwidth{1.14in} % width cleared for the biography photo area
+\def\@IEEEBIOhangdepth{1.25in} % depth cleared for the biography photo area
+ % actual depth will be a multiple of
+ % \baselineskip, rounded up
+\def\@IEEEBIOskipN{4\baselineskip}% nominal value of the vskip above the biography
+
+\newenvironment{IEEEbiography}[2][]{\normalfont\@IEEEcompsoconly{\sffamily}\footnotesize%
+\unitlength 1in\parskip=0pt\par\parindent 1em\interlinepenalty500%
+% we need enough space to support the hanging indent
+% the nominal value of the spacer
+% and one extra line for good measure
+\@IEEEtrantmpdimenA=\@IEEEBIOhangdepth%
+\advance\@IEEEtrantmpdimenA by \@IEEEBIOskipN%
+\advance\@IEEEtrantmpdimenA by 1\baselineskip%
+% if this page does not have enough space, break it and lets start
+% with a new one
+\@IEEEtranneedspace{\@IEEEtrantmpdimenA}{\relax}%
+% nominal spacer can strech, not shrink use 1fil so user can out stretch with \vfill
+\vskip \@IEEEBIOskipN plus 1fil minus 0\baselineskip%
+% the default box for where the photo goes
+\def\@IEEEtempbiographybox{{\setlength{\fboxsep}{0pt}\framebox{%
+\begin{minipage}[b][\@IEEEBIOphotodepth][c]{\@IEEEBIOphotowidth}\centering PLACE\\ PHOTO\\ HERE \end{minipage}}}}%
+%
+% detect if the optional argument was supplied, this requires the
+% \@ifmtarg command as defined in the appendix section above
+% and if so, override the default box with what they want
+\@ifmtarg{#1}{\relax}{\def\@IEEEtempbiographybox{\mbox{\begin{minipage}[b][\@IEEEBIOphotodepth][c]{\@IEEEBIOphotowidth}%
+\centering%
+#1%
+\end{minipage}}}}% end if optional argument supplied
+% Make an entry into the table of contents only if we have not done so before
+\if@IEEEbiographyTOCentrynotmade%
+% link labels to the biography counter so hyperref will jump
+% to the biography, not the previous section
+\setcounter{IEEEbiography}{-1}%
+\refstepcounter{IEEEbiography}%
+\addcontentsline{toc}{section}{Biographies}%
+\global\@IEEEbiographyTOCentrynotmadefalse%
+\fi%
+% one more biography
+\refstepcounter{IEEEbiography}%
+% Make an entry for this name into the table of contents
+\addcontentsline{toc}{subsection}{#2}%
+% V1.6 properly handle if a new paragraph should occur while the
+% hanging indent is still active. Do this by redefining \par so
+% that it will not start a new paragraph. (But it will appear to the
+% user as if it did.) Also, strip any leading pars, newlines, or spaces.
+\let\@IEEEBIOORGparCMD=\par% save the original \par command
+\edef\par{\hfil\break\indent}% the new \par will not be a "real" \par
+\settoheight{\@IEEEtrantmpdimenA}{\@IEEEtempbiographybox}% get height of biography box
+\@IEEEtrantmpdimenB=\@IEEEBIOhangdepth%
+\@IEEEtrantmpcountA=\@IEEEtrantmpdimenB% countA has the hang depth
+\divide\@IEEEtrantmpcountA by \baselineskip% calculates lines needed to produce the hang depth
+\advance\@IEEEtrantmpcountA by 1% ensure we overestimate
+% set the hanging indent
+\hangindent\@IEEEBIOhangwidth%
+\hangafter-\@IEEEtrantmpcountA%
+% reference the top of the photo area to the top of a capital T
+\settoheight{\@IEEEtrantmpdimenB}{\mbox{T}}%
+% set the photo box, give it zero width and height so as not to disturb anything
+\noindent\makebox[0pt][l]{\hspace{-\@IEEEBIOhangwidth}\raisebox{\@IEEEtrantmpdimenB}[0pt][0pt]{%
+\raisebox{-\@IEEEBIOphotodepth}[0pt][0pt]{\@IEEEtempbiographybox}}}%
+% now place the author name and begin the bio text
+\noindent\textbf{#2\ }\@IEEEgobbleleadPARNLSP}{\relax\let\par=\@IEEEBIOORGparCMD\par%
+% 7/2001 V1.5 detect when the biography text is shorter than the photo area
+% and pad the unused area - preventing a collision from the next biography entry
+% MDS
+\ifnum \prevgraf <\@IEEEtrantmpcountA\relax% detect when the biography text is shorter than the photo
+ \advance\@IEEEtrantmpcountA by -\prevgraf% calculate how many lines we need to pad
+ \advance\@IEEEtrantmpcountA by -1\relax% we compensate for the fact that we indented an extra line
+ \@IEEEtrantmpdimenA=\baselineskip% calculate the length of the padding
+ \multiply\@IEEEtrantmpdimenA by \@IEEEtrantmpcountA%
+ \noindent\rule{0pt}{\@IEEEtrantmpdimenA}% insert an invisible support strut
+\fi%
+\par\normalfont}
+
+
+
+% V1.6
+% added biography without a photo environment
+\newenvironment{IEEEbiographynophoto}[1]{%
+% Make an entry into the table of contents only if we have not done so before
+\if@IEEEbiographyTOCentrynotmade%
+% link labels to the biography counter so hyperref will jump
+% to the biography, not the previous section
+\setcounter{IEEEbiography}{-1}%
+\refstepcounter{IEEEbiography}%
+\addcontentsline{toc}{section}{Biographies}%
+\global\@IEEEbiographyTOCentrynotmadefalse%
+\fi%
+% one more biography
+\refstepcounter{IEEEbiography}%
+% Make an entry for this name into the table of contents
+\addcontentsline{toc}{subsection}{#1}%
+\normalfont\@IEEEcompsoconly{\sffamily}\footnotesize\interlinepenalty500%
+\vskip 4\baselineskip plus 1fil minus 0\baselineskip%
+\parskip=0pt\par%
+\noindent\textbf{#1\ }\@IEEEgobbleleadPARNLSP}{\relax\par\normalfont}
+
+
+% provide the user with some old font commands
+% got this from article.cls
+\DeclareOldFontCommand{\rm}{\normalfont\rmfamily}{\mathrm}
+\DeclareOldFontCommand{\sf}{\normalfont\sffamily}{\mathsf}
+\DeclareOldFontCommand{\tt}{\normalfont\ttfamily}{\mathtt}
+\DeclareOldFontCommand{\bf}{\normalfont\bfseries}{\mathbf}
+\DeclareOldFontCommand{\it}{\normalfont\itshape}{\mathit}
+\DeclareOldFontCommand{\sl}{\normalfont\slshape}{\@nomath\sl}
+\DeclareOldFontCommand{\sc}{\normalfont\scshape}{\@nomath\sc}
+\DeclareRobustCommand*\cal{\@fontswitch\relax\mathcal}
+\DeclareRobustCommand*\mit{\@fontswitch\relax\mathnormal}
+
+
+% SPECIAL PAPER NOTICE COMMANDS
+%
+% holds the special notice text
+\def\@IEEEspecialpapernotice{\relax}
+
+% for special papers, like invited papers, the user can do:
+% \IEEEspecialpapernotice{(Invited Paper)} before \maketitle
+\def\IEEEspecialpapernotice#1{\ifCLASSOPTIONconference%
+\def\@IEEEspecialpapernotice{{\sublargesize\textit{#1}\vspace*{1em}}}%
+\else%
+\def\@IEEEspecialpapernotice{{\\*[1.5ex]\sublargesize\textit{#1}}\vspace*{-2ex}}%
+\fi}
+
+
+
+
+% PUBLISHER ID COMMANDS
+% to insert a publisher's ID footer
+% V1.6 \IEEEpubid has been changed so that the change in page size and style
+% occurs in \maketitle. \IEEEpubid must now be issued prior to \maketitle
+% use \IEEEpubidadjcol as before - in the second column of the title page
+% These changes allow \maketitle to take the reduced page height into
+% consideration when dynamically setting the space between the author
+% names and the maintext.
+%
+% the amount the main text is pulled up to make room for the
+% publisher's ID footer
+% IEEE uses about 1.3\baselineskip for journals,
+% dynamic title spacing will clean up the fraction
+\def\@IEEEpubidpullup{1.3\baselineskip}
+\ifCLASSOPTIONtechnote
+% for technotes it must be an integer of baselineskip as there can be no
+% dynamic title spacing for two column mode technotes (the title is in the
+% in first column) and we should maintain an integer number of lines in the
+% second column
+% There are some examples (such as older issues of "Transactions on
+% Information Theory") in which IEEE really pulls the text off the ID for
+% technotes - about 0.55in (or 4\baselineskip). We'll use 2\baselineskip
+% and call it even.
+\def\@IEEEpubidpullup{2\baselineskip}
+\fi
+
+% V1.7 compsoc does not use a pullup
+\ifCLASSOPTIONcompsoc
+\def\@IEEEpubidpullup{0pt}
+\fi
+
+% holds the ID text
+\def\@IEEEpubid{\relax}
+
+% flag so \maketitle can tell if \IEEEpubid was called
+\newif\if@IEEEusingpubid
+\global\@IEEEusingpubidfalse
+% issue this command in the page to have the ID at the bottom
+% V1.6 use before \maketitle
+\def\IEEEpubid#1{\def\@IEEEpubid{#1}\global\@IEEEusingpubidtrue}
+
+
+% command which will pull up (shorten) the column it is executed in
+% to make room for the publisher ID. Place in the second column of
+% the title page when using \IEEEpubid
+% Is smart enough not to do anything when in single column text or
+% if the user hasn't called \IEEEpubid
+% currently needed in for the second column of a page with the
+% publisher ID. If not needed in future releases, please provide this
+% command and define it as \relax for backward compatibility
+% v1.6b do not allow command to operate if the peer review option has been
+% selected because \IEEEpubidadjcol will not be on the cover page.
+% V1.7 do nothing if compsoc
+\def\IEEEpubidadjcol{\ifCLASSOPTIONcompsoc\else\ifCLASSOPTIONpeerreview\else
+\if@twocolumn\if@IEEEusingpubid\enlargethispage{-\@IEEEpubidpullup}\fi\fi\fi\fi}
+
+% Special thanks to Peter Wilson, Daniel Luecking, and the other
+% gurus at comp.text.tex, for helping me to understand how best to
+% implement the IEEEpubid command in LaTeX.
+
+
+
+%% Lockout some commands under various conditions
+
+% general purpose bit bucket
+\newsavebox{\@IEEEtranrubishbin}
+
+% flags to prevent multiple warning messages
+\newif\if@IEEEWARNthanks
+\newif\if@IEEEWARNIEEEPARstart
+\newif\if@IEEEWARNIEEEbiography
+\newif\if@IEEEWARNIEEEbiographynophoto
+\newif\if@IEEEWARNIEEEpubid
+\newif\if@IEEEWARNIEEEpubidadjcol
+\newif\if@IEEEWARNIEEEmembership
+\newif\if@IEEEWARNIEEEaftertitletext
+\@IEEEWARNthankstrue
+\@IEEEWARNIEEEPARstarttrue
+\@IEEEWARNIEEEbiographytrue
+\@IEEEWARNIEEEbiographynophototrue
+\@IEEEWARNIEEEpubidtrue
+\@IEEEWARNIEEEpubidadjcoltrue
+\@IEEEWARNIEEEmembershiptrue
+\@IEEEWARNIEEEaftertitletexttrue
+
+
+%% Lockout some commands when in various modes, but allow them to be restored if needed
+%%
+% save commands which might be locked out
+% so that the user can later restore them if needed
+\let\@IEEESAVECMDthanks\thanks
+\let\@IEEESAVECMDIEEEPARstart\IEEEPARstart
+\let\@IEEESAVECMDIEEEbiography\IEEEbiography
+\let\@IEEESAVECMDendIEEEbiography\endIEEEbiography
+\let\@IEEESAVECMDIEEEbiographynophoto\IEEEbiographynophoto
+\let\@IEEESAVECMDendIEEEbiographynophoto\endIEEEbiographynophoto
+\let\@IEEESAVECMDIEEEpubid\IEEEpubid
+\let\@IEEESAVECMDIEEEpubidadjcol\IEEEpubidadjcol
+\let\@IEEESAVECMDIEEEmembership\IEEEmembership
+\let\@IEEESAVECMDIEEEaftertitletext\IEEEaftertitletext
+
+
+% disable \IEEEPARstart when in draft mode
+% This may have originally been done because the pre-V1.6 drop letter
+% algorithm had problems with a non-unity baselinestretch
+% At any rate, it seems too formal to have a drop letter in a draft
+% paper.
+\ifCLASSOPTIONdraftcls
+\def\IEEEPARstart#1#2{#1#2\if@IEEEWARNIEEEPARstart\typeout{** ATTENTION: \noexpand\IEEEPARstart
+ is disabled in draft mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEPARstartfalse}
+\fi
+% and for technotes
+\ifCLASSOPTIONtechnote
+\def\IEEEPARstart#1#2{#1#2\if@IEEEWARNIEEEPARstart\typeout{** WARNING: \noexpand\IEEEPARstart
+ is locked out for technotes (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEPARstartfalse}
+\fi
+
+
+% lockout unneeded commands when in conference mode
+\ifCLASSOPTIONconference
+% when locked out, \thanks, \IEEEbiography, \IEEEbiographynophoto, \IEEEpubid,
+% \IEEEmembership and \IEEEaftertitletext will all swallow their given text.
+% \IEEEPARstart will output a normal character instead
+% warn the user about these commands only once to prevent the console screen
+% from filling up with redundant messages
+\def\thanks#1{\if@IEEEWARNthanks\typeout{** WARNING: \noexpand\thanks
+ is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNthanksfalse}
+\def\IEEEPARstart#1#2{#1#2\if@IEEEWARNIEEEPARstart\typeout{** WARNING: \noexpand\IEEEPARstart
+ is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEPARstartfalse}
+
+
+% LaTeX treats environments and commands with optional arguments differently.
+% the actual ("internal") command is stored as \\commandname
+% (accessed via \csname\string\commandname\endcsname )
+% the "external" command \commandname is a macro with code to determine
+% whether or not the optional argument is presented and to provide the
+% default if it is absent. So, in order to save and restore such a command
+% we would have to save and restore \\commandname as well. But, if LaTeX
+% ever changes the way it names the internal names, the trick would break.
+% Instead let us just define a new environment so that the internal
+% name can be left undisturbed.
+\newenvironment{@IEEEbogusbiography}[2][]{\if@IEEEWARNIEEEbiography\typeout{** WARNING: \noexpand\IEEEbiography
+ is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEbiographyfalse%
+\setbox\@IEEEtranrubishbin\vbox\bgroup}{\egroup\relax}
+% and make biography point to our bogus biography
+\let\IEEEbiography=\@IEEEbogusbiography
+\let\endIEEEbiography=\end@IEEEbogusbiography
+
+\renewenvironment{IEEEbiographynophoto}[1]{\if@IEEEWARNIEEEbiographynophoto\typeout{** WARNING: \noexpand\IEEEbiographynophoto
+ is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEbiographynophotofalse%
+\setbox\@IEEEtranrubishbin\vbox\bgroup}{\egroup\relax}
+
+\def\IEEEpubid#1{\if@IEEEWARNIEEEpubid\typeout{** WARNING: \noexpand\IEEEpubid
+ is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEpubidfalse}
+\def\IEEEpubidadjcol{\if@IEEEWARNIEEEpubidadjcol\typeout{** WARNING: \noexpand\IEEEpubidadjcol
+ is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEpubidadjcolfalse}
+\def\IEEEmembership#1{\if@IEEEWARNIEEEmembership\typeout{** WARNING: \noexpand\IEEEmembership
+ is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEmembershipfalse}
+\def\IEEEaftertitletext#1{\if@IEEEWARNIEEEaftertitletext\typeout{** WARNING: \noexpand\IEEEaftertitletext
+ is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEaftertitletextfalse}
+\fi
+
+
+% provide a way to restore the commands that are locked out
+\def\IEEEoverridecommandlockouts{%
+\typeout{** ATTENTION: Overriding command lockouts (line \the\inputlineno).}%
+\let\thanks\@IEEESAVECMDthanks%
+\let\IEEEPARstart\@IEEESAVECMDIEEEPARstart%
+\let\IEEEbiography\@IEEESAVECMDIEEEbiography%
+\let\endIEEEbiography\@IEEESAVECMDendIEEEbiography%
+\let\IEEEbiographynophoto\@IEEESAVECMDIEEEbiographynophoto%
+\let\endIEEEbiographynophoto\@IEEESAVECMDendIEEEbiographynophoto%
+\let\IEEEpubid\@IEEESAVECMDIEEEpubid%
+\let\IEEEpubidadjcol\@IEEESAVECMDIEEEpubidadjcol%
+\let\IEEEmembership\@IEEESAVECMDIEEEmembership%
+\let\IEEEaftertitletext\@IEEESAVECMDIEEEaftertitletext}
+
+
+
+% need a backslash character for typeout output
+{\catcode`\|=0 \catcode`\\=12
+|xdef|@IEEEbackslash{\}}
+
+
+% hook to allow easy disabling of all legacy warnings
+\def\@IEEElegacywarn#1#2{\typeout{** ATTENTION: \@IEEEbackslash #1 is deprecated (line \the\inputlineno).
+Use \@IEEEbackslash #2 instead.}}
+
+
+% provide for legacy commands
+\def\authorblockA{\@IEEElegacywarn{authorblockA}{IEEEauthorblockA}\IEEEauthorblockA}
+\def\authorblockN{\@IEEElegacywarn{authorblockN}{IEEEauthorblockN}\IEEEauthorblockN}
+\def\authorrefmark{\@IEEElegacywarn{authorrefmark}{IEEEauthorrefmark}\IEEEauthorrefmark}
+\def\PARstart{\@IEEElegacywarn{PARstart}{IEEEPARstart}\IEEEPARstart}
+\def\pubid{\@IEEElegacywarn{pubid}{IEEEpubid}\IEEEpubid}
+\def\pubidadjcol{\@IEEElegacywarn{pubidadjcol}{IEEEpubidadjcol}\IEEEpubidadjcol}
+\def\QED{\@IEEElegacywarn{QED}{IEEEQED}\IEEEQED}
+\def\QEDclosed{\@IEEElegacywarn{QEDclosed}{IEEEQEDclosed}\IEEEQEDclosed}
+\def\QEDopen{\@IEEElegacywarn{QEDopen}{IEEEQEDopen}\IEEEQEDopen}
+\def\specialpapernotice{\@IEEElegacywarn{specialpapernotice}{IEEEspecialpapernotice}\IEEEspecialpapernotice}
+
+
+
+% provide for legacy environments
+\def\biography{\@IEEElegacywarn{biography}{IEEEbiography}\IEEEbiography}
+\def\biographynophoto{\@IEEElegacywarn{biographynophoto}{IEEEbiographynophoto}\IEEEbiographynophoto}
+\def\keywords{\@IEEElegacywarn{keywords}{IEEEkeywords}\IEEEkeywords}
+\def\endbiography{\endIEEEbiography}
+\def\endbiographynophoto{\endIEEEbiographynophoto}
+\def\endkeywords{\endIEEEkeywords}
+
+
+% provide for legacy IED commands/lengths when possible
+\let\labelindent\IEEElabelindent
+\def\calcleftmargin{\@IEEElegacywarn{calcleftmargin}{IEEEcalcleftmargin}\IEEEcalcleftmargin}
+\def\setlabelwidth{\@IEEElegacywarn{setlabelwidth}{IEEEsetlabelwidth}\IEEEsetlabelwidth}
+\def\usemathlabelsep{\@IEEElegacywarn{usemathlabelsep}{IEEEusemathlabelsep}\IEEEusemathlabelsep}
+\def\iedlabeljustifyc{\@IEEElegacywarn{iedlabeljustifyc}{IEEEiedlabeljustifyc}\IEEEiedlabeljustifyc}
+\def\iedlabeljustifyl{\@IEEElegacywarn{iedlabeljustifyl}{IEEEiedlabeljustifyl}\IEEEiedlabeljustifyl}
+\def\iedlabeljustifyr{\@IEEElegacywarn{iedlabeljustifyr}{IEEEiedlabeljustifyr}\IEEEiedlabeljustifyr}
+
+
+
+% let \proof use the IEEEtran version even after amsthm is loaded
+% \proof is now deprecated in favor of \IEEEproof
+\AtBeginDocument{\def\proof{\@IEEElegacywarn{proof}{IEEEproof}\IEEEproof}\def\endproof{\endIEEEproof}}
+
+% V1.7 \overrideIEEEmargins is no longer supported.
+\def\overrideIEEEmargins{%
+\typeout{** WARNING: \string\overrideIEEEmargins \space no longer supported (line \the\inputlineno).}%
+\typeout{** Use the \string\CLASSINPUTinnersidemargin, \string\CLASSINPUToutersidemargin \space controls instead.}}
+
+
+\endinput
+
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%% End of IEEEtran.cls %%%%%%%%%%%%%%%%%%%%%%%%%%%%
+% That's all folks!
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Paper/document/llncs.cls Fri Jan 18 11:40:01 2013 +0000
@@ -0,0 +1,1189 @@
+% LLNCS DOCUMENT CLASS -- version 2.13 (28-Jan-2002)
+% Springer Verlag LaTeX2e support for Lecture Notes in Computer Science
+%
+%%
+%% \CharacterTable
+%% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
+%% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
+%% Digits \0\1\2\3\4\5\6\7\8\9
+%% Exclamation \! Double quote \" Hash (number) \#
+%% Dollar \$ Percent \% Ampersand \&
+%% Acute accent \' Left paren \( Right paren \)
+%% Asterisk \* Plus \+ Comma \,
+%% Minus \- Point \. Solidus \/
+%% Colon \: Semicolon \; Less than \<
+%% Equals \= Greater than \> Question mark \?
+%% Commercial at \@ Left bracket \[ Backslash \\
+%% Right bracket \] Circumflex \^ Underscore \_
+%% Grave accent \` Left brace \{ Vertical bar \|
+%% Right brace \} Tilde \~}
+%%
+\NeedsTeXFormat{LaTeX2e}[1995/12/01]
+\ProvidesClass{llncs}[2002/01/28 v2.13
+^^J LaTeX document class for Lecture Notes in Computer Science]
+% Options
+\let\if@envcntreset\iffalse
+\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue}
+\DeclareOption{citeauthoryear}{\let\citeauthoryear=Y}
+\DeclareOption{oribibl}{\let\oribibl=Y}
+\let\if@custvec\iftrue
+\DeclareOption{orivec}{\let\if@custvec\iffalse}
+\let\if@envcntsame\iffalse
+\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue}
+\let\if@envcntsect\iffalse
+\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue}
+\let\if@runhead\iffalse
+\DeclareOption{runningheads}{\let\if@runhead\iftrue}
+
+\let\if@openbib\iffalse
+\DeclareOption{openbib}{\let\if@openbib\iftrue}
+
+% languages
+\let\switcht@@therlang\relax
+\def\ds@deutsch{\def\switcht@@therlang{\switcht@deutsch}}
+\def\ds@francais{\def\switcht@@therlang{\switcht@francais}}
+
+\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}}
+
+\ProcessOptions
+
+\LoadClass[twoside]{article}
+\RequirePackage{multicol} % needed for the list of participants, index
+
+\setlength{\textwidth}{12.2cm}
+\setlength{\textheight}{19.3cm}
+\renewcommand\@pnumwidth{2em}
+\renewcommand\@tocrmarg{3.5em}
+%
+\def\@dottedtocline#1#2#3#4#5{%
+ \ifnum #1>\c@tocdepth \else
+ \vskip \z@ \@plus.2\p@
+ {\leftskip #2\relax \rightskip \@tocrmarg \advance\rightskip by 0pt plus 2cm
+ \parfillskip -\rightskip \pretolerance=10000
+ \parindent #2\relax\@afterindenttrue
+ \interlinepenalty\@M
+ \leavevmode
+ \@tempdima #3\relax
+ \advance\leftskip \@tempdima \null\nobreak\hskip -\leftskip
+ {#4}\nobreak
+ \leaders\hbox{$\m@th
+ \mkern \@dotsep mu\hbox{.}\mkern \@dotsep
+ mu$}\hfill
+ \nobreak
+ \hb@xt@\@pnumwidth{\hfil\normalfont \normalcolor #5}%
+ \par}%
+ \fi}
+%
+\def\switcht@albion{%
+\def\abstractname{Abstract.}
+\def\ackname{Acknowledgement.}
+\def\andname{and}
+\def\lastandname{\unskip, and}
+\def\appendixname{Appendix}
+\def\chaptername{Chapter}
+\def\claimname{Claim}
+\def\conjecturename{Conjecture}
+\def\contentsname{Table of Contents}
+\def\corollaryname{Corollary}
+\def\definitionname{Definition}
+\def\examplename{Example}
+\def\exercisename{Exercise}
+\def\figurename{Fig.}
+\def\keywordname{{\bf Key words:}}
+\def\indexname{Index}
+\def\lemmaname{Lemma}
+\def\contriblistname{List of Contributors}
+\def\listfigurename{List of Figures}
+\def\listtablename{List of Tables}
+\def\mailname{{\it Correspondence to\/}:}
+\def\noteaddname{Note added in proof}
+\def\notename{Note}
+\def\partname{Part}
+\def\problemname{Problem}
+\def\proofname{Proof}
+\def\propertyname{Property}
+\def\propositionname{Proposition}
+\def\questionname{Question}
+\def\remarkname{Remark}
+\def\seename{see}
+\def\solutionname{Solution}
+\def\subclassname{{\it Subject Classifications\/}:}
+\def\tablename{Table}
+\def\theoremname{Theorem}}
+\switcht@albion
+% Names of theorem like environments are already defined
+% but must be translated if another language is chosen
+%
+% French section
+\def\switcht@francais{%\typeout{On parle francais.}%
+ \def\abstractname{R\'esum\'e.}%
+ \def\ackname{Remerciements.}%
+ \def\andname{et}%
+ \def\lastandname{ et}%
+ \def\appendixname{Appendice}
+ \def\chaptername{Chapitre}%
+ \def\claimname{Pr\'etention}%
+ \def\conjecturename{Hypoth\`ese}%
+ \def\contentsname{Table des mati\`eres}%
+ \def\corollaryname{Corollaire}%
+ \def\definitionname{D\'efinition}%
+ \def\examplename{Exemple}%
+ \def\exercisename{Exercice}%
+ \def\figurename{Fig.}%
+ \def\keywordname{{\bf Mots-cl\'e:}}
+ \def\indexname{Index}
+ \def\lemmaname{Lemme}%
+ \def\contriblistname{Liste des contributeurs}
+ \def\listfigurename{Liste des figures}%
+ \def\listtablename{Liste des tables}%
+ \def\mailname{{\it Correspondence to\/}:}
+ \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}%
+ \def\notename{Remarque}%
+ \def\partname{Partie}%
+ \def\problemname{Probl\`eme}%
+ \def\proofname{Preuve}%
+ \def\propertyname{Caract\'eristique}%
+%\def\propositionname{Proposition}%
+ \def\questionname{Question}%
+ \def\remarkname{Remarque}%
+ \def\seename{voir}
+ \def\solutionname{Solution}%
+ \def\subclassname{{\it Subject Classifications\/}:}
+ \def\tablename{Tableau}%
+ \def\theoremname{Th\'eor\`eme}%
+}
+%
+% German section
+\def\switcht@deutsch{%\typeout{Man spricht deutsch.}%
+ \def\abstractname{Zusammenfassung.}%
+ \def\ackname{Danksagung.}%
+ \def\andname{und}%
+ \def\lastandname{ und}%
+ \def\appendixname{Anhang}%
+ \def\chaptername{Kapitel}%
+ \def\claimname{Behauptung}%
+ \def\conjecturename{Hypothese}%
+ \def\contentsname{Inhaltsverzeichnis}%
+ \def\corollaryname{Korollar}%
+%\def\definitionname{Definition}%
+ \def\examplename{Beispiel}%
+ \def\exercisename{\"Ubung}%
+ \def\figurename{Abb.}%
+ \def\keywordname{{\bf Schl\"usselw\"orter:}}
+ \def\indexname{Index}
+%\def\lemmaname{Lemma}%
+ \def\contriblistname{Mitarbeiter}
+ \def\listfigurename{Abbildungsverzeichnis}%
+ \def\listtablename{Tabellenverzeichnis}%
+ \def\mailname{{\it Correspondence to\/}:}
+ \def\noteaddname{Nachtrag}%
+ \def\notename{Anmerkung}%
+ \def\partname{Teil}%
+%\def\problemname{Problem}%
+ \def\proofname{Beweis}%
+ \def\propertyname{Eigenschaft}%
+%\def\propositionname{Proposition}%
+ \def\questionname{Frage}%
+ \def\remarkname{Anmerkung}%
+ \def\seename{siehe}
+ \def\solutionname{L\"osung}%
+ \def\subclassname{{\it Subject Classifications\/}:}
+ \def\tablename{Tabelle}%
+%\def\theoremname{Theorem}%
+}
+
+% Ragged bottom for the actual page
+\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil
+\global\let\@textbottom\relax}}
+
+\renewcommand\small{%
+ \@setfontsize\small\@ixpt{11}%
+ \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
+ \abovedisplayshortskip \z@ \@plus2\p@
+ \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
+ \def\@listi{\leftmargin\leftmargini
+ \parsep 0\p@ \@plus1\p@ \@minus\p@
+ \topsep 8\p@ \@plus2\p@ \@minus4\p@
+ \itemsep0\p@}%
+ \belowdisplayskip \abovedisplayskip
+}
+
+\frenchspacing
+\widowpenalty=10000
+\clubpenalty=10000
+
+\setlength\oddsidemargin {63\p@}
+\setlength\evensidemargin {63\p@}
+\setlength\marginparwidth {90\p@}
+
+\setlength\headsep {16\p@}
+
+\setlength\footnotesep{7.7\p@}
+\setlength\textfloatsep{8mm\@plus 2\p@ \@minus 4\p@}
+\setlength\intextsep {8mm\@plus 2\p@ \@minus 2\p@}
+
+\setcounter{secnumdepth}{2}
+
+\newcounter {chapter}
+\renewcommand\thechapter {\@arabic\c@chapter}
+
+\newif\if@mainmatter \@mainmattertrue
+\newcommand\frontmatter{\cleardoublepage
+ \@mainmatterfalse\pagenumbering{Roman}}
+\newcommand\mainmatter{\cleardoublepage
+ \@mainmattertrue\pagenumbering{arabic}}
+\newcommand\backmatter{\if@openright\cleardoublepage\else\clearpage\fi
+ \@mainmatterfalse}
+
+\renewcommand\part{\cleardoublepage
+ \thispagestyle{empty}%
+ \if@twocolumn
+ \onecolumn
+ \@tempswatrue
+ \else
+ \@tempswafalse
+ \fi
+ \null\vfil
+ \secdef\@part\@spart}
+
+\def\@part[#1]#2{%
+ \ifnum \c@secnumdepth >-2\relax
+ \refstepcounter{part}%
+ \addcontentsline{toc}{part}{\thepart\hspace{1em}#1}%
+ \else
+ \addcontentsline{toc}{part}{#1}%
+ \fi
+ \markboth{}{}%
+ {\centering
+ \interlinepenalty \@M
+ \normalfont
+ \ifnum \c@secnumdepth >-2\relax
+ \huge\bfseries \partname~\thepart
+ \par
+ \vskip 20\p@
+ \fi
+ \Huge \bfseries #2\par}%
+ \@endpart}
+\def\@spart#1{%
+ {\centering
+ \interlinepenalty \@M
+ \normalfont
+ \Huge \bfseries #1\par}%
+ \@endpart}
+\def\@endpart{\vfil\newpage
+ \if@twoside
+ \null
+ \thispagestyle{empty}%
+ \newpage
+ \fi
+ \if@tempswa
+ \twocolumn
+ \fi}
+
+\newcommand\chapter{\clearpage
+ \thispagestyle{empty}%
+ \global\@topnum\z@
+ \@afterindentfalse
+ \secdef\@chapter\@schapter}
+\def\@chapter[#1]#2{\ifnum \c@secnumdepth >\m@ne
+ \if@mainmatter
+ \refstepcounter{chapter}%
+ \typeout{\@chapapp\space\thechapter.}%
+ \addcontentsline{toc}{chapter}%
+ {\protect\numberline{\thechapter}#1}%
+ \else
+ \addcontentsline{toc}{chapter}{#1}%
+ \fi
+ \else
+ \addcontentsline{toc}{chapter}{#1}%
+ \fi
+ \chaptermark{#1}%
+ \addtocontents{lof}{\protect\addvspace{10\p@}}%
+ \addtocontents{lot}{\protect\addvspace{10\p@}}%
+ \if@twocolumn
+ \@topnewpage[\@makechapterhead{#2}]%
+ \else
+ \@makechapterhead{#2}%
+ \@afterheading
+ \fi}
+\def\@makechapterhead#1{%
+% \vspace*{50\p@}%
+ {\centering
+ \ifnum \c@secnumdepth >\m@ne
+ \if@mainmatter
+ \large\bfseries \@chapapp{} \thechapter
+ \par\nobreak
+ \vskip 20\p@
+ \fi
+ \fi
+ \interlinepenalty\@M
+ \Large \bfseries #1\par\nobreak
+ \vskip 40\p@
+ }}
+\def\@schapter#1{\if@twocolumn
+ \@topnewpage[\@makeschapterhead{#1}]%
+ \else
+ \@makeschapterhead{#1}%
+ \@afterheading
+ \fi}
+\def\@makeschapterhead#1{%
+% \vspace*{50\p@}%
+ {\centering
+ \normalfont
+ \interlinepenalty\@M
+ \Large \bfseries #1\par\nobreak
+ \vskip 40\p@
+ }}
+
+\renewcommand\section{\@startsection{section}{1}{\z@}%
+ {-18\p@ \@plus -4\p@ \@minus -4\p@}%
+ {12\p@ \@plus 4\p@ \@minus 4\p@}%
+ {\normalfont\large\bfseries\boldmath
+ \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
+\renewcommand\subsection{\@startsection{subsection}{2}{\z@}%
+ {-18\p@ \@plus -4\p@ \@minus -4\p@}%
+ {8\p@ \@plus 4\p@ \@minus 4\p@}%
+ {\normalfont\normalsize\bfseries\boldmath
+ \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
+\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}%
+ {-18\p@ \@plus -4\p@ \@minus -4\p@}%
+ {-0.5em \@plus -0.22em \@minus -0.1em}%
+ {\normalfont\normalsize\bfseries\boldmath}}
+\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}%
+ {-12\p@ \@plus -4\p@ \@minus -4\p@}%
+ {-0.5em \@plus -0.22em \@minus -0.1em}%
+ {\normalfont\normalsize\itshape}}
+\renewcommand\subparagraph[1]{\typeout{LLNCS warning: You should not use
+ \string\subparagraph\space with this class}\vskip0.5cm
+You should not use \verb|\subparagraph| with this class.\vskip0.5cm}
+
+\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00}
+\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01}
+\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02}
+\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03}
+\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04}
+\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05}
+\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06}
+\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07}
+\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08}
+\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09}
+\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A}
+
+\let\footnotesize\small
+
+\if@custvec
+\def\vec#1{\mathchoice{\mbox{\boldmath$\displaystyle#1$}}
+{\mbox{\boldmath$\textstyle#1$}}
+{\mbox{\boldmath$\scriptstyle#1$}}
+{\mbox{\boldmath$\scriptscriptstyle#1$}}}
+\fi
+
+\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}}
+\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil
+\penalty50\hskip1em\null\nobreak\hfil\squareforqed
+\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi}
+
+\def\getsto{\mathrel{\mathchoice {\vcenter{\offinterlineskip
+\halign{\hfil
+$\displaystyle##$\hfil\cr\gets\cr\to\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr\gets
+\cr\to\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr\gets
+\cr\to\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
+\gets\cr\to\cr}}}}}
+\def\lid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
+$\displaystyle##$\hfil\cr<\cr\noalign{\vskip1.2pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr<\cr
+\noalign{\vskip1.2pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr<\cr
+\noalign{\vskip1pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
+<\cr
+\noalign{\vskip0.9pt}=\cr}}}}}
+\def\gid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
+$\displaystyle##$\hfil\cr>\cr\noalign{\vskip1.2pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr>\cr
+\noalign{\vskip1.2pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr>\cr
+\noalign{\vskip1pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
+>\cr
+\noalign{\vskip0.9pt}=\cr}}}}}
+\def\grole{\mathrel{\mathchoice {\vcenter{\offinterlineskip
+\halign{\hfil
+$\displaystyle##$\hfil\cr>\cr\noalign{\vskip-1pt}<\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr
+>\cr\noalign{\vskip-1pt}<\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr
+>\cr\noalign{\vskip-0.8pt}<\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
+>\cr\noalign{\vskip-0.3pt}<\cr}}}}}
+\def\bbbr{{\rm I\!R}} %reelle Zahlen
+\def\bbbm{{\rm I\!M}}
+\def\bbbn{{\rm I\!N}} %natuerliche Zahlen
+\def\bbbf{{\rm I\!F}}
+\def\bbbh{{\rm I\!H}}
+\def\bbbk{{\rm I\!K}}
+\def\bbbp{{\rm I\!P}}
+\def\bbbone{{\mathchoice {\rm 1\mskip-4mu l} {\rm 1\mskip-4mu l}
+{\rm 1\mskip-4.5mu l} {\rm 1\mskip-5mu l}}}
+\def\bbbc{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm C$}\hbox{\hbox
+to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\textstyle\rm C$}\hbox{\hbox
+to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptstyle\rm C$}\hbox{\hbox
+to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptscriptstyle\rm C$}\hbox{\hbox
+to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}}}
+\def\bbbq{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
+Q$}\hbox{\raise
+0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
+{\setbox0=\hbox{$\textstyle\rm Q$}\hbox{\raise
+0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptstyle\rm Q$}\hbox{\raise
+0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptscriptstyle\rm Q$}\hbox{\raise
+0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}}}
+\def\bbbt{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
+T$}\hbox{\hbox to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\textstyle\rm T$}\hbox{\hbox
+to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptstyle\rm T$}\hbox{\hbox
+to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptscriptstyle\rm T$}\hbox{\hbox
+to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}}}
+\def\bbbs{{\mathchoice
+{\setbox0=\hbox{$\displaystyle \rm S$}\hbox{\raise0.5\ht0\hbox
+to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
+to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
+{\setbox0=\hbox{$\textstyle \rm S$}\hbox{\raise0.5\ht0\hbox
+to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
+to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptstyle \rm S$}\hbox{\raise0.5\ht0\hbox
+to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
+to0pt{\kern0.5\wd0\vrule height0.45\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptscriptstyle\rm S$}\hbox{\raise0.5\ht0\hbox
+to0pt{\kern0.4\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
+to0pt{\kern0.55\wd0\vrule height0.45\ht0\hss}\box0}}}}
+\def\bbbz{{\mathchoice {\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
+{\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
+{\hbox{$\mathsf\scriptstyle Z\kern-0.3em Z$}}
+{\hbox{$\mathsf\scriptscriptstyle Z\kern-0.2em Z$}}}}
+
+\let\ts\,
+
+\setlength\leftmargini {17\p@}
+\setlength\leftmargin {\leftmargini}
+\setlength\leftmarginii {\leftmargini}
+\setlength\leftmarginiii {\leftmargini}
+\setlength\leftmarginiv {\leftmargini}
+\setlength \labelsep {.5em}
+\setlength \labelwidth{\leftmargini}
+\addtolength\labelwidth{-\labelsep}
+
+\def\@listI{\leftmargin\leftmargini
+ \parsep 0\p@ \@plus1\p@ \@minus\p@
+ \topsep 8\p@ \@plus2\p@ \@minus4\p@
+ \itemsep0\p@}
+\let\@listi\@listI
+\@listi
+\def\@listii {\leftmargin\leftmarginii
+ \labelwidth\leftmarginii
+ \advance\labelwidth-\labelsep
+ \topsep 0\p@ \@plus2\p@ \@minus\p@}
+\def\@listiii{\leftmargin\leftmarginiii
+ \labelwidth\leftmarginiii
+ \advance\labelwidth-\labelsep
+ \topsep 0\p@ \@plus\p@\@minus\p@
+ \parsep \z@
+ \partopsep \p@ \@plus\z@ \@minus\p@}
+
+\renewcommand\labelitemi{\normalfont\bfseries --}
+\renewcommand\labelitemii{$\m@th\bullet$}
+
+\setlength\arraycolsep{1.4\p@}
+\setlength\tabcolsep{1.4\p@}
+
+\def\tableofcontents{\chapter*{\contentsname\@mkboth{{\contentsname}}%
+ {{\contentsname}}}
+ \def\authcount##1{\setcounter{auco}{##1}\setcounter{@auth}{1}}
+ \def\lastand{\ifnum\value{auco}=2\relax
+ \unskip{} \andname\
+ \else
+ \unskip \lastandname\
+ \fi}%
+ \def\and{\stepcounter{@auth}\relax
+ \ifnum\value{@auth}=\value{auco}%
+ \lastand
+ \else
+ \unskip,
+ \fi}%
+ \@starttoc{toc}\if@restonecol\twocolumn\fi}
+
+\def\l@part#1#2{\addpenalty{\@secpenalty}%
+ \addvspace{2em plus\p@}% % space above part line
+ \begingroup
+ \parindent \z@
+ \rightskip \z@ plus 5em
+ \hrule\vskip5pt
+ \large % same size as for a contribution heading
+ \bfseries\boldmath % set line in boldface
+ \leavevmode % TeX command to enter horizontal mode.
+ #1\par
+ \vskip5pt
+ \hrule
+ \vskip1pt
+ \nobreak % Never break after part entry
+ \endgroup}
+
+\def\@dotsep{2}
+
+\def\hyperhrefextend{\ifx\hyper@anchor\@undefined\else
+{chapter.\thechapter}\fi}
+
+\def\addnumcontentsmark#1#2#3{%
+\addtocontents{#1}{\protect\contentsline{#2}{\protect\numberline
+ {\thechapter}#3}{\thepage}\hyperhrefextend}}
+\def\addcontentsmark#1#2#3{%
+\addtocontents{#1}{\protect\contentsline{#2}{#3}{\thepage}\hyperhrefextend}}
+\def\addcontentsmarkwop#1#2#3{%
+\addtocontents{#1}{\protect\contentsline{#2}{#3}{0}\hyperhrefextend}}
+
+\def\@adcmk[#1]{\ifcase #1 \or
+\def\@gtempa{\addnumcontentsmark}%
+ \or \def\@gtempa{\addcontentsmark}%
+ \or \def\@gtempa{\addcontentsmarkwop}%
+ \fi\@gtempa{toc}{chapter}}
+\def\addtocmark{\@ifnextchar[{\@adcmk}{\@adcmk[3]}}
+
+\def\l@chapter#1#2{\addpenalty{-\@highpenalty}
+ \vskip 1.0em plus 1pt \@tempdima 1.5em \begingroup
+ \parindent \z@ \rightskip \@tocrmarg
+ \advance\rightskip by 0pt plus 2cm
+ \parfillskip -\rightskip \pretolerance=10000
+ \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
+ {\large\bfseries\boldmath#1}\ifx0#2\hfil\null
+ \else
+ \nobreak
+ \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
+ \@dotsep mu$}\hfill
+ \nobreak\hbox to\@pnumwidth{\hss #2}%
+ \fi\par
+ \penalty\@highpenalty \endgroup}
+
+\def\l@title#1#2{\addpenalty{-\@highpenalty}
+ \addvspace{8pt plus 1pt}
+ \@tempdima \z@
+ \begingroup
+ \parindent \z@ \rightskip \@tocrmarg
+ \advance\rightskip by 0pt plus 2cm
+ \parfillskip -\rightskip \pretolerance=10000
+ \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
+ #1\nobreak
+ \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
+ \@dotsep mu$}\hfill
+ \nobreak\hbox to\@pnumwidth{\hss #2}\par
+ \penalty\@highpenalty \endgroup}
+
+\def\l@author#1#2{\addpenalty{\@highpenalty}
+ \@tempdima=\z@ %15\p@
+ \begingroup
+ \parindent \z@ \rightskip \@tocrmarg
+ \advance\rightskip by 0pt plus 2cm
+ \pretolerance=10000
+ \leavevmode \advance\leftskip\@tempdima %\hskip -\leftskip
+ \textit{#1}\par
+ \penalty\@highpenalty \endgroup}
+
+%\setcounter{tocdepth}{0}
+\newdimen\tocchpnum
+\newdimen\tocsecnum
+\newdimen\tocsectotal
+\newdimen\tocsubsecnum
+\newdimen\tocsubsectotal
+\newdimen\tocsubsubsecnum
+\newdimen\tocsubsubsectotal
+\newdimen\tocparanum
+\newdimen\tocparatotal
+\newdimen\tocsubparanum
+\tocchpnum=\z@ % no chapter numbers
+\tocsecnum=15\p@ % section 88. plus 2.222pt
+\tocsubsecnum=23\p@ % subsection 88.8 plus 2.222pt
+\tocsubsubsecnum=27\p@ % subsubsection 88.8.8 plus 1.444pt
+\tocparanum=35\p@ % paragraph 88.8.8.8 plus 1.666pt
+\tocsubparanum=43\p@ % subparagraph 88.8.8.8.8 plus 1.888pt
+\def\calctocindent{%
+\tocsectotal=\tocchpnum
+\advance\tocsectotal by\tocsecnum
+\tocsubsectotal=\tocsectotal
+\advance\tocsubsectotal by\tocsubsecnum
+\tocsubsubsectotal=\tocsubsectotal
+\advance\tocsubsubsectotal by\tocsubsubsecnum
+\tocparatotal=\tocsubsubsectotal
+\advance\tocparatotal by\tocparanum}
+\calctocindent
+
+\def\l@section{\@dottedtocline{1}{\tocchpnum}{\tocsecnum}}
+\def\l@subsection{\@dottedtocline{2}{\tocsectotal}{\tocsubsecnum}}
+\def\l@subsubsection{\@dottedtocline{3}{\tocsubsectotal}{\tocsubsubsecnum}}
+\def\l@paragraph{\@dottedtocline{4}{\tocsubsubsectotal}{\tocparanum}}
+\def\l@subparagraph{\@dottedtocline{5}{\tocparatotal}{\tocsubparanum}}
+
+\def\listoffigures{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
+ \fi\section*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}}
+ \@starttoc{lof}\if@restonecol\twocolumn\fi}
+\def\l@figure{\@dottedtocline{1}{0em}{1.5em}}
+
+\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
+ \fi\section*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}}
+ \@starttoc{lot}\if@restonecol\twocolumn\fi}
+\let\l@table\l@figure
+
+\renewcommand\listoffigures{%
+ \section*{\listfigurename
+ \@mkboth{\listfigurename}{\listfigurename}}%
+ \@starttoc{lof}%
+ }
+
+\renewcommand\listoftables{%
+ \section*{\listtablename
+ \@mkboth{\listtablename}{\listtablename}}%
+ \@starttoc{lot}%
+ }
+
+\ifx\oribibl\undefined
+\ifx\citeauthoryear\undefined
+\renewenvironment{thebibliography}[1]
+ {\section*{\refname}
+ \def\@biblabel##1{##1.}
+ \small
+ \list{\@biblabel{\@arabic\c@enumiv}}%
+ {\settowidth\labelwidth{\@biblabel{#1}}%
+ \leftmargin\labelwidth
+ \advance\leftmargin\labelsep
+ \if@openbib
+ \advance\leftmargin\bibindent
+ \itemindent -\bibindent
+ \listparindent \itemindent
+ \parsep \z@
+ \fi
+ \usecounter{enumiv}%
+ \let\p@enumiv\@empty
+ \renewcommand\theenumiv{\@arabic\c@enumiv}}%
+ \if@openbib
+ \renewcommand\newblock{\par}%
+ \else
+ \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
+ \fi
+ \sloppy\clubpenalty4000\widowpenalty4000%
+ \sfcode`\.=\@m}
+ {\def\@noitemerr
+ {\@latex@warning{Empty `thebibliography' environment}}%
+ \endlist}
+\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw
+ {\let\protect\noexpand\immediate
+ \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
+\newcount\@tempcntc
+\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi
+ \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do
+ {\@ifundefined
+ {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries
+ ?}\@warning
+ {Citation `\@citeb' on page \thepage \space undefined}}%
+ {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}%
+ \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne
+ \@citea\def\@citea{,}\hbox{\csname b@\@citeb\endcsname}%
+ \else
+ \advance\@tempcntb\@ne
+ \ifnum\@tempcntb=\@tempcntc
+ \else\advance\@tempcntb\m@ne\@citeo
+ \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}}
+\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else
+ \@citea\def\@citea{,\,\hskip\z@skip}%
+ \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else
+ {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else
+ \def\@citea{--}\fi
+ \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi}
+\else
+\renewenvironment{thebibliography}[1]
+ {\section*{\refname}
+ \small
+ \list{}%
+ {\settowidth\labelwidth{}%
+ \leftmargin\parindent
+ \itemindent=-\parindent
+ \labelsep=\z@
+ \if@openbib
+ \advance\leftmargin\bibindent
+ \itemindent -\bibindent
+ \listparindent \itemindent
+ \parsep \z@
+ \fi
+ \usecounter{enumiv}%
+ \let\p@enumiv\@empty
+ \renewcommand\theenumiv{}}%
+ \if@openbib
+ \renewcommand\newblock{\par}%
+ \else
+ \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
+ \fi
+ \sloppy\clubpenalty4000\widowpenalty4000%
+ \sfcode`\.=\@m}
+ {\def\@noitemerr
+ {\@latex@warning{Empty `thebibliography' environment}}%
+ \endlist}
+ \def\@cite#1{#1}%
+ \def\@lbibitem[#1]#2{\item[]\if@filesw
+ {\def\protect##1{\string ##1\space}\immediate
+ \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
+ \fi
+\else
+\@cons\@openbib@code{\noexpand\small}
+\fi
+
+\def\idxquad{\hskip 10\p@}% space that divides entry from number
+
+\def\@idxitem{\par\hangindent 10\p@}
+
+\def\subitem{\par\setbox0=\hbox{--\enspace}% second order
+ \noindent\hangindent\wd0\box0}% index entry
+
+\def\subsubitem{\par\setbox0=\hbox{--\,--\enspace}% third
+ \noindent\hangindent\wd0\box0}% order index entry
+
+\def\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax}
+
+\renewenvironment{theindex}
+ {\@mkboth{\indexname}{\indexname}%
+ \thispagestyle{empty}\parindent\z@
+ \parskip\z@ \@plus .3\p@\relax
+ \let\item\par
+ \def\,{\relax\ifmmode\mskip\thinmuskip
+ \else\hskip0.2em\ignorespaces\fi}%
+ \normalfont\small
+ \begin{multicols}{2}[\@makeschapterhead{\indexname}]%
+ }
+ {\end{multicols}}
+
+\renewcommand\footnoterule{%
+ \kern-3\p@
+ \hrule\@width 2truecm
+ \kern2.6\p@}
+ \newdimen\fnindent
+ \fnindent1em
+\long\def\@makefntext#1{%
+ \parindent \fnindent%
+ \leftskip \fnindent%
+ \noindent
+ \llap{\hb@xt@1em{\hss\@makefnmark\ }}\ignorespaces#1}
+
+\long\def\@makecaption#1#2{%
+ \vskip\abovecaptionskip
+ \sbox\@tempboxa{{\bfseries #1.} #2}%
+ \ifdim \wd\@tempboxa >\hsize
+ {\bfseries #1.} #2\par
+ \else
+ \global \@minipagefalse
+ \hb@xt@\hsize{\hfil\box\@tempboxa\hfil}%
+ \fi
+ \vskip\belowcaptionskip}
+
+\def\fps@figure{htbp}
+\def\fnum@figure{\figurename\thinspace\thefigure}
+\def \@floatboxreset {%
+ \reset@font
+ \small
+ \@setnobreak
+ \@setminipage
+}
+\def\fps@table{htbp}
+\def\fnum@table{\tablename~\thetable}
+\renewenvironment{table}
+ {\setlength\abovecaptionskip{0\p@}%
+ \setlength\belowcaptionskip{10\p@}%
+ \@float{table}}
+ {\end@float}
+\renewenvironment{table*}
+ {\setlength\abovecaptionskip{0\p@}%
+ \setlength\belowcaptionskip{10\p@}%
+ \@dblfloat{table}}
+ {\end@dblfloat}
+
+\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname
+ ext@#1\endcsname}{#1}{\protect\numberline{\csname
+ the#1\endcsname}{\ignorespaces #2}}\begingroup
+ \@parboxrestore
+ \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
+ \endgroup}
+
+% LaTeX does not provide a command to enter the authors institute
+% addresses. The \institute command is defined here.
+
+\newcounter{@inst}
+\newcounter{@auth}
+\newcounter{auco}
+\newdimen\instindent
+\newbox\authrun
+\newtoks\authorrunning
+\newtoks\tocauthor
+\newbox\titrun
+\newtoks\titlerunning
+\newtoks\toctitle
+
+\def\clearheadinfo{\gdef\@author{No Author Given}%
+ \gdef\@title{No Title Given}%
+ \gdef\@subtitle{}%
+ \gdef\@institute{No Institute Given}%
+ \gdef\@thanks{}%
+ \global\titlerunning={}\global\authorrunning={}%
+ \global\toctitle={}\global\tocauthor={}}
+
+\def\institute#1{\gdef\@institute{#1}}
+
+\def\institutename{\par
+ \begingroup
+ \parskip=\z@
+ \parindent=\z@
+ \setcounter{@inst}{1}%
+ \def\and{\par\stepcounter{@inst}%
+ \noindent$^{\the@inst}$\enspace\ignorespaces}%
+ \setbox0=\vbox{\def\thanks##1{}\@institute}%
+ \ifnum\c@@inst=1\relax
+ \gdef\fnnstart{0}%
+ \else
+ \xdef\fnnstart{\c@@inst}%
+ \setcounter{@inst}{1}%
+ \noindent$^{\the@inst}$\enspace
+ \fi
+ \ignorespaces
+ \@institute\par
+ \endgroup}
+
+\def\@fnsymbol#1{\ensuremath{\ifcase#1\or\star\or{\star\star}\or
+ {\star\star\star}\or \dagger\or \ddagger\or
+ \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger
+ \or \ddagger\ddagger \else\@ctrerr\fi}}
+
+\def\inst#1{\unskip$^{#1}$}
+\def\fnmsep{\unskip$^,$}
+\def\email#1{{\tt#1}}
+\AtBeginDocument{\@ifundefined{url}{\def\url#1{#1}}{}%
+\@ifpackageloaded{babel}{%
+\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}%
+\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}%
+\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}%
+}{\switcht@@therlang}%
+}
+\def\homedir{\~{ }}
+
+\def\subtitle#1{\gdef\@subtitle{#1}}
+\clearheadinfo
+
+\renewcommand\maketitle{\newpage
+ \refstepcounter{chapter}%
+ \stepcounter{section}%
+ \setcounter{section}{0}%
+ \setcounter{subsection}{0}%
+ \setcounter{figure}{0}
+ \setcounter{table}{0}
+ \setcounter{equation}{0}
+ \setcounter{footnote}{0}%
+ \begingroup
+ \parindent=\z@
+ \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
+ \if@twocolumn
+ \ifnum \col@number=\@ne
+ \@maketitle
+ \else
+ \twocolumn[\@maketitle]%
+ \fi
+ \else
+ \newpage
+ \global\@topnum\z@ % Prevents figures from going at top of page.
+ \@maketitle
+ \fi
+ \thispagestyle{empty}\@thanks
+%
+ \def\\{\unskip\ \ignorespaces}\def\inst##1{\unskip{}}%
+ \def\thanks##1{\unskip{}}\def\fnmsep{\unskip}%
+ \instindent=\hsize
+ \advance\instindent by-\headlineindent
+% \if!\the\toctitle!\addcontentsline{toc}{title}{\@title}\else
+% \addcontentsline{toc}{title}{\the\toctitle}\fi
+ \if@runhead
+ \if!\the\titlerunning!\else
+ \edef\@title{\the\titlerunning}%
+ \fi
+ \global\setbox\titrun=\hbox{\small\rm\unboldmath\ignorespaces\@title}%
+ \ifdim\wd\titrun>\instindent
+ \typeout{Title too long for running head. Please supply}%
+ \typeout{a shorter form with \string\titlerunning\space prior to
+ \string\maketitle}%
+ \global\setbox\titrun=\hbox{\small\rm
+ Title Suppressed Due to Excessive Length}%
+ \fi
+ \xdef\@title{\copy\titrun}%
+ \fi
+%
+ \if!\the\tocauthor!\relax
+ {\def\and{\noexpand\protect\noexpand\and}%
+ \protected@xdef\toc@uthor{\@author}}%
+ \else
+ \def\\{\noexpand\protect\noexpand\newline}%
+ \protected@xdef\scratch{\the\tocauthor}%
+ \protected@xdef\toc@uthor{\scratch}%
+ \fi
+% \addcontentsline{toc}{author}{\toc@uthor}%
+ \if@runhead
+ \if!\the\authorrunning!
+ \value{@inst}=\value{@auth}%
+ \setcounter{@auth}{1}%
+ \else
+ \edef\@author{\the\authorrunning}%
+ \fi
+ \global\setbox\authrun=\hbox{\small\unboldmath\@author\unskip}%
+ \ifdim\wd\authrun>\instindent
+ \typeout{Names of authors too long for running head. Please supply}%
+ \typeout{a shorter form with \string\authorrunning\space prior to
+ \string\maketitle}%
+ \global\setbox\authrun=\hbox{\small\rm
+ Authors Suppressed Due to Excessive Length}%
+ \fi
+ \xdef\@author{\copy\authrun}%
+ \markboth{\@author}{\@title}%
+ \fi
+ \endgroup
+ \setcounter{footnote}{\fnnstart}%
+ \clearheadinfo}
+%
+\def\@maketitle{\newpage
+ \markboth{}{}%
+ \def\lastand{\ifnum\value{@inst}=2\relax
+ \unskip{} \andname\
+ \else
+ \unskip \lastandname\
+ \fi}%
+ \def\and{\stepcounter{@auth}\relax
+ \ifnum\value{@auth}=\value{@inst}%
+ \lastand
+ \else
+ \unskip,
+ \fi}%
+ \begin{center}%
+ \let\newline\\
+ {\Large \bfseries\boldmath
+ \pretolerance=10000
+ \@title \par}\vskip .8cm
+\if!\@subtitle!\else {\large \bfseries\boldmath
+ \vskip -.65cm
+ \pretolerance=10000
+ \@subtitle \par}\vskip .8cm\fi
+ \setbox0=\vbox{\setcounter{@auth}{1}\def\and{\stepcounter{@auth}}%
+ \def\thanks##1{}\@author}%
+ \global\value{@inst}=\value{@auth}%
+ \global\value{auco}=\value{@auth}%
+ \setcounter{@auth}{1}%
+{\lineskip .5em
+\noindent\ignorespaces
+\@author\vskip.35cm}
+ {\small\institutename}
+ \end{center}%
+ }
+
+% definition of the "\spnewtheorem" command.
+%
+% Usage:
+%
+% \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font}
+% or \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font}
+% or \spnewtheorem*{env_nam}{caption}{cap_font}{body_font}
+%
+% New is "cap_font" and "body_font". It stands for
+% fontdefinition of the caption and the text itself.
+%
+% "\spnewtheorem*" gives a theorem without number.
+%
+% A defined spnewthoerem environment is used as described
+% by Lamport.
+%
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+
+\def\@thmcountersep{}
+\def\@thmcounterend{.}
+
+\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}}
+
+% definition of \spnewtheorem with number
+
+\def\@spnthm#1#2{%
+ \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}}
+\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}}
+
+\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname
+ {\@definecounter{#1}\@addtoreset{#1}{#3}%
+ \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand
+ \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}%
+ \expandafter\xdef\csname #1name\endcsname{#2}%
+ \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
+ \global\@namedef{end#1}{\@endtheorem}}}
+
+\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
+ {\@definecounter{#1}%
+ \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
+ \expandafter\xdef\csname #1name\endcsname{#2}%
+ \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}%
+ \global\@namedef{end#1}{\@endtheorem}}}
+
+\def\@spothm#1[#2]#3#4#5{%
+ \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}%
+ {\expandafter\@ifdefinable\csname #1\endcsname
+ {\global\@namedef{the#1}{\@nameuse{the#2}}%
+ \expandafter\xdef\csname #1name\endcsname{#3}%
+ \global\@namedef{#1}{\@spthm{#2}{\csname #1name\endcsname}{#4}{#5}}%
+ \global\@namedef{end#1}{\@endtheorem}}}}
+
+\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@
+\refstepcounter{#1}%
+\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}}
+
+\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}%
+ \ignorespaces}
+
+\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname
+ the#1\endcsname}{#5}{#3}{#4}\ignorespaces}
+
+\def\@spbegintheorem#1#2#3#4{\trivlist
+ \item[\hskip\labelsep{#3#1\ #2\@thmcounterend}]#4}
+
+\def\@spopargbegintheorem#1#2#3#4#5{\trivlist
+ \item[\hskip\labelsep{#4#1\ #2}]{#4(#3)\@thmcounterend\ }#5}
+
+% definition of \spnewtheorem* without number
+
+\def\@sthm#1#2{\@Ynthm{#1}{#2}}
+
+\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
+ {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}%
+ \expandafter\xdef\csname #1name\endcsname{#2}%
+ \global\@namedef{end#1}{\@endtheorem}}}
+
+\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@
+\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}}
+
+\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces}
+
+\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1}
+ {#4}{#2}{#3}\ignorespaces}
+
+\def\@Begintheorem#1#2#3{#3\trivlist
+ \item[\hskip\labelsep{#2#1\@thmcounterend}]}
+
+\def\@Opargbegintheorem#1#2#3#4{#4\trivlist
+ \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }}
+
+\if@envcntsect
+ \def\@thmcountersep{.}
+ \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape}
+\else
+ \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape}
+ \if@envcntreset
+ \@addtoreset{theorem}{section}
+ \else
+ \@addtoreset{theorem}{chapter}
+ \fi
+\fi
+
+%definition of divers theorem environments
+\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily}
+\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily}
+\if@envcntsame % alle Umgebungen wie Theorem.
+ \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}}
+\else % alle Umgebungen mit eigenem Zaehler
+ \if@envcntsect % mit section numeriert
+ \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}}
+ \else % nicht mit section numeriert
+ \if@envcntreset
+ \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
+ \@addtoreset{#1}{section}}
+ \else
+ \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
+ \@addtoreset{#1}{chapter}}%
+ \fi
+ \fi
+\fi
+\spn@wtheorem{case}{Case}{\itshape}{\rmfamily}
+\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily}
+\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape}
+\spn@wtheorem{definition}{Definition}{\bfseries}{\itshape}
+\spn@wtheorem{example}{Example}{\itshape}{\rmfamily}
+\spn@wtheorem{exercise}{Exercise}{\itshape}{\rmfamily}
+\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape}
+\spn@wtheorem{note}{Note}{\itshape}{\rmfamily}
+\spn@wtheorem{problem}{Problem}{\itshape}{\rmfamily}
+\spn@wtheorem{property}{Property}{\itshape}{\rmfamily}
+\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape}
+\spn@wtheorem{question}{Question}{\itshape}{\rmfamily}
+\spn@wtheorem{solution}{Solution}{\itshape}{\rmfamily}
+\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily}
+
+\def\@takefromreset#1#2{%
+ \def\@tempa{#1}%
+ \let\@tempd\@elt
+ \def\@elt##1{%
+ \def\@tempb{##1}%
+ \ifx\@tempa\@tempb\else
+ \@addtoreset{##1}{#2}%
+ \fi}%
+ \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
+ \expandafter\def\csname cl@#2\endcsname{}%
+ \@tempc
+ \let\@elt\@tempd}
+
+\def\theopargself{\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
+ \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5}
+ \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
+ \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }}
+ }
+
+\renewenvironment{abstract}{%
+ \list{}{\advance\topsep by0.35cm\relax\small
+ \leftmargin=1cm
+ \labelwidth=\z@
+ \listparindent=\z@
+ \itemindent\listparindent
+ \rightmargin\leftmargin}\item[\hskip\labelsep
+ \bfseries\abstractname]}
+ {\endlist}
+
+\newdimen\headlineindent % dimension for space between
+\headlineindent=1.166cm % number and text of headings.
+
+\def\ps@headings{\let\@mkboth\@gobbletwo
+ \let\@oddfoot\@empty\let\@evenfoot\@empty
+ \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
+ \leftmark\hfil}
+ \def\@oddhead{\normalfont\small\hfil\rightmark\hspace{\headlineindent}%
+ \llap{\thepage}}
+ \def\chaptermark##1{}%
+ \def\sectionmark##1{}%
+ \def\subsectionmark##1{}}
+
+\def\ps@titlepage{\let\@mkboth\@gobbletwo
+ \let\@oddfoot\@empty\let\@evenfoot\@empty
+ \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
+ \hfil}
+ \def\@oddhead{\normalfont\small\hfil\hspace{\headlineindent}%
+ \llap{\thepage}}
+ \def\chaptermark##1{}%
+ \def\sectionmark##1{}%
+ \def\subsectionmark##1{}}
+
+\if@runhead\ps@headings\else
+\ps@empty\fi
+
+\setlength\arraycolsep{1.4\p@}
+\setlength\tabcolsep{1.4\p@}
+
+\endinput
+%end of file llncs.cls
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Paper/document/mathpartir.sty Fri Jan 18 11:40:01 2013 +0000
@@ -0,0 +1,446 @@
+% Mathpartir --- Math Paragraph for Typesetting Inference Rules
+%
+% Copyright (C) 2001, 2002, 2003, 2004, 2005 Didier Rémy
+%
+% Author : Didier Remy
+% Version : 1.2.0
+% Bug Reports : to author
+% Web Site : http://pauillac.inria.fr/~remy/latex/
+%
+% Mathpartir is free software; you can redistribute it and/or modify
+% it under the terms of the GNU General Public License as published by
+% the Free Software Foundation; either version 2, or (at your option)
+% any later version.
+%
+% Mathpartir is distributed in the hope that it will be useful,
+% but WITHOUT ANY WARRANTY; without even the implied warranty of
+% MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
+% GNU General Public License for more details
+% (http://pauillac.inria.fr/~remy/license/GPL).
+%
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+% File mathpartir.sty (LaTeX macros)
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+
+\NeedsTeXFormat{LaTeX2e}
+\ProvidesPackage{mathpartir}
+ [2005/12/20 version 1.2.0 Math Paragraph for Typesetting Inference Rules]
+
+%%
+
+%% Identification
+%% Preliminary declarations
+
+\RequirePackage {keyval}
+
+%% Options
+%% More declarations
+
+%% PART I: Typesetting maths in paragraphe mode
+
+%% \newdimen \mpr@tmpdim
+%% Dimens are a precious ressource. Uses seems to be local.
+\let \mpr@tmpdim \@tempdima
+
+% To ensure hevea \hva compatibility, \hva should expands to nothing
+% in mathpar or in inferrule
+\let \mpr@hva \empty
+
+%% normal paragraph parametters, should rather be taken dynamically
+\def \mpr@savepar {%
+ \edef \MathparNormalpar
+ {\noexpand \lineskiplimit \the\lineskiplimit
+ \noexpand \lineskip \the\lineskip}%
+ }
+
+\def \mpr@rulelineskip {\lineskiplimit=0.3em\lineskip=0.2em plus 0.1em}
+\def \mpr@lesslineskip {\lineskiplimit=0.6em\lineskip=0.5em plus 0.2em}
+\def \mpr@lineskip {\lineskiplimit=1.2em\lineskip=1.2em plus 0.2em}
+\let \MathparLineskip \mpr@lineskip
+\def \mpr@paroptions {\MathparLineskip}
+\let \mpr@prebindings \relax
+
+\newskip \mpr@andskip \mpr@andskip 2em plus 0.5fil minus 0.5em
+
+\def \mpr@goodbreakand
+ {\hskip -\mpr@andskip \penalty -1000\hskip \mpr@andskip}
+\def \mpr@and {\hskip \mpr@andskip}
+\def \mpr@andcr {\penalty 50\mpr@and}
+\def \mpr@cr {\penalty -10000\mpr@and}
+\def \mpr@eqno #1{\mpr@andcr #1\hskip 0em plus -1fil \penalty 10}
+
+\def \mpr@bindings {%
+ \let \and \mpr@andcr
+ \let \par \mpr@andcr
+ \let \\\mpr@cr
+ \let \eqno \mpr@eqno
+ \let \hva \mpr@hva
+ }
+\let \MathparBindings \mpr@bindings
+
+% \@ifundefined {ignorespacesafterend}
+% {\def \ignorespacesafterend {\aftergroup \ignorespaces}
+
+\newenvironment{mathpar}[1][]
+ {$$\mpr@savepar \parskip 0em \hsize \linewidth \centering
+ \vbox \bgroup \mpr@prebindings \mpr@paroptions #1\ifmmode $\else
+ \noindent $\displaystyle\fi
+ \MathparBindings}
+ {\unskip \ifmmode $\fi\egroup $$\ignorespacesafterend}
+
+\newenvironment{mathparpagebreakable}[1][]
+ {\begingroup
+ \par
+ \mpr@savepar \parskip 0em \hsize \linewidth \centering
+ \mpr@prebindings \mpr@paroptions #1%
+ \vskip \abovedisplayskip \vskip -\lineskip%
+ \ifmmode \else $\displaystyle\fi
+ \MathparBindings
+ }
+ {\unskip
+ \ifmmode $\fi \par\endgroup
+ \vskip \belowdisplayskip
+ \noindent
+ \ignorespacesafterend}
+
+% \def \math@mathpar #1{\setbox0 \hbox {$\displaystyle #1$}\ifnum
+% \wd0 < \hsize $$\box0$$\else \bmathpar #1\emathpar \fi}
+
+%%% HOV BOXES
+
+\def \mathvbox@ #1{\hbox \bgroup \mpr@normallineskip
+ \vbox \bgroup \tabskip 0em \let \\ \cr
+ \halign \bgroup \hfil $##$\hfil\cr #1\crcr \egroup \egroup
+ \egroup}
+
+\def \mathhvbox@ #1{\setbox0 \hbox {\let \\\qquad $#1$}\ifnum \wd0 < \hsize
+ \box0\else \mathvbox {#1}\fi}
+
+
+%% Part II -- operations on lists
+
+\newtoks \mpr@lista
+\newtoks \mpr@listb
+
+\long \def\mpr@cons #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter
+{#2}\edef #2{\the \mpr@lista \the \mpr@listb}}
+
+\long \def\mpr@snoc #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter
+{#2}\edef #2{\the \mpr@listb\the\mpr@lista}}
+
+\long \def \mpr@concat#1=#2\mpr@to#3{\mpr@lista \expandafter {#2}\mpr@listb
+\expandafter {#3}\edef #1{\the \mpr@listb\the\mpr@lista}}
+
+\def \mpr@head #1\mpr@to #2{\expandafter \mpr@head@ #1\mpr@head@ #1#2}
+\long \def \mpr@head@ #1#2\mpr@head@ #3#4{\def #4{#1}\def#3{#2}}
+
+\def \mpr@flatten #1\mpr@to #2{\expandafter \mpr@flatten@ #1\mpr@flatten@ #1#2}
+\long \def \mpr@flatten@ \\#1\\#2\mpr@flatten@ #3#4{\def #4{#1}\def #3{\\#2}}
+
+\def \mpr@makelist #1\mpr@to #2{\def \mpr@all {#1}%
+ \mpr@lista {\\}\mpr@listb \expandafter {\mpr@all}\edef \mpr@all {\the
+ \mpr@lista \the \mpr@listb \the \mpr@lista}\let #2\empty
+ \def \mpr@stripof ##1##2\mpr@stripend{\def \mpr@stripped{##2}}\loop
+ \mpr@flatten \mpr@all \mpr@to \mpr@one
+ \expandafter \mpr@snoc \mpr@one \mpr@to #2\expandafter \mpr@stripof
+ \mpr@all \mpr@stripend
+ \ifx \mpr@stripped \empty \let \mpr@isempty 0\else \let \mpr@isempty 1\fi
+ \ifx 1\mpr@isempty
+ \repeat
+}
+
+\def \mpr@rev #1\mpr@to #2{\let \mpr@tmp \empty
+ \def \\##1{\mpr@cons ##1\mpr@to \mpr@tmp}#1\let #2\mpr@tmp}
+
+%% Part III -- Type inference rules
+
+\newif \if@premisse
+\newbox \mpr@hlist
+\newbox \mpr@vlist
+\newif \ifmpr@center \mpr@centertrue
+\def \mpr@htovlist {%
+ \setbox \mpr@hlist
+ \hbox {\strut
+ \ifmpr@center \hskip -0.5\wd\mpr@hlist\fi
+ \unhbox \mpr@hlist}%
+ \setbox \mpr@vlist
+ \vbox {\if@premisse \box \mpr@hlist \unvbox \mpr@vlist
+ \else \unvbox \mpr@vlist \box \mpr@hlist
+ \fi}%
+}
+% OLD version
+% \def \mpr@htovlist {%
+% \setbox \mpr@hlist
+% \hbox {\strut \hskip -0.5\wd\mpr@hlist \unhbox \mpr@hlist}%
+% \setbox \mpr@vlist
+% \vbox {\if@premisse \box \mpr@hlist \unvbox \mpr@vlist
+% \else \unvbox \mpr@vlist \box \mpr@hlist
+% \fi}%
+% }
+
+\def \mpr@item #1{$\displaystyle #1$}
+\def \mpr@sep{2em}
+\def \mpr@blank { }
+\def \mpr@hovbox #1#2{\hbox
+ \bgroup
+ \ifx #1T\@premissetrue
+ \else \ifx #1B\@premissefalse
+ \else
+ \PackageError{mathpartir}
+ {Premisse orientation should either be T or B}
+ {Fatal error in Package}%
+ \fi \fi
+ \def \@test {#2}\ifx \@test \mpr@blank\else
+ \setbox \mpr@hlist \hbox {}%
+ \setbox \mpr@vlist \vbox {}%
+ \if@premisse \let \snoc \mpr@cons \else \let \snoc \mpr@snoc \fi
+ \let \@hvlist \empty \let \@rev \empty
+ \mpr@tmpdim 0em
+ \expandafter \mpr@makelist #2\mpr@to \mpr@flat
+ \if@premisse \mpr@rev \mpr@flat \mpr@to \@rev \else \let \@rev \mpr@flat \fi
+ \def \\##1{%
+ \def \@test {##1}\ifx \@test \empty
+ \mpr@htovlist
+ \mpr@tmpdim 0em %%% last bug fix not extensively checked
+ \else
+ \setbox0 \hbox{\mpr@item {##1}}\relax
+ \advance \mpr@tmpdim by \wd0
+ %\mpr@tmpdim 1.02\mpr@tmpdim
+ \ifnum \mpr@tmpdim < \hsize
+ \ifnum \wd\mpr@hlist > 0
+ \if@premisse
+ \setbox \mpr@hlist
+ \hbox {\unhbox0 \hskip \mpr@sep \unhbox \mpr@hlist}%
+ \else
+ \setbox \mpr@hlist
+ \hbox {\unhbox \mpr@hlist \hskip \mpr@sep \unhbox0}%
+ \fi
+ \else
+ \setbox \mpr@hlist \hbox {\unhbox0}%
+ \fi
+ \else
+ \ifnum \wd \mpr@hlist > 0
+ \mpr@htovlist
+ \mpr@tmpdim \wd0
+ \fi
+ \setbox \mpr@hlist \hbox {\unhbox0}%
+ \fi
+ \advance \mpr@tmpdim by \mpr@sep
+ \fi
+ }%
+ \@rev
+ \mpr@htovlist
+ \ifmpr@center \hskip \wd\mpr@vlist\fi \box \mpr@vlist
+ \fi
+ \egroup
+}
+
+%%% INFERENCE RULES
+
+\@ifundefined{@@over}{%
+ \let\@@over\over % fallback if amsmath is not loaded
+ \let\@@overwithdelims\overwithdelims
+ \let\@@atop\atop \let\@@atopwithdelims\atopwithdelims
+ \let\@@above\above \let\@@abovewithdelims\abovewithdelims
+ }{}
+
+%% The default
+
+\def \mpr@@fraction #1#2{\hbox {\advance \hsize by -0.5em
+ $\displaystyle {#1\mpr@over #2}$}}
+\def \mpr@@nofraction #1#2{\hbox {\advance \hsize by -0.5em
+ $\displaystyle {#1\@@atop #2}$}}
+
+\let \mpr@fraction \mpr@@fraction
+
+%% A generic solution to arrow
+
+\def \mpr@make@fraction #1#2#3#4#5{\hbox {%
+ \def \mpr@tail{#1}%
+ \def \mpr@body{#2}%
+ \def \mpr@head{#3}%
+ \setbox1=\hbox{$#4$}\setbox2=\hbox{$#5$}%
+ \setbox3=\hbox{$\mkern -3mu\mpr@body\mkern -3mu$}%
+ \setbox3=\hbox{$\mkern -3mu \mpr@body\mkern -3mu$}%
+ \dimen0=\dp1\advance\dimen0 by \ht3\relax\dp1\dimen0\relax
+ \dimen0=\ht2\advance\dimen0 by \dp3\relax\ht2\dimen0\relax
+ \setbox0=\hbox {$\box1 \@@atop \box2$}%
+ \dimen0=\wd0\box0
+ \box0 \hskip -\dimen0\relax
+ \hbox to \dimen0 {$%
+ \mathrel{\mpr@tail}\joinrel
+ \xleaders\hbox{\copy3}\hfil\joinrel\mathrel{\mpr@head}%
+ $}}}
+
+%% Old stuff should be removed in next version
+\def \mpr@@nothing #1#2
+ {$\lower 0.01pt \mpr@@nofraction {#1}{#2}$}
+\def \mpr@@reduce #1#2{\hbox
+ {$\lower 0.01pt \mpr@@fraction {#1}{#2}\mkern -15mu\rightarrow$}}
+\def \mpr@@rewrite #1#2#3{\hbox
+ {$\lower 0.01pt \mpr@@fraction {#2}{#3}\mkern -8mu#1$}}
+\def \mpr@infercenter #1{\vcenter {\mpr@hovbox{T}{#1}}}
+
+\def \mpr@empty {}
+\def \mpr@inferrule
+ {\bgroup
+ \ifnum \linewidth<\hsize \hsize \linewidth\fi
+ \mpr@rulelineskip
+ \let \and \qquad
+ \let \hva \mpr@hva
+ \let \@rulename \mpr@empty
+ \let \@rule@options \mpr@empty
+ \let \mpr@over \@@over
+ \mpr@inferrule@}
+\newcommand {\mpr@inferrule@}[3][]
+ {\everymath={\displaystyle}%
+ \def \@test {#2}\ifx \empty \@test
+ \setbox0 \hbox {$\vcenter {\mpr@hovbox{B}{#3}}$}%
+ \else
+ \def \@test {#3}\ifx \empty \@test
+ \setbox0 \hbox {$\vcenter {\mpr@hovbox{T}{#2}}$}%
+ \else
+ \setbox0 \mpr@fraction {\mpr@hovbox{T}{#2}}{\mpr@hovbox{B}{#3}}%
+ \fi \fi
+ \def \@test {#1}\ifx \@test\empty \box0
+ \else \vbox
+%%% Suggestion de Francois pour les etiquettes longues
+%%% {\hbox to \wd0 {\RefTirName {#1}\hfil}\box0}\fi
+ {\hbox {\RefTirName {#1}}\box0}\fi
+ \egroup}
+
+\def \mpr@vdotfil #1{\vbox to #1{\leaders \hbox{$\cdot$} \vfil}}
+
+% They are two forms
+% \inferrule [label]{[premisses}{conclusions}
+% or
+% \inferrule* [options]{[premisses}{conclusions}
+%
+% Premisses and conclusions are lists of elements separated by \\
+% Each \\ produces a break, attempting horizontal breaks if possible,
+% and vertical breaks if needed.
+%
+% An empty element obtained by \\\\ produces a vertical break in all cases.
+%
+% The former rule is aligned on the fraction bar.
+% The optional label appears on top of the rule
+% The second form to be used in a derivation tree is aligned on the last
+% line of its conclusion
+%
+% The second form can be parameterized, using the key=val interface. The
+% folloiwng keys are recognized:
+%
+% width set the width of the rule to val
+% narrower set the width of the rule to val\hsize
+% before execute val at the beginning/left
+% lab put a label [Val] on top of the rule
+% lskip add negative skip on the right
+% left put a left label [Val]
+% Left put a left label [Val], ignoring its width
+% right put a right label [Val]
+% Right put a right label [Val], ignoring its width
+% leftskip skip negative space on the left-hand side
+% rightskip skip negative space on the right-hand side
+% vdots lift the rule by val and fill vertical space with dots
+% after execute val at the end/right
+%
+% Note that most options must come in this order to avoid strange
+% typesetting (in particular leftskip must preceed left and Left and
+% rightskip must follow Right or right; vdots must come last
+% or be only followed by rightskip.
+%
+
+%% Keys that make sence in all kinds of rules
+\def \mprset #1{\setkeys{mprset}{#1}}
+\define@key {mprset}{andskip}[]{\mpr@andskip=#1}
+\define@key {mprset}{lineskip}[]{\lineskip=#1}
+\define@key {mprset}{flushleft}[]{\mpr@centerfalse}
+\define@key {mprset}{center}[]{\mpr@centertrue}
+\define@key {mprset}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite}
+\define@key {mprset}{atop}[]{\let \mpr@fraction \mpr@@nofraction}
+\define@key {mprset}{myfraction}[]{\let \mpr@fraction #1}
+\define@key {mprset}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}}
+\define@key {mprset}{sep}{\def\mpr@sep{#1}}
+
+\newbox \mpr@right
+\define@key {mpr}{flushleft}[]{\mpr@centerfalse}
+\define@key {mpr}{center}[]{\mpr@centertrue}
+\define@key {mpr}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite}
+\define@key {mpr}{myfraction}[]{\let \mpr@fraction #1}
+\define@key {mpr}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}}
+\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax
+ \advance \hsize by -\wd0\box0}
+\define@key {mpr}{width}{\hsize #1}
+\define@key {mpr}{sep}{\def\mpr@sep{#1}}
+\define@key {mpr}{before}{#1}
+\define@key {mpr}{lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}}
+\define@key {mpr}{Lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}}
+\define@key {mpr}{narrower}{\hsize #1\hsize}
+\define@key {mpr}{leftskip}{\hskip -#1}
+\define@key {mpr}{reduce}[]{\let \mpr@fraction \mpr@@reduce}
+\define@key {mpr}{rightskip}
+ {\setbox \mpr@right \hbox {\unhbox \mpr@right \hskip -#1}}
+\define@key {mpr}{LEFT}{\setbox0 \hbox {$#1$}\relax
+ \advance \hsize by -\wd0\box0}
+\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax
+ \advance \hsize by -\wd0\box0}
+\define@key {mpr}{Left}{\llap{$\TirName {#1}\;$}}
+\define@key {mpr}{right}
+ {\setbox0 \hbox {$\;\TirName {#1}$}\relax \advance \hsize by -\wd0
+ \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}}
+\define@key {mpr}{RIGHT}
+ {\setbox0 \hbox {$#1$}\relax \advance \hsize by -\wd0
+ \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}}
+\define@key {mpr}{Right}
+ {\setbox \mpr@right \hbox {\unhbox \mpr@right \rlap {$\;\TirName {#1}$}}}
+\define@key {mpr}{vdots}{\def \mpr@vdots {\@@atop \mpr@vdotfil{#1}}}
+\define@key {mpr}{after}{\edef \mpr@after {\mpr@after #1}}
+
+\newcommand \mpr@inferstar@ [3][]{\setbox0
+ \hbox {\let \mpr@rulename \mpr@empty \let \mpr@vdots \relax
+ \setbox \mpr@right \hbox{}%
+ $\setkeys{mpr}{#1}%
+ \ifx \mpr@rulename \mpr@empty \mpr@inferrule {#2}{#3}\else
+ \mpr@inferrule [{\mpr@rulename}]{#2}{#3}\fi
+ \box \mpr@right \mpr@vdots$}
+ \setbox1 \hbox {\strut}
+ \@tempdima \dp0 \advance \@tempdima by -\dp1
+ \raise \@tempdima \box0}
+
+\def \mpr@infer {\@ifnextchar *{\mpr@inferstar}{\mpr@inferrule}}
+\newcommand \mpr@err@skipargs[3][]{}
+\def \mpr@inferstar*{\ifmmode
+ \let \@do \mpr@inferstar@
+ \else
+ \let \@do \mpr@err@skipargs
+ \PackageError {mathpartir}
+ {\string\inferrule* can only be used in math mode}{}%
+ \fi \@do}
+
+
+%%% Exports
+
+% Envirnonment mathpar
+
+\let \inferrule \mpr@infer
+
+% make a short name \infer is not already defined
+\@ifundefined {infer}{\let \infer \mpr@infer}{}
+
+\def \TirNameStyle #1{\small \textsc{#1}}
+\def \tir@name #1{\hbox {\small \TirNameStyle{#1}}}
+\let \TirName \tir@name
+\let \DefTirName \TirName
+\let \RefTirName \TirName
+
+%%% Other Exports
+
+% \let \listcons \mpr@cons
+% \let \listsnoc \mpr@snoc
+% \let \listhead \mpr@head
+% \let \listmake \mpr@makelist
+
+
+
+
+\endinput
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Paper/document/root.bib Fri Jan 18 11:40:01 2013 +0000
@@ -0,0 +1,68 @@
+@article{UrbanCheneyBerghofer11,
+ author = {C.~Urban and J.~Cheney and S.~Berghofer},
+ title = {{M}echanizing the {M}etatheory of {LF}},
+ journal = {ACM Transactions on Computational Logic},
+ volume = {12},
+ issue = {2},
+ year = {2011},
+ pages = {15:1--15:42}
+}
+
+@inproceedings{Norrish11,
+ author = {M.~Norrish},
+ title = {{M}echanised {C}omputability {T}heory},
+ booktitle = {Proc.~of the 2nd Conference on Interactive Theorem Proving (ITP)},
+ year = {2011},
+ series = {LNCS},
+ volume = {6898},
+ pages = {297--311}
+}
+
+@inproceedings{AspertiRicciotti12,
+ author = {A.~Asperti and W.~Ricciotti},
+ title = {{F}ormalizing {T}uring {M}achines},
+ booktitle = {Proc.~of the 19th International Workshop on Logic, Language,
+ Information and Computation (WoLLIC)},
+ year = {2012},
+ pages = {1-25},
+ series = {LNCS},
+ volume = {7456}
+}
+
+
+@Unpublished{WuZhangUrban12,
+ author = {C.~Wu and X.~Zhang and C.~Urban},
+ title = {???},
+ note = {Submitted},
+ year = {2012}
+}
+
+@book{Boolos87,
+ author = {G.~Boolos and J.~P.~Burgess and R.~C.~Jeffrey},
+ title = {{C}omputability and {L}ogic (5th~ed.)},
+ publisher = {Cambridge University Press},
+ year = {2007}
+}
+
+@inproceedings{WuZhangUrban11,
+ author = {C.~Wu and X.~Zhang and C.~Urban},
+ title = {{A} {F}ormalisation of the {M}yhill-{N}erode {T}heorem based on {R}egular {E}xpressions
+ ({P}roof {P}earl)},
+ booktitle = {Proc.~of the 2nd Conference on Interactive Theorem Proving},
+ year = {2011},
+ pages = {341--356},
+ series = {LNCS},
+ volume = {6898}
+}
+
+
+@Article{Post36,
+ author = {E.~Post},
+ title = {{F}inite {C}ombinatory {P}rocesses-{F}ormulation 1},
+ journal = {Journal of Symbolic Logic},
+ year = {1936},
+ volume = {1},
+ number = {3},
+ pages = {103--105}
+}
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Paper/document/root.tex Fri Jan 18 11:40:01 2013 +0000
@@ -0,0 +1,62 @@
+\documentclass[runningheads]{llncs}
+%\documentclass[10pt, conference, compsocconf]{IEEEtran}
+\usepackage{isabelle}
+\usepackage{isabellesym}
+\usepackage{times}
+\usepackage{amssymb}
+\usepackage{amsmath}
+\usepackage{mathpartir}
+\usepackage{pdfsetup}
+\usepackage{tikz}
+\usepackage{pgf}
+
+% urls in roman style, theory text in math-similar italics
+\urlstyle{rm}
+\isabellestyle{it}
+
+% for uniform font size
+%\renewcommand{\isastyle}{\isastyleminor}
+
+\def\dn{\,\stackrel{\mbox{\scriptsize def}}{=}\,}
+\renewcommand{\isasymequiv}{$\dn$}
+\renewcommand{\isasymemptyset}{$\varnothing$}
+\renewcommand{\isacharunderscore}{\mbox{$\_$}}
+\renewcommand{\isasymiota}{}
+\newcommand{\isasymulcorner}{$\ulcorner$}
+\newcommand{\isasymurcorner}{$\urcorner$}
+\begin{document}
+
+
+\title{Formalising Computability Theory in Isabelle/HOL}
+\author{Jian Xu\inst{1} \and Xingyuan Zhang\inst{1} \and Christian Urban\inst{2}}
+\institute{PLA University of Science and Technology, China \and King's College London, UK}
+
+\maketitle
+
+
+\begin{abstract}
+We present a formalised theory of computability in the
+theorem prover Isabelle/HOL. This theorem prover is based on classical
+logic which precludes \emph{direct} reasoning about computability: every
+boolean predicate is either true or false because of the law of excluded
+middle. The only way to reason about computability in a classical theorem
+prover is to formalise a concrete model for computation.
+We formalise Turing machines and relate them to abacus machines and recursive
+functions. Our theory can be used to formalise other computability results:
+{\it we give one example about the undecidability of Wang's tiling problem, whose proof uses
+the notion of a universal Turing machine.}
+\end{abstract}
+
+% generated text of all theories
+\input{session}
+
+% optional bibliography
+\bibliographystyle{abbrv}
+\bibliography{root}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
--- a/ROOT1.ML Thu Jan 17 11:51:00 2013 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-(*
- turing_basic.thy : The basic definitions of Turing Machine.
- uncomputable.thy : The existence of Turing uncomputable functions.
- abacus.thy : The basic definitions of Abacus machine (An intermediate langauge underneath recursive functions) and
- the compilation of Abacus machines into Turing Machines.
- rec_def.thy: The basic definitions of Recursive Functions.
- recursive.thy : The compilation of Recursive Functions into
- Abacus machines.
- UF.thy : The construction of Universal Function, named "rec_F" and the proof of its correctness.
- UTM.thy: Obtaining Uinversal Turing Machine by scarfolding the Turing Machine compiled from "rec_F" with some
- initialization and termination processing Turing Machines.
-*)
-
-no_document
-use_thys ["turing_basic",
- "uncomputable",
- "abacus",
- "rec_def",
- "recursive",
- "UF",
- "UTM"];
-
-
-use_thys ["Paper"]
--- a/document/IEEEtran.cls Thu Jan 17 11:51:00 2013 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4722 +0,0 @@
-%%
-%% IEEEtran.cls 2007/03/05 version V1.7a
-%%
-%%
-%% This is the official IEEE LaTeX class for authors of the Institute of
-%% Electrical and Electronics Engineers (IEEE) Transactions journals and
-%% conferences.
-%%
-%% Support sites:
-%% http://www.michaelshell.org/tex/ieeetran/
-%% http://www.ctan.org/tex-archive/macros/latex/contrib/IEEEtran/
-%% and
-%% http://www.ieee.org/
-%%
-%% Based on the original 1993 IEEEtran.cls, but with many bug fixes
-%% and enhancements (from both JVH and MDS) over the 1996/7 version.
-%%
-%%
-%% Contributors:
-%% Gerry Murray (1993), Silvano Balemi (1993),
-%% Jon Dixon (1996), Peter N"uchter (1996),
-%% Juergen von Hagen (2000), and Michael Shell (2001-2007)
-%%
-%%
-%% Copyright (c) 1993-2000 by Gerry Murray, Silvano Balemi,
-%% Jon Dixon, Peter N"uchter,
-%% Juergen von Hagen
-%% and
-%% Copyright (c) 2001-2007 by Michael Shell
-%%
-%% Current maintainer (V1.3 to V1.7): Michael Shell
-%% See:
-%% http://www.michaelshell.org/
-%% for current contact information.
-%%
-%% Special thanks to Peter Wilson (CUA) and Donald Arseneau
-%% for allowing the inclusion of the \@ifmtarg command
-%% from their ifmtarg LaTeX package.
-%%
-%%*************************************************************************
-%% Legal Notice:
-%% This code is offered as-is without any warranty either expressed or
-%% implied; without even the implied warranty of MERCHANTABILITY or
-%% FITNESS FOR A PARTICULAR PURPOSE!
-%% User assumes all risk.
-%% In no event shall IEEE or any contributor to this code be liable for
-%% any damages or losses, including, but not limited to, incidental,
-%% consequential, or any other damages, resulting from the use or misuse
-%% of any information contained here.
-%%
-%% All comments are the opinions of their respective authors and are not
-%% necessarily endorsed by the IEEE.
-%%
-%% This work is distributed under the LaTeX Project Public License (LPPL)
-%% ( http://www.latex-project.org/ ) version 1.3, and may be freely used,
-%% distributed and modified. A copy of the LPPL, version 1.3, is included
-%% in the base LaTeX documentation of all distributions of LaTeX released
-%% 2003/12/01 or later.
-%% Retain all contribution notices and credits.
-%% ** Modified files should be clearly indicated as such, including **
-%% ** renaming them and changing author support contact information. **
-%%
-%% File list of work: IEEEtran.cls, IEEEtran_HOWTO.pdf, bare_adv.tex,
-%% bare_conf.tex, bare_jrnl.tex, bare_jrnl_compsoc.tex
-%%
-%% Major changes to the user interface should be indicated by an
-%% increase in the version numbers. If a version is a beta, it will
-%% be indicated with a BETA suffix, i.e., 1.4 BETA.
-%% Small changes can be indicated by appending letters to the version
-%% such as "IEEEtran_v14a.cls".
-%% In all cases, \Providesclass, any \typeout messages to the user,
-%% \IEEEtransversionmajor and \IEEEtransversionminor must reflect the
-%% correct version information.
-%% The changes should also be documented via source comments.
-%%*************************************************************************
-%%
-%
-% Available class options
-% e.g., \documentclass[10pt,conference]{IEEEtran}
-%
-% *** choose only one from each category ***
-%
-% 9pt, 10pt, 11pt, 12pt
-% Sets normal font size. The default is 10pt.
-%
-% conference, journal, technote, peerreview, peerreviewca
-% determines format mode - conference papers, journal papers,
-% correspondence papers (technotes), or peer review papers. The user
-% should also select 9pt when using technote. peerreview is like
-% journal mode, but provides for a single-column "cover" title page for
-% anonymous peer review. The paper title (without the author names) is
-% repeated at the top of the page after the cover page. For peer review
-% papers, the \IEEEpeerreviewmaketitle command must be executed (will
-% automatically be ignored for non-peerreview modes) at the place the
-% cover page is to end, usually just after the abstract (keywords are
-% not normally used with peer review papers). peerreviewca is like
-% peerreview, but allows the author names to be entered and formatted
-% as with conference mode so that author affiliation and contact
-% information can be easily seen on the cover page.
-% The default is journal.
-%
-% draft, draftcls, draftclsnofoot, final
-% determines if paper is formatted as a widely spaced draft (for
-% handwritten editor comments) or as a properly typeset final version.
-% draftcls restricts draft mode to the class file while all other LaTeX
-% packages (i.e., \usepackage{graphicx}) will behave as final - allows
-% for a draft paper with visible figures, etc. draftclsnofoot is like
-% draftcls, but does not display the date and the word "DRAFT" at the foot
-% of the pages. If using one of the draft modes, the user will probably
-% also want to select onecolumn.
-% The default is final.
-%
-% letterpaper, a4paper
-% determines paper size: 8.5in X 11in or 210mm X 297mm. CHANGING THE PAPER
-% SIZE WILL NOT ALTER THE TYPESETTING OF THE DOCUMENT - ONLY THE MARGINS
-% WILL BE AFFECTED. In particular, documents using the a4paper option will
-% have reduced side margins (A4 is narrower than US letter) and a longer
-% bottom margin (A4 is longer than US letter). For both cases, the top
-% margins will be the same and the text will be horizontally centered.
-% For final submission to IEEE, authors should use US letter (8.5 X 11in)
-% paper. Note that authors should ensure that all post-processing
-% (ps, pdf, etc.) uses the same paper specificiation as the .tex document.
-% Problems here are by far the number one reason for incorrect margins.
-% IEEEtran will automatically set the default paper size under pdflatex
-% (without requiring a change to pdftex.cfg), so this issue is more
-% important to dvips users. Fix config.ps, config.pdf, or ~/.dvipsrc for
-% dvips, or use the dvips -t papersize option instead as needed. See the
-% testflow documentation
-% http://www.ctan.org/tex-archive/macros/latex/contrib/IEEEtran/testflow
-% for more details on dvips paper size configuration.
-% The default is letterpaper.
-%
-% oneside, twoside
-% determines if layout follows single sided or two sided (duplex)
-% printing. The only notable change is with the headings at the top of
-% the pages.
-% The default is oneside.
-%
-% onecolumn, twocolumn
-% determines if text is organized into one or two columns per page. One
-% column mode is usually used only with draft papers.
-% The default is twocolumn.
-%
-% compsoc
-% Use the format of the IEEE Computer Society.
-%
-% compsocconf
-% Use the format of IEEE Computer Society conferencs (CPS)
-%
-% romanappendices
-% Use the "Appendix I" convention when numbering appendices. IEEEtran.cls
-% now defaults to Alpha "Appendix A" convention - the opposite of what
-% v1.6b and earlier did.
-%
-% captionsoff
-% disables the display of the figure/table captions. Some IEEE journals
-% request that captions be removed and figures/tables be put on pages
-% of their own at the end of an initial paper submission. The endfloat
-% package can be used with this class option to achieve this format.
-%
-% nofonttune
-% turns off tuning of the font interword spacing. Maybe useful to those
-% not using the standard Times fonts or for those who have already "tuned"
-% their fonts.
-% The default is to enable IEEEtran to tune font parameters.
-%
-%
-%----------
-% Available CLASSINPUTs provided (all are macros unless otherwise noted):
-% \CLASSINPUTbaselinestretch
-% \CLASSINPUTinnersidemargin
-% \CLASSINPUToutersidemargin
-% \CLASSINPUTtoptextmargin
-% \CLASSINPUTbottomtextmargin
-%
-% Available CLASSINFOs provided:
-% \ifCLASSINFOpdf (TeX if conditional)
-% \CLASSINFOpaperwidth (macro)
-% \CLASSINFOpaperheight (macro)
-% \CLASSINFOnormalsizebaselineskip (length)
-% \CLASSINFOnormalsizeunitybaselineskip (length)
-%
-% Available CLASSOPTIONs provided:
-% all class option flags (TeX if conditionals) unless otherwise noted,
-% e.g., \ifCLASSOPTIONcaptionsoff
-% point size options provided as a single macro:
-% \CLASSOPTIONpt
-% which will be defined as 9, 10, 11, or 12 depending on the document's
-% normalsize point size.
-% also, class option peerreviewca implies the use of class option peerreview
-% and classoption draft implies the use of class option draftcls
-
-
-
-
-
-\ProvidesClass{IEEEtran}[2007/03/05 V1.7a by Michael Shell]
-\typeout{-- See the "IEEEtran_HOWTO" manual for usage information.}
-\typeout{-- http://www.michaelshell.org/tex/ieeetran/}
-\NeedsTeXFormat{LaTeX2e}
-
-% IEEEtran.cls version numbers, provided as of V1.3
-% These values serve as a way a .tex file can
-% determine if the new features are provided.
-% The version number of this IEEEtrans.cls can be obtained from
-% these values. i.e., V1.4
-% KEEP THESE AS INTEGERS! i.e., NO {4a} or anything like that-
-% (no need to enumerate "a" minor changes here)
-\def\IEEEtransversionmajor{1}
-\def\IEEEtransversionminor{7}
-
-% These do nothing, but provide them like in article.cls
-\newif\if@restonecol
-\newif\if@titlepage
-
-
-% class option conditionals
-\newif\ifCLASSOPTIONonecolumn \CLASSOPTIONonecolumnfalse
-\newif\ifCLASSOPTIONtwocolumn \CLASSOPTIONtwocolumntrue
-
-\newif\ifCLASSOPTIONoneside \CLASSOPTIONonesidetrue
-\newif\ifCLASSOPTIONtwoside \CLASSOPTIONtwosidefalse
-
-\newif\ifCLASSOPTIONfinal \CLASSOPTIONfinaltrue
-\newif\ifCLASSOPTIONdraft \CLASSOPTIONdraftfalse
-\newif\ifCLASSOPTIONdraftcls \CLASSOPTIONdraftclsfalse
-\newif\ifCLASSOPTIONdraftclsnofoot \CLASSOPTIONdraftclsnofootfalse
-
-\newif\ifCLASSOPTIONpeerreview \CLASSOPTIONpeerreviewfalse
-\newif\ifCLASSOPTIONpeerreviewca \CLASSOPTIONpeerreviewcafalse
-
-\newif\ifCLASSOPTIONjournal \CLASSOPTIONjournaltrue
-\newif\ifCLASSOPTIONconference \CLASSOPTIONconferencefalse
-\newif\ifCLASSOPTIONtechnote \CLASSOPTIONtechnotefalse
-
-\newif\ifCLASSOPTIONnofonttune \CLASSOPTIONnofonttunefalse
-
-\newif\ifCLASSOPTIONcaptionsoff \CLASSOPTIONcaptionsofffalse
-
-\newif\ifCLASSOPTIONcompsoc \CLASSOPTIONcompsocfalse
-
-\newif\ifCLASSOPTIONcompsocconf \CLASSOPTIONcompsocconffalse
-
-\newif\ifCLASSOPTIONromanappendices \CLASSOPTIONromanappendicesfalse
-
-
-% class info conditionals
-
-% indicates if pdf (via pdflatex) output
-\newif\ifCLASSINFOpdf \CLASSINFOpdffalse
-
-
-% V1.6b internal flag to show if using a4paper
-\newif\if@IEEEusingAfourpaper \@IEEEusingAfourpaperfalse
-
-
-
-% IEEEtran class scratch pad registers
-% dimen
-\newdimen\@IEEEtrantmpdimenA
-\newdimen\@IEEEtrantmpdimenB
-% count
-\newcount\@IEEEtrantmpcountA
-\newcount\@IEEEtrantmpcountB
-% token list
-\newtoks\@IEEEtrantmptoksA
-
-% we use \CLASSOPTIONpt so that we can ID the point size (even for 9pt docs)
-% as well as LaTeX's \@ptsize to retain some compatability with some
-% external packages
-\def\@ptsize{0}
-% LaTeX does not support 9pt, so we set \@ptsize to 0 - same as that of 10pt
-\DeclareOption{9pt}{\def\CLASSOPTIONpt{9}\def\@ptsize{0}}
-\DeclareOption{10pt}{\def\CLASSOPTIONpt{10}\def\@ptsize{0}}
-\DeclareOption{11pt}{\def\CLASSOPTIONpt{11}\def\@ptsize{1}}
-\DeclareOption{12pt}{\def\CLASSOPTIONpt{12}\def\@ptsize{2}}
-
-
-
-\DeclareOption{letterpaper}{\setlength{\paperheight}{11in}%
- \setlength{\paperwidth}{8.5in}%
- \@IEEEusingAfourpaperfalse
- \def\CLASSOPTIONpaper{letter}%
- \def\CLASSINFOpaperwidth{8.5in}%
- \def\CLASSINFOpaperheight{11in}}
-
-
-\DeclareOption{a4paper}{\setlength{\paperheight}{297mm}%
- \setlength{\paperwidth}{210mm}%
- \@IEEEusingAfourpapertrue
- \def\CLASSOPTIONpaper{a4}%
- \def\CLASSINFOpaperwidth{210mm}%
- \def\CLASSINFOpaperheight{297mm}}
-
-\DeclareOption{oneside}{\@twosidefalse\@mparswitchfalse
- \CLASSOPTIONonesidetrue\CLASSOPTIONtwosidefalse}
-\DeclareOption{twoside}{\@twosidetrue\@mparswitchtrue
- \CLASSOPTIONtwosidetrue\CLASSOPTIONonesidefalse}
-
-\DeclareOption{onecolumn}{\CLASSOPTIONonecolumntrue\CLASSOPTIONtwocolumnfalse}
-\DeclareOption{twocolumn}{\CLASSOPTIONtwocolumntrue\CLASSOPTIONonecolumnfalse}
-
-% If the user selects draft, then this class AND any packages
-% will go into draft mode.
-\DeclareOption{draft}{\CLASSOPTIONdrafttrue\CLASSOPTIONdraftclstrue
- \CLASSOPTIONdraftclsnofootfalse}
-% draftcls is for a draft mode which will not affect any packages
-% used by the document.
-\DeclareOption{draftcls}{\CLASSOPTIONdraftfalse\CLASSOPTIONdraftclstrue
- \CLASSOPTIONdraftclsnofootfalse}
-% draftclsnofoot is like draftcls, but without the footer.
-\DeclareOption{draftclsnofoot}{\CLASSOPTIONdraftfalse\CLASSOPTIONdraftclstrue
- \CLASSOPTIONdraftclsnofoottrue}
-\DeclareOption{final}{\CLASSOPTIONdraftfalse\CLASSOPTIONdraftclsfalse
- \CLASSOPTIONdraftclsnofootfalse}
-
-\DeclareOption{journal}{\CLASSOPTIONpeerreviewfalse\CLASSOPTIONpeerreviewcafalse
- \CLASSOPTIONjournaltrue\CLASSOPTIONconferencefalse\CLASSOPTIONtechnotefalse}
-
-\DeclareOption{conference}{\CLASSOPTIONpeerreviewfalse\CLASSOPTIONpeerreviewcafalse
- \CLASSOPTIONjournalfalse\CLASSOPTIONconferencetrue\CLASSOPTIONtechnotefalse}
-
-\DeclareOption{technote}{\CLASSOPTIONpeerreviewfalse\CLASSOPTIONpeerreviewcafalse
- \CLASSOPTIONjournalfalse\CLASSOPTIONconferencefalse\CLASSOPTIONtechnotetrue}
-
-\DeclareOption{peerreview}{\CLASSOPTIONpeerreviewtrue\CLASSOPTIONpeerreviewcafalse
- \CLASSOPTIONjournalfalse\CLASSOPTIONconferencefalse\CLASSOPTIONtechnotefalse}
-
-\DeclareOption{peerreviewca}{\CLASSOPTIONpeerreviewtrue\CLASSOPTIONpeerreviewcatrue
- \CLASSOPTIONjournalfalse\CLASSOPTIONconferencefalse\CLASSOPTIONtechnotefalse}
-
-\DeclareOption{nofonttune}{\CLASSOPTIONnofonttunetrue}
-
-\DeclareOption{captionsoff}{\CLASSOPTIONcaptionsofftrue}
-
-\DeclareOption{compsoc}{\CLASSOPTIONcompsoctrue}
-
-\DeclareOption{compsocconf}{\CLASSOPTIONcompsocconftrue}
-
-\DeclareOption{romanappendices}{\CLASSOPTIONromanappendicestrue}
-
-
-% default to US letter paper, 10pt, twocolumn, one sided, final, journal
-\ExecuteOptions{letterpaper,10pt,twocolumn,oneside,final,journal}
-% overrride these defaults per user requests
-\ProcessOptions
-
-
-
-% Computer Society conditional execution command
-\long\def\@IEEEcompsoconly#1{\relax\ifCLASSOPTIONcompsoc\relax#1\relax\fi\relax}
-% inverse
-\long\def\@IEEEnotcompsoconly#1{\relax\ifCLASSOPTIONcompsoc\else\relax#1\relax\fi\relax}
-% compsoc conference
-\long\def\@IEEEcompsocconfonly#1{\relax\ifCLASSOPTIONcompsocconf\ifCLASSOPTIONconference\relax#1\relax\fi\fi\relax}
-% compsoc not conference
-\long\def\@IEEEcompsocnotconfonly#1{\relax\ifCLASSOPTIONcompsoc\ifCLASSOPTIONconference\else\relax#1\relax\fi\fi\relax}
-
-
-% IEEE uses Times Roman font, so we'll default to Times.
-% These three commands make up the entire times.sty package.
-\renewcommand{\sfdefault}{phv}
-\renewcommand{\rmdefault}{ptm}
-\renewcommand{\ttdefault}{pcr}
-
-\@IEEEcompsoconly{\typeout{-- Using IEEE Computer Society mode.}}
-
-% V1.7 compsoc nonconference papers, use Palatino/Palladio as the main text font,
-% not Times Roman.
-\@IEEEcompsocnotconfonly{\renewcommand{\rmdefault}{ppl}}
-
-% enable Times/Palatino main text font
-\normalfont\selectfont
-
-
-
-
-
-% V1.7 conference notice message hook
-\def\@IEEEconsolenoticeconference{\typeout{}%
-\typeout{** Conference Paper **}%
-\typeout{Before submitting the final camera ready copy, remember to:}%
-\typeout{}%
-\typeout{ 1. Manually equalize the lengths of two columns on the last page}%
-\typeout{ of your paper;}%
-\typeout{}%
-\typeout{ 2. Ensure that any PostScript and/or PDF output post-processing}%
-\typeout{ uses only Type 1 fonts and that every step in the generation}%
-\typeout{ process uses the appropriate paper size.}%
-\typeout{}}
-
-
-% we can send console reminder messages to the user here
-\AtEndDocument{\ifCLASSOPTIONconference\@IEEEconsolenoticeconference\fi}
-
-
-% warn about the use of single column other than for draft mode
-\ifCLASSOPTIONtwocolumn\else%
- \ifCLASSOPTIONdraftcls\else%
- \typeout{** ATTENTION: Single column mode is not typically used with IEEE publications.}%
- \fi%
-\fi
-
-
-% V1.7 improved paper size setting code.
-% Set pdfpage and dvips paper sizes. Conditional tests are similar to that
-% of ifpdf.sty. Retain within {} to ensure tested macros are never altered,
-% even if only effect is to set them to \relax.
-% if \pdfoutput is undefined or equal to relax, output a dvips special
-{\@ifundefined{pdfoutput}{\AtBeginDvi{\special{papersize=\CLASSINFOpaperwidth,\CLASSINFOpaperheight}}}{%
-% pdfoutput is defined and not equal to \relax
-% check for pdfpageheight existence just in case someone sets pdfoutput
-% under non-pdflatex. If exists, set them regardless of value of \pdfoutput.
-\@ifundefined{pdfpageheight}{\relax}{\global\pdfpagewidth\paperwidth
-\global\pdfpageheight\paperheight}%
-% if using \pdfoutput=0 under pdflatex, send dvips papersize special
-\ifcase\pdfoutput
-\AtBeginDvi{\special{papersize=\CLASSINFOpaperwidth,\CLASSINFOpaperheight}}%
-\else
-% we are using pdf output, set CLASSINFOpdf flag
-\global\CLASSINFOpdftrue
-\fi}}
-
-% let the user know the selected papersize
-\typeout{-- Using \CLASSINFOpaperwidth\space x \CLASSINFOpaperheight\space
-(\CLASSOPTIONpaper)\space paper.}
-
-\ifCLASSINFOpdf
-\typeout{-- Using PDF output.}
-\else
-\typeout{-- Using DVI output.}
-\fi
-
-
-% The idea hinted here is for LaTeX to generate markleft{} and markright{}
-% automatically for you after you enter \author{}, \journal{},
-% \journaldate{}, journalvol{}, \journalnum{}, etc.
-% However, there may be some backward compatibility issues here as
-% well as some special applications for IEEEtran.cls and special issues
-% that may require the flexible \markleft{}, \markright{} and/or \markboth{}.
-% We'll leave this as an open future suggestion.
-%\newcommand{\journal}[1]{\def\@journal{#1}}
-%\def\@journal{}
-
-
-
-% pointsize values
-% used with ifx to determine the document's normal size
-\def\@IEEEptsizenine{9}
-\def\@IEEEptsizeten{10}
-\def\@IEEEptsizeeleven{11}
-\def\@IEEEptsizetwelve{12}
-
-
-
-% FONT DEFINITIONS (No sizexx.clo file needed)
-% V1.6 revised font sizes, displayskip values and
-% revised normalsize baselineskip to reduce underfull vbox problems
-% on the 58pc = 696pt = 9.5in text height we want
-% normalsize #lines/column baselineskip (aka leading)
-% 9pt 63 11.0476pt (truncated down)
-% 10pt 58 12pt (exact)
-% 11pt 52 13.3846pt (truncated down)
-% 12pt 50 13.92pt (exact)
-%
-
-% we need to store the nominal baselineskip for the given font size
-% in case baselinestretch ever changes.
-% this is a dimen, so it will not hold stretch or shrink
-\newdimen\@IEEEnormalsizeunitybaselineskip
-\@IEEEnormalsizeunitybaselineskip\baselineskip
-
-\ifx\CLASSOPTIONpt\@IEEEptsizenine
-\typeout{-- This is a 9 point document.}
-\def\normalsize{\@setfontsize{\normalsize}{9}{11.0476pt}}%
-\setlength{\@IEEEnormalsizeunitybaselineskip}{11.0476pt}%
-\normalsize
-\abovedisplayskip 1.5ex plus3pt minus1pt%
-\belowdisplayskip \abovedisplayskip%
-\abovedisplayshortskip 0pt plus3pt%
-\belowdisplayshortskip 1.5ex plus3pt minus1pt
-\def\small{\@setfontsize{\small}{8.5}{10pt}}
-\def\footnotesize{\@setfontsize{\footnotesize}{8}{9pt}}
-\def\scriptsize{\@setfontsize{\scriptsize}{7}{8pt}}
-\def\tiny{\@setfontsize{\tiny}{5}{6pt}}
-% sublargesize is the same as large - 10pt
-\def\sublargesize{\@setfontsize{\sublargesize}{10}{12pt}}
-\def\large{\@setfontsize{\large}{10}{12pt}}
-\def\Large{\@setfontsize{\Large}{12}{14pt}}
-\def\LARGE{\@setfontsize{\LARGE}{14}{17pt}}
-\def\huge{\@setfontsize{\huge}{17}{20pt}}
-\def\Huge{\@setfontsize{\Huge}{20}{24pt}}
-\fi
-
-
-% Check if we have selected 10 points
-\ifx\CLASSOPTIONpt\@IEEEptsizeten
-\typeout{-- This is a 10 point document.}
-\def\normalsize{\@setfontsize{\normalsize}{10}{12.00pt}}%
-\setlength{\@IEEEnormalsizeunitybaselineskip}{12pt}%
-\normalsize
-\abovedisplayskip 1.5ex plus4pt minus2pt%
-\belowdisplayskip \abovedisplayskip%
-\abovedisplayshortskip 0pt plus4pt%
-\belowdisplayshortskip 1.5ex plus4pt minus2pt
-\def\small{\@setfontsize{\small}{9}{10pt}}
-\def\footnotesize{\@setfontsize{\footnotesize}{8}{9pt}}
-\def\scriptsize{\@setfontsize{\scriptsize}{7}{8pt}}
-\def\tiny{\@setfontsize{\tiny}{5}{6pt}}
-% sublargesize is a tad smaller than large - 11pt
-\def\sublargesize{\@setfontsize{\sublargesize}{11}{13.4pt}}
-\def\large{\@setfontsize{\large}{12}{14pt}}
-\def\Large{\@setfontsize{\Large}{14}{17pt}}
-\def\LARGE{\@setfontsize{\LARGE}{17}{20pt}}
-\def\huge{\@setfontsize{\huge}{20}{24pt}}
-\def\Huge{\@setfontsize{\Huge}{24}{28pt}}
-\fi
-
-
-% Check if we have selected 11 points
-\ifx\CLASSOPTIONpt\@IEEEptsizeeleven
-\typeout{-- This is an 11 point document.}
-\def\normalsize{\@setfontsize{\normalsize}{11}{13.3846pt}}%
-\setlength{\@IEEEnormalsizeunitybaselineskip}{13.3846pt}%
-\normalsize
-\abovedisplayskip 1.5ex plus5pt minus3pt%
-\belowdisplayskip \abovedisplayskip%
-\abovedisplayshortskip 0pt plus5pt%
-\belowdisplayshortskip 1.5ex plus5pt minus3pt
-\def\small{\@setfontsize{\small}{10}{12pt}}
-\def\footnotesize{\@setfontsize{\footnotesize}{9}{10.5pt}}
-\def\scriptsize{\@setfontsize{\scriptsize}{8}{9pt}}
-\def\tiny{\@setfontsize{\tiny}{6}{7pt}}
-% sublargesize is the same as large - 12pt
-\def\sublargesize{\@setfontsize{\sublargesize}{12}{14pt}}
-\def\large{\@setfontsize{\large}{12}{14pt}}
-\def\Large{\@setfontsize{\Large}{14}{17pt}}
-\def\LARGE{\@setfontsize{\LARGE}{17}{20pt}}
-\def\huge{\@setfontsize{\huge}{20}{24pt}}
-\def\Huge{\@setfontsize{\Huge}{24}{28pt}}
-\fi
-
-
-% Check if we have selected 12 points
-\ifx\CLASSOPTIONpt\@IEEEptsizetwelve
-\typeout{-- This is a 12 point document.}
-\def\normalsize{\@setfontsize{\normalsize}{12}{13.92pt}}%
-\setlength{\@IEEEnormalsizeunitybaselineskip}{13.92pt}%
-\normalsize
-\abovedisplayskip 1.5ex plus6pt minus4pt%
-\belowdisplayskip \abovedisplayskip%
-\abovedisplayshortskip 0pt plus6pt%
-\belowdisplayshortskip 1.5ex plus6pt minus4pt
-\def\small{\@setfontsize{\small}{10}{12pt}}
-\def\footnotesize{\@setfontsize{\footnotesize}{9}{10.5pt}}
-\def\scriptsize{\@setfontsize{\scriptsize}{8}{9pt}}
-\def\tiny{\@setfontsize{\tiny}{6}{7pt}}
-% sublargesize is the same as large - 14pt
-\def\sublargesize{\@setfontsize{\sublargesize}{14}{17pt}}
-\def\large{\@setfontsize{\large}{14}{17pt}}
-\def\Large{\@setfontsize{\Large}{17}{20pt}}
-\def\LARGE{\@setfontsize{\LARGE}{20}{24pt}}
-\def\huge{\@setfontsize{\huge}{22}{26pt}}
-\def\Huge{\@setfontsize{\Huge}{24}{28pt}}
-\fi
-
-
-% V1.6 The Computer Modern Fonts will issue a substitution warning for
-% 24pt titles (24.88pt is used instead) increase the substitution
-% tolerance to turn off this warning
-\def\fontsubfuzz{.9pt}
-% However, the default (and correct) Times font will scale exactly as needed.
-
-
-% warn the user in case they forget to use the 9pt option with
-% technote
-\ifCLASSOPTIONtechnote%
- \ifx\CLASSOPTIONpt\@IEEEptsizenine\else%
- \typeout{** ATTENTION: Technotes are normally 9pt documents.}%
- \fi%
-\fi
-
-
-% V1.7
-% Improved \textunderscore to provide a much better fake _ when used with
-% OT1 encoding. Under OT1, detect use of pcr or cmtt \ttfamily and use
-% available true _ glyph for those two typewriter fonts.
-\def\@IEEEstringptm{ptm} % Times Roman family
-\def\@IEEEstringppl{ppl} % Palatino Roman family
-\def\@IEEEstringphv{phv} % Helvetica Sans Serif family
-\def\@IEEEstringpcr{pcr} % Courier typewriter family
-\def\@IEEEstringcmtt{cmtt} % Computer Modern typewriter family
-\DeclareTextCommandDefault{\textunderscore}{\leavevmode
-\ifx\f@family\@IEEEstringpcr\string_\else
-\ifx\f@family\@IEEEstringcmtt\string_\else
-\ifx\f@family\@IEEEstringptm\kern 0em\vbox{\hrule\@width 0.5em\@height 0.5pt\kern -0.3ex}\else
-\ifx\f@family\@IEEEstringppl\kern 0em\vbox{\hrule\@width 0.5em\@height 0.5pt\kern -0.3ex}\else
-\ifx\f@family\@IEEEstringphv\kern -0.03em\vbox{\hrule\@width 0.62em\@height 0.52pt\kern -0.33ex}\kern -0.03em\else
-\kern 0.09em\vbox{\hrule\@width 0.6em\@height 0.44pt\kern -0.63pt\kern -0.42ex}\kern 0.09em\fi\fi\fi\fi\fi\relax}
-
-
-
-
-% set the default \baselinestretch
-\def\baselinestretch{1}
-\ifCLASSOPTIONdraftcls
- \def\baselinestretch{1.5}% default baselinestretch for draft modes
-\fi
-
-
-% process CLASSINPUT baselinestretch
-\ifx\CLASSINPUTbaselinestretch\@IEEEundefined
-\else
- \edef\baselinestretch{\CLASSINPUTbaselinestretch} % user CLASSINPUT override
- \typeout{** ATTENTION: Overriding \string\baselinestretch\space to
- \baselinestretch\space via \string\CLASSINPUT.}
-\fi
-
-\normalsize % make \baselinestretch take affect
-
-
-
-
-% store the normalsize baselineskip
-\newdimen\CLASSINFOnormalsizebaselineskip
-\CLASSINFOnormalsizebaselineskip=\baselineskip\relax
-% and the normalsize unity (baselinestretch=1) baselineskip
-% we could save a register by giving the user access to
-% \@IEEEnormalsizeunitybaselineskip. However, let's protect
-% its read only internal status
-\newdimen\CLASSINFOnormalsizeunitybaselineskip
-\CLASSINFOnormalsizeunitybaselineskip=\@IEEEnormalsizeunitybaselineskip\relax
-% store the nominal value of jot
-\newdimen\IEEEnormaljot
-\IEEEnormaljot=0.25\baselineskip\relax
-
-% set \jot
-\jot=\IEEEnormaljot\relax
-
-
-
-
-% V1.6, we are now going to fine tune the interword spacing
-% The default interword glue for Times under TeX appears to use a
-% nominal interword spacing of 25% (relative to the font size, i.e., 1em)
-% a maximum of 40% and a minimum of 19%.
-% For example, 10pt text uses an interword glue of:
-%
-% 2.5pt plus 1.49998pt minus 0.59998pt
-%
-% However, IEEE allows for a more generous range which reduces the need
-% for hyphenation, especially for two column text. Furthermore, IEEE
-% tends to use a little bit more nominal space between the words.
-% IEEE's interword spacing percentages appear to be:
-% 35% nominal
-% 23% minimum
-% 50% maximum
-% (They may even be using a tad more for the largest fonts such as 24pt.)
-%
-% for bold text, IEEE increases the spacing a little more:
-% 37.5% nominal
-% 23% minimum
-% 55% maximum
-
-% here are the interword spacing ratios we'll use
-% for medium (normal weight)
-\def\@IEEEinterspaceratioM{0.35}
-\def\@IEEEinterspaceMINratioM{0.23}
-\def\@IEEEinterspaceMAXratioM{0.50}
-
-% for bold
-\def\@IEEEinterspaceratioB{0.375}
-\def\@IEEEinterspaceMINratioB{0.23}
-\def\@IEEEinterspaceMAXratioB{0.55}
-
-
-% command to revise the interword spacing for the current font under TeX:
-% \fontdimen2 = nominal interword space
-% \fontdimen3 = interword stretch
-% \fontdimen4 = interword shrink
-% since all changes to the \fontdimen are global, we can enclose these commands
-% in braces to confine any font attribute or length changes
-\def\@@@IEEEsetfontdimens#1#2#3{{%
-\setlength{\@IEEEtrantmpdimenB}{\f@size pt}% grab the font size in pt, could use 1em instead.
-\setlength{\@IEEEtrantmpdimenA}{#1\@IEEEtrantmpdimenB}%
-\fontdimen2\font=\@IEEEtrantmpdimenA\relax
-\addtolength{\@IEEEtrantmpdimenA}{-#2\@IEEEtrantmpdimenB}%
-\fontdimen3\font=-\@IEEEtrantmpdimenA\relax
-\setlength{\@IEEEtrantmpdimenA}{#1\@IEEEtrantmpdimenB}%
-\addtolength{\@IEEEtrantmpdimenA}{-#3\@IEEEtrantmpdimenB}%
-\fontdimen4\font=\@IEEEtrantmpdimenA\relax}}
-
-% revise the interword spacing for each font weight
-\def\@@IEEEsetfontdimens{{%
-\mdseries
-\@@@IEEEsetfontdimens{\@IEEEinterspaceratioM}{\@IEEEinterspaceMAXratioM}{\@IEEEinterspaceMINratioM}%
-\bfseries
-\@@@IEEEsetfontdimens{\@IEEEinterspaceratioB}{\@IEEEinterspaceMAXratioB}{\@IEEEinterspaceMINratioB}%
-}}
-
-% revise the interword spacing for each font shape
-% \slshape is not often used for IEEE work and is not altered here. The \scshape caps are
-% already a tad too large in the free LaTeX fonts (as compared to what IEEE uses) so we
-% won't alter these either.
-\def\@IEEEsetfontdimens{{%
-\normalfont
-\@@IEEEsetfontdimens
-\normalfont\itshape
-\@@IEEEsetfontdimens
-}}
-
-% command to revise the interword spacing for each font size (and shape
-% and weight). Only the \rmfamily is done here as \ttfamily uses a
-% fixed spacing and \sffamily is not used as the main text of IEEE papers.
-\def\@IEEEtunefonts{{\selectfont\rmfamily
-\tiny\@IEEEsetfontdimens
-\scriptsize\@IEEEsetfontdimens
-\footnotesize\@IEEEsetfontdimens
-\small\@IEEEsetfontdimens
-\normalsize\@IEEEsetfontdimens
-\sublargesize\@IEEEsetfontdimens
-\large\@IEEEsetfontdimens
-\LARGE\@IEEEsetfontdimens
-\huge\@IEEEsetfontdimens
-\Huge\@IEEEsetfontdimens}}
-
-% if the nofonttune class option is not given, revise the interword spacing
-% now - in case IEEEtran makes any default length measurements, and make
-% sure all the default fonts are loaded
-\ifCLASSOPTIONnofonttune\else
-\@IEEEtunefonts
-\fi
-
-% and again at the start of the document in case the user loaded different fonts
-\AtBeginDocument{\ifCLASSOPTIONnofonttune\else\@IEEEtunefonts\fi}
-
-
-
-% V1.6
-% LaTeX is a little to quick to use hyphenations
-% So, we increase the penalty for their use and raise
-% the badness level that triggers an underfull hbox
-% warning. The author may still have to tweak things,
-% but the appearance will be much better "right out
-% of the box" than that under V1.5 and prior.
-% TeX default is 50
-\hyphenpenalty=750
-% If we didn't adjust the interword spacing, 2200 might be better.
-% The TeX default is 1000
-\hbadness=1350
-% IEEE does not use extra spacing after punctuation
-\frenchspacing
-
-% V1.7 increase this a tad to discourage equation breaks
-\binoppenalty=1000 % default 700
-\relpenalty=800 % default 500
-
-
-% margin note stuff
-\marginparsep 10pt
-\marginparwidth 20pt
-\marginparpush 25pt
-
-
-% if things get too close, go ahead and let them touch
-\lineskip 0pt
-\normallineskip 0pt
-\lineskiplimit 0pt
-\normallineskiplimit 0pt
-
-% The distance from the lower edge of the text body to the
-% footline
-\footskip 0.4in
-
-% normally zero, should be relative to font height.
-% put in a little rubber to help stop some bad breaks (underfull vboxes)
-\parskip 0ex plus 0.2ex minus 0.1ex
-
-\parindent 1.0em
-
-\topmargin -49.0pt
-\headheight 12pt
-\headsep 0.25in
-
-% use the normal font baselineskip
-% so that \topskip is unaffected by changes in \baselinestretch
-\topskip=\@IEEEnormalsizeunitybaselineskip
-\textheight 58pc % 9.63in, 696pt
-% Tweak textheight to a perfect integer number of lines/page.
-% The normal baselineskip for each document point size is used
-% to determine these values.
-\ifx\CLASSOPTIONpt\@IEEEptsizenine\textheight=63\@IEEEnormalsizeunitybaselineskip\fi % 63 lines/page
-\ifx\CLASSOPTIONpt\@IEEEptsizeten\textheight=58\@IEEEnormalsizeunitybaselineskip\fi % 58 lines/page
-\ifx\CLASSOPTIONpt\@IEEEptsizeeleven\textheight=52\@IEEEnormalsizeunitybaselineskip\fi % 52 lines/page
-\ifx\CLASSOPTIONpt\@IEEEptsizetwelve\textheight=50\@IEEEnormalsizeunitybaselineskip\fi % 50 lines/page
-
-
-\columnsep 1pc
-\textwidth 43pc % 2 x 21pc + 1pc = 43pc
-
-
-% the default side margins are equal
-\if@IEEEusingAfourpaper
-\oddsidemargin 19.05mm
-\evensidemargin 19.05mm
-\else
-\oddsidemargin 0.680in
-\evensidemargin 0.680in
-\fi
-% compensate for LaTeX's 1in offset
-\addtolength{\oddsidemargin}{-1in}
-\addtolength{\evensidemargin}{-1in}
-
-
-
-% adjust margins for conference mode
-\ifCLASSOPTIONconference
- \topmargin -0.25in
- % we retain the reserved, but unused space for headers
- \addtolength{\topmargin}{-\headheight}
- \addtolength{\topmargin}{-\headsep}
- \textheight 9.25in % The standard for conferences (668.4975pt)
- % Tweak textheight to a perfect integer number of lines/page.
- \ifx\CLASSOPTIONpt\@IEEEptsizenine\textheight=61\@IEEEnormalsizeunitybaselineskip\fi % 61 lines/page
- \ifx\CLASSOPTIONpt\@IEEEptsizeten\textheight=56\@IEEEnormalsizeunitybaselineskip\fi % 56 lines/page
- \ifx\CLASSOPTIONpt\@IEEEptsizeeleven\textheight=50\@IEEEnormalsizeunitybaselineskip\fi % 50 lines/page
- \ifx\CLASSOPTIONpt\@IEEEptsizetwelve\textheight=48\@IEEEnormalsizeunitybaselineskip\fi % 48 lines/page
-\fi
-
-
-% compsoc conference
-\ifCLASSOPTIONcompsocconf
-\ifCLASSOPTIONconference
- % compsoc conference use a larger value for columnsep
- \columnsep 0.25in
- % compsoc conferences want 1in top margin, 1.125in bottom margin
- \topmargin 0in
- %\addtolength{\topmargin}{-6pt}% we tweak this a tad to better comply with top of line stuff
- % we retain the reserved, but unused space for headers
- \addtolength{\topmargin}{-\headheight}
- \addtolength{\topmargin}{-\headsep}
- \textheight 9.0in % (641.39625pt)
-
- % Tweak textheight to a perfect integer number of lines/page.
- \ifx\CLASSOPTIONpt\@IEEEptsizenine\textheight=58\@IEEEnormalsizeunitybaselineskip\fi % 58 lines/page
- \ifx\CLASSOPTIONpt\@IEEEptsizeten\textheight=54\@IEEEnormalsizeunitybaselineskip\fi % 54 lines/page
- \ifx\CLASSOPTIONpt\@IEEEptsizeeleven\textheight=48\@IEEEnormalsizeunitybaselineskip\fi % 48 lines/page
- \ifx\CLASSOPTIONpt\@IEEEptsizetwelve\textheight=46\@IEEEnormalsizeunitybaselineskip\fi % 46 lines/page
- \textwidth 7in
-
-
- %adjust text h/w for A4 paper
- \if@IEEEusingAfourpaper
- \textheight 9.69in
- \textwidth 6.77in
- \fi
-
- % the default side margins are equal
- \if@IEEEusingAfourpaper
- \oddsidemargin 19.05mm
- \evensidemargin 19.05mm
- \else
- \oddsidemargin 0.75in
- \evensidemargin 0.75in
- \fi
- % compensate for LaTeX's 1in offset
- \addtolength{\oddsidemargin}{-1in}
- \addtolength{\evensidemargin}{-1in}
-\fi\fi
-
-
-
-% draft mode settings override that of all other modes
-% provides a nice 1in margin all around the paper and extra
-% space between the lines for editor's comments
-\ifCLASSOPTIONdraftcls
- % want 1in from top of paper to text
- \setlength{\topmargin}{-\headsep}%
- \addtolength{\topmargin}{-\headheight}%
- % we want 1in side margins regardless of paper type
- \oddsidemargin 0in
- \evensidemargin 0in
- % set the text width
- \setlength{\textwidth}{\paperwidth}%
- \addtolength{\textwidth}{-2.0in}%
- \setlength{\textheight}{\paperheight}%
- \addtolength{\textheight}{-2.0in}%
- % digitize textheight to be an integer number of lines.
- % this may cause the bottom margin to be off a tad
- \addtolength{\textheight}{-1\topskip}%
- \divide\textheight by \baselineskip%
- \multiply\textheight by \baselineskip%
- \addtolength{\textheight}{\topskip}%
-\fi
-
-
-
-% process CLASSINPUT inner/outer margin
-% if inner margin defined, but outer margin not, set outer to inner.
-\ifx\CLASSINPUTinnersidemargin\@IEEEundefined
-\else
- \ifx\CLASSINPUToutersidemargin\@IEEEundefined
- \edef\CLASSINPUToutersidemargin{\CLASSINPUTinnersidemargin}
- \fi
-\fi
-
-\ifx\CLASSINPUToutersidemargin\@IEEEundefined
-\else
- % if outer margin defined, but inner margin not, set inner to outer.
- \ifx\CLASSINPUTinnersidemargin\@IEEEundefined
- \edef\CLASSINPUTinnersidemargin{\CLASSINPUToutersidemargin}
- \fi
- \setlength{\oddsidemargin}{\CLASSINPUTinnersidemargin}
- \ifCLASSOPTIONtwoside
- \setlength{\evensidemargin}{\CLASSINPUToutersidemargin}
- \else
- \setlength{\evensidemargin}{\CLASSINPUTinnersidemargin}
- \fi
- \addtolength{\oddsidemargin}{-1in}
- \addtolength{\evensidemargin}{-1in}
- \setlength{\textwidth}{\paperwidth}
- \addtolength{\textwidth}{-\CLASSINPUTinnersidemargin}
- \addtolength{\textwidth}{-\CLASSINPUToutersidemargin}
- \typeout{** ATTENTION: Overriding inner side margin to \CLASSINPUTinnersidemargin\space and
- outer side margin to \CLASSINPUToutersidemargin\space via \string\CLASSINPUT.}
-\fi
-
-
-
-% process CLASSINPUT top/bottom text margin
-% if toptext margin defined, but bottomtext margin not, set bottomtext to toptext margin
-\ifx\CLASSINPUTtoptextmargin\@IEEEundefined
-\else
- \ifx\CLASSINPUTbottomtextmargin\@IEEEundefined
- \edef\CLASSINPUTbottomtextmargin{\CLASSINPUTtoptextmargin}
- \fi
-\fi
-
-\ifx\CLASSINPUTbottomtextmargin\@IEEEundefined
-\else
- % if bottomtext margin defined, but toptext margin not, set toptext to bottomtext margin
- \ifx\CLASSINPUTtoptextmargin\@IEEEundefined
- \edef\CLASSINPUTtoptextmargin{\CLASSINPUTbottomtextmargin}
- \fi
- \setlength{\topmargin}{\CLASSINPUTtoptextmargin}
- \addtolength{\topmargin}{-1in}
- \addtolength{\topmargin}{-\headheight}
- \addtolength{\topmargin}{-\headsep}
- \setlength{\textheight}{\paperheight}
- \addtolength{\textheight}{-\CLASSINPUTtoptextmargin}
- \addtolength{\textheight}{-\CLASSINPUTbottomtextmargin}
- % in the default format we use the normal baselineskip as topskip
- % we only need 0.7 of this to clear typical top text and we need
- % an extra 0.3 spacing at the bottom for descenders. This will
- % correct for both.
- \addtolength{\topmargin}{-0.3\@IEEEnormalsizeunitybaselineskip}
- \typeout{** ATTENTION: Overriding top text margin to \CLASSINPUTtoptextmargin\space and
- bottom text margin to \CLASSINPUTbottomtextmargin\space via \string\CLASSINPUT.}
-\fi
-
-
-
-
-
-
-
-% LIST SPACING CONTROLS
-
-% Controls the amount of EXTRA spacing
-% above and below \trivlist
-% Both \list and IED lists override this.
-% However, \trivlist will use this as will most
-% things built from \trivlist like the \center
-% environment.
-\topsep 0.5\baselineskip
-
-% Controls the additional spacing around lists preceded
-% or followed by blank lines. IEEE does not increase
-% spacing before or after paragraphs so it is set to zero.
-% \z@ is the same as zero, but faster.
-\partopsep \z@
-
-% Controls the spacing between paragraphs in lists.
-% IEEE does not increase spacing before or after paragraphs
-% so this is also zero.
-% With IEEEtran.cls, global changes to
-% this value DO affect lists (but not IED lists).
-\parsep \z@
-
-% Controls the extra spacing between list items.
-% IEEE does not put extra spacing between items.
-% With IEEEtran.cls, global changes to this value DO affect
-% lists (but not IED lists).
-\itemsep \z@
-
-% \itemindent is the amount to indent the FIRST line of a list
-% item. It is auto set to zero within the \list environment. To alter
-% it, you have to do so when you call the \list.
-% However, IEEE uses this for the theorem environment
-% There is an alternative value for this near \leftmargini below
-\itemindent -1em
-
-% \leftmargin, the spacing from the left margin of the main text to
-% the left of the main body of a list item is set by \list.
-% Hence this statement does nothing for lists.
-% But, quote and verse do use it for indention.
-\leftmargin 2em
-
-% we retain this stuff from the older IEEEtran.cls so that \list
-% will work the same way as before. However, itemize, enumerate and
-% description (IED) could care less about what these are as they
-% all are overridden.
-\leftmargini 2em
-%\itemindent 2em % Alternative values: sometimes used.
-%\leftmargini 0em
-\leftmarginii 1em
-\leftmarginiii 1.5em
-\leftmarginiv 1.5em
-\leftmarginv 1.0em
-\leftmarginvi 1.0em
-\labelsep 0.5em
-\labelwidth \z@
-
-
-% The old IEEEtran.cls behavior of \list is retained.
-% However, the new V1.3 IED list environments override all the
-% @list stuff (\@listX is called within \list for the
-% appropriate level just before the user's list_decl is called).
-% \topsep is now 2pt as IEEE puts a little extra space around
-% lists - used by those non-IED macros that depend on \list.
-% Note that \parsep and \itemsep are not redefined as in
-% the sizexx.clo \@listX (which article.cls uses) so global changes
-% of these values DO affect \list
-%
-\def\@listi{\leftmargin\leftmargini \topsep 2pt plus 1pt minus 1pt}
-\let\@listI\@listi
-\def\@listii{\leftmargin\leftmarginii\labelwidth\leftmarginii%
- \advance\labelwidth-\labelsep \topsep 2pt}
-\def\@listiii{\leftmargin\leftmarginiii\labelwidth\leftmarginiii%
- \advance\labelwidth-\labelsep \topsep 2pt}
-\def\@listiv{\leftmargin\leftmarginiv\labelwidth\leftmarginiv%
- \advance\labelwidth-\labelsep \topsep 2pt}
-\def\@listv{\leftmargin\leftmarginv\labelwidth\leftmarginv%
- \advance\labelwidth-\labelsep \topsep 2pt}
-\def\@listvi{\leftmargin\leftmarginvi\labelwidth\leftmarginvi%
- \advance\labelwidth-\labelsep \topsep 2pt}
-
-
-% IEEE uses 5) not 5.
-\def\labelenumi{\theenumi)} \def\theenumi{\arabic{enumi}}
-
-% IEEE uses a) not (a)
-\def\labelenumii{\theenumii)} \def\theenumii{\alph{enumii}}
-
-% IEEE uses iii) not iii.
-\def\labelenumiii{\theenumiii)} \def\theenumiii{\roman{enumiii}}
-
-% IEEE uses A) not A.
-\def\labelenumiv{\theenumiv)} \def\theenumiv{\Alph{enumiv}}
-
-% exactly the same as in article.cls
-\def\p@enumii{\theenumi}
-\def\p@enumiii{\theenumi(\theenumii)}
-\def\p@enumiv{\p@enumiii\theenumiii}
-
-% itemized list label styles
-\def\labelitemi{$\scriptstyle\bullet$}
-\def\labelitemii{\textbf{--}}
-\def\labelitemiii{$\ast$}
-\def\labelitemiv{$\cdot$}
-
-
-
-% **** V1.3 ENHANCEMENTS ****
-% Itemize, Enumerate and Description (IED) List Controls
-% ***************************
-%
-%
-% IEEE seems to use at least two different values by
-% which ITEMIZED list labels are indented to the right
-% For The Journal of Lightwave Technology (JLT) and The Journal
-% on Selected Areas in Communications (JSAC), they tend to use
-% an indention equal to \parindent. For Transactions on Communications
-% they tend to indent ITEMIZED lists a little more--- 1.3\parindent.
-% We'll provide both values here for you so that you can choose
-% which one you like in your document using a command such as:
-% setlength{\IEEEilabelindent}{\IEEEilabelindentB}
-\newdimen\IEEEilabelindentA
-\IEEEilabelindentA \parindent
-
-\newdimen\IEEEilabelindentB
-\IEEEilabelindentB 1.3\parindent
-% However, we'll default to using \parindent
-% which makes more sense to me
-\newdimen\IEEEilabelindent
-\IEEEilabelindent \IEEEilabelindentA
-
-
-% This controls the default amount the enumerated list labels
-% are indented to the right.
-% Normally, this is the same as the paragraph indention
-\newdimen\IEEEelabelindent
-\IEEEelabelindent \parindent
-
-% This controls the default amount the description list labels
-% are indented to the right.
-% Normally, this is the same as the paragraph indention
-\newdimen\IEEEdlabelindent
-\IEEEdlabelindent \parindent
-
-% This is the value actually used within the IED lists.
-% The IED environments automatically set its value to
-% one of the three values above, so global changes do
-% not have any effect
-\newdimen\IEEElabelindent
-\IEEElabelindent \parindent
-
-% The actual amount labels will be indented is
-% \IEEElabelindent multiplied by the factor below
-% corresponding to the level of nesting depth
-% This provides a means by which the user can
-% alter the effective \IEEElabelindent for deeper
-% levels
-% There may not be such a thing as correct "standard IEEE"
-% values. What IEEE actually does may depend on the specific
-% circumstances.
-% The first list level almost always has full indention.
-% The second levels I've seen have only 75% of the normal indentation
-% Three level or greater nestings are very rare. I am guessing
-% that they don't use any indentation.
-\def\IEEElabelindentfactori{1.0} % almost always one
-\def\IEEElabelindentfactorii{0.75} % 0.0 or 1.0 may be used in some cases
-\def\IEEElabelindentfactoriii{0.0} % 0.75? 0.5? 0.0?
-\def\IEEElabelindentfactoriv{0.0}
-\def\IEEElabelindentfactorv{0.0}
-\def\IEEElabelindentfactorvi{0.0}
-
-% value actually used within IED lists, it is auto
-% set to one of the 6 values above
-% global changes here have no effect
-\def\IEEElabelindentfactor{1.0}
-
-% This controls the default spacing between the end of the IED
-% list labels and the list text, when normal text is used for
-% the labels.
-\newdimen\IEEEiednormlabelsep
-\IEEEiednormlabelsep 0.6em
-
-% This controls the default spacing between the end of the IED
-% list labels and the list text, when math symbols are used for
-% the labels (nomenclature lists). IEEE usually increases the
-% spacing in these cases
-\newdimen\IEEEiedmathlabelsep
-\IEEEiedmathlabelsep 1.2em
-
-% This controls the extra vertical separation put above and
-% below each IED list. IEEE usually puts a little extra spacing
-% around each list. However, this spacing is barely noticeable.
-\newskip\IEEEiedtopsep
-\IEEEiedtopsep 2pt plus 1pt minus 1pt
-
-
-% This command is executed within each IED list environment
-% at the beginning of the list. You can use this to set the
-% parameters for some/all your IED list(s) without disturbing
-% global parameters that affect things other than lists.
-% i.e., renewcommand{\IEEEiedlistdecl}{\setlength{\labelsep}{5em}}
-% will alter the \labelsep for the next list(s) until
-% \IEEEiedlistdecl is redefined.
-\def\IEEEiedlistdecl{\relax}
-
-% This command provides an easy way to set \leftmargin based
-% on the \labelwidth, \labelsep and the argument \IEEElabelindent
-% Usage: \IEEEcalcleftmargin{width-to-indent-the-label}
-% output is in the \leftmargin variable, i.e., effectively:
-% \leftmargin = argument + \labelwidth + \labelsep
-% Note controlled spacing here, shield end of lines with %
-\def\IEEEcalcleftmargin#1{\setlength{\leftmargin}{#1}%
-\addtolength{\leftmargin}{\labelwidth}%
-\addtolength{\leftmargin}{\labelsep}}
-
-% This command provides an easy way to set \labelwidth to the
-% width of the given text. It is the same as
-% \settowidth{\labelwidth}{label-text}
-% and useful as a shorter alternative.
-% Typically used to set \labelwidth to be the width
-% of the longest label in the list
-\def\IEEEsetlabelwidth#1{\settowidth{\labelwidth}{#1}}
-
-% When this command is executed, IED lists will use the
-% IEEEiedmathlabelsep label separation rather than the normal
-% spacing. To have an effect, this command must be executed via
-% the \IEEEiedlistdecl or within the option of the IED list
-% environments.
-\def\IEEEusemathlabelsep{\setlength{\labelsep}{\IEEEiedmathlabelsep}}
-
-% A flag which controls whether the IED lists automatically
-% calculate \leftmargin from \IEEElabelindent, \labelwidth and \labelsep
-% Useful if you want to specify your own \leftmargin
-% This flag must be set (\IEEEnocalcleftmargintrue or \IEEEnocalcleftmarginfalse)
-% via the \IEEEiedlistdecl or within the option of the IED list
-% environments to have an effect.
-\newif\ifIEEEnocalcleftmargin
-\IEEEnocalcleftmarginfalse
-
-% A flag which controls whether \IEEElabelindent is multiplied by
-% the \IEEElabelindentfactor for each list level.
-% This flag must be set via the \IEEEiedlistdecl or within the option
-% of the IED list environments to have an effect.
-\newif\ifIEEEnolabelindentfactor
-\IEEEnolabelindentfactorfalse
-
-
-% internal variable to indicate type of IED label
-% justification
-% 0 - left; 1 - center; 2 - right
-\def\@IEEEiedjustify{0}
-
-
-% commands to allow the user to control IED
-% label justifications. Use these commands within
-% the IED environment option or in the \IEEEiedlistdecl
-% Note that changing the normal list justifications
-% is nonstandard and IEEE may not like it if you do so!
-% I include these commands as they may be helpful to
-% those who are using these enhanced list controls for
-% other non-IEEE related LaTeX work.
-% itemize and enumerate automatically default to right
-% justification, description defaults to left.
-\def\IEEEiedlabeljustifyl{\def\@IEEEiedjustify{0}}%left
-\def\IEEEiedlabeljustifyc{\def\@IEEEiedjustify{1}}%center
-\def\IEEEiedlabeljustifyr{\def\@IEEEiedjustify{2}}%right
-
-
-
-
-% commands to save to and restore from the list parameter copies
-% this allows us to set all the list parameters within
-% the list_decl and prevent \list (and its \@list)
-% from overriding any of our parameters
-% V1.6 use \edefs instead of dimen's to conserve dimen registers
-% Note controlled spacing here, shield end of lines with %
-\def\@IEEEsavelistparams{\edef\@IEEEiedtopsep{\the\topsep}%
-\edef\@IEEEiedlabelwidth{\the\labelwidth}%
-\edef\@IEEEiedlabelsep{\the\labelsep}%
-\edef\@IEEEiedleftmargin{\the\leftmargin}%
-\edef\@IEEEiedpartopsep{\the\partopsep}%
-\edef\@IEEEiedparsep{\the\parsep}%
-\edef\@IEEEieditemsep{\the\itemsep}%
-\edef\@IEEEiedrightmargin{\the\rightmargin}%
-\edef\@IEEEiedlistparindent{\the\listparindent}%
-\edef\@IEEEieditemindent{\the\itemindent}}
-
-% Note controlled spacing here
-\def\@IEEErestorelistparams{\topsep\@IEEEiedtopsep\relax%
-\labelwidth\@IEEEiedlabelwidth\relax%
-\labelsep\@IEEEiedlabelsep\relax%
-\leftmargin\@IEEEiedleftmargin\relax%
-\partopsep\@IEEEiedpartopsep\relax%
-\parsep\@IEEEiedparsep\relax%
-\itemsep\@IEEEieditemsep\relax%
-\rightmargin\@IEEEiedrightmargin\relax%
-\listparindent\@IEEEiedlistparindent\relax%
-\itemindent\@IEEEieditemindent\relax}
-
-
-% v1.6b provide original LaTeX IED list environments
-% note that latex.ltx defines \itemize and \enumerate, but not \description
-% which must be created by the base classes
-% save original LaTeX itemize and enumerate
-\let\LaTeXitemize\itemize
-\let\endLaTeXitemize\enditemize
-\let\LaTeXenumerate\enumerate
-\let\endLaTeXenumerate\endenumerate
-
-% provide original LaTeX description environment from article.cls
-\newenvironment{LaTeXdescription}
- {\list{}{\labelwidth\z@ \itemindent-\leftmargin
- \let\makelabel\descriptionlabel}}
- {\endlist}
-\newcommand*\descriptionlabel[1]{\hspace\labelsep
- \normalfont\bfseries #1}
-
-
-% override LaTeX's default IED lists
-\def\itemize{\@IEEEitemize}
-\def\enditemize{\@endIEEEitemize}
-\def\enumerate{\@IEEEenumerate}
-\def\endenumerate{\@endIEEEenumerate}
-\def\description{\@IEEEdescription}
-\def\enddescription{\@endIEEEdescription}
-
-% provide the user with aliases - may help those using packages that
-% override itemize, enumerate, or description
-\def\IEEEitemize{\@IEEEitemize}
-\def\endIEEEitemize{\@endIEEEitemize}
-\def\IEEEenumerate{\@IEEEenumerate}
-\def\endIEEEenumerate{\@endIEEEenumerate}
-\def\IEEEdescription{\@IEEEdescription}
-\def\endIEEEdescription{\@endIEEEdescription}
-
-
-% V1.6 we want to keep the IEEEtran IED list definitions as our own internal
-% commands so they are protected against redefinition
-\def\@IEEEitemize{\@ifnextchar[{\@@IEEEitemize}{\@@IEEEitemize[\relax]}}
-\def\@IEEEenumerate{\@ifnextchar[{\@@IEEEenumerate}{\@@IEEEenumerate[\relax]}}
-\def\@IEEEdescription{\@ifnextchar[{\@@IEEEdescription}{\@@IEEEdescription[\relax]}}
-\def\@endIEEEitemize{\endlist}
-\def\@endIEEEenumerate{\endlist}
-\def\@endIEEEdescription{\endlist}
-
-
-% DO NOT ALLOW BLANK LINES TO BE IN THESE IED ENVIRONMENTS
-% AS THIS WILL FORCE NEW PARAGRAPHS AFTER THE IED LISTS
-% IEEEtran itemized list MDS 1/2001
-% Note controlled spacing here, shield end of lines with %
-\def\@@IEEEitemize[#1]{%
- \ifnum\@itemdepth>3\relax\@toodeep\else%
- \ifnum\@listdepth>5\relax\@toodeep\else%
- \advance\@itemdepth\@ne%
- \edef\@itemitem{labelitem\romannumeral\the\@itemdepth}%
- % get the labelindentfactor for this level
- \advance\@listdepth\@ne% we need to know what the level WILL be
- \edef\IEEElabelindentfactor{\csname IEEElabelindentfactor\romannumeral\the\@listdepth\endcsname}%
- \advance\@listdepth-\@ne% undo our increment
- \def\@IEEEiedjustify{2}% right justified labels are default
- % set other defaults
- \IEEEnocalcleftmarginfalse%
- \IEEEnolabelindentfactorfalse%
- \topsep\IEEEiedtopsep%
- \IEEElabelindent\IEEEilabelindent%
- \labelsep\IEEEiednormlabelsep%
- \partopsep 0ex%
- \parsep 0ex%
- \itemsep 0ex%
- \rightmargin 0em%
- \listparindent 0em%
- \itemindent 0em%
- % calculate the label width
- % the user can override this later if
- % they specified a \labelwidth
- \settowidth{\labelwidth}{\csname labelitem\romannumeral\the\@itemdepth\endcsname}%
- \@IEEEsavelistparams% save our list parameters
- \list{\csname\@itemitem\endcsname}{%
- \@IEEErestorelistparams% override any list{} changes
- % to our globals
- \let\makelabel\@IEEEiedmakelabel% v1.6b setup \makelabel
- \IEEEiedlistdecl% let user alter parameters
- #1\relax%
- % If the user has requested not to use the
- % labelindent factor, don't revise \labelindent
- \ifIEEEnolabelindentfactor\relax%
- \else\IEEElabelindent=\IEEElabelindentfactor\labelindent%
- \fi%
- % Unless the user has requested otherwise,
- % calculate our left margin based
- % on \IEEElabelindent, \labelwidth and
- % \labelsep
- \ifIEEEnocalcleftmargin\relax%
- \else\IEEEcalcleftmargin{\IEEElabelindent}%
- \fi}\fi\fi}%
-
-
-% DO NOT ALLOW BLANK LINES TO BE IN THESE IED ENVIRONMENTS
-% AS THIS WILL FORCE NEW PARAGRAPHS AFTER THE IED LISTS
-% IEEEtran enumerate list MDS 1/2001
-% Note controlled spacing here, shield end of lines with %
-\def\@@IEEEenumerate[#1]{%
- \ifnum\@enumdepth>3\relax\@toodeep\else%
- \ifnum\@listdepth>5\relax\@toodeep\else%
- \advance\@enumdepth\@ne%
- \edef\@enumctr{enum\romannumeral\the\@enumdepth}%
- % get the labelindentfactor for this level
- \advance\@listdepth\@ne% we need to know what the level WILL be
- \edef\IEEElabelindentfactor{\csname IEEElabelindentfactor\romannumeral\the\@listdepth\endcsname}%
- \advance\@listdepth-\@ne% undo our increment
- \def\@IEEEiedjustify{2}% right justified labels are default
- % set other defaults
- \IEEEnocalcleftmarginfalse%
- \IEEEnolabelindentfactorfalse%
- \topsep\IEEEiedtopsep%
- \IEEElabelindent\IEEEelabelindent%
- \labelsep\IEEEiednormlabelsep%
- \partopsep 0ex%
- \parsep 0ex%
- \itemsep 0ex%
- \rightmargin 0em%
- \listparindent 0em%
- \itemindent 0em%
- % calculate the label width
- % We'll set it to the width suitable for all labels using
- % normalfont 1) to 9)
- % The user can override this later
- \settowidth{\labelwidth}{9)}%
- \@IEEEsavelistparams% save our list parameters
- \list{\csname label\@enumctr\endcsname}{\usecounter{\@enumctr}%
- \@IEEErestorelistparams% override any list{} changes
- % to our globals
- \let\makelabel\@IEEEiedmakelabel% v1.6b setup \makelabel
- \IEEEiedlistdecl% let user alter parameters
- #1\relax%
- % If the user has requested not to use the
- % IEEElabelindent factor, don't revise \IEEElabelindent
- \ifIEEEnolabelindentfactor\relax%
- \else\IEEElabelindent=\IEEElabelindentfactor\IEEElabelindent%
- \fi%
- % Unless the user has requested otherwise,
- % calculate our left margin based
- % on \IEEElabelindent, \labelwidth and
- % \labelsep
- \ifIEEEnocalcleftmargin\relax%
- \else\IEEEcalcleftmargin{\IEEElabelindent}%
- \fi}\fi\fi}%
-
-
-% DO NOT ALLOW BLANK LINES TO BE IN THESE IED ENVIRONMENTS
-% AS THIS WILL FORCE NEW PARAGRAPHS AFTER THE IED LISTS
-% IEEEtran description list MDS 1/2001
-% Note controlled spacing here, shield end of lines with %
-\def\@@IEEEdescription[#1]{%
- \ifnum\@listdepth>5\relax\@toodeep\else%
- % get the labelindentfactor for this level
- \advance\@listdepth\@ne% we need to know what the level WILL be
- \edef\IEEElabelindentfactor{\csname IEEElabelindentfactor\romannumeral\the\@listdepth\endcsname}%
- \advance\@listdepth-\@ne% undo our increment
- \def\@IEEEiedjustify{0}% left justified labels are default
- % set other defaults
- \IEEEnocalcleftmarginfalse%
- \IEEEnolabelindentfactorfalse%
- \topsep\IEEEiedtopsep%
- \IEEElabelindent\IEEEdlabelindent%
- % assume normal labelsep
- \labelsep\IEEEiednormlabelsep%
- \partopsep 0ex%
- \parsep 0ex%
- \itemsep 0ex%
- \rightmargin 0em%
- \listparindent 0em%
- \itemindent 0em%
- % Bogus label width in case the user forgets
- % to set it.
- % TIP: If you want to see what a variable's width is you
- % can use the TeX command \showthe\width-variable to
- % display it on the screen during compilation
- % (This might be helpful to know when you need to find out
- % which label is the widest)
- \settowidth{\labelwidth}{Hello}%
- \@IEEEsavelistparams% save our list parameters
- \list{}{\@IEEErestorelistparams% override any list{} changes
- % to our globals
- \let\makelabel\@IEEEiedmakelabel% v1.6b setup \makelabel
- \IEEEiedlistdecl% let user alter parameters
- #1\relax%
- % If the user has requested not to use the
- % labelindent factor, don't revise \IEEElabelindent
- \ifIEEEnolabelindentfactor\relax%
- \else\IEEElabelindent=\IEEElabelindentfactor\IEEElabelindent%
- \fi%
- % Unless the user has requested otherwise,
- % calculate our left margin based
- % on \IEEElabelindent, \labelwidth and
- % \labelsep
- \ifIEEEnocalcleftmargin\relax%
- \else\IEEEcalcleftmargin{\IEEElabelindent}\relax%
- \fi}\fi}
-
-% v1.6b we use one makelabel that does justification as needed.
-\def\@IEEEiedmakelabel#1{\relax\if\@IEEEiedjustify 0\relax
-\makebox[\labelwidth][l]{\normalfont #1}\else
-\if\@IEEEiedjustify 1\relax
-\makebox[\labelwidth][c]{\normalfont #1}\else
-\makebox[\labelwidth][r]{\normalfont #1}\fi\fi}
-
-
-% VERSE and QUOTE
-% V1.7 define environments with newenvironment
-\newenvironment{verse}{\let\\=\@centercr
- \list{}{\itemsep\z@ \itemindent -1.5em \listparindent \itemindent
- \rightmargin\leftmargin\advance\leftmargin 1.5em}\item\relax}
- {\endlist}
-\newenvironment{quotation}{\list{}{\listparindent 1.5em \itemindent\listparindent
- \rightmargin\leftmargin \parsep 0pt plus 1pt}\item\relax}
- {\endlist}
-\newenvironment{quote}{\list{}{\rightmargin\leftmargin}\item\relax}
- {\endlist}
-
-
-% \titlepage
-% provided only for backward compatibility. \maketitle is the correct
-% way to create the title page.
-\newif\if@restonecol
-\def\titlepage{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \else \newpage \fi \thispagestyle{empty}\c@page\z@}
-\def\endtitlepage{\if@restonecol\twocolumn \else \newpage \fi}
-
-% standard values from article.cls
-\arraycolsep 5pt
-\arrayrulewidth .4pt
-\doublerulesep 2pt
-
-\tabcolsep 6pt
-\tabbingsep 0.5em
-
-
-%% FOOTNOTES
-%
-%\skip\footins 10pt plus 4pt minus 2pt
-% V1.6 respond to changes in font size
-% space added above the footnotes (if present)
-\skip\footins 0.9\baselineskip plus 0.4\baselineskip minus 0.2\baselineskip
-
-% V1.6, we need to make \footnotesep responsive to changes
-% in \baselineskip or strange spacings will result when in
-% draft mode. Here is a little LaTeX secret - \footnotesep
-% determines the height of an invisible strut that is placed
-% *above* the baseline of footnotes after the first. Since
-% LaTeX considers the space for characters to be 0.7/baselineskip
-% above the baseline and 0.3/baselineskip below it, we need to
-% use 0.7/baselineskip as a \footnotesep to maintain equal spacing
-% between all the lines of the footnotes. IEEE often uses a tad
-% more, so use 0.8\baselineskip. This slightly larger value also helps
-% the text to clear the footnote marks. Note that \thanks in IEEEtran
-% uses its own value of \footnotesep which is set in \maketitle.
-{\footnotesize
-\global\footnotesep 0.8\baselineskip}
-
-
-\skip\@mpfootins = \skip\footins
-\fboxsep = 3pt
-\fboxrule = .4pt
-% V1.6 use 1em, then use LaTeX2e's \@makefnmark
-% Note that IEEE normally *left* aligns the footnote marks, so we don't need
-% box resizing tricks here.
-\long\def\@makefntext#1{\parindent 1em\indent\hbox{\@makefnmark}#1}% V1.6 use 1em
-% V1.7 compsoc does not use superscipts for footnote marks
-\ifCLASSOPTIONcompsoc
-\def\@IEEEcompsocmakefnmark{\hbox{\normalfont\@thefnmark.\ }}
-\long\def\@makefntext#1{\parindent 1em\indent\hbox{\@IEEEcompsocmakefnmark}#1}
-\fi
-
-% IEEE does not use footnote rules
-\def\footnoterule{}
-
-% V1.7 for compsoc, IEEE uses a footnote rule only for \thanks. We devise a "one-shot"
-% system to implement this.
-\newif\if@IEEEenableoneshotfootnoterule
-\@IEEEenableoneshotfootnoterulefalse
-\ifCLASSOPTIONcompsoc
-\def\footnoterule{\relax\if@IEEEenableoneshotfootnoterule
-\kern-5pt
-\hbox to \columnwidth{\hfill\vrule width 0.5\columnwidth height 0.4pt\hfill}
-\kern4.6pt
-\global\@IEEEenableoneshotfootnoterulefalse
-\else
-\relax
-\fi}
-\fi
-
-% V1.6 do not allow LaTeX to break a footnote across multiple pages
-\interfootnotelinepenalty=10000
-
-% V1.6 discourage breaks within equations
-% Note that amsmath normally sets this to 10000,
-% but LaTeX2e normally uses 100.
-\interdisplaylinepenalty=2500
-
-% default allows section depth up to /paragraph
-\setcounter{secnumdepth}{4}
-
-% technotes do not allow /paragraph
-\ifCLASSOPTIONtechnote
- \setcounter{secnumdepth}{3}
-\fi
-% neither do compsoc conferences
-\@IEEEcompsocconfonly{\setcounter{secnumdepth}{3}}
-
-
-\newcounter{section}
-\newcounter{subsection}[section]
-\newcounter{subsubsection}[subsection]
-\newcounter{paragraph}[subsubsection]
-
-% used only by IEEEtran's IEEEeqnarray as other packages may
-% have their own, different, implementations
-\newcounter{IEEEsubequation}[equation]
-
-% as shown when called by user from \ref, \label and in table of contents
-\def\theequation{\arabic{equation}} % 1
-\def\theIEEEsubequation{\theequation\alph{IEEEsubequation}} % 1a (used only by IEEEtran's IEEEeqnarray)
-\ifCLASSOPTIONcompsoc
-% compsoc is all arabic
-\def\thesection{\arabic{section}}
-\def\thesubsection{\thesection.\arabic{subsection}}
-\def\thesubsubsection{\thesubsection.\arabic{subsubsection}}
-\def\theparagraph{\thesubsubsection.\arabic{paragraph}}
-\else
-\def\thesection{\Roman{section}} % I
-% V1.7, \mbox prevents breaks around -
-\def\thesubsection{\mbox{\thesection-\Alph{subsection}}} % I-A
-% V1.7 use I-A1 format used by IEEE rather than I-A.1
-\def\thesubsubsection{\thesubsection\arabic{subsubsection}} % I-A1
-\def\theparagraph{\thesubsubsection\alph{paragraph}} % I-A1a
-\fi
-
-% From Heiko Oberdiek. Because of the \mbox in \thesubsection, we need to
-% tell hyperref to disable the \mbox command when making PDF bookmarks.
-% This done already with hyperref.sty version 6.74o and later, but
-% it will not hurt to do it here again for users of older versions.
-\@ifundefined{pdfstringdefPreHook}{\let\pdfstringdefPreHook\@empty}{}%
-\g@addto@macro\pdfstringdefPreHook{\let\mbox\relax}
-
-
-% Main text forms (how shown in main text headings)
-% V1.6, using \thesection in \thesectiondis allows changes
-% in the former to automatically appear in the latter
-\ifCLASSOPTIONcompsoc
- \ifCLASSOPTIONconference% compsoc conference
- \def\thesectiondis{\thesection.}
- \def\thesubsectiondis{\thesectiondis\arabic{subsection}.}
- \def\thesubsubsectiondis{\thesubsectiondis\arabic{subsubsection}.}
- \def\theparagraphdis{\thesubsubsectiondis\arabic{paragraph}.}
- \else% compsoc not conferencs
- \def\thesectiondis{\thesection}
- \def\thesubsectiondis{\thesectiondis.\arabic{subsection}}
- \def\thesubsubsectiondis{\thesubsectiondis.\arabic{subsubsection}}
- \def\theparagraphdis{\thesubsubsectiondis.\arabic{paragraph}}
- \fi
-\else% not compsoc
- \def\thesectiondis{\thesection.} % I.
- \def\thesubsectiondis{\Alph{subsection}.} % B.
- \def\thesubsubsectiondis{\arabic{subsubsection})} % 3)
- \def\theparagraphdis{\alph{paragraph})} % d)
-\fi
-
-% just like LaTeX2e's \@eqnnum
-\def\theequationdis{{\normalfont \normalcolor (\theequation)}}% (1)
-% IEEEsubequation used only by IEEEtran's IEEEeqnarray
-\def\theIEEEsubequationdis{{\normalfont \normalcolor (\theIEEEsubequation)}}% (1a)
-% redirect LaTeX2e's equation number display and all that depend on
-% it, through IEEEtran's \theequationdis
-\def\@eqnnum{\theequationdis}
-
-
-
-% V1.7 provide string macros as article.cls does
-\def\contentsname{Contents}
-\def\listfigurename{List of Figures}
-\def\listtablename{List of Tables}
-\def\refname{References}
-\def\indexname{Index}
-\def\figurename{Fig.}
-\def\tablename{TABLE}
-\@IEEEcompsocconfonly{\def\figurename{Figure}\def\tablename{Table}}
-\def\partname{Part}
-\def\appendixname{Appendix}
-\def\abstractname{Abstract}
-% IEEE specific names
-\def\IEEEkeywordsname{Keywords}
-\def\IEEEproofname{Proof}
-
-
-% LIST OF FIGURES AND TABLES AND TABLE OF CONTENTS
-%
-\def\@pnumwidth{1.55em}
-\def\@tocrmarg{2.55em}
-\def\@dotsep{4.5}
-\setcounter{tocdepth}{3}
-
-% adjusted some spacings here so that section numbers will not easily
-% collide with the section titles.
-% VIII; VIII-A; and VIII-A.1 are usually the worst offenders.
-% MDS 1/2001
-\def\tableofcontents{\section*{\contentsname}\@starttoc{toc}}
-\def\l@section#1#2{\addpenalty{\@secpenalty}\addvspace{1.0em plus 1pt}%
- \@tempdima 2.75em \begingroup \parindent \z@ \rightskip \@pnumwidth%
- \parfillskip-\@pnumwidth {\bfseries\leavevmode #1}\hfil\hbox to\@pnumwidth{\hss #2}\par%
- \endgroup}
-% argument format #1:level, #2:labelindent,#3:labelsep
-\def\l@subsection{\@dottedtocline{2}{2.75em}{3.75em}}
-\def\l@subsubsection{\@dottedtocline{3}{6.5em}{4.5em}}
-% must provide \l@ defs for ALL sublevels EVEN if tocdepth
-% is such as they will not appear in the table of contents
-% these defs are how TOC knows what level these things are!
-\def\l@paragraph{\@dottedtocline{4}{6.5em}{5.5em}}
-\def\l@subparagraph{\@dottedtocline{5}{6.5em}{6.5em}}
-\def\listoffigures{\section*{\listfigurename}\@starttoc{lof}}
-\def\l@figure{\@dottedtocline{1}{0em}{2.75em}}
-\def\listoftables{\section*{\listtablename}\@starttoc{lot}}
-\let\l@table\l@figure
-
-
-%% Definitions for floats
-%%
-%% Normal Floats
-\floatsep 1\baselineskip plus 0.2\baselineskip minus 0.2\baselineskip
-\textfloatsep 1.7\baselineskip plus 0.2\baselineskip minus 0.4\baselineskip
-\@fptop 0pt plus 1fil
-\@fpsep 0.75\baselineskip plus 2fil
-\@fpbot 0pt plus 1fil
-\def\topfraction{0.9}
-\def\bottomfraction{0.4}
-\def\floatpagefraction{0.8}
-% V1.7, let top floats approach 90% of page
-\def\textfraction{0.1}
-
-%% Double Column Floats
-\dblfloatsep 1\baselineskip plus 0.2\baselineskip minus 0.2\baselineskip
-
-\dbltextfloatsep 1.7\baselineskip plus 0.2\baselineskip minus 0.4\baselineskip
-% Note that it would be nice if the rubber here actually worked in LaTeX2e.
-% There is a long standing limitation in LaTeX, first discovered (to the best
-% of my knowledge) by Alan Jeffrey in 1992. LaTeX ignores the stretchable
-% portion of \dbltextfloatsep, and as a result, double column figures can and
-% do result in an non-integer number of lines in the main text columns with
-% underfull vbox errors as a consequence. A post to comp.text.tex
-% by Donald Arseneau confirms that this had not yet been fixed in 1998.
-% IEEEtran V1.6 will fix this problem for you in the titles, but it doesn't
-% protect you from other double floats. Happy vspace'ing.
-
-\@dblfptop 0pt plus 1fil
-\@dblfpsep 0.75\baselineskip plus 2fil
-\@dblfpbot 0pt plus 1fil
-\def\dbltopfraction{0.8}
-\def\dblfloatpagefraction{0.8}
-\setcounter{dbltopnumber}{4}
-
-\intextsep 1\baselineskip plus 0.2\baselineskip minus 0.2\baselineskip
-\setcounter{topnumber}{2}
-\setcounter{bottomnumber}{2}
-\setcounter{totalnumber}{4}
-
-
-
-% article class provides these, we should too.
-\newlength\abovecaptionskip
-\newlength\belowcaptionskip
-% but only \abovecaptionskip is used above figure captions and *below* table
-% captions
-\setlength\abovecaptionskip{0.5\baselineskip}
-\setlength\belowcaptionskip{0pt}
-% V1.6 create hooks in case the caption spacing ever needs to be
-% overridden by a user
-\def\@IEEEfigurecaptionsepspace{\vskip\abovecaptionskip\relax}%
-\def\@IEEEtablecaptionsepspace{\vskip\abovecaptionskip\relax}%
-
-
-% 1.6b revise caption system so that \@makecaption uses two arguments
-% as with LaTeX2e. Otherwise, there will be problems when using hyperref.
-\def\@IEEEtablestring{table}
-
-\ifCLASSOPTIONcompsoc
-% V1.7 compsoc \@makecaption
-\ifCLASSOPTIONconference% compsoc conference
-\long\def\@makecaption#1#2{%
-% test if is a for a figure or table
-\ifx\@captype\@IEEEtablestring%
-% if a table, do table caption
-\normalsize\begin{center}{\normalfont\sffamily\normalsize {#1.}~ #2}\end{center}%
-\@IEEEtablecaptionsepspace
-% if not a table, format it as a figure
-\else
-\@IEEEfigurecaptionsepspace
-\setbox\@tempboxa\hbox{\normalfont\sffamily\normalsize {#1.}~ #2}%
-\ifdim \wd\@tempboxa >\hsize%
-% if caption is longer than a line, let it wrap around
-\setbox\@tempboxa\hbox{\normalfont\sffamily\normalsize {#1.}~ }%
-\parbox[t]{\hsize}{\normalfont\sffamily\normalsize \noindent\unhbox\@tempboxa#2}%
-% if caption is shorter than a line, center
-\else%
-\hbox to\hsize{\normalfont\sffamily\normalsize\hfil\box\@tempboxa\hfil}%
-\fi\fi}
-\else% nonconference compsoc
-\long\def\@makecaption#1#2{%
-% test if is a for a figure or table
-\ifx\@captype\@IEEEtablestring%
-% if a table, do table caption
-\normalsize\begin{center}{\normalfont\sffamily\normalsize #1}\\{\normalfont\sffamily\normalsize #2}\end{center}%
-\@IEEEtablecaptionsepspace
-% if not a table, format it as a figure
-\else
-\@IEEEfigurecaptionsepspace
-\setbox\@tempboxa\hbox{\normalfont\sffamily\normalsize {#1.}~ #2}%
-\ifdim \wd\@tempboxa >\hsize%
-% if caption is longer than a line, let it wrap around
-\setbox\@tempboxa\hbox{\normalfont\sffamily\normalsize {#1.}~ }%
-\parbox[t]{\hsize}{\normalfont\sffamily\normalsize \noindent\unhbox\@tempboxa#2}%
-% if caption is shorter than a line, left justify
-\else%
-\hbox to\hsize{\normalfont\sffamily\normalsize\box\@tempboxa\hfil}%
-\fi\fi}
-\fi
-
-\else% traditional noncompsoc \@makecaption
-\long\def\@makecaption#1#2{%
-% test if is a for a figure or table
-\ifx\@captype\@IEEEtablestring%
-% if a table, do table caption
-\footnotesize\begin{center}{\normalfont\footnotesize #1}\\{\normalfont\footnotesize\scshape #2}\end{center}%
-\@IEEEtablecaptionsepspace
-% if not a table, format it as a figure
-\else
-\@IEEEfigurecaptionsepspace
-% 3/2001 use footnotesize, not small; use two nonbreaking spaces, not one
-\setbox\@tempboxa\hbox{\normalfont\footnotesize {#1.}~~ #2}%
-\ifdim \wd\@tempboxa >\hsize%
-% if caption is longer than a line, let it wrap around
-\setbox\@tempboxa\hbox{\normalfont\footnotesize {#1.}~~ }%
-\parbox[t]{\hsize}{\normalfont\footnotesize\noindent\unhbox\@tempboxa#2}%
-% if caption is shorter than a line, center if conference, left justify otherwise
-\else%
-\ifCLASSOPTIONconference \hbox to\hsize{\normalfont\footnotesize\hfil\box\@tempboxa\hfil}%
-\else \hbox to\hsize{\normalfont\footnotesize\box\@tempboxa\hfil}%
-\fi\fi\fi}
-\fi
-
-
-
-% V1.7 disable captions class option, do so in a way that retains operation of \label
-% within \caption
-\ifCLASSOPTIONcaptionsoff
-\long\def\@makecaption#1#2{\vspace*{2em}\footnotesize\begin{center}{\footnotesize #1}\end{center}%
-\let\@IEEEtemporiglabeldefsave\label
-\let\@IEEEtemplabelargsave\relax
-\def\label##1{\gdef\@IEEEtemplabelargsave{##1}}%
-\setbox\@tempboxa\hbox{#2}%
-\let\label\@IEEEtemporiglabeldefsave
-\ifx\@IEEEtemplabelargsave\relax\else\label{\@IEEEtemplabelargsave}\fi}
-\fi
-
-
-% V1.7 define end environments with \def not \let so as to work OK with
-% preview-latex
-\newcounter{figure}
-\def\thefigure{\@arabic\c@figure}
-\def\fps@figure{tbp}
-\def\ftype@figure{1}
-\def\ext@figure{lof}
-\def\fnum@figure{\figurename~\thefigure}
-\def\figure{\@float{figure}}
-\def\endfigure{\end@float}
-\@namedef{figure*}{\@dblfloat{figure}}
-\@namedef{endfigure*}{\end@dblfloat}
-\newcounter{table}
-\ifCLASSOPTIONcompsoc
-\def\thetable{\arabic{table}}
-\else
-\def\thetable{\@Roman\c@table}
-\fi
-\def\fps@table{tbp}
-\def\ftype@table{2}
-\def\ext@table{lot}
-\def\fnum@table{\tablename~\thetable}
-% V1.6 IEEE uses 8pt text for tables
-% to default to footnotesize, we hack into LaTeX2e's \@floatboxreset and pray
-\def\table{\def\@floatboxreset{\reset@font\footnotesize\@setminipage}\@float{table}}
-\def\endtable{\end@float}
-% v1.6b double column tables need to default to footnotesize as well.
-\@namedef{table*}{\def\@floatboxreset{\reset@font\footnotesize\@setminipage}\@dblfloat{table}}
-\@namedef{endtable*}{\end@dblfloat}
-
-
-
-
-%%
-%% START OF IEEEeqnarry DEFINITIONS
-%%
-%% Inspired by the concepts, examples, and previous works of LaTeX
-%% coders and developers such as Donald Arseneau, Fred Bartlett,
-%% David Carlisle, Tony Liu, Frank Mittelbach, Piet van Oostrum,
-%% Roland Winkler and Mark Wooding.
-%% I don't make the claim that my work here is even near their calibre. ;)
-
-
-% hook to allow easy changeover to IEEEtran.cls/tools.sty error reporting
-\def\@IEEEclspkgerror{\ClassError{IEEEtran}}
-
-\newif\if@IEEEeqnarraystarform% flag to indicate if the environment was called as the star form
-\@IEEEeqnarraystarformfalse
-
-\newif\if@advanceIEEEeqncolcnt% tracks if the environment should advance the col counter
-% allows a way to make an \IEEEeqnarraybox that can be used within an \IEEEeqnarray
-% used by IEEEeqnarraymulticol so that it can work properly in both
-\@advanceIEEEeqncolcnttrue
-
-\newcount\@IEEEeqnnumcols % tracks how many IEEEeqnarray cols are defined
-\newcount\@IEEEeqncolcnt % tracks how many IEEEeqnarray cols the user actually used
-
-
-% The default math style used by the columns
-\def\IEEEeqnarraymathstyle{\displaystyle}
-% The default text style used by the columns
-% default to using the current font
-\def\IEEEeqnarraytextstyle{\relax}
-
-% like the iedlistdecl but for \IEEEeqnarray
-\def\IEEEeqnarraydecl{\relax}
-\def\IEEEeqnarrayboxdecl{\relax}
-
-% \yesnumber is the opposite of \nonumber
-% a novel concept with the same def as the equationarray package
-% However, we give IEEE versions too since some LaTeX packages such as
-% the MDWtools mathenv.sty redefine \nonumber to something else.
-\providecommand{\yesnumber}{\global\@eqnswtrue}
-\def\IEEEyesnumber{\global\@eqnswtrue}
-\def\IEEEnonumber{\global\@eqnswfalse}
-
-
-\def\IEEEyessubnumber{\global\@IEEEissubequationtrue\global\@eqnswtrue%
-\if@IEEEeqnarrayISinner% only do something inside an IEEEeqnarray
-\if@IEEElastlinewassubequation\addtocounter{equation}{-1}\else\setcounter{IEEEsubequation}{1}\fi%
-\def\@currentlabel{\p@IEEEsubequation\theIEEEsubequation}\fi}
-
-% flag to indicate that an equation is a sub equation
-\newif\if@IEEEissubequation%
-\@IEEEissubequationfalse
-
-% allows users to "push away" equations that get too close to the equation numbers
-\def\IEEEeqnarraynumspace{\hphantom{\if@IEEEissubequation\theIEEEsubequationdis\else\theequationdis\fi}}
-
-% provides a way to span multiple columns within IEEEeqnarray environments
-% will consider \if@advanceIEEEeqncolcnt before globally advancing the
-% column counter - so as to work within \IEEEeqnarraybox
-% usage: \IEEEeqnarraymulticol{number cols. to span}{col type}{cell text}
-\long\def\IEEEeqnarraymulticol#1#2#3{\multispan{#1}%
-% check if column is defined
-\relax\expandafter\ifx\csname @IEEEeqnarraycolDEF#2\endcsname\@IEEEeqnarraycolisdefined%
-\csname @IEEEeqnarraycolPRE#2\endcsname#3\relax\relax\relax\relax\relax%
-\relax\relax\relax\relax\relax\csname @IEEEeqnarraycolPOST#2\endcsname%
-\else% if not, error and use default type
-\@IEEEclspkgerror{Invalid column type "#2" in \string\IEEEeqnarraymulticol.\MessageBreak
-Using a default centering column instead}%
-{You must define IEEEeqnarray column types before use.}%
-\csname @IEEEeqnarraycolPRE@IEEEdefault\endcsname#3\relax\relax\relax\relax\relax%
-\relax\relax\relax\relax\relax\csname @IEEEeqnarraycolPOST@IEEEdefault\endcsname%
-\fi%
-% advance column counter only if the IEEEeqnarray environment wants it
-\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by #1\relax\fi}
-
-% like \omit, but maintains track of the column counter for \IEEEeqnarray
-\def\IEEEeqnarrayomit{\omit\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by 1\relax\fi}
-
-
-% provides a way to define a letter referenced column type
-% usage: \IEEEeqnarraydefcol{col. type letter/name}{pre insertion text}{post insertion text}
-\def\IEEEeqnarraydefcol#1#2#3{\expandafter\def\csname @IEEEeqnarraycolPRE#1\endcsname{#2}%
-\expandafter\def\csname @IEEEeqnarraycolPOST#1\endcsname{#3}%
-\expandafter\def\csname @IEEEeqnarraycolDEF#1\endcsname{1}}
-
-
-% provides a way to define a numerically referenced inter-column glue types
-% usage: \IEEEeqnarraydefcolsep{col. glue number}{glue definition}
-\def\IEEEeqnarraydefcolsep#1#2{\expandafter\def\csname @IEEEeqnarraycolSEP\romannumeral #1\endcsname{#2}%
-\expandafter\def\csname @IEEEeqnarraycolSEPDEF\romannumeral #1\endcsname{1}}
-
-
-\def\@IEEEeqnarraycolisdefined{1}% just a macro for 1, used for checking undefined column types
-
-
-% expands and appends the given argument to the \@IEEEtrantmptoksA token list
-% used to build up the \halign preamble
-\def\@IEEEappendtoksA#1{\edef\@@IEEEappendtoksA{\@IEEEtrantmptoksA={\the\@IEEEtrantmptoksA #1}}%
-\@@IEEEappendtoksA}
-
-% also appends to \@IEEEtrantmptoksA, but does not expand the argument
-% uses \toks8 as a scratchpad register
-\def\@IEEEappendNOEXPANDtoksA#1{\toks8={#1}%
-\edef\@@IEEEappendNOEXPANDtoksA{\@IEEEtrantmptoksA={\the\@IEEEtrantmptoksA\the\toks8}}%
-\@@IEEEappendNOEXPANDtoksA}
-
-% define some common column types for the user
-% math
-\IEEEeqnarraydefcol{l}{$\IEEEeqnarraymathstyle}{$\hfil}
-\IEEEeqnarraydefcol{c}{\hfil$\IEEEeqnarraymathstyle}{$\hfil}
-\IEEEeqnarraydefcol{r}{\hfil$\IEEEeqnarraymathstyle}{$}
-\IEEEeqnarraydefcol{L}{$\IEEEeqnarraymathstyle{}}{{}$\hfil}
-\IEEEeqnarraydefcol{C}{\hfil$\IEEEeqnarraymathstyle{}}{{}$\hfil}
-\IEEEeqnarraydefcol{R}{\hfil$\IEEEeqnarraymathstyle{}}{{}$}
-% text
-\IEEEeqnarraydefcol{s}{\IEEEeqnarraytextstyle}{\hfil}
-\IEEEeqnarraydefcol{t}{\hfil\IEEEeqnarraytextstyle}{\hfil}
-\IEEEeqnarraydefcol{u}{\hfil\IEEEeqnarraytextstyle}{}
-
-% vertical rules
-\IEEEeqnarraydefcol{v}{}{\vrule width\arrayrulewidth}
-\IEEEeqnarraydefcol{vv}{\vrule width\arrayrulewidth\hfil}{\hfil\vrule width\arrayrulewidth}
-\IEEEeqnarraydefcol{V}{}{\vrule width\arrayrulewidth\hskip\doublerulesep\vrule width\arrayrulewidth}
-\IEEEeqnarraydefcol{VV}{\vrule width\arrayrulewidth\hskip\doublerulesep\vrule width\arrayrulewidth\hfil}%
-{\hfil\vrule width\arrayrulewidth\hskip\doublerulesep\vrule width\arrayrulewidth}
-
-% horizontal rules
-\IEEEeqnarraydefcol{h}{}{\leaders\hrule height\arrayrulewidth\hfil}
-\IEEEeqnarraydefcol{H}{}{\leaders\vbox{\hrule width\arrayrulewidth\vskip\doublerulesep\hrule width\arrayrulewidth}\hfil}
-
-% plain
-\IEEEeqnarraydefcol{x}{}{}
-\IEEEeqnarraydefcol{X}{$}{$}
-
-% the default column type to use in the event a column type is not defined
-\IEEEeqnarraydefcol{@IEEEdefault}{\hfil$\IEEEeqnarraymathstyle}{$\hfil}
-
-
-% a zero tabskip (used for "-" col types)
-\def\@IEEEeqnarraycolSEPzero{0pt plus 0pt minus 0pt}
-% a centering tabskip (used for "+" col types)
-\def\@IEEEeqnarraycolSEPcenter{1000pt plus 0pt minus 1000pt}
-
-% top level default tabskip glues for the start, end, and inter-column
-% may be reset within environments not always at the top level, e.g., \IEEEeqnarraybox
-\edef\@IEEEeqnarraycolSEPdefaultstart{\@IEEEeqnarraycolSEPcenter}% default start glue
-\edef\@IEEEeqnarraycolSEPdefaultend{\@IEEEeqnarraycolSEPcenter}% default end glue
-\edef\@IEEEeqnarraycolSEPdefaultmid{\@IEEEeqnarraycolSEPzero}% default inter-column glue
-
-
-
-% creates a vertical rule that extends from the bottom to the top a a cell
-% Provided in case other packages redefine \vline some other way.
-% usage: \IEEEeqnarrayvrule[rule thickness]
-% If no argument is provided, \arrayrulewidth will be used for the rule thickness.
-\newcommand\IEEEeqnarrayvrule[1][\arrayrulewidth]{\vrule\@width#1\relax}
-
-% creates a blank separator row
-% usage: \IEEEeqnarrayseprow[separation length][font size commands]
-% default is \IEEEeqnarrayseprow[0.25\normalbaselineskip][\relax]
-% blank arguments inherit the default values
-% uses \skip5 as a scratch register - calls \@IEEEeqnarraystrutsize which uses more scratch registers
-\def\IEEEeqnarrayseprow{\relax\@ifnextchar[{\@IEEEeqnarrayseprow}{\@IEEEeqnarrayseprow[0.25\normalbaselineskip]}}
-\def\@IEEEeqnarrayseprow[#1]{\relax\@ifnextchar[{\@@IEEEeqnarrayseprow[#1]}{\@@IEEEeqnarrayseprow[#1][\relax]}}
-\def\@@IEEEeqnarrayseprow[#1][#2]{\def\@IEEEeqnarrayseprowARGONE{#1}%
-\ifx\@IEEEeqnarrayseprowARGONE\@empty%
-% get the skip value, based on the font commands
-% use skip5 because \IEEEeqnarraystrutsize uses \skip0, \skip2, \skip3
-% assign within a bogus box to confine the font changes
-{\setbox0=\hbox{#2\relax\global\skip5=0.25\normalbaselineskip}}%
-\else%
-{\setbox0=\hbox{#2\relax\global\skip5=#1}}%
-\fi%
-\@IEEEeqnarrayhoptolastcolumn\IEEEeqnarraystrutsize{\skip5}{0pt}[\relax]\relax}
-
-% creates a blank separator row, but omits all the column templates
-% usage: \IEEEeqnarrayseprowcut[separation length][font size commands]
-% default is \IEEEeqnarrayseprowcut[0.25\normalbaselineskip][\relax]
-% blank arguments inherit the default values
-% uses \skip5 as a scratch register - calls \@IEEEeqnarraystrutsize which uses more scratch registers
-\def\IEEEeqnarrayseprowcut{\multispan{\@IEEEeqnnumcols}\relax% span all the cols
-% advance column counter only if the IEEEeqnarray environment wants it
-\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by \@IEEEeqnnumcols\relax\fi%
-\@ifnextchar[{\@IEEEeqnarrayseprowcut}{\@IEEEeqnarrayseprowcut[0.25\normalbaselineskip]}}
-\def\@IEEEeqnarrayseprowcut[#1]{\relax\@ifnextchar[{\@@IEEEeqnarrayseprowcut[#1]}{\@@IEEEeqnarrayseprowcut[#1][\relax]}}
-\def\@@IEEEeqnarrayseprowcut[#1][#2]{\def\@IEEEeqnarrayseprowARGONE{#1}%
-\ifx\@IEEEeqnarrayseprowARGONE\@empty%
-% get the skip value, based on the font commands
-% use skip5 because \IEEEeqnarraystrutsize uses \skip0, \skip2, \skip3
-% assign within a bogus box to confine the font changes
-{\setbox0=\hbox{#2\relax\global\skip5=0.25\normalbaselineskip}}%
-\else%
-{\setbox0=\hbox{#2\relax\global\skip5=#1}}%
-\fi%
-\IEEEeqnarraystrutsize{\skip5}{0pt}[\relax]\relax}
-
-
-
-% draws a single rule across all the columns optional
-% argument determines the rule width, \arrayrulewidth is the default
-% updates column counter as needed and turns off struts
-% usage: \IEEEeqnarrayrulerow[rule line thickness]
-\def\IEEEeqnarrayrulerow{\multispan{\@IEEEeqnnumcols}\relax% span all the cols
-% advance column counter only if the IEEEeqnarray environment wants it
-\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by \@IEEEeqnnumcols\relax\fi%
-\@ifnextchar[{\@IEEEeqnarrayrulerow}{\@IEEEeqnarrayrulerow[\arrayrulewidth]}}
-\def\@IEEEeqnarrayrulerow[#1]{\leaders\hrule height#1\hfil\relax% put in our rule
-% turn off any struts
-\IEEEeqnarraystrutsize{0pt}{0pt}[\relax]\relax}
-
-
-% draws a double rule by using a single rule row, a separator row, and then
-% another single rule row
-% first optional argument determines the rule thicknesses, \arrayrulewidth is the default
-% second optional argument determines the rule spacing, \doublerulesep is the default
-% usage: \IEEEeqnarraydblrulerow[rule line thickness][rule spacing]
-\def\IEEEeqnarraydblrulerow{\multispan{\@IEEEeqnnumcols}\relax% span all the cols
-% advance column counter only if the IEEEeqnarray environment wants it
-\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by \@IEEEeqnnumcols\relax\fi%
-\@ifnextchar[{\@IEEEeqnarraydblrulerow}{\@IEEEeqnarraydblrulerow[\arrayrulewidth]}}
-\def\@IEEEeqnarraydblrulerow[#1]{\relax\@ifnextchar[{\@@IEEEeqnarraydblrulerow[#1]}%
-{\@@IEEEeqnarraydblrulerow[#1][\doublerulesep]}}
-\def\@@IEEEeqnarraydblrulerow[#1][#2]{\def\@IEEEeqnarraydblrulerowARG{#1}%
-% we allow the user to say \IEEEeqnarraydblrulerow[][]
-\ifx\@IEEEeqnarraydblrulerowARG\@empty%
-\@IEEEeqnarrayrulerow[\arrayrulewidth]%
-\else%
-\@IEEEeqnarrayrulerow[#1]\relax%
-\fi%
-\def\@IEEEeqnarraydblrulerowARG{#2}%
-\ifx\@IEEEeqnarraydblrulerowARG\@empty%
-\\\IEEEeqnarrayseprow[\doublerulesep][\relax]%
-\else%
-\\\IEEEeqnarrayseprow[#2][\relax]%
-\fi%
-\\\multispan{\@IEEEeqnnumcols}%
-% advance column counter only if the IEEEeqnarray environment wants it
-\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by \@IEEEeqnnumcols\relax\fi%
-\def\@IEEEeqnarraydblrulerowARG{#1}%
-\ifx\@IEEEeqnarraydblrulerowARG\@empty%
-\@IEEEeqnarrayrulerow[\arrayrulewidth]%
-\else%
-\@IEEEeqnarrayrulerow[#1]%
-\fi%
-}
-
-% draws a double rule by using a single rule row, a separator (cutting) row, and then
-% another single rule row
-% first optional argument determines the rule thicknesses, \arrayrulewidth is the default
-% second optional argument determines the rule spacing, \doublerulesep is the default
-% usage: \IEEEeqnarraydblrulerow[rule line thickness][rule spacing]
-\def\IEEEeqnarraydblrulerowcut{\multispan{\@IEEEeqnnumcols}\relax% span all the cols
-% advance column counter only if the IEEEeqnarray environment wants it
-\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by \@IEEEeqnnumcols\relax\fi%
-\@ifnextchar[{\@IEEEeqnarraydblrulerowcut}{\@IEEEeqnarraydblrulerowcut[\arrayrulewidth]}}
-\def\@IEEEeqnarraydblrulerowcut[#1]{\relax\@ifnextchar[{\@@IEEEeqnarraydblrulerowcut[#1]}%
-{\@@IEEEeqnarraydblrulerowcut[#1][\doublerulesep]}}
-\def\@@IEEEeqnarraydblrulerowcut[#1][#2]{\def\@IEEEeqnarraydblrulerowARG{#1}%
-% we allow the user to say \IEEEeqnarraydblrulerow[][]
-\ifx\@IEEEeqnarraydblrulerowARG\@empty%
-\@IEEEeqnarrayrulerow[\arrayrulewidth]%
-\else%
-\@IEEEeqnarrayrulerow[#1]%
-\fi%
-\def\@IEEEeqnarraydblrulerowARG{#2}%
-\ifx\@IEEEeqnarraydblrulerowARG\@empty%
-\\\IEEEeqnarrayseprowcut[\doublerulesep][\relax]%
-\else%
-\\\IEEEeqnarrayseprowcut[#2][\relax]%
-\fi%
-\\\multispan{\@IEEEeqnnumcols}%
-% advance column counter only if the IEEEeqnarray environment wants it
-\if@advanceIEEEeqncolcnt\global\advance\@IEEEeqncolcnt by \@IEEEeqnnumcols\relax\fi%
-\def\@IEEEeqnarraydblrulerowARG{#1}%
-\ifx\@IEEEeqnarraydblrulerowARG\@empty%
-\@IEEEeqnarrayrulerow[\arrayrulewidth]%
-\else%
-\@IEEEeqnarrayrulerow[#1]%
-\fi%
-}
-
-
-
-% inserts a full row's worth of &'s
-% relies on \@IEEEeqnnumcols to provide the correct number of columns
-% uses \@IEEEtrantmptoksA, \count0 as scratch registers
-\def\@IEEEeqnarrayhoptolastcolumn{\@IEEEtrantmptoksA={}\count0=1\relax%
-\loop% add cols if the user did not use them all
-\ifnum\count0<\@IEEEeqnnumcols\relax%
-\@IEEEappendtoksA{&}%
-\advance\count0 by 1\relax% update the col count
-\repeat%
-\the\@IEEEtrantmptoksA%execute the &'s
-}
-
-
-
-\newif\if@IEEEeqnarrayISinner % flag to indicate if we are within the lines
-\@IEEEeqnarrayISinnerfalse % of an IEEEeqnarray - after the IEEEeqnarraydecl
-
-\edef\@IEEEeqnarrayTHEstrutheight{0pt} % height and depth of IEEEeqnarray struts
-\edef\@IEEEeqnarrayTHEstrutdepth{0pt}
-
-\edef\@IEEEeqnarrayTHEmasterstrutheight{0pt} % default height and depth of
-\edef\@IEEEeqnarrayTHEmasterstrutdepth{0pt} % struts within an IEEEeqnarray
-
-\edef\@IEEEeqnarrayTHEmasterstrutHSAVE{0pt} % saved master strut height
-\edef\@IEEEeqnarrayTHEmasterstrutDSAVE{0pt} % and depth
-
-\newif\if@IEEEeqnarrayusemasterstrut % flag to indicate that the master strut value
-\@IEEEeqnarrayusemasterstruttrue % is to be used
-
-
-
-% saves the strut height and depth of the master strut
-\def\@IEEEeqnarraymasterstrutsave{\relax%
-\expandafter\skip0=\@IEEEeqnarrayTHEmasterstrutheight\relax%
-\expandafter\skip2=\@IEEEeqnarrayTHEmasterstrutdepth\relax%
-% remove stretchability
-\dimen0\skip0\relax%
-\dimen2\skip2\relax%
-% save values
-\edef\@IEEEeqnarrayTHEmasterstrutHSAVE{\the\dimen0}%
-\edef\@IEEEeqnarrayTHEmasterstrutDSAVE{\the\dimen2}}
-
-% restores the strut height and depth of the master strut
-\def\@IEEEeqnarraymasterstrutrestore{\relax%
-\expandafter\skip0=\@IEEEeqnarrayTHEmasterstrutHSAVE\relax%
-\expandafter\skip2=\@IEEEeqnarrayTHEmasterstrutDSAVE\relax%
-% remove stretchability
-\dimen0\skip0\relax%
-\dimen2\skip2\relax%
-% restore values
-\edef\@IEEEeqnarrayTHEmasterstrutheight{\the\dimen0}%
-\edef\@IEEEeqnarrayTHEmasterstrutdepth{\the\dimen2}}
-
-
-% globally restores the strut height and depth to the
-% master values and sets the master strut flag to true
-\def\@IEEEeqnarraystrutreset{\relax%
-\expandafter\skip0=\@IEEEeqnarrayTHEmasterstrutheight\relax%
-\expandafter\skip2=\@IEEEeqnarrayTHEmasterstrutdepth\relax%
-% remove stretchability
-\dimen0\skip0\relax%
-\dimen2\skip2\relax%
-% restore values
-\xdef\@IEEEeqnarrayTHEstrutheight{\the\dimen0}%
-\xdef\@IEEEeqnarrayTHEstrutdepth{\the\dimen2}%
-\global\@IEEEeqnarrayusemasterstruttrue}
-
-
-% if the master strut is not to be used, make the current
-% values of \@IEEEeqnarrayTHEstrutheight, \@IEEEeqnarrayTHEstrutdepth
-% and the use master strut flag, global
-% this allows user strut commands issued in the last column to be carried
-% into the isolation/strut column
-\def\@IEEEeqnarrayglobalizestrutstatus{\relax%
-\if@IEEEeqnarrayusemasterstrut\else%
-\xdef\@IEEEeqnarrayTHEstrutheight{\@IEEEeqnarrayTHEstrutheight}%
-\xdef\@IEEEeqnarrayTHEstrutdepth{\@IEEEeqnarrayTHEstrutdepth}%
-\global\@IEEEeqnarrayusemasterstrutfalse%
-\fi}
-
-
-
-% usage: \IEEEeqnarraystrutsize{height}{depth}[font size commands]
-% If called outside the lines of an IEEEeqnarray, sets the height
-% and depth of both the master and local struts. If called inside
-% an IEEEeqnarray line, sets the height and depth of the local strut
-% only and sets the flag to indicate the use of the local strut
-% values. If the height or depth is left blank, 0.7\normalbaselineskip
-% and 0.3\normalbaselineskip will be used, respectively.
-% The optional argument can be used to evaluate the lengths under
-% a different font size and styles. If none is specified, the current
-% font is used.
-% uses scratch registers \skip0, \skip2, \skip3, \dimen0, \dimen2
-\def\IEEEeqnarraystrutsize#1#2{\relax\@ifnextchar[{\@IEEEeqnarraystrutsize{#1}{#2}}{\@IEEEeqnarraystrutsize{#1}{#2}[\relax]}}
-\def\@IEEEeqnarraystrutsize#1#2[#3]{\def\@IEEEeqnarraystrutsizeARG{#1}%
-\ifx\@IEEEeqnarraystrutsizeARG\@empty%
-{\setbox0=\hbox{#3\relax\global\skip3=0.7\normalbaselineskip}}%
-\skip0=\skip3\relax%
-\else% arg one present
-{\setbox0=\hbox{#3\relax\global\skip3=#1\relax}}%
-\skip0=\skip3\relax%
-\fi% if null arg
-\def\@IEEEeqnarraystrutsizeARG{#2}%
-\ifx\@IEEEeqnarraystrutsizeARG\@empty%
-{\setbox0=\hbox{#3\relax\global\skip3=0.3\normalbaselineskip}}%
-\skip2=\skip3\relax%
-\else% arg two present
-{\setbox0=\hbox{#3\relax\global\skip3=#2\relax}}%
-\skip2=\skip3\relax%
-\fi% if null arg
-% remove stretchability, just to be safe
-\dimen0\skip0\relax%
-\dimen2\skip2\relax%
-% dimen0 = height, dimen2 = depth
-\if@IEEEeqnarrayISinner% inner does not touch master strut size
-\edef\@IEEEeqnarrayTHEstrutheight{\the\dimen0}%
-\edef\@IEEEeqnarrayTHEstrutdepth{\the\dimen2}%
-\@IEEEeqnarrayusemasterstrutfalse% do not use master
-\else% outer, have to set master strut too
-\edef\@IEEEeqnarrayTHEmasterstrutheight{\the\dimen0}%
-\edef\@IEEEeqnarrayTHEmasterstrutdepth{\the\dimen2}%
-\edef\@IEEEeqnarrayTHEstrutheight{\the\dimen0}%
-\edef\@IEEEeqnarrayTHEstrutdepth{\the\dimen2}%
-\@IEEEeqnarrayusemasterstruttrue% use master strut
-\fi}
-
-
-% usage: \IEEEeqnarraystrutsizeadd{added height}{added depth}[font size commands]
-% If called outside the lines of an IEEEeqnarray, adds the given height
-% and depth to both the master and local struts.
-% If called inside an IEEEeqnarray line, adds the given height and depth
-% to the local strut only and sets the flag to indicate the use
-% of the local strut values.
-% In both cases, if a height or depth is left blank, 0pt is used instead.
-% The optional argument can be used to evaluate the lengths under
-% a different font size and styles. If none is specified, the current
-% font is used.
-% uses scratch registers \skip0, \skip2, \skip3, \dimen0, \dimen2
-\def\IEEEeqnarraystrutsizeadd#1#2{\relax\@ifnextchar[{\@IEEEeqnarraystrutsizeadd{#1}{#2}}{\@IEEEeqnarraystrutsizeadd{#1}{#2}[\relax]}}
-\def\@IEEEeqnarraystrutsizeadd#1#2[#3]{\def\@IEEEeqnarraystrutsizearg{#1}%
-\ifx\@IEEEeqnarraystrutsizearg\@empty%
-\skip0=0pt\relax%
-\else% arg one present
-{\setbox0=\hbox{#3\relax\global\skip3=#1}}%
-\skip0=\skip3\relax%
-\fi% if null arg
-\def\@IEEEeqnarraystrutsizearg{#2}%
-\ifx\@IEEEeqnarraystrutsizearg\@empty%
-\skip2=0pt\relax%
-\else% arg two present
-{\setbox0=\hbox{#3\relax\global\skip3=#2}}%
-\skip2=\skip3\relax%
-\fi% if null arg
-% remove stretchability, just to be safe
-\dimen0\skip0\relax%
-\dimen2\skip2\relax%
-% dimen0 = height, dimen2 = depth
-\if@IEEEeqnarrayISinner% inner does not touch master strut size
-% get local strut size
-\expandafter\skip0=\@IEEEeqnarrayTHEstrutheight\relax%
-\expandafter\skip2=\@IEEEeqnarrayTHEstrutdepth\relax%
-% add it to the user supplied values
-\advance\dimen0 by \skip0\relax%
-\advance\dimen2 by \skip2\relax%
-% update the local strut size
-\edef\@IEEEeqnarrayTHEstrutheight{\the\dimen0}%
-\edef\@IEEEeqnarrayTHEstrutdepth{\the\dimen2}%
-\@IEEEeqnarrayusemasterstrutfalse% do not use master
-\else% outer, have to set master strut too
-% get master strut size
-\expandafter\skip0=\@IEEEeqnarrayTHEmasterstrutheight\relax%
-\expandafter\skip2=\@IEEEeqnarrayTHEmasterstrutdepth\relax%
-% add it to the user supplied values
-\advance\dimen0 by \skip0\relax%
-\advance\dimen2 by \skip2\relax%
-% update the local and master strut sizes
-\edef\@IEEEeqnarrayTHEmasterstrutheight{\the\dimen0}%
-\edef\@IEEEeqnarrayTHEmasterstrutdepth{\the\dimen2}%
-\edef\@IEEEeqnarrayTHEstrutheight{\the\dimen0}%
-\edef\@IEEEeqnarrayTHEstrutdepth{\the\dimen2}%
-\@IEEEeqnarrayusemasterstruttrue% use master strut
-\fi}
-
-
-% allow user a way to see the struts
-\newif\ifIEEEvisiblestruts
-\IEEEvisiblestrutsfalse
-
-% inserts an invisible strut using the master or local strut values
-% uses scratch registers \skip0, \skip2, \dimen0, \dimen2
-\def\@IEEEeqnarrayinsertstrut{\relax%
-\if@IEEEeqnarrayusemasterstrut
-% get master strut size
-\expandafter\skip0=\@IEEEeqnarrayTHEmasterstrutheight\relax%
-\expandafter\skip2=\@IEEEeqnarrayTHEmasterstrutdepth\relax%
-\else%
-% get local strut size
-\expandafter\skip0=\@IEEEeqnarrayTHEstrutheight\relax%
-\expandafter\skip2=\@IEEEeqnarrayTHEstrutdepth\relax%
-\fi%
-% remove stretchability, probably not needed
-\dimen0\skip0\relax%
-\dimen2\skip2\relax%
-% dimen0 = height, dimen2 = depth
-% allow user to see struts if desired
-\ifIEEEvisiblestruts%
-\vrule width0.2pt height\dimen0 depth\dimen2\relax%
-\else%
-\vrule width0pt height\dimen0 depth\dimen2\relax\fi}
-
-
-% creates an invisible strut, useable even outside \IEEEeqnarray
-% if \IEEEvisiblestrutstrue, the strut will be visible and 0.2pt wide.
-% usage: \IEEEstrut[height][depth][font size commands]
-% default is \IEEEstrut[0.7\normalbaselineskip][0.3\normalbaselineskip][\relax]
-% blank arguments inherit the default values
-% uses \dimen0, \dimen2, \skip0, \skip2
-\def\IEEEstrut{\relax\@ifnextchar[{\@IEEEstrut}{\@IEEEstrut[0.7\normalbaselineskip]}}
-\def\@IEEEstrut[#1]{\relax\@ifnextchar[{\@@IEEEstrut[#1]}{\@@IEEEstrut[#1][0.3\normalbaselineskip]}}
-\def\@@IEEEstrut[#1][#2]{\relax\@ifnextchar[{\@@@IEEEstrut[#1][#2]}{\@@@IEEEstrut[#1][#2][\relax]}}
-\def\@@@IEEEstrut[#1][#2][#3]{\mbox{#3\relax%
-\def\@IEEEstrutARG{#1}%
-\ifx\@IEEEstrutARG\@empty%
-\skip0=0.7\normalbaselineskip\relax%
-\else%
-\skip0=#1\relax%
-\fi%
-\def\@IEEEstrutARG{#2}%
-\ifx\@IEEEstrutARG\@empty%
-\skip2=0.3\normalbaselineskip\relax%
-\else%
-\skip2=#2\relax%
-\fi%
-% remove stretchability, probably not needed
-\dimen0\skip0\relax%
-\dimen2\skip2\relax%
-\ifIEEEvisiblestruts%
-\vrule width0.2pt height\dimen0 depth\dimen2\relax%
-\else%
-\vrule width0.0pt height\dimen0 depth\dimen2\relax\fi}}
-
-
-% enables strut mode by setting a default strut size and then zeroing the
-% \baselineskip, \lineskip, \lineskiplimit and \jot
-\def\IEEEeqnarraystrutmode{\IEEEeqnarraystrutsize{0.7\normalbaselineskip}{0.3\normalbaselineskip}[\relax]%
-\baselineskip=0pt\lineskip=0pt\lineskiplimit=0pt\jot=0pt}
-
-
-
-\def\IEEEeqnarray{\@IEEEeqnarraystarformfalse\@IEEEeqnarray}
-\def\endIEEEeqnarray{\end@IEEEeqnarray}
-
-\@namedef{IEEEeqnarray*}{\@IEEEeqnarraystarformtrue\@IEEEeqnarray}
-\@namedef{endIEEEeqnarray*}{\end@IEEEeqnarray}
-
-
-% \IEEEeqnarray is an enhanced \eqnarray.
-% The star form defaults to not putting equation numbers at the end of each row.
-% usage: \IEEEeqnarray[decl]{cols}
-\def\@IEEEeqnarray{\relax\@ifnextchar[{\@@IEEEeqnarray}{\@@IEEEeqnarray[\relax]}}
-\def\@@IEEEeqnarray[#1]#2{%
- % default to showing the equation number or not based on whether or not
- % the star form was involked
- \if@IEEEeqnarraystarform\global\@eqnswfalse
- \else% not the star form
- \global\@eqnswtrue
- \fi% if star form
- \@IEEEissubequationfalse% default to no subequations
- \@IEEElastlinewassubequationfalse% assume last line is not a sub equation
- \@IEEEeqnarrayISinnerfalse% not yet within the lines of the halign
- \@IEEEeqnarraystrutsize{0pt}{0pt}[\relax]% turn off struts by default
- \@IEEEeqnarrayusemasterstruttrue% use master strut till user asks otherwise
- \IEEEvisiblestrutsfalse% diagnostic mode defaults to off
- % no extra space unless the user specifically requests it
- \lineskip=0pt\relax
- \lineskiplimit=0pt\relax
- \baselineskip=\normalbaselineskip\relax%
- \jot=\IEEEnormaljot\relax%
- \mathsurround\z@\relax% no extra spacing around math
- \@advanceIEEEeqncolcnttrue% advance the col counter for each col the user uses,
- % used in \IEEEeqnarraymulticol and in the preamble build
- \stepcounter{equation}% advance equation counter before first line
- \setcounter{IEEEsubequation}{0}% no subequation yet
- \def\@currentlabel{\p@equation\theequation}% redefine the ref label
- \IEEEeqnarraydecl\relax% allow a way for the user to make global overrides
- #1\relax% allow user to override defaults
- \let\\\@IEEEeqnarraycr% replace newline with one that can put in eqn. numbers
- \global\@IEEEeqncolcnt\z@% col. count = 0 for first line
- \@IEEEbuildpreamble #2\end\relax% build the preamble and put it into \@IEEEtrantmptoksA
- % put in the column for the equation number
- \ifnum\@IEEEeqnnumcols>0\relax\@IEEEappendtoksA{&}\fi% col separator for those after the first
- \toks0={##}%
- % advance the \@IEEEeqncolcnt for the isolation col, this helps with error checking
- \@IEEEappendtoksA{\global\advance\@IEEEeqncolcnt by 1\relax}%
- % add the isolation column
- \@IEEEappendtoksA{\tabskip\z@skip\bgroup\the\toks0\egroup}%
- % advance the \@IEEEeqncolcnt for the equation number col, this helps with error checking
- \@IEEEappendtoksA{&\global\advance\@IEEEeqncolcnt by 1\relax}%
- % add the equation number col to the preamble
- \@IEEEappendtoksA{\tabskip\z@skip\hb@xt@\z@\bgroup\hss\the\toks0\egroup}%
- % note \@IEEEeqnnumcols does not count the equation col or isolation col
- % set the starting tabskip glue as determined by the preamble build
- \tabskip=\@IEEEBPstartglue\relax
- % begin the display alignment
- \@IEEEeqnarrayISinnertrue% commands are now within the lines
- $$\everycr{}\halign to\displaywidth\bgroup
- % "exspand" the preamble
- \span\the\@IEEEtrantmptoksA\cr}
-
-% enter isolation/strut column (or the next column if the user did not use
-% every column), record the strut status, complete the columns, do the strut if needed,
-% restore counters to correct values and exit
-\def\end@IEEEeqnarray{\@IEEEeqnarrayglobalizestrutstatus&\@@IEEEeqnarraycr\egroup%
-\if@IEEElastlinewassubequation\global\advance\c@IEEEsubequation\m@ne\fi%
-\global\advance\c@equation\m@ne%
-$$\@ignoretrue}
-
-% need a way to remember if last line is a subequation
-\newif\if@IEEElastlinewassubequation%
-\@IEEElastlinewassubequationfalse
-
-% IEEEeqnarray uses a modifed \\ instead of the plain \cr to
-% end rows. This allows for things like \\*[vskip amount]
-% This "cr" macros are modified versions those for LaTeX2e's eqnarray
-% the {\ifnum0=`} braces must be kept away from the last column to avoid
-% altering spacing of its math, so we use & to advance to the next column
-% as there is an isolation/strut column after the user's columns
-\def\@IEEEeqnarraycr{\@IEEEeqnarrayglobalizestrutstatus&% save strut status and advance to next column
- {\ifnum0=`}\fi
- \@ifstar{%
- \global\@eqpen\@M\@IEEEeqnarrayYCR
- }{%
- \global\@eqpen\interdisplaylinepenalty \@IEEEeqnarrayYCR
- }%
-}
-
-\def\@IEEEeqnarrayYCR{\@testopt\@IEEEeqnarrayXCR\z@skip}
-
-\def\@IEEEeqnarrayXCR[#1]{%
- \ifnum0=`{\fi}%
- \@@IEEEeqnarraycr
- \noalign{\penalty\@eqpen\vskip\jot\vskip #1\relax}}%
-
-\def\@@IEEEeqnarraycr{\@IEEEtrantmptoksA={}% clear token register
- \advance\@IEEEeqncolcnt by -1\relax% adjust col count because of the isolation column
- \ifnum\@IEEEeqncolcnt>\@IEEEeqnnumcols\relax
- \@IEEEclspkgerror{Too many columns within the IEEEeqnarray\MessageBreak
- environment}%
- {Use fewer \string &'s or put more columns in the IEEEeqnarry column\MessageBreak
- specifications.}\relax%
- \else
- \loop% add cols if the user did not use them all
- \ifnum\@IEEEeqncolcnt<\@IEEEeqnnumcols\relax
- \@IEEEappendtoksA{&}%
- \advance\@IEEEeqncolcnt by 1\relax% update the col count
- \repeat
- % this number of &'s will take us the the isolation column
- \fi
- % execute the &'s
- \the\@IEEEtrantmptoksA%
- % handle the strut/isolation column
- \@IEEEeqnarrayinsertstrut% do the strut if needed
- \@IEEEeqnarraystrutreset% reset the strut system for next line or IEEEeqnarray
- &% and enter the equation number column
- % is this line needs an equation number, display it and advance the
- % (sub)equation counters, record what type this line was
- \if@eqnsw%
- \if@IEEEissubequation\theIEEEsubequationdis\addtocounter{equation}{1}\stepcounter{IEEEsubequation}%
- \global\@IEEElastlinewassubequationtrue%
- \else% display a standard equation number, initialize the IEEEsubequation counter
- \theequationdis\stepcounter{equation}\setcounter{IEEEsubequation}{0}%
- \global\@IEEElastlinewassubequationfalse\fi%
- \fi%
- % reset the eqnsw flag to indicate default preference of the display of equation numbers
- \if@IEEEeqnarraystarform\global\@eqnswfalse\else\global\@eqnswtrue\fi
- \global\@IEEEissubequationfalse% reset the subequation flag
- % reset the number of columns the user actually used
- \global\@IEEEeqncolcnt\z@\relax
- % the real end of the line
- \cr}
-
-
-
-
-
-% \IEEEeqnarraybox is like \IEEEeqnarray except the box form puts everything
-% inside a vtop, vbox, or vcenter box depending on the letter in the second
-% optional argument (t,b,c). Vbox is the default. Unlike \IEEEeqnarray,
-% equation numbers are not displayed and \IEEEeqnarraybox can be nested.
-% \IEEEeqnarrayboxm is for math mode (like \array) and does not put the vbox
-% within an hbox.
-% \IEEEeqnarrayboxt is for text mode (like \tabular) and puts the vbox within
-% a \hbox{$ $} construct.
-% \IEEEeqnarraybox will auto detect whether to use \IEEEeqnarrayboxm or
-% \IEEEeqnarrayboxt depending on the math mode.
-% The third optional argument specifies the width this box is to be set to -
-% natural width is the default.
-% The * forms do not add \jot line spacing
-% usage: \IEEEeqnarraybox[decl][pos][width]{cols}
-\def\IEEEeqnarrayboxm{\@IEEEeqnarraystarformfalse\@IEEEeqnarrayboxHBOXSWfalse\@IEEEeqnarraybox}
-\def\endIEEEeqnarrayboxm{\end@IEEEeqnarraybox}
-\@namedef{IEEEeqnarrayboxm*}{\@IEEEeqnarraystarformtrue\@IEEEeqnarrayboxHBOXSWfalse\@IEEEeqnarraybox}
-\@namedef{endIEEEeqnarrayboxm*}{\end@IEEEeqnarraybox}
-
-\def\IEEEeqnarrayboxt{\@IEEEeqnarraystarformfalse\@IEEEeqnarrayboxHBOXSWtrue\@IEEEeqnarraybox}
-\def\endIEEEeqnarrayboxt{\end@IEEEeqnarraybox}
-\@namedef{IEEEeqnarrayboxt*}{\@IEEEeqnarraystarformtrue\@IEEEeqnarrayboxHBOXSWtrue\@IEEEeqnarraybox}
-\@namedef{endIEEEeqnarrayboxt*}{\end@IEEEeqnarraybox}
-
-\def\IEEEeqnarraybox{\@IEEEeqnarraystarformfalse\ifmmode\@IEEEeqnarrayboxHBOXSWfalse\else\@IEEEeqnarrayboxHBOXSWtrue\fi%
-\@IEEEeqnarraybox}
-\def\endIEEEeqnarraybox{\end@IEEEeqnarraybox}
-
-\@namedef{IEEEeqnarraybox*}{\@IEEEeqnarraystarformtrue\ifmmode\@IEEEeqnarrayboxHBOXSWfalse\else\@IEEEeqnarrayboxHBOXSWtrue\fi%
-\@IEEEeqnarraybox}
-\@namedef{endIEEEeqnarraybox*}{\end@IEEEeqnarraybox}
-
-% flag to indicate if the \IEEEeqnarraybox needs to put things into an hbox{$ $}
-% for \vcenter in non-math mode
-\newif\if@IEEEeqnarrayboxHBOXSW%
-\@IEEEeqnarrayboxHBOXSWfalse
-
-\def\@IEEEeqnarraybox{\relax\@ifnextchar[{\@@IEEEeqnarraybox}{\@@IEEEeqnarraybox[\relax]}}
-\def\@@IEEEeqnarraybox[#1]{\relax\@ifnextchar[{\@@@IEEEeqnarraybox[#1]}{\@@@IEEEeqnarraybox[#1][b]}}
-\def\@@@IEEEeqnarraybox[#1][#2]{\relax\@ifnextchar[{\@@@@IEEEeqnarraybox[#1][#2]}{\@@@@IEEEeqnarraybox[#1][#2][\relax]}}
-
-% #1 = decl; #2 = t,b,c; #3 = width, #4 = col specs
-\def\@@@@IEEEeqnarraybox[#1][#2][#3]#4{\@IEEEeqnarrayISinnerfalse % not yet within the lines of the halign
- \@IEEEeqnarraymasterstrutsave% save current master strut values
- \@IEEEeqnarraystrutsize{0pt}{0pt}[\relax]% turn off struts by default
- \@IEEEeqnarrayusemasterstruttrue% use master strut till user asks otherwise
- \IEEEvisiblestrutsfalse% diagnostic mode defaults to off
- % no extra space unless the user specifically requests it
- \lineskip=0pt\relax%
- \lineskiplimit=0pt\relax%
- \baselineskip=\normalbaselineskip\relax%
- \jot=\IEEEnormaljot\relax%
- \mathsurround\z@\relax% no extra spacing around math
- % the default end glues are zero for an \IEEEeqnarraybox
- \edef\@IEEEeqnarraycolSEPdefaultstart{\@IEEEeqnarraycolSEPzero}% default start glue
- \edef\@IEEEeqnarraycolSEPdefaultend{\@IEEEeqnarraycolSEPzero}% default end glue
- \edef\@IEEEeqnarraycolSEPdefaultmid{\@IEEEeqnarraycolSEPzero}% default inter-column glue
- \@advanceIEEEeqncolcntfalse% do not advance the col counter for each col the user uses,
- % used in \IEEEeqnarraymulticol and in the preamble build
- \IEEEeqnarrayboxdecl\relax% allow a way for the user to make global overrides
- #1\relax% allow user to override defaults
- \let\\\@IEEEeqnarrayboxcr% replace newline with one that allows optional spacing
- \@IEEEbuildpreamble #4\end\relax% build the preamble and put it into \@IEEEtrantmptoksA
- % add an isolation column to the preamble to stop \\'s {} from getting into the last col
- \ifnum\@IEEEeqnnumcols>0\relax\@IEEEappendtoksA{&}\fi% col separator for those after the first
- \toks0={##}%
- % add the isolation column to the preamble
- \@IEEEappendtoksA{\tabskip\z@skip\bgroup\the\toks0\egroup}%
- % set the starting tabskip glue as determined by the preamble build
- \tabskip=\@IEEEBPstartglue\relax
- % begin the alignment
- \everycr{}%
- % use only the very first token to determine the positioning
- % this stops some problems when the user uses more than one letter,
- % but is probably not worth the effort
- % \noindent is used as a delimiter
- \def\@IEEEgrabfirstoken##1##2\noindent{\let\@IEEEgrabbedfirstoken=##1}%
- \@IEEEgrabfirstoken#2\relax\relax\noindent
- % \@IEEEgrabbedfirstoken has the first token, the rest are discarded
- % if we need to put things into and hbox and go into math mode, do so now
- \if@IEEEeqnarrayboxHBOXSW \leavevmode \hbox \bgroup $\fi%
- % use the appropriate vbox type
- \if\@IEEEgrabbedfirstoken t\relax\vtop\else\if\@IEEEgrabbedfirstoken c\relax%
- \vcenter\else\vbox\fi\fi\bgroup%
- \@IEEEeqnarrayISinnertrue% commands are now within the lines
- \ifx#3\relax\halign\else\halign to #3\relax\fi%
- \bgroup
- % "exspand" the preamble
- \span\the\@IEEEtrantmptoksA\cr}
-
-% carry strut status and enter the isolation/strut column,
-% exit from math mode if needed, and exit
-\def\end@IEEEeqnarraybox{\@IEEEeqnarrayglobalizestrutstatus% carry strut status
-&% enter isolation/strut column
-\@IEEEeqnarrayinsertstrut% do strut if needed
-\@IEEEeqnarraymasterstrutrestore% restore the previous master strut values
-% reset the strut system for next IEEEeqnarray
-% (sets local strut values back to previous master strut values)
-\@IEEEeqnarraystrutreset%
-% ensure last line, exit from halign, close vbox
-\crcr\egroup\egroup%
-% exit from math mode and close hbox if needed
-\if@IEEEeqnarrayboxHBOXSW $\egroup\fi}
-
-
-
-% IEEEeqnarraybox uses a modifed \\ instead of the plain \cr to
-% end rows. This allows for things like \\[vskip amount]
-% This "cr" macros are modified versions those for LaTeX2e's eqnarray
-% For IEEEeqnarraybox, \\* is the same as \\
-% the {\ifnum0=`} braces must be kept away from the last column to avoid
-% altering spacing of its math, so we use & to advance to the isolation/strut column
-% carry strut status into isolation/strut column
-\def\@IEEEeqnarrayboxcr{\@IEEEeqnarrayglobalizestrutstatus% carry strut status
-&% enter isolation/strut column
-\@IEEEeqnarrayinsertstrut% do strut if needed
-% reset the strut system for next line or IEEEeqnarray
-\@IEEEeqnarraystrutreset%
-{\ifnum0=`}\fi%
-\@ifstar{\@IEEEeqnarrayboxYCR}{\@IEEEeqnarrayboxYCR}}
-
-% test and setup the optional argument to \\[]
-\def\@IEEEeqnarrayboxYCR{\@testopt\@IEEEeqnarrayboxXCR\z@skip}
-
-% IEEEeqnarraybox does not automatically increase line spacing by \jot
-\def\@IEEEeqnarrayboxXCR[#1]{\ifnum0=`{\fi}%
-\cr\noalign{\if@IEEEeqnarraystarform\else\vskip\jot\fi\vskip#1\relax}}
-
-
-
-% starts the halign preamble build
-\def\@IEEEbuildpreamble{\@IEEEtrantmptoksA={}% clear token register
-\let\@IEEEBPcurtype=u%current column type is not yet known
-\let\@IEEEBPprevtype=s%the previous column type was the start
-\let\@IEEEBPnexttype=u%next column type is not yet known
-% ensure these are valid
-\def\@IEEEBPcurglue={0pt plus 0pt minus 0pt}%
-\def\@IEEEBPcurcolname{@IEEEdefault}% name of current column definition
-% currently acquired numerically referenced glue
-% use a name that is easier to remember
-\let\@IEEEBPcurnum=\@IEEEtrantmpcountA%
-\@IEEEBPcurnum=0%
-% tracks number of columns in the preamble
-\@IEEEeqnnumcols=0%
-% record the default end glues
-\edef\@IEEEBPstartglue{\@IEEEeqnarraycolSEPdefaultstart}%
-\edef\@IEEEBPendglue{\@IEEEeqnarraycolSEPdefaultend}%
-% now parse the user's column specifications
-\@@IEEEbuildpreamble}
-
-
-% parses and builds the halign preamble
-\def\@@IEEEbuildpreamble#1#2{\let\@@nextIEEEbuildpreamble=\@@IEEEbuildpreamble%
-% use only the very first token to check the end
-% \noindent is used as a delimiter as \end can be present here
-\def\@IEEEgrabfirstoken##1##2\noindent{\let\@IEEEgrabbedfirstoken=##1}%
-\@IEEEgrabfirstoken#1\relax\relax\noindent
-\ifx\@IEEEgrabbedfirstoken\end\let\@@nextIEEEbuildpreamble=\@@IEEEfinishpreamble\else%
-% identify current and next token type
-\@IEEEgetcoltype{#1}{\@IEEEBPcurtype}{1}% current, error on invalid
-\@IEEEgetcoltype{#2}{\@IEEEBPnexttype}{0}% next, no error on invalid next
-% if curtype is a glue, get the glue def
-\if\@IEEEBPcurtype g\@IEEEgetcurglue{#1}{\@IEEEBPcurglue}\fi%
-% if curtype is a column, get the column def and set the current column name
-\if\@IEEEBPcurtype c\@IEEEgetcurcol{#1}\fi%
-% if curtype is a numeral, acquire the user defined glue
-\if\@IEEEBPcurtype n\@IEEEprocessNcol{#1}\fi%
-% process the acquired glue
-\if\@IEEEBPcurtype g\@IEEEprocessGcol\fi%
-% process the acquired col
-\if\@IEEEBPcurtype c\@IEEEprocessCcol\fi%
-% ready prevtype for next col spec.
-\let\@IEEEBPprevtype=\@IEEEBPcurtype%
-% be sure and put back the future token(s) as a group
-\fi\@@nextIEEEbuildpreamble{#2}}
-
-
-% executed just after preamble build is completed
-% warn about zero cols, and if prevtype type = u, put in end tabskip glue
-\def\@@IEEEfinishpreamble#1{\ifnum\@IEEEeqnnumcols<1\relax
-\@IEEEclspkgerror{No column specifiers declared for IEEEeqnarray}%
-{At least one column type must be declared for each IEEEeqnarray.}%
-\fi%num cols less than 1
-%if last type undefined, set default end tabskip glue
-\if\@IEEEBPprevtype u\@IEEEappendtoksA{\tabskip=\@IEEEBPendglue}\fi}
-
-
-% Identify and return the column specifier's type code
-\def\@IEEEgetcoltype#1#2#3{%
-% use only the very first token to determine the type
-% \noindent is used as a delimiter as \end can be present here
-\def\@IEEEgrabfirstoken##1##2\noindent{\let\@IEEEgrabbedfirstoken=##1}%
-\@IEEEgrabfirstoken#1\relax\relax\noindent
-% \@IEEEgrabfirstoken has the first token, the rest are discarded
-% n = number
-% g = glue (any other char in catagory 12)
-% c = letter
-% e = \end
-% u = undefined
-% third argument: 0 = no error message, 1 = error on invalid char
-\let#2=u\relax% assume invalid until know otherwise
-\ifx\@IEEEgrabbedfirstoken\end\let#2=e\else
-\ifcat\@IEEEgrabbedfirstoken\relax\else% screen out control sequences
-\if0\@IEEEgrabbedfirstoken\let#2=n\else
-\if1\@IEEEgrabbedfirstoken\let#2=n\else
-\if2\@IEEEgrabbedfirstoken\let#2=n\else
-\if3\@IEEEgrabbedfirstoken\let#2=n\else
-\if4\@IEEEgrabbedfirstoken\let#2=n\else
-\if5\@IEEEgrabbedfirstoken\let#2=n\else
-\if6\@IEEEgrabbedfirstoken\let#2=n\else
-\if7\@IEEEgrabbedfirstoken\let#2=n\else
-\if8\@IEEEgrabbedfirstoken\let#2=n\else
-\if9\@IEEEgrabbedfirstoken\let#2=n\else
-\ifcat,\@IEEEgrabbedfirstoken\let#2=g\relax
-\else\ifcat a\@IEEEgrabbedfirstoken\let#2=c\relax\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi
-\if#2u\relax
-\if0\noexpand#3\relax\else\@IEEEclspkgerror{Invalid character in column specifications}%
-{Only letters, numerals and certain other symbols are allowed \MessageBreak
-as IEEEeqnarray column specifiers.}\fi\fi}
-
-
-% identify the current letter referenced column
-% if invalid, use a default column
-\def\@IEEEgetcurcol#1{\expandafter\ifx\csname @IEEEeqnarraycolDEF#1\endcsname\@IEEEeqnarraycolisdefined%
-\def\@IEEEBPcurcolname{#1}\else% invalid column name
-\@IEEEclspkgerror{Invalid column type "#1" in column specifications.\MessageBreak
-Using a default centering column instead}%
-{You must define IEEEeqnarray column types before use.}%
-\def\@IEEEBPcurcolname{@IEEEdefault}\fi}
-
-
-% identify and return the predefined (punctuation) glue value
-\def\@IEEEgetcurglue#1#2{%
-% ! = \! (neg small) -0.16667em (-3/18 em)
-% , = \, (small) 0.16667em ( 3/18 em)
-% : = \: (med) 0.22222em ( 4/18 em)
-% ; = \; (large) 0.27778em ( 5/18 em)
-% ' = \quad 1em
-% " = \qquad 2em
-% . = 0.5\arraycolsep
-% / = \arraycolsep
-% ? = 2\arraycolsep
-% * = 1fil
-% + = \@IEEEeqnarraycolSEPcenter
-% - = \@IEEEeqnarraycolSEPzero
-% Note that all em values are referenced to the math font (textfont2) fontdimen6
-% value for 1em.
-%
-% use only the very first token to determine the type
-% this prevents errant tokens from getting in the main text
-% \noindent is used as a delimiter here
-\def\@IEEEgrabfirstoken##1##2\noindent{\let\@IEEEgrabbedfirstoken=##1}%
-\@IEEEgrabfirstoken#1\relax\relax\noindent
-% get the math font 1em value
-% LaTeX2e's NFSS2 does not preload the fonts, but \IEEEeqnarray needs
-% to gain access to the math (\textfont2) font's spacing parameters.
-% So we create a bogus box here that uses the math font to ensure
-% that \textfont2 is loaded and ready. If this is not done,
-% the \textfont2 stuff here may not work.
-% Thanks to Bernd Raichle for his 1997 post on this topic.
-{\setbox0=\hbox{$\displaystyle\relax$}}%
-% fontdimen6 has the width of 1em (a quad).
-\@IEEEtrantmpdimenA=\fontdimen6\textfont2\relax%
-% identify the glue value based on the first token
-% we discard anything after the first
-\if!\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=-0.16667\@IEEEtrantmpdimenA\edef#2{\the\@IEEEtrantmpdimenA}\else
-\if,\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=0.16667\@IEEEtrantmpdimenA\edef#2{\the\@IEEEtrantmpdimenA}\else
-\if:\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=0.22222\@IEEEtrantmpdimenA\edef#2{\the\@IEEEtrantmpdimenA}\else
-\if;\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=0.27778\@IEEEtrantmpdimenA\edef#2{\the\@IEEEtrantmpdimenA}\else
-\if'\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=1\@IEEEtrantmpdimenA\edef#2{\the\@IEEEtrantmpdimenA}\else
-\if"\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=2\@IEEEtrantmpdimenA\edef#2{\the\@IEEEtrantmpdimenA}\else
-\if.\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=0.5\arraycolsep\edef#2{\the\@IEEEtrantmpdimenA}\else
-\if/\@IEEEgrabbedfirstoken\edef#2{\the\arraycolsep}\else
-\if?\@IEEEgrabbedfirstoken\@IEEEtrantmpdimenA=2\arraycolsep\edef#2{\the\@IEEEtrantmpdimenA}\else
-\if *\@IEEEgrabbedfirstoken\edef#2{0pt plus 1fil minus 0pt}\else
-\if+\@IEEEgrabbedfirstoken\edef#2{\@IEEEeqnarraycolSEPcenter}\else
-\if-\@IEEEgrabbedfirstoken\edef#2{\@IEEEeqnarraycolSEPzero}\else
-\edef#2{\@IEEEeqnarraycolSEPzero}%
-\@IEEEclspkgerror{Invalid predefined inter-column glue type "#1" in\MessageBreak
-column specifications. Using a default value of\MessageBreak
-0pt instead}%
-{Only !,:;'"./?*+ and - are valid predefined glue types in the\MessageBreak
-IEEEeqnarray column specifications.}\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi\fi}
-
-
-
-% process a numerical digit from the column specification
-% and look up the corresponding user defined glue value
-% can transform current type from n to g or a as the user defined glue is acquired
-\def\@IEEEprocessNcol#1{\if\@IEEEBPprevtype g%
-\@IEEEclspkgerror{Back-to-back inter-column glue specifiers in column\MessageBreak
-specifications. Ignoring consecutive glue specifiers\MessageBreak
-after the first}%
-{You cannot have two or more glue types next to each other\MessageBreak
-in the IEEEeqnarray column specifications.}%
-\let\@IEEEBPcurtype=a% abort this glue, future digits will be discarded
-\@IEEEBPcurnum=0\relax%
-\else% if we previously aborted a glue
-\if\@IEEEBPprevtype a\@IEEEBPcurnum=0\let\@IEEEBPcurtype=a%maintain digit abortion
-\else%acquire this number
-% save the previous type before the numerical digits started
-\if\@IEEEBPprevtype n\else\let\@IEEEBPprevsavedtype=\@IEEEBPprevtype\fi%
-\multiply\@IEEEBPcurnum by 10\relax%
-\advance\@IEEEBPcurnum by #1\relax% add in number, \relax is needed to stop TeX's number scan
-\if\@IEEEBPnexttype n\else%close acquisition
-\expandafter\ifx\csname @IEEEeqnarraycolSEPDEF\expandafter\romannumeral\number\@IEEEBPcurnum\endcsname\@IEEEeqnarraycolisdefined%
-\edef\@IEEEBPcurglue{\csname @IEEEeqnarraycolSEP\expandafter\romannumeral\number\@IEEEBPcurnum\endcsname}%
-\else%user glue not defined
-\@IEEEclspkgerror{Invalid user defined inter-column glue type "\number\@IEEEBPcurnum" in\MessageBreak
-column specifications. Using a default value of\MessageBreak
-0pt instead}%
-{You must define all IEEEeqnarray numerical inter-column glue types via\MessageBreak
-\string\IEEEeqnarraydefcolsep \space before they are used in column specifications.}%
-\edef\@IEEEBPcurglue{\@IEEEeqnarraycolSEPzero}%
-\fi% glue defined or not
-\let\@IEEEBPcurtype=g% change the type to reflect the acquired glue
-\let\@IEEEBPprevtype=\@IEEEBPprevsavedtype% restore the prev type before this number glue
-\@IEEEBPcurnum=0\relax%ready for next acquisition
-\fi%close acquisition, get glue
-\fi%discard or acquire number
-\fi%prevtype glue or not
-}
-
-
-% process an acquired glue
-% add any acquired column/glue pair to the preamble
-\def\@IEEEprocessGcol{\if\@IEEEBPprevtype a\let\@IEEEBPcurtype=a%maintain previous glue abortions
-\else
-% if this is the start glue, save it, but do nothing else
-% as this is not used in the preamble, but before
-\if\@IEEEBPprevtype s\edef\@IEEEBPstartglue{\@IEEEBPcurglue}%
-\else%not the start glue
-\if\@IEEEBPprevtype g%ignore if back to back glues
-\@IEEEclspkgerror{Back-to-back inter-column glue specifiers in column\MessageBreak
-specifications. Ignoring consecutive glue specifiers\MessageBreak
-after the first}%
-{You cannot have two or more glue types next to each other\MessageBreak
-in the IEEEeqnarray column specifications.}%
-\let\@IEEEBPcurtype=a% abort this glue
-\else% not a back to back glue
-\if\@IEEEBPprevtype c\relax% if the previoustype was a col, add column/glue pair to preamble
-\ifnum\@IEEEeqnnumcols>0\relax\@IEEEappendtoksA{&}\fi
-\toks0={##}%
-% make preamble advance col counter if this environment needs this
-\if@advanceIEEEeqncolcnt\@IEEEappendtoksA{\global\advance\@IEEEeqncolcnt by 1\relax}\fi
-% insert the column defintion into the preamble, being careful not to expand
-% the column definition
-\@IEEEappendtoksA{\tabskip=\@IEEEBPcurglue}%
-\@IEEEappendNOEXPANDtoksA{\begingroup\csname @IEEEeqnarraycolPRE}%
-\@IEEEappendtoksA{\@IEEEBPcurcolname}%
-\@IEEEappendNOEXPANDtoksA{\endcsname}%
-\@IEEEappendtoksA{\the\toks0}%
-\@IEEEappendNOEXPANDtoksA{\relax\relax\relax\relax\relax%
-\relax\relax\relax\relax\relax\csname @IEEEeqnarraycolPOST}%
-\@IEEEappendtoksA{\@IEEEBPcurcolname}%
-\@IEEEappendNOEXPANDtoksA{\endcsname\relax\relax\relax\relax\relax%
-\relax\relax\relax\relax\relax\endgroup}%
-\advance\@IEEEeqnnumcols by 1\relax%one more column in the preamble
-\else% error: non-start glue with no pending column
-\@IEEEclspkgerror{Inter-column glue specifier without a prior column\MessageBreak
-type in the column specifications. Ignoring this glue\MessageBreak
-specifier}%
-{Except for the first and last positions, glue can be placed only\MessageBreak
-between column types.}%
-\let\@IEEEBPcurtype=a% abort this glue
-\fi% previous was a column
-\fi% back-to-back glues
-\fi% is start column glue
-\fi% prev type not a
-}
-
-
-% process an acquired letter referenced column and, if necessary, add it to the preamble
-\def\@IEEEprocessCcol{\if\@IEEEBPnexttype g\else
-\if\@IEEEBPnexttype n\else
-% we have a column followed by something other than a glue (or numeral glue)
-% so we must add this column to the preamble now
-\ifnum\@IEEEeqnnumcols>0\relax\@IEEEappendtoksA{&}\fi%col separator for those after the first
-\if\@IEEEBPnexttype e\@IEEEappendtoksA{\tabskip=\@IEEEBPendglue\relax}\else%put in end glue
-\@IEEEappendtoksA{\tabskip=\@IEEEeqnarraycolSEPdefaultmid\relax}\fi% or default mid glue
-\toks0={##}%
-% make preamble advance col counter if this environment needs this
-\if@advanceIEEEeqncolcnt\@IEEEappendtoksA{\global\advance\@IEEEeqncolcnt by 1\relax}\fi
-% insert the column definition into the preamble, being careful not to expand
-% the column definition
-\@IEEEappendNOEXPANDtoksA{\begingroup\csname @IEEEeqnarraycolPRE}%
-\@IEEEappendtoksA{\@IEEEBPcurcolname}%
-\@IEEEappendNOEXPANDtoksA{\endcsname}%
-\@IEEEappendtoksA{\the\toks0}%
-\@IEEEappendNOEXPANDtoksA{\relax\relax\relax\relax\relax%
-\relax\relax\relax\relax\relax\csname @IEEEeqnarraycolPOST}%
-\@IEEEappendtoksA{\@IEEEBPcurcolname}%
-\@IEEEappendNOEXPANDtoksA{\endcsname\relax\relax\relax\relax\relax%
-\relax\relax\relax\relax\relax\endgroup}%
-\advance\@IEEEeqnnumcols by 1\relax%one more column in the preamble
-\fi%next type not numeral
-\fi%next type not glue
-}
-
-
-%%
-%% END OF IEEEeqnarry DEFINITIONS
-%%
-
-
-
-
-% set up the running headings, this complex because of all the different
-% modes IEEEtran supports
-\if@twoside
- \ifCLASSOPTIONtechnote
- \def\ps@headings{%
- \def\@oddhead{\hbox{}\scriptsize\leftmark \hfil \thepage}
- \def\@evenhead{\scriptsize\thepage \hfil \leftmark\hbox{}}
- \ifCLASSOPTIONdraftcls
- \ifCLASSOPTIONdraftclsnofoot
- \def\@oddfoot{}\def\@evenfoot{}%
- \else
- \def\@oddfoot{\scriptsize\@date\hfil DRAFT}
- \def\@evenfoot{\scriptsize DRAFT\hfil\@date}
- \fi
- \else
- \def\@oddfoot{}\def\@evenfoot{}
- \fi}
- \else % not a technote
- \def\ps@headings{%
- \ifCLASSOPTIONconference
- \def\@oddhead{}
- \def\@evenhead{}
- \else
- \def\@oddhead{\hbox{}\scriptsize\rightmark \hfil \thepage}
- \def\@evenhead{\scriptsize\thepage \hfil \leftmark\hbox{}}
- \fi
- \ifCLASSOPTIONdraftcls
- \def\@oddhead{\hbox{}\scriptsize\rightmark \hfil \thepage}
- \def\@evenhead{\scriptsize\thepage \hfil \leftmark\hbox{}}
- \ifCLASSOPTIONdraftclsnofoot
- \def\@oddfoot{}\def\@evenfoot{}%
- \else
- \def\@oddfoot{\scriptsize\@date\hfil DRAFT}
- \def\@evenfoot{\scriptsize DRAFT\hfil\@date}
- \fi
- \else
- \def\@oddfoot{}\def\@evenfoot{}%
- \fi}
- \fi
-\else % single side
-\def\ps@headings{%
- \ifCLASSOPTIONconference
- \def\@oddhead{}
- \def\@evenhead{}
- \else
- \def\@oddhead{\hbox{}\scriptsize\leftmark \hfil \thepage}
- \def\@evenhead{}
- \fi
- \ifCLASSOPTIONdraftcls
- \def\@oddhead{\hbox{}\scriptsize\leftmark \hfil \thepage}
- \def\@evenhead{}
- \ifCLASSOPTIONdraftclsnofoot
- \def\@oddfoot{}
- \else
- \def\@oddfoot{\scriptsize \@date \hfil DRAFT}
- \fi
- \else
- \def\@oddfoot{}
- \fi
- \def\@evenfoot{}}
-\fi
-
-
-% title page style
-\def\ps@IEEEtitlepagestyle{\def\@oddfoot{}\def\@evenfoot{}%
-\ifCLASSOPTIONconference
- \def\@oddhead{}%
- \def\@evenhead{}%
-\else
- \def\@oddhead{\hbox{}\scriptsize\leftmark \hfil \thepage}%
- \def\@evenhead{\scriptsize\thepage \hfil \leftmark\hbox{}}%
-\fi
-\ifCLASSOPTIONdraftcls
- \def\@oddhead{\hbox{}\scriptsize\leftmark \hfil \thepage}%
- \def\@evenhead{\scriptsize\thepage \hfil \leftmark\hbox{}}%
- \ifCLASSOPTIONdraftclsnofoot\else
- \def\@oddfoot{\scriptsize \@date\hfil DRAFT}%
- \def\@evenfoot{\scriptsize DRAFT\hfil \@date}%
- \fi
-\else
- % all non-draft mode footers
- \if@IEEEusingpubid
- % for title pages that are using a pubid
- % do not repeat pubid if using peer review option
- \ifCLASSOPTIONpeerreview
- \else
- \footskip 0pt%
- \ifCLASSOPTIONcompsocconf
- \def\@oddfoot{\hss\normalfont\scriptsize\raisebox{-1.5\@IEEEnormalsizeunitybaselineskip}[0ex][0ex]{\@IEEEpubid}\hss}%
- \def\@evenfoot{\hss\normalfont\scriptsize\raisebox{-1.5\@IEEEnormalsizeunitybaselineskip}[0ex][0ex]{\@IEEEpubid}\hss}%
- \else
- \def\@oddfoot{\hss\normalfont\footnotesize\raisebox{1.5ex}[1.5ex]{\@IEEEpubid}\hss}%
- \def\@evenfoot{\hss\normalfont\footnotesize\raisebox{1.5ex}[1.5ex]{\@IEEEpubid}\hss}%
- \fi
- \fi
- \fi
-\fi}
-
-
-% peer review cover page style
-\def\ps@IEEEpeerreviewcoverpagestyle{%
-\def\@oddhead{}\def\@evenhead{}%
-\def\@oddfoot{}\def\@evenfoot{}%
-\ifCLASSOPTIONdraftcls
- \ifCLASSOPTIONdraftclsnofoot\else
- \def\@oddfoot{\scriptsize \@date\hfil DRAFT}%
- \def\@evenfoot{\scriptsize DRAFT\hfil \@date}%
- \fi
-\else
- % non-draft mode footers
- \if@IEEEusingpubid
- \footskip 0pt%
- \ifCLASSOPTIONcompsoc
- \def\@oddfoot{\hss\normalfont\scriptsize\raisebox{-1.5\@IEEEnormalsizeunitybaselineskip}[0ex][0ex]{\@IEEEpubid}\hss}%
- \def\@evenfoot{\hss\normalfont\scriptsize\raisebox{-1.5\@IEEEnormalsizeunitybaselineskip}[0ex][0ex]{\@IEEEpubid}\hss}%
- \else
- \def\@oddfoot{\hss\normalfont\footnotesize\raisebox{1.5ex}[1.5ex]{\@IEEEpubid}\hss}%
- \def\@evenfoot{\hss\normalfont\footnotesize\raisebox{1.5ex}[1.5ex]{\@IEEEpubid}\hss}%
- \fi
- \fi
-\fi}
-
-
-% start with empty headings
-\def\rightmark{}\def\leftmark{}
-
-
-%% Defines the command for putting the header. \footernote{TEXT} is the same
-%% as \markboth{TEXT}{TEXT}.
-%% Note that all the text is forced into uppercase, if you have some text
-%% that needs to be in lower case, for instance et. al., then either manually
-%% set \leftmark and \rightmark or use \MakeLowercase{et. al.} within the
-%% arguments to \markboth.
-\def\markboth#1#2{\def\leftmark{\@IEEEcompsoconly{\sffamily}\MakeUppercase{#1}}%
-\def\rightmark{\@IEEEcompsoconly{\sffamily}\MakeUppercase{#2}}}
-\def\footernote#1{\markboth{#1}{#1}}
-
-\def\today{\ifcase\month\or
- January\or February\or March\or April\or May\or June\or
- July\or August\or September\or October\or November\or December\fi
- \space\number\day, \number\year}
-
-
-
-
-%% CITATION AND BIBLIOGRAPHY COMMANDS
-%%
-%% V1.6 no longer supports the older, nonstandard \shortcite and \citename setup stuff
-%
-%
-% Modify Latex2e \@citex to separate citations with "], ["
-\def\@citex[#1]#2{%
- \let\@citea\@empty
- \@cite{\@for\@citeb:=#2\do
- {\@citea\def\@citea{], [}%
- \edef\@citeb{\expandafter\@firstofone\@citeb\@empty}%
- \if@filesw\immediate\write\@auxout{\string\citation{\@citeb}}\fi
- \@ifundefined{b@\@citeb}{\mbox{\reset@font\bfseries ?}%
- \G@refundefinedtrue
- \@latex@warning
- {Citation `\@citeb' on page \thepage \space undefined}}%
- {\hbox{\csname b@\@citeb\endcsname}}}}{#1}}
-
-% V1.6 we create hooks for the optional use of Donald Arseneau's
-% cite.sty package. cite.sty is "smart" and will notice that the
-% following format controls are already defined and will not
-% redefine them. The result will be the proper sorting of the
-% citation numbers and auto detection of 3 or more entry "ranges" -
-% all in IEEE style: [1], [2], [5]--[7], [12]
-% This also allows for an optional note, i.e., \cite[mynote]{..}.
-% If the \cite with note has more than one reference, the note will
-% be applied to the last of the listed references. It is generally
-% desired that if a note is given, only one reference is listed in
-% that \cite.
-% Thanks to Mr. Arseneau for providing the required format arguments
-% to produce the IEEE style.
-\def\citepunct{], [}
-\def\citedash{]--[}
-
-% V1.7 default to using same font for urls made by url.sty
-\AtBeginDocument{\csname url@samestyle\endcsname}
-
-% V1.6 class files should always provide these
-\def\newblock{\hskip .11em\@plus.33em\@minus.07em}
-\let\@openbib@code\@empty
-
-
-% Provide support for the control entries of IEEEtran.bst V1.00 and later.
-% V1.7 optional argument allows for a different aux file to be specified in
-% order to handle multiple bibliographies. For example, with multibib.sty:
-% \newcites{sec}{Secondary Literature}
-% \bstctlcite[@auxoutsec]{BSTcontrolhak}
-\def\bstctlcite{\@ifnextchar[{\@bstctlcite}{\@bstctlcite[@auxout]}}
-\def\@bstctlcite[#1]#2{\@bsphack
- \@for\@citeb:=#2\do{%
- \edef\@citeb{\expandafter\@firstofone\@citeb}%
- \if@filesw\immediate\write\csname #1\endcsname{\string\citation{\@citeb}}\fi}%
- \@esphack}
-
-% V1.6 provide a way for a user to execute a command just before
-% a given reference number - used to insert a \newpage to balance
-% the columns on the last page
-\edef\@IEEEtriggerrefnum{0} % the default of zero means that
- % the command is not executed
-\def\@IEEEtriggercmd{\newpage}
-
-% allow the user to alter the triggered command
-\long\def\IEEEtriggercmd#1{\long\def\@IEEEtriggercmd{#1}}
-
-% allow user a way to specify the reference number just before the
-% command is executed
-\def\IEEEtriggeratref#1{\@IEEEtrantmpcountA=#1%
-\edef\@IEEEtriggerrefnum{\the\@IEEEtrantmpcountA}}%
-
-% trigger command at the given reference
-\def\@IEEEbibitemprefix{\@IEEEtrantmpcountA=\@IEEEtriggerrefnum\relax%
-\advance\@IEEEtrantmpcountA by -1\relax%
-\ifnum\c@enumiv=\@IEEEtrantmpcountA\relax\@IEEEtriggercmd\relax\fi}
-
-
-\def\@biblabel#1{[#1]}
-
-% compsoc journals left align the reference numbers
-\@IEEEcompsocnotconfonly{\def\@biblabel#1{[#1]\hfill}}
-
-% controls bib item spacing
-\def\IEEEbibitemsep{0pt plus .5pt}
-
-\@IEEEcompsocconfonly{\def\IEEEbibitemsep{1\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip}}
-
-
-\def\thebibliography#1{\section*{\refname}%
- \addcontentsline{toc}{section}{\refname}%
- % V1.6 add some rubber space here and provide a command trigger
- \footnotesize\@IEEEcompsocconfonly{\small}\vskip 0.3\baselineskip plus 0.1\baselineskip minus 0.1\baselineskip%
- \list{\@biblabel{\@arabic\c@enumiv}}%
- {\settowidth\labelwidth{\@biblabel{#1}}%
- \leftmargin\labelwidth
- \advance\leftmargin\labelsep\relax
- \itemsep \IEEEbibitemsep\relax
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{\@arabic\c@enumiv}}%
- \let\@IEEElatexbibitem\bibitem%
- \def\bibitem{\@IEEEbibitemprefix\@IEEElatexbibitem}%
-\def\newblock{\hskip .11em plus .33em minus .07em}%
-% originally:
-% \sloppy\clubpenalty4000\widowpenalty4000%
-% by adding the \interlinepenalty here, we make it more
-% difficult, but not impossible, for LaTeX to break within a reference.
-% IEEE almost never breaks a reference (but they do it more often with
-% technotes). You may get an underfull vbox warning around the bibliography,
-% but the final result will be much more like what IEEE will publish.
-% MDS 11/2000
-\ifCLASSOPTIONtechnote\sloppy\clubpenalty4000\widowpenalty4000\interlinepenalty100%
-\else\sloppy\clubpenalty4000\widowpenalty4000\interlinepenalty500\fi%
- \sfcode`\.=1000\relax}
-\let\endthebibliography=\endlist
-
-
-
-
-% TITLE PAGE COMMANDS
-%
-%
-% \IEEEmembership is used to produce the sublargesize italic font used to indicate author
-% IEEE membership. compsoc uses a large size sans slant font
-\def\IEEEmembership#1{{\@IEEEnotcompsoconly{\sublargesize}\normalfont\@IEEEcompsoconly{\sffamily}\textit{#1}}}
-
-
-% \IEEEauthorrefmark{} produces a footnote type symbol to indicate author affiliation.
-% When given an argument of 1 to 9, \IEEEauthorrefmark{} follows the standard LaTeX footnote
-% symbol sequence convention. However, for arguments 10 and above, \IEEEauthorrefmark{}
-% reverts to using lower case roman numerals, so it cannot overflow. Do note that you
-% cannot use \footnotemark[] in place of \IEEEauthorrefmark{} within \author as the footnote
-% symbols will have been turned off to prevent \thanks from creating footnote marks.
-% \IEEEauthorrefmark{} produces a symbol that appears to LaTeX as having zero vertical
-% height - this allows for a more compact line packing, but the user must ensure that
-% the interline spacing is large enough to prevent \IEEEauthorrefmark{} from colliding
-% with the text above.
-% V1.7 make this a robust command
-\DeclareRobustCommand*{\IEEEauthorrefmark}[1]{\raisebox{0pt}[0pt][0pt]{\textsuperscript{\footnotesize\ensuremath{\ifcase#1\or *\or \dagger\or \ddagger\or%
- \mathsection\or \mathparagraph\or \|\or **\or \dagger\dagger%
- \or \ddagger\ddagger \else\textsuperscript{\expandafter\romannumeral#1}\fi}}}}
-
-
-% FONT CONTROLS AND SPACINGS FOR CONFERENCE MODE AUTHOR NAME AND AFFILIATION BLOCKS
-%
-% The default font styles for the author name and affiliation blocks (confmode)
-%\def\@IEEEauthorblockNstyle{\normalfont\@IEEEcompsocnotconfonly{\sffamily}\sublargesize\@IEEEcompsocconfonly{\large}}
-%\def\@IEEEauthorblockAstyle{\normalfont\@IEEEcompsocnotconfonly{\sffamily}\@IEEEcompsocconfonly{\itshape}\normalsize\@IEEEcompsocconfonly{\large}}
-\def\@IEEEauthorblockNstyle{\normalfont\normalsize}
-\def\@IEEEauthorblockAstyle{\normalfont\@IEEEcompsocnotconfonly{\sffamily}\@IEEEcompsocconfonly{\itshape}\normalsize}
-
-% The default if the user does not use an author block
-\def\@IEEEauthordefaulttextstyle{\normalfont\@IEEEcompsocnotconfonly{\sffamily}\sublargesize}
-
-% spacing from title (or special paper notice) to author name blocks (confmode)
-% can be negative
-\def\@IEEEauthorblockconfadjspace{-0.25em}
-% compsoc conferences need more space here
-\@IEEEcompsocconfonly{\def\@IEEEauthorblockconfadjspace{0.75\@IEEEnormalsizeunitybaselineskip}}
-
-% spacing between name and affiliation blocks (confmode)
-% This can be negative.
-% IEEE doesn't want any added spacing here, but I will leave these
-% controls in place in case they ever change their mind.
-% Personally, I like 0.75ex.
-%\def\@IEEEauthorblockNtopspace{0.75ex}
-%\def\@IEEEauthorblockAtopspace{0.75ex}
-\def\@IEEEauthorblockNtopspace{0.0ex}
-\def\@IEEEauthorblockAtopspace{0.0ex}
-% baseline spacing within name and affiliation blocks (confmode)
-% must be positive, spacings below certain values will make
-% the position of line of text sensitive to the contents of the
-% line above it i.e., whether or not the prior line has descenders,
-% subscripts, etc. For this reason it is a good idea to keep
-% these above 2.6ex
-\def\@IEEEauthorblockNinterlinespace{2.6ex}
-\def\@IEEEauthorblockAinterlinespace{2.75ex}
-
-% This tracks the required strut size.
-% See the \@IEEEauthorhalign command for the actual default value used.
-\def\@IEEEauthorblockXinterlinespace{2.7ex}
-
-% variables to retain font size and style across groups
-% values given here have no effect as they will be overwritten later
-\gdef\@IEEESAVESTATEfontsize{10}
-\gdef\@IEEESAVESTATEfontbaselineskip{12}
-\gdef\@IEEESAVESTATEfontencoding{OT1}
-\gdef\@IEEESAVESTATEfontfamily{ptm}
-\gdef\@IEEESAVESTATEfontseries{m}
-\gdef\@IEEESAVESTATEfontshape{n}
-
-% saves the current font attributes
-\def\@IEEEcurfontSAVE{\global\let\@IEEESAVESTATEfontsize\f@size%
-\global\let\@IEEESAVESTATEfontbaselineskip\f@baselineskip%
-\global\let\@IEEESAVESTATEfontencoding\f@encoding%
-\global\let\@IEEESAVESTATEfontfamily\f@family%
-\global\let\@IEEESAVESTATEfontseries\f@series%
-\global\let\@IEEESAVESTATEfontshape\f@shape}
-
-% restores the saved font attributes
-\def\@IEEEcurfontRESTORE{\fontsize{\@IEEESAVESTATEfontsize}{\@IEEESAVESTATEfontbaselineskip}%
-\fontencoding{\@IEEESAVESTATEfontencoding}%
-\fontfamily{\@IEEESAVESTATEfontfamily}%
-\fontseries{\@IEEESAVESTATEfontseries}%
-\fontshape{\@IEEESAVESTATEfontshape}%
-\selectfont}
-
-
-% variable to indicate if the current block is the first block in the column
-\newif\if@IEEEprevauthorblockincol \@IEEEprevauthorblockincolfalse
-
-
-% the command places a strut with height and depth = \@IEEEauthorblockXinterlinespace
-% we use this technique to have complete manual control over the spacing of the lines
-% within the halign environment.
-% We set the below baseline portion at 30%, the above
-% baseline portion at 70% of the total length.
-% Responds to changes in the document's \baselinestretch
-\def\@IEEEauthorstrutrule{\@IEEEtrantmpdimenA\@IEEEauthorblockXinterlinespace%
-\@IEEEtrantmpdimenA=\baselinestretch\@IEEEtrantmpdimenA%
-\rule[-0.3\@IEEEtrantmpdimenA]{0pt}{\@IEEEtrantmpdimenA}}
-
-
-% blocks to hold the authors' names and affilations.
-% Makes formatting easy for conferences
-%
-% use real definitions in conference mode
-% name block
-\def\IEEEauthorblockN#1{\relax\@IEEEauthorblockNstyle% set the default text style
-\gdef\@IEEEauthorblockXinterlinespace{0pt}% disable strut for spacer row
-% the \expandafter hides the \cr in conditional tex, see the array.sty docs
-% for details, probably not needed here as the \cr is in a macro
-% do a spacer row if needed
-\if@IEEEprevauthorblockincol\expandafter\@IEEEauthorblockNtopspaceline\fi
-\global\@IEEEprevauthorblockincoltrue% we now have a block in this column
-%restore the correct strut value
-\gdef\@IEEEauthorblockXinterlinespace{\@IEEEauthorblockNinterlinespace}%
-% input the author names
-#1%
-% end the row if the user did not already
-\crcr}
-% spacer row for names
-\def\@IEEEauthorblockNtopspaceline{\cr\noalign{\vskip\@IEEEauthorblockNtopspace}}
-%
-% affiliation block
-\def\IEEEauthorblockA#1{\relax\@IEEEauthorblockAstyle% set the default text style
-\gdef\@IEEEauthorblockXinterlinespace{0pt}%disable strut for spacer row
-% the \expandafter hides the \cr in conditional tex, see the array.sty docs
-% for details, probably not needed here as the \cr is in a macro
-% do a spacer row if needed
-\if@IEEEprevauthorblockincol\expandafter\@IEEEauthorblockAtopspaceline\fi
-\global\@IEEEprevauthorblockincoltrue% we now have a block in this column
-%restore the correct strut value
-\gdef\@IEEEauthorblockXinterlinespace{\@IEEEauthorblockAinterlinespace}%
-% input the author affiliations
-#1%
-% end the row if the user did not already
-\crcr}
-% spacer row for affiliations
-\def\@IEEEauthorblockAtopspaceline{\cr\noalign{\vskip\@IEEEauthorblockAtopspace}}
-
-
-% allow papers to compile even if author blocks are used in modes other
-% than conference or peerreviewca. For such cases, we provide dummy blocks.
-\ifCLASSOPTIONconference
-\else
- \ifCLASSOPTIONpeerreviewca\else
- % not conference or peerreviewca mode
- \def\IEEEauthorblockN#1{#1}%
- \def\IEEEauthorblockA#1{#1}%
- \fi
-\fi
-
-
-
-% we provide our own halign so as not to have to depend on tabular
-\def\@IEEEauthorhalign{\@IEEEauthordefaulttextstyle% default text style
- \lineskip=0pt\relax% disable line spacing
- \lineskiplimit=0pt\relax%
- \baselineskip=0pt\relax%
- \@IEEEcurfontSAVE% save the current font
- \mathsurround\z@\relax% no extra spacing around math
- \let\\\@IEEEauthorhaligncr% replace newline with halign friendly one
- \tabskip=0pt\relax% no column spacing
- \everycr{}% ensure no problems here
- \@IEEEprevauthorblockincolfalse% no author blocks yet
- \def\@IEEEauthorblockXinterlinespace{2.7ex}% default interline space
- \vtop\bgroup%vtop box
- \halign\bgroup&\relax\hfil\@IEEEcurfontRESTORE\relax ##\relax
- \hfil\@IEEEcurfontSAVE\@IEEEauthorstrutrule\cr}
-
-% ensure last line, exit from halign, close vbox
-\def\end@IEEEauthorhalign{\crcr\egroup\egroup}
-
-% handle bogus star form
-\def\@IEEEauthorhaligncr{{\ifnum0=`}\fi\@ifstar{\@@IEEEauthorhaligncr}{\@@IEEEauthorhaligncr}}
-
-% test and setup the optional argument to \\[]
-\def\@@IEEEauthorhaligncr{\@testopt\@@@IEEEauthorhaligncr\z@skip}
-
-% end the line and do the optional spacer
-\def\@@@IEEEauthorhaligncr[#1]{\ifnum0=`{\fi}\cr\noalign{\vskip#1\relax}}
-
-
-
-% flag to prevent multiple \and warning messages
-\newif\if@IEEEWARNand
-\@IEEEWARNandtrue
-
-% if in conference or peerreviewca modes, we support the use of \and as \author is a
-% tabular environment, otherwise we warn the user that \and is invalid
-% outside of conference or peerreviewca modes.
-\def\and{\relax} % provide a bogus \and that we will then override
-
-\renewcommand{\and}[1][\relax]{\if@IEEEWARNand\typeout{** WARNING: \noexpand\and is valid only
- when in conference or peerreviewca}\typeout{modes (line \the\inputlineno).}\fi\global\@IEEEWARNandfalse}
-
-\ifCLASSOPTIONconference%
-\renewcommand{\and}[1][\hfill]{\end{@IEEEauthorhalign}#1\begin{@IEEEauthorhalign}}%
-\fi
-\ifCLASSOPTIONpeerreviewca
-\renewcommand{\and}[1][\hfill]{\end{@IEEEauthorhalign}#1\begin{@IEEEauthorhalign}}%
-\fi
-
-
-% page clearing command
-% based on LaTeX2e's \cleardoublepage, but allows different page styles
-% for the inserted blank pages
-\def\@IEEEcleardoublepage#1{\clearpage\if@twoside\ifodd\c@page\else
-\hbox{}\thispagestyle{#1}\newpage\if@twocolumn\hbox{}\thispagestyle{#1}\newpage\fi\fi\fi}
-
-
-% user command to invoke the title page
-\def\maketitle{\par%
- \begingroup%
- \normalfont%
- \def\thefootnote{}% the \thanks{} mark type is empty
- \def\footnotemark{}% and kill space from \thanks within author
- \let\@makefnmark\relax% V1.7, must *really* kill footnotemark to remove all \textsuperscript spacing as well.
- \footnotesize% equal spacing between thanks lines
- \footnotesep 0.7\baselineskip%see global setting of \footnotesep for more info
- % V1.7 disable \thanks note indention for compsoc
- \@IEEEcompsoconly{\long\def\@makefntext##1{\parindent 1em\noindent\hbox{\@makefnmark}##1}}%
- \normalsize%
- \ifCLASSOPTIONpeerreview
- \newpage\global\@topnum\z@ \@maketitle\@IEEEstatictitlevskip\@IEEEaftertitletext%
- \thispagestyle{IEEEpeerreviewcoverpagestyle}\@thanks%
- \else
- \if@twocolumn%
- \ifCLASSOPTIONtechnote%
- \newpage\global\@topnum\z@ \@maketitle\@IEEEstatictitlevskip\@IEEEaftertitletext%
- \else
- \twocolumn[\@maketitle\@IEEEdynamictitlevspace\@IEEEaftertitletext]%
- \fi
- \else
- \newpage\global\@topnum\z@ \@maketitle\@IEEEstatictitlevskip\@IEEEaftertitletext%
- \fi
- \thispagestyle{IEEEtitlepagestyle}\@thanks%
- \fi
- % pullup page for pubid if used.
- \if@IEEEusingpubid
- \enlargethispage{-\@IEEEpubidpullup}%
- \fi
- \endgroup
- \setcounter{footnote}{0}\let\maketitle\relax\let\@maketitle\relax
- \gdef\@thanks{}%
- % v1.6b do not clear these as we will need the title again for peer review papers
- % \gdef\@author{}\gdef\@title{}%
- \let\thanks\relax}
-
-
-
-% V1.7 parbox to format \@IEEEcompsoctitleabstractindextext
-\long\def\@IEEEcompsoctitleabstractindextextbox#1{\parbox{0.915\textwidth}{#1}}
-
-
-% formats the Title, authors names, affiliations and special paper notice
-% THIS IS A CONTROLLED SPACING COMMAND! Do not allow blank lines or unintentional
-% spaces to enter the definition - use % at the end of each line
-\def\@maketitle{\newpage
-\begin{center}%
-\ifCLASSOPTIONtechnote% technotes
- {\bfseries\large\@IEEEcompsoconly{\sffamily}\@title\par}\vskip 1.3em{\lineskip .5em\@IEEEcompsoconly{\sffamily}\@author
- \@IEEEspecialpapernotice\par{\@IEEEcompsoconly{\vskip 1.5em\relax
- \@IEEEcompsoctitleabstractindextextbox{\@IEEEcompsoctitleabstractindextext}\par
- \hfill\@IEEEcompsocdiamondline\hfill\hbox{}\par}}}\relax
-\else% not a technote
- \vskip0.2em{\Huge\@IEEEcompsoconly{\sffamily}\@IEEEcompsocconfonly{\normalfont\normalsize\vskip 0\@IEEEnormalsizeunitybaselineskip
- \bfseries\Large}\@title\par}\vskip 1.0em\par%
- % V1.6 handle \author differently if in conference mode
- \ifCLASSOPTIONconference%
- {\@IEEEspecialpapernotice\mbox{}\vskip\@IEEEauthorblockconfadjspace%
- \mbox{}\hfill\begin{@IEEEauthorhalign}\@author\end{@IEEEauthorhalign}\hfill\mbox{}\par}\relax
- \else% peerreviewca, peerreview or journal
- \ifCLASSOPTIONpeerreviewca
- % peerreviewca handles author names just like conference mode
- {\@IEEEcompsoconly{\sffamily}\@IEEEspecialpapernotice\mbox{}\vskip\@IEEEauthorblockconfadjspace%
- \mbox{}\hfill\begin{@IEEEauthorhalign}\@author\end{@IEEEauthorhalign}\hfill\mbox{}\par
- {\@IEEEcompsoconly{\vskip 1.5em\relax
- \@IEEEcompsoctitleabstractindextextbox{\@IEEEcompsoctitleabstractindextext}\par\hfill
- \@IEEEcompsocdiamondline\hfill\hbox{}\par}}}\relax
- \else% journal or peerreview
- {\lineskip.5em\@IEEEcompsoconly{\sffamily}\sublargesize\@author\@IEEEspecialpapernotice\par
- {\@IEEEcompsoconly{\vskip 1.5em\relax
- \@IEEEcompsoctitleabstractindextextbox{\@IEEEcompsoctitleabstractindextext}\par\hfill
- \@IEEEcompsocdiamondline\hfill\hbox{}\par}}}\relax
- \fi
- \fi
-\fi\end{center}}
-
-
-
-% V1.7 Computer Society "diamond line" which follows index terms for nonconference papers
-\def\@IEEEcompsocdiamondline{\vrule depth 0pt height 0.5pt width 4cm\hspace{7.5pt}%
-\raisebox{-3.5pt}{\fontfamily{pzd}\fontencoding{U}\fontseries{m}\fontshape{n}\fontsize{11}{12}\selectfont\char70}%
-\hspace{7.5pt}\vrule depth 0pt height 0.5pt width 4cm\relax}
-
-% V1.7 standard LateX2e \thanks, but with \itshape under compsoc. Also make it a \long\def
-% We also need to trigger the one-shot footnote rule
-\def\@IEEEtriggeroneshotfootnoterule{\global\@IEEEenableoneshotfootnoteruletrue}
-
-
-\long\def\thanks#1{\footnotemark
- \protected@xdef\@thanks{\@thanks
- \protect\footnotetext[\the\c@footnote]{\@IEEEcompsoconly{\itshape
- \protect\@IEEEtriggeroneshotfootnoterule\relax}\ignorespaces#1}}}
-\let\@thanks\@empty
-
-% V1.7 allow \author to contain \par's. This is needed to allow \thanks to contain \par.
-\long\def\author#1{\gdef\@author{#1}}
-
-
-% in addition to setting up IEEEitemize, we need to remove a baselineskip space above and
-% below it because \list's \pars introduce blank lines because of the footnote struts.
-\def\@IEEEsetupcompsocitemizelist{\def\labelitemi{$\bullet$}%
-\setlength{\IEEElabelindent}{0pt}\setlength{\parskip}{0pt}%
-\setlength{\partopsep}{0pt}\setlength{\topsep}{0.5\baselineskip}\vspace{-1\baselineskip}\relax}
-
-
-% flag for fake non-compsoc \IEEEcompsocthanksitem - prevents line break on very first item
-\newif\if@IEEEbreakcompsocthanksitem \@IEEEbreakcompsocthanksitemfalse
-
-\ifCLASSOPTIONcompsoc
-% V1.7 compsoc bullet item \thanks
-% also, we need to redefine this to destroy the argument in \@IEEEdynamictitlevspace
-\long\def\IEEEcompsocitemizethanks#1{\relax\@IEEEbreakcompsocthanksitemfalse\footnotemark
- \protected@xdef\@thanks{\@thanks
- \protect\footnotetext[\the\c@footnote]{\itshape\protect\@IEEEtriggeroneshotfootnoterule
- {\let\IEEEiedlistdecl\relax\protect\begin{IEEEitemize}[\protect\@IEEEsetupcompsocitemizelist]\ignorespaces#1\relax
- \protect\end{IEEEitemize}}\protect\vspace{-1\baselineskip}}}}
-\DeclareRobustCommand*{\IEEEcompsocthanksitem}{\item}
-\else
-% non-compsoc, allow for dual compilation via rerouting to normal \thanks
-\long\def\IEEEcompsocitemizethanks#1{\thanks{#1}}
-% redirect to "pseudo-par" \hfil\break\indent after swallowing [] from \IEEEcompsocthanksitem[]
-\DeclareRobustCommand{\IEEEcompsocthanksitem}{\@ifnextchar [{\@IEEEthanksswallowoptionalarg}%
-{\@IEEEthanksswallowoptionalarg[\relax]}}
-% be sure and break only after first item, be sure and ignore spaces after optional argument
-\def\@IEEEthanksswallowoptionalarg[#1]{\relax\if@IEEEbreakcompsocthanksitem\hfil\break
-\indent\fi\@IEEEbreakcompsocthanksitemtrue\ignorespaces}
-\fi
-
-
-% V1.6b define the \IEEEpeerreviewmaketitle as needed
-\ifCLASSOPTIONpeerreview
-\def\IEEEpeerreviewmaketitle{\@IEEEcleardoublepage{empty}%
-\ifCLASSOPTIONtwocolumn
-\twocolumn[\@IEEEpeerreviewmaketitle\@IEEEdynamictitlevspace]
-\else
-\newpage\@IEEEpeerreviewmaketitle\@IEEEstatictitlevskip
-\fi
-\thispagestyle{IEEEtitlepagestyle}}
-\else
-% \IEEEpeerreviewmaketitle does nothing if peer review option has not been selected
-\def\IEEEpeerreviewmaketitle{\relax}
-\fi
-
-% peerreview formats the repeated title like the title in journal papers.
-\def\@IEEEpeerreviewmaketitle{\begin{center}\@IEEEcompsoconly{\sffamily}%
-\normalfont\normalsize\vskip0.2em{\Huge\@title\par}\vskip1.0em\par
-\end{center}}
-
-
-
-% V1.6
-% this is a static rubber spacer between the title/authors and the main text
-% used for single column text, or when the title appears in the first column
-% of two column text (technotes).
-\def\@IEEEstatictitlevskip{{\normalfont\normalsize
-% adjust spacing to next text
-% v1.6b handle peer review papers
-\ifCLASSOPTIONpeerreview
-% for peer review papers, the same value is used for both title pages
-% regardless of the other paper modes
- \vskip 1\baselineskip plus 0.375\baselineskip minus 0.1875\baselineskip
-\else
- \ifCLASSOPTIONconference% conference
- \vskip 1\baselineskip plus 0.375\baselineskip minus 0.1875\baselineskip%
- \else%
- \ifCLASSOPTIONtechnote% technote
- \vskip 1\baselineskip plus 0.375\baselineskip minus 0.1875\baselineskip%
- \else% journal uses more space
- \vskip 2.5\baselineskip plus 0.75\baselineskip minus 0.375\baselineskip%
- \fi
- \fi
-\fi}}
-
-
-% V1.6
-% This is a dynamically determined rigid spacer between the title/authors
-% and the main text. This is used only for single column titles over two
-% column text (most common)
-% This is bit tricky because we have to ensure that the textheight of the
-% main text is an integer multiple of \baselineskip
-% otherwise underfull vbox problems may develop in the second column of the
-% text on the titlepage
-% The possible use of \IEEEpubid must also be taken into account.
-\def\@IEEEdynamictitlevspace{{%
- % we run within a group so that all the macros can be forgotten when we are done
- \long\def\thanks##1{\relax}%don't allow \thanks to run when we evaluate the vbox height
- \long\def\IEEEcompsocitemizethanks##1{\relax}%don't allow \IEEEcompsocitemizethanks to run when we evaluate the vbox height
- \normalfont\normalsize% we declare more descriptive variable names
- \let\@IEEEmaintextheight=\@IEEEtrantmpdimenA%height of the main text columns
- \let\@IEEEINTmaintextheight=\@IEEEtrantmpdimenB%height of the main text columns with integer # lines
- % set the nominal and minimum values for the title spacer
- % the dynamic algorithm will not allow the spacer size to
- % become less than \@IEEEMINtitlevspace - instead it will be
- % lengthened
- % default to journal values
- \def\@IEEENORMtitlevspace{2.5\baselineskip}%
- \def\@IEEEMINtitlevspace{2\baselineskip}%
- % conferences and technotes need tighter spacing
- \ifCLASSOPTIONconference%conference
- \def\@IEEENORMtitlevspace{1\baselineskip}%
- \def\@IEEEMINtitlevspace{0.75\baselineskip}%
- \fi
- \ifCLASSOPTIONtechnote%technote
- \def\@IEEENORMtitlevspace{1\baselineskip}%
- \def\@IEEEMINtitlevspace{0.75\baselineskip}%
- \fi%
- % get the height that the title will take up
- \ifCLASSOPTIONpeerreview
- \settoheight{\@IEEEmaintextheight}{\vbox{\hsize\textwidth \@IEEEpeerreviewmaketitle}}%
- \else
- \settoheight{\@IEEEmaintextheight}{\vbox{\hsize\textwidth \@maketitle}}%
- \fi
- \@IEEEmaintextheight=-\@IEEEmaintextheight% title takes away from maintext, so reverse sign
- % add the height of the page textheight
- \advance\@IEEEmaintextheight by \textheight%
- % correct for title pages using pubid
- \ifCLASSOPTIONpeerreview\else
- % peerreview papers use the pubid on the cover page only.
- % And the cover page uses a static spacer.
- \if@IEEEusingpubid\advance\@IEEEmaintextheight by -\@IEEEpubidpullup\fi
- \fi%
- % subtract off the nominal value of the title bottom spacer
- \advance\@IEEEmaintextheight by -\@IEEENORMtitlevspace%
- % \topskip takes away some too
- \advance\@IEEEmaintextheight by -\topskip%
- % calculate the column height of the main text for lines
- % now we calculate the main text height as if holding
- % an integer number of \normalsize lines after the first
- % and discard any excess fractional remainder
- % we subtracted the first line, because the first line
- % is placed \topskip into the maintext, not \baselineskip like the
- % rest of the lines.
- \@IEEEINTmaintextheight=\@IEEEmaintextheight%
- \divide\@IEEEINTmaintextheight by \baselineskip%
- \multiply\@IEEEINTmaintextheight by \baselineskip%
- % now we calculate how much the title spacer height will
- % have to be reduced from nominal (\@IEEEREDUCEmaintextheight is always
- % a positive value) so that the maintext area will contain an integer
- % number of normal size lines
- % we change variable names here (to avoid confusion) as we no longer
- % need \@IEEEINTmaintextheight and can reuse its dimen register
- \let\@IEEEREDUCEmaintextheight=\@IEEEINTmaintextheight%
- \advance\@IEEEREDUCEmaintextheight by -\@IEEEmaintextheight%
- \advance\@IEEEREDUCEmaintextheight by \baselineskip%
- % this is the calculated height of the spacer
- % we change variable names here (to avoid confusion) as we no longer
- % need \@IEEEmaintextheight and can reuse its dimen register
- \let\@IEEECOMPENSATElen=\@IEEEmaintextheight%
- \@IEEECOMPENSATElen=\@IEEENORMtitlevspace% set the nominal value
- % we go with the reduced length if it is smaller than an increase
- \ifdim\@IEEEREDUCEmaintextheight < 0.5\baselineskip\relax%
- \advance\@IEEECOMPENSATElen by -\@IEEEREDUCEmaintextheight%
- % if the resulting spacer is too small back out and go with an increase instead
- \ifdim\@IEEECOMPENSATElen<\@IEEEMINtitlevspace\relax%
- \advance\@IEEECOMPENSATElen by \baselineskip%
- \fi%
- \else%
- % go with an increase because it is closer to the nominal than a decrease
- \advance\@IEEECOMPENSATElen by -\@IEEEREDUCEmaintextheight%
- \advance\@IEEECOMPENSATElen by \baselineskip%
- \fi%
- % set the calculated rigid spacer
- \vspace{\@IEEECOMPENSATElen}}}
-
-
-
-% V1.6
-% we allow the user access to the last part of the title area
-% useful in emergencies such as when a different spacing is needed
-% This text is NOT compensated for in the dynamic sizer.
-\let\@IEEEaftertitletext=\relax
-\long\def\IEEEaftertitletext#1{\def\@IEEEaftertitletext{#1}}
-
-% V1.7 provide a way for users to enter abstract and keywords
-% into the onecolumn title are. This text is compensated for
-% in the dynamic sizer.
-\let\@IEEEcompsoctitleabstractindextext=\relax
-\long\def\IEEEcompsoctitleabstractindextext#1{\def\@IEEEcompsoctitleabstractindextext{#1}}
-% V1.7 provide a way for users to get the \@IEEEcompsoctitleabstractindextext if
-% not in compsoc journal mode - this way abstract and keywords can be placed
-% in their conventional position if not in compsoc mode.
-\def\IEEEdisplaynotcompsoctitleabstractindextext{%
-\ifCLASSOPTIONcompsoc% display if compsoc conf
-\ifCLASSOPTIONconference\@IEEEcompsoctitleabstractindextext\fi
-\else% or if not compsoc
-\@IEEEcompsoctitleabstractindextext\fi}
-
-
-% command to allow alteration of baselinestretch, but only if the current
-% baselineskip is unity. Used to tweak the compsoc abstract and keywords line spacing.
-\def\@IEEEtweakunitybaselinestretch#1{{\def\baselinestretch{1}\selectfont
-\global\@tempskipa\baselineskip}\ifnum\@tempskipa=\baselineskip%
-\def\baselinestretch{#1}\selectfont\fi\relax}
-
-
-% abstract and keywords are in \small, except
-% for 9pt docs in which they are in \footnotesize
-% Because 9pt docs use an 8pt footnotesize, \small
-% becomes a rather awkward 8.5pt
-\def\@IEEEabskeysecsize{\small}
-\ifx\CLASSOPTIONpt\@IEEEptsizenine
- \def\@IEEEabskeysecsize{\footnotesize}
-\fi
-
-% compsoc journals use \footnotesize, compsoc conferences use normalsize
-\@IEEEcompsoconly{\def\@IEEEabskeysecsize{\footnotesize}}
-%\@IEEEcompsocconfonly{\def\@IEEEabskeysecsize{\normalsize}}
-
-
-
-
-% V1.6 have abstract and keywords strip leading spaces, pars and newlines
-% so that spacing is more tightly controlled.
-\def\abstract{\normalfont
- \if@twocolumn
- \@IEEEabskeysecsize\bfseries\textit{\abstractname}---\relax
- \else
- \begin{center}\vspace{-1.78ex}\@IEEEabskeysecsize\textbf{\abstractname}\end{center}\quotation\@IEEEabskeysecsize
- \fi\@IEEEgobbleleadPARNLSP}
-% V1.6 IEEE wants only 1 pica from end of abstract to introduction heading when in
-% conference mode (the heading already has this much above it)
-\def\endabstract{\relax\ifCLASSOPTIONconference\vspace{1.34ex}\else\vspace{1.34ex}\fi\par\if@twocolumn\else\endquotation\fi
- \normalfont\normalsize}
-
-\def\IEEEkeywords{\normalfont
- \if@twocolumn
- \@IEEEabskeysecsize\bfseries\textit{\IEEEkeywordsname}-\relax
- \else
- \begin{center}\@IEEEabskeysecsize\textbf{\IEEEkeywordsname}\end{center}\quotation\@IEEEabskeysecsize
- \fi\@IEEEgobbleleadPARNLSP}
-\def\endIEEEkeywords{\relax\ifCLASSOPTIONtechnote\vspace{1.34ex}\else\vspace{0.67ex}\fi
- \par\if@twocolumn\else\endquotation\fi%
- \normalfont\normalsize}
-
-% V1.7 compsoc keywords index terms
-\ifCLASSOPTIONcompsoc
- \ifCLASSOPTIONconference% compsoc conference
-\def\abstract{\normalfont
- \begin{center}\@IEEEabskeysecsize\textbf{\large\abstractname}\end{center}\vskip 0.5\baselineskip plus 0.1\baselineskip minus 0.1\baselineskip
- \if@twocolumn\else\quotation\fi\itshape\@IEEEabskeysecsize%
- \par\@IEEEgobbleleadPARNLSP}
-\def\IEEEkeywords{\normalfont\vskip 1.5\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip
- \begin{center}\@IEEEabskeysecsize\textbf{\large\IEEEkeywordsname}\end{center}\vskip 0.5\baselineskip plus 0.1\baselineskip minus 0.1\baselineskip
- \if@twocolumn\else\quotation\fi\itshape\@IEEEabskeysecsize%
- \par\@IEEEgobbleleadPARNLSP}
- \else% compsoc not conference
-\def\abstract{\normalfont\@IEEEtweakunitybaselinestretch{1.15}\sffamily
- \if@twocolumn
- \@IEEEabskeysecsize\noindent\textbf{\abstractname}---\relax
- \else
- \begin{center}\vspace{-1.78ex}\@IEEEabskeysecsize\textbf{\abstractname}\end{center}\quotation\@IEEEabskeysecsize%
- \fi\@IEEEgobbleleadPARNLSP}
-\def\IEEEkeywords{\normalfont\@IEEEtweakunitybaselinestretch{1.15}\sffamily
- \if@twocolumn
- \@IEEEabskeysecsize\vskip 0.5\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip\noindent
- \textbf{\IEEEkeywordsname}---\relax
- \else
- \begin{center}\@IEEEabskeysecsize\textbf{\IEEEkeywordsname}\end{center}\quotation\@IEEEabskeysecsize%
- \fi\@IEEEgobbleleadPARNLSP}
- \fi
-\fi
-
-
-
-% gobbles all leading \, \\ and \par, upon finding first token that
-% is not a \ , \\ or a \par, it ceases and returns that token
-%
-% used to strip leading \, \\ and \par from the input
-% so that such things in the beginning of an environment will not
-% affect the formatting of the text
-\long\def\@IEEEgobbleleadPARNLSP#1{\let\@IEEEswallowthistoken=0%
-\let\@IEEEgobbleleadPARNLSPtoken#1%
-\let\@IEEEgobbleleadPARtoken=\par%
-\let\@IEEEgobbleleadNLtoken=\\%
-\let\@IEEEgobbleleadSPtoken=\ %
-\def\@IEEEgobbleleadSPMACRO{\ }%
-\ifx\@IEEEgobbleleadPARNLSPtoken\@IEEEgobbleleadPARtoken%
-\let\@IEEEswallowthistoken=1%
-\fi%
-\ifx\@IEEEgobbleleadPARNLSPtoken\@IEEEgobbleleadNLtoken%
-\let\@IEEEswallowthistoken=1%
-\fi%
-\ifx\@IEEEgobbleleadPARNLSPtoken\@IEEEgobbleleadSPtoken%
-\let\@IEEEswallowthistoken=1%
-\fi%
-% a control space will come in as a macro
-% when it is the last one on a line
-\ifx\@IEEEgobbleleadPARNLSPtoken\@IEEEgobbleleadSPMACRO%
-\let\@IEEEswallowthistoken=1%
-\fi%
-% if we have to swallow this token, do so and taste the next one
-% else spit it out and stop gobbling
-\ifx\@IEEEswallowthistoken 1\let\@IEEEnextgobbleleadPARNLSP=\@IEEEgobbleleadPARNLSP\else%
-\let\@IEEEnextgobbleleadPARNLSP=#1\fi%
-\@IEEEnextgobbleleadPARNLSP}%
-
-
-
-
-% TITLING OF SECTIONS
-\def\@IEEEsectpunct{:\ \,} % Punctuation after run-in section heading (headings which are
- % part of the paragraphs), need little bit more than a single space
- % spacing from section number to title
-% compsoc conferences use regular period/space punctuation
-\ifCLASSOPTIONcompsoc
-\ifCLASSOPTIONconference
-\def\@IEEEsectpunct{.\ }
-\fi\fi
-
-
-\def\@seccntformat#1{\csname the#1dis\endcsname\hskip 0.5em\relax}
-
-\ifCLASSOPTIONcompsoc
-% compsoc journals need extra spacing
-\ifCLASSOPTIONconference\else
-\def\@seccntformat#1{\csname the#1dis\endcsname\hskip 1em\relax}
-\fi\fi
-
-%v1.7 put {} after #6 to allow for some types of user font control
-%and use \@@par rather than \par
-\def\@sect#1#2#3#4#5#6[#7]#8{%
- \ifnum #2>\c@secnumdepth
- \let\@svsec\@empty
- \else
- \refstepcounter{#1}%
- % load section label and spacer into \@svsec
- \protected@edef\@svsec{\@seccntformat{#1}\relax}%
- \fi%
- \@tempskipa #5\relax
- \ifdim \@tempskipa>\z@% tempskipa determines whether is treated as a high
- \begingroup #6{\relax% or low level heading
- \noindent % subsections are NOT indented
- % print top level headings. \@svsec is label, #8 is heading title
- % IEEE does not block indent the section title text, it flows like normal
- {\hskip #3\relax\@svsec}{\interlinepenalty \@M #8\@@par}}%
- \endgroup
- \addcontentsline{toc}{#1}{\ifnum #2>\c@secnumdepth\relax\else
- \protect\numberline{\csname the#1\endcsname}\fi#7}%
- \else % printout low level headings
- % svsechd seems to swallow the trailing space, protect it with \mbox{}
- % got rid of sectionmark stuff
- \def\@svsechd{#6{\hskip #3\relax\@svsec #8\@IEEEsectpunct\mbox{}}%
- \addcontentsline{toc}{#1}{\ifnum #2>\c@secnumdepth\relax\else
- \protect\numberline{\csname the#1\endcsname}\fi#7}}%
- \fi%skip down
- \@xsect{#5}}
-
-
-% section* handler
-%v1.7 put {} after #4 to allow for some types of user font control
-%and use \@@par rather than \par
-\def\@ssect#1#2#3#4#5{\@tempskipa #3\relax
- \ifdim \@tempskipa>\z@
- %\begingroup #4\@hangfrom{\hskip #1}{\interlinepenalty \@M #5\par}\endgroup
- % IEEE does not block indent the section title text, it flows like normal
- \begingroup \noindent #4{\relax{\hskip #1}{\interlinepenalty \@M #5\@@par}}\endgroup
- % svsechd swallows the trailing space, protect it with \mbox{}
- \else \def\@svsechd{#4{\hskip #1\relax #5\@IEEEsectpunct\mbox{}}}\fi
- \@xsect{#3}}
-
-
-%% SECTION heading spacing and font
-%%
-% arguments are: #1 - sectiontype name
-% (for \@sect) #2 - section level
-% #3 - section heading indent
-% #4 - top separation (absolute value used, neg indicates not to indent main text)
-% If negative, make stretch parts negative too!
-% #5 - (absolute value used) positive: bottom separation after heading,
-% negative: amount to indent main text after heading
-% Both #4 and #5 negative means to indent main text and use negative top separation
-% #6 - font control
-% You've got to have \normalfont\normalsize in the font specs below to prevent
-% trouble when you do something like:
-% \section{Note}{\ttfamily TT-TEXT} is known to ...
-% IEEE sometimes REALLY stretches the area before a section
-% heading by up to about 0.5in. However, it may not be a good
-% idea to let LaTeX have quite this much rubber.
-\ifCLASSOPTIONconference%
-% IEEE wants section heading spacing to decrease for conference mode
-\def\section{\@startsection{section}{1}{\z@}{1.5ex plus 1.5ex minus 0.5ex}%
-{0.7ex plus 1ex minus 0ex}{\normalfont\normalsize\centering\scshape}}%
-\def\subsection{\@startsection{subsection}{2}{\z@}{1.5ex plus 1.5ex minus 0.5ex}%
-{0.7ex plus .5ex minus 0ex}{\normalfont\normalsize\itshape}}%
-\else % for journals
-\def\section{\@startsection{section}{1}{\z@}{3.0ex plus 1.5ex minus 1.5ex}% V1.6 3.0ex from 3.5ex
-{0.7ex plus 1ex minus 0ex}{\normalfont\normalsize\centering\scshape}}%
-\def\subsection{\@startsection{subsection}{2}{\z@}{3.5ex plus 1.5ex minus 1.5ex}%
-{0.7ex plus .5ex minus 0ex}{\normalfont\normalsize\itshape}}%
-\fi
-
-% for both journals and conferences
-% decided to put in a little rubber above the section, might help somebody
-\def\subsubsection{\@startsection{subsubsection}{3}{\parindent}{0ex plus 0.1ex minus 0.1ex}%
-{0ex}{\normalfont\normalsize\itshape}}%
-\def\paragraph{\@startsection{paragraph}{4}{2\parindent}{0ex plus 0.1ex minus 0.1ex}%
-{0ex}{\normalfont\normalsize\itshape}}%
-
-
-% compsoc
-\ifCLASSOPTIONcompsoc
-\ifCLASSOPTIONconference
-% compsoc conference
-\def\section{\@startsection{section}{1}{\z@}{1\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip}%
-{1\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip}{\normalfont\large\bfseries}}%
-\def\subsection{\@startsection{subsection}{2}{\z@}{1\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip}%
-{1\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip}{\normalfont\sublargesize\bfseries}}%
-\def\subsubsection{\@startsection{subsubsection}{3}{\z@}{1\baselineskip plus 0.25\baselineskip minus 0.25\baselineskip}%
-{0ex}{\normalfont\normalsize\bfseries}}%
-\def\paragraph{\@startsection{paragraph}{4}{2\parindent}{0ex plus 0.1ex minus 0.1ex}%
-{0ex}{\normalfont\normalsize}}%
-\else% compsoc journals
-% use negative top separation as compsoc journals do not indent paragraphs after section titles
-\def\section{\@startsection{section}{1}{\z@}{-3ex plus -2ex minus -1.5ex}%
-{0.7ex plus 1ex minus 0ex}{\normalfont\large\sffamily\bfseries\scshape}}%
-% Note that subsection and smaller may not be correct for the Computer Society,
-% I have to look up an example.
-\def\subsection{\@startsection{subsection}{2}{\z@}{-3.5ex plus -1.5ex minus -1.5ex}%
-{0.7ex plus .5ex minus 0ex}{\normalfont\normalsize\sffamily\bfseries}}%
-\def\subsubsection{\@startsection{subsubsection}{3}{\z@}{-2.5ex plus -1ex minus -1ex}%
-{0.5ex plus 0.5ex minus 0ex}{\normalfont\normalsize\sffamily\itshape}}%
-\def\paragraph{\@startsection{paragraph}{4}{2\parindent}{-0ex plus -0.1ex minus -0.1ex}%
-{0ex}{\normalfont\normalsize}}%
-\fi\fi
-
-
-
-
-%% ENVIRONMENTS
-% "box" symbols at end of proofs
-\def\IEEEQEDclosed{\mbox{\rule[0pt]{1.3ex}{1.3ex}}} % for a filled box
-% V1.6 some journals use an open box instead that will just fit around a closed one
-\def\IEEEQEDopen{{\setlength{\fboxsep}{0pt}\setlength{\fboxrule}{0.2pt}\fbox{\rule[0pt]{0pt}{1.3ex}\rule[0pt]{1.3ex}{0pt}}}}
-\ifCLASSOPTIONcompsoc
-\def\IEEEQED{\IEEEQEDopen} % default to open for compsoc
-\else
-\def\IEEEQED{\IEEEQEDclosed} % otherwise default to closed
-\fi
-
-% v1.7 name change to avoid namespace collision with amsthm. Also add support
-% for an optional argument.
-\def\IEEEproof{\@ifnextchar[{\@IEEEproof}{\@IEEEproof[\IEEEproofname]}}
-\def\@IEEEproof[#1]{\par\noindent\hspace{2em}{\itshape #1: }}
-\def\endIEEEproof{\hspace*{\fill}~\IEEEQED\par}
-
-
-%\itemindent is set to \z@ by list, so define new temporary variable
-\newdimen\@IEEEtmpitemindent
-\def\@begintheorem#1#2{\@IEEEtmpitemindent\itemindent\topsep 0pt\rmfamily\trivlist%
- \item[\hskip \labelsep{\indent\itshape #1\ #2:}]\itemindent\@IEEEtmpitemindent}
-\def\@opargbegintheorem#1#2#3{\@IEEEtmpitemindent\itemindent\topsep 0pt\rmfamily \trivlist%
-% V1.6 IEEE is back to using () around theorem names which are also in italics
-% Thanks to Christian Peel for reporting this.
- \item[\hskip\labelsep{\indent\itshape #1\ #2\ (#3):}]\itemindent\@IEEEtmpitemindent}
-% V1.7 remove bogus \unskip that caused equations in theorems to collide with
-% lines below.
-\def\@endtheorem{\endtrivlist}
-
-% V1.6
-% display command for the section the theorem is in - so that \thesection
-% is not used as this will be in Roman numerals when we want arabic.
-% LaTeX2e uses \def\@thmcounter#1{\noexpand\arabic{#1}} for the theorem number
-% (second part) display and \def\@thmcountersep{.} as a separator.
-% V1.7 intercept calls to the section counter and reroute to \@IEEEthmcounterinsection
-% to allow \appendix(ices} to override as needed.
-%
-% special handler for sections, allows appendix(ices) to override
-\gdef\@IEEEthmcounterinsection#1{\arabic{#1}}
-% string macro
-\edef\@IEEEstringsection{section}
-
-% redefine the #1#2[#3] form of newtheorem to use a hook to \@IEEEthmcounterinsection
-% if section in_counter is used
-\def\@xnthm#1#2[#3]{%
- \expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}\@newctr{#1}[#3]%
- \edef\@IEEEstringtmp{#3}
- \ifx\@IEEEstringtmp\@IEEEstringsection
- \expandafter\xdef\csname the#1\endcsname{%
- \noexpand\@IEEEthmcounterinsection{#3}\@thmcountersep
- \@thmcounter{#1}}%
- \else
- \expandafter\xdef\csname the#1\endcsname{%
- \expandafter\noexpand\csname the#3\endcsname \@thmcountersep
- \@thmcounter{#1}}%
- \fi
- \global\@namedef{#1}{\@thm{#1}{#2}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-
-
-%% SET UP THE DEFAULT PAGESTYLE
-\ps@headings
-\pagenumbering{arabic}
-
-% normally the page counter starts at 1
-\setcounter{page}{1}
-% however, for peerreview the cover sheet is page 0 or page -1
-% (for duplex printing)
-\ifCLASSOPTIONpeerreview
- \if@twoside
- \setcounter{page}{-1}
- \else
- \setcounter{page}{0}
- \fi
-\fi
-
-% standard book class behavior - let bottom line float up and down as
-% needed when single sided
-\ifCLASSOPTIONtwoside\else\raggedbottom\fi
-% if two column - turn on twocolumn, allow word spacings to stretch more and
-% enforce a rigid position for the last lines
-\ifCLASSOPTIONtwocolumn
-% the peer review option delays invoking twocolumn
- \ifCLASSOPTIONpeerreview\else
- \twocolumn
- \fi
-\sloppy
-\flushbottom
-\fi
-
-
-
-
-% \APPENDIX and \APPENDICES definitions
-
-% This is the \@ifmtarg command from the LaTeX ifmtarg package
-% by Peter Wilson (CUA) and Donald Arseneau
-% \@ifmtarg is used to determine if an argument to a command
-% is present or not.
-% For instance:
-% \@ifmtarg{#1}{\typeout{empty}}{\typeout{has something}}
-% \@ifmtarg is used with our redefined \section command if
-% \appendices is invoked.
-% The command \section will behave slightly differently depending
-% on whether the user specifies a title:
-% \section{My appendix title}
-% or not:
-% \section{}
-% This way, we can eliminate the blank lines where the title
-% would be, and the unneeded : after Appendix in the table of
-% contents
-\begingroup
-\catcode`\Q=3
-\long\gdef\@ifmtarg#1{\@xifmtarg#1QQ\@secondoftwo\@firstoftwo\@nil}
-\long\gdef\@xifmtarg#1#2Q#3#4#5\@nil{#4}
-\endgroup
-% end of \@ifmtarg defs
-
-
-% V1.7
-% command that allows the one time saving of the original definition
-% of section to \@IEEEappendixsavesection for \appendix or \appendices
-% we don't save \section here as it may be redefined later by other
-% packages (hyperref.sty, etc.)
-\def\@IEEEsaveoriginalsectiononce{\let\@IEEEappendixsavesection\section
-\let\@IEEEsaveoriginalsectiononce\relax}
-
-% neat trick to grab and process the argument from \section{argument}
-% we process differently if the user invoked \section{} with no
-% argument (title)
-% note we reroute the call to the old \section*
-\def\@IEEEprocessthesectionargument#1{%
-\@ifmtarg{#1}{%
-\@IEEEappendixsavesection*{\appendixname~\thesectiondis}%
-\addcontentsline{toc}{section}{\appendixname~\thesection}}{%
-\@IEEEappendixsavesection*{\appendixname~\thesectiondis \\* #1}%
-\addcontentsline{toc}{section}{\appendixname~\thesection: #1}}}
-
-% we use this if the user calls \section{} after
-% \appendix-- which has no meaning. So, we ignore the
-% command and its argument. Then, warn the user.
-\def\@IEEEdestroythesectionargument#1{\typeout{** WARNING: Ignoring useless
-\protect\section\space in Appendix (line \the\inputlineno).}}
-
-
-% remember \thesection forms will be displayed in \ref calls
-% and in the Table of Contents.
-% The \sectiondis form is used in the actual heading itself
-
-% appendix command for one single appendix
-% normally has no heading. However, if you want a
-% heading, you can do so via the optional argument:
-% \appendix[Optional Heading]
-\def\appendix{\relax}
-\renewcommand{\appendix}[1][]{\@IEEEsaveoriginalsectiononce\par
- % v1.6 keep hyperref's identifiers unique
- \gdef\theHsection{Appendix.A}%
- % v1.6 adjust hyperref's string name for the section
- \xdef\Hy@chapapp{appendix}%
- \setcounter{section}{0}%
- \setcounter{subsection}{0}%
- \setcounter{subsubsection}{0}%
- \setcounter{paragraph}{0}%
- \gdef\thesection{A}%
- \gdef\thesectiondis{}%
- \gdef\thesubsection{\Alph{subsection}}%
- \gdef\@IEEEthmcounterinsection##1{A}
- \refstepcounter{section}% update the \ref counter
- \@ifmtarg{#1}{\@IEEEappendixsavesection*{\appendixname}%
- \addcontentsline{toc}{section}{\appendixname}}{%
- \@IEEEappendixsavesection*{\appendixname~\\* #1}%
- \addcontentsline{toc}{section}{\appendixname: #1}}%
- % redefine \section command for appendix
- % leave \section* as is
- \def\section{\@ifstar{\@IEEEappendixsavesection*}{%
- \@IEEEdestroythesectionargument}}% throw out the argument
- % of the normal form
-}
-
-
-
-% appendices command for multiple appendices
-% user then calls \section with an argument (possibly empty) to
-% declare the individual appendices
-\def\appendices{\@IEEEsaveoriginalsectiononce\par
- % v1.6 keep hyperref's identifiers unique
- \gdef\theHsection{Appendix.\Alph{section}}%
- % v1.6 adjust hyperref's string name for the section
- \xdef\Hy@chapapp{appendix}%
- \setcounter{section}{-1}% we want \refstepcounter to use section 0
- \setcounter{subsection}{0}%
- \setcounter{subsubsection}{0}%
- \setcounter{paragraph}{0}%
- \ifCLASSOPTIONromanappendices%
- \gdef\thesection{\Roman{section}}%
- \gdef\thesectiondis{\Roman{section}}%
- \@IEEEcompsocconfonly{\gdef\thesectiondis{\Roman{section}.}}%
- \gdef\@IEEEthmcounterinsection##1{A\arabic{##1}}
- \else%
- \gdef\thesection{\Alph{section}}%
- \gdef\thesectiondis{\Alph{section}}%
- \@IEEEcompsocconfonly{\gdef\thesectiondis{\Alph{section}.}}%
- \gdef\@IEEEthmcounterinsection##1{\Alph{##1}}
- \fi%
- \refstepcounter{section}% update the \ref counter
- \setcounter{section}{0}% NEXT \section will be the FIRST appendix
- % redefine \section command for appendices
- % leave \section* as is
- \def\section{\@ifstar{\@IEEEappendixsavesection*}{% process the *-form
- \refstepcounter{section}% or is a new section so,
- \@IEEEprocessthesectionargument}}% process the argument
- % of the normal form
-}
-
-
-
-% \IEEEPARstart
-% Definition for the big two line drop cap letter at the beginning of the
-% first paragraph of journal papers. The first argument is the first letter
-% of the first word, the second argument is the remaining letters of the
-% first word which will be rendered in upper case.
-% In V1.6 this has been completely rewritten to:
-%
-% 1. no longer have problems when the user begins an environment
-% within the paragraph that uses \IEEEPARstart.
-% 2. auto-detect and use the current font family
-% 3. revise handling of the space at the end of the first word so that
-% interword glue will now work as normal.
-% 4. produce correctly aligned edges for the (two) indented lines.
-%
-% We generalize things via control macros - playing with these is fun too.
-%
-% V1.7 added more control macros to make it easy for IEEEtrantools.sty users
-% to change the font style.
-%
-% the number of lines that are indented to clear it
-% may need to increase if using decenders
-\def\@IEEEPARstartDROPLINES{2}
-% minimum number of lines left on a page to allow a \@IEEEPARstart
-% Does not take into consideration rubber shrink, so it tends to
-% be overly cautious
-\def\@IEEEPARstartMINPAGELINES{2}
-% V1.7 the height of the drop cap is adjusted to match the height of this text
-% in the current font (when \IEEEPARstart is called).
-\def\@IEEEPARstartHEIGHTTEXT{T}
-% the depth the letter is lowered below the baseline
-% the height (and size) of the letter is determined by the sum
-% of this value and the height of the \@IEEEPARstartHEIGHTTEXT in the current
-% font. It is a good idea to set this value in terms of the baselineskip
-% so that it can respond to changes therein.
-\def\@IEEEPARstartDROPDEPTH{1.1\baselineskip}
-% V1.7 the font the drop cap will be rendered in,
-% can take zero or one argument.
-\def\@IEEEPARstartFONTSTYLE{\bfseries}
-% V1.7 any additional, non-font related commands needed to modify
-% the drop cap letter, can take zero or one argument.
-\def\@IEEEPARstartCAPSTYLE{\MakeUppercase}
-% V1.7 the font that will be used to render the rest of the word,
-% can take zero or one argument.
-\def\@IEEEPARstartWORDFONTSTYLE{\relax}
-% V1.7 any additional, non-font related commands needed to modify
-% the rest of the word, can take zero or one argument.
-\def\@IEEEPARstartWORDCAPSTYLE{\MakeUppercase}
-% This is the horizontal separation distance from the drop letter to the main text.
-% Lengths that depend on the font (e.g., ex, em, etc.) will be referenced
-% to the font that is active when \IEEEPARstart is called.
-\def\@IEEEPARstartSEP{0.15em}
-% V1.7 horizontal offset applied to the left of the drop cap.
-\def\@IEEEPARstartHOFFSET{0em}
-% V1.7 Italic correction command applied at the end of the drop cap.
-\def\@IEEEPARstartITLCORRECT{\/}
-
-% V1.7 compoc uses nonbold drop cap and small caps word style
-\ifCLASSOPTIONcompsoc
-\def\@IEEEPARstartFONTSTYLE{\mdseries}
-\def\@IEEEPARstartWORDFONTSTYLE{\scshape}
-\def\@IEEEPARstartWORDCAPSTYLE{\relax}
-\fi
-
-% definition of \IEEEPARstart
-% THIS IS A CONTROLLED SPACING AREA, DO NOT ALLOW SPACES WITHIN THESE LINES
-%
-% The token \@IEEEPARstartfont will be globally defined after the first use
-% of \IEEEPARstart and will be a font command which creates the big letter
-% The first argument is the first letter of the first word and the second
-% argument is the rest of the first word(s).
-\def\IEEEPARstart#1#2{\par{%
-% if this page does not have enough space, break it and lets start
-% on a new one
-\@IEEEtranneedspace{\@IEEEPARstartMINPAGELINES\baselineskip}{\relax}%
-% V1.7 move this up here in case user uses \textbf for \@IEEEPARstartFONTSTYLE
-% which uses command \leavevmode which causes an unwanted \indent to be issued
-\noindent
-% calculate the desired height of the big letter
-% it extends from the top of \@IEEEPARstartHEIGHTTEXT in the current font
-% down to \@IEEEPARstartDROPDEPTH below the current baseline
-\settoheight{\@IEEEtrantmpdimenA}{\@IEEEPARstartHEIGHTTEXT}%
-\addtolength{\@IEEEtrantmpdimenA}{\@IEEEPARstartDROPDEPTH}%
-% extract the name of the current font in bold
-% and place it in \@IEEEPARstartFONTNAME
-\def\@IEEEPARstartGETFIRSTWORD##1 ##2\relax{##1}%
-{\@IEEEPARstartFONTSTYLE{\selectfont\edef\@IEEEPARstartFONTNAMESPACE{\fontname\font\space}%
-\xdef\@IEEEPARstartFONTNAME{\expandafter\@IEEEPARstartGETFIRSTWORD\@IEEEPARstartFONTNAMESPACE\relax}}}%
-% define a font based on this name with a point size equal to the desired
-% height of the drop letter
-\font\@IEEEPARstartsubfont\@IEEEPARstartFONTNAME\space at \@IEEEtrantmpdimenA\relax%
-% save this value as a counter (integer) value (sp points)
-\@IEEEtrantmpcountA=\@IEEEtrantmpdimenA%
-% now get the height of the actual letter produced by this font size
-\settoheight{\@IEEEtrantmpdimenB}{\@IEEEPARstartsubfont\@IEEEPARstartCAPSTYLE{#1}}%
-% If something bogus happens like the first argument is empty or the
-% current font is strange, do not allow a zero height.
-\ifdim\@IEEEtrantmpdimenB=0pt\relax%
-\typeout{** WARNING: IEEEPARstart drop letter has zero height! (line \the\inputlineno)}%
-\typeout{ Forcing the drop letter font size to 10pt.}%
-\@IEEEtrantmpdimenB=10pt%
-\fi%
-% and store it as a counter
-\@IEEEtrantmpcountB=\@IEEEtrantmpdimenB%
-% Since a font size doesn't exactly correspond to the height of the capital
-% letters in that font, the actual height of the letter, \@IEEEtrantmpcountB,
-% will be less than that desired, \@IEEEtrantmpcountA
-% we need to raise the font size, \@IEEEtrantmpdimenA
-% by \@IEEEtrantmpcountA / \@IEEEtrantmpcountB
-% But, TeX doesn't have floating point division, so we have to use integer
-% division. Hence the use of the counters.
-% We need to reduce the denominator so that the loss of the remainder will
-% have minimal affect on the accuracy of the result
-\divide\@IEEEtrantmpcountB by 200%
-\divide\@IEEEtrantmpcountA by \@IEEEtrantmpcountB%
-% Then reequalize things when we use TeX's ability to multiply by
-% floating point values
-\@IEEEtrantmpdimenB=0.005\@IEEEtrantmpdimenA%
-\multiply\@IEEEtrantmpdimenB by \@IEEEtrantmpcountA%
-% \@IEEEPARstartfont is globaly set to the calculated font of the big letter
-% We need to carry this out of the local calculation area to to create the
-% big letter.
-\global\font\@IEEEPARstartfont\@IEEEPARstartFONTNAME\space at \@IEEEtrantmpdimenB%
-% Now set \@IEEEtrantmpdimenA to the width of the big letter
-% We need to carry this out of the local calculation area to set the
-% hanging indent
-\settowidth{\global\@IEEEtrantmpdimenA}{\@IEEEPARstartfont
-\@IEEEPARstartCAPSTYLE{#1\@IEEEPARstartITLCORRECT}}}%
-% end of the isolated calculation environment
-% add in the extra clearance we want
-\advance\@IEEEtrantmpdimenA by \@IEEEPARstartSEP\relax%
-% add in the optional offset
-\advance\@IEEEtrantmpdimenA by \@IEEEPARstartHOFFSET\relax%
-% V1.7 don't allow negative offsets to produce negative hanging indents
-\@IEEEtrantmpdimenB\@IEEEtrantmpdimenA
-\ifnum\@IEEEtrantmpdimenB < 0 \@IEEEtrantmpdimenB 0pt\fi
-% \@IEEEtrantmpdimenA has the width of the big letter plus the
-% separation space and \@IEEEPARstartfont is the font we need to use
-% Now, we make the letter and issue the hanging indent command
-% The letter is placed in a box of zero width and height so that other
-% text won't be displaced by it.
-\hangindent\@IEEEtrantmpdimenB\hangafter=-\@IEEEPARstartDROPLINES%
-\makebox[0pt][l]{\hspace{-\@IEEEtrantmpdimenA}%
-\raisebox{-\@IEEEPARstartDROPDEPTH}[0pt][0pt]{\hspace{\@IEEEPARstartHOFFSET}%
-\@IEEEPARstartfont\@IEEEPARstartCAPSTYLE{#1\@IEEEPARstartITLCORRECT}%
-\hspace{\@IEEEPARstartSEP}}}%
-{\@IEEEPARstartWORDFONTSTYLE{\@IEEEPARstartWORDCAPSTYLE{\selectfont#2}}}}
-
-
-
-
-
-
-% determines if the space remaining on a given page is equal to or greater
-% than the specified space of argument one
-% if not, execute argument two (only if the remaining space is greater than zero)
-% and issue a \newpage
-%
-% example: \@IEEEtranneedspace{2in}{\vfill}
-%
-% Does not take into consideration rubber shrinkage, so it tends to
-% be overly cautious
-% Based on an example posted by Donald Arseneau
-% Note this macro uses \@IEEEtrantmpdimenB internally for calculations,
-% so DO NOT PASS \@IEEEtrantmpdimenB to this routine
-% if you need a dimen register, import with \@IEEEtrantmpdimenA instead
-\def\@IEEEtranneedspace#1#2{\penalty-100\begingroup%shield temp variable
-\@IEEEtrantmpdimenB\pagegoal\advance\@IEEEtrantmpdimenB-\pagetotal% space left
-\ifdim #1>\@IEEEtrantmpdimenB\relax% not enough space left
-\ifdim\@IEEEtrantmpdimenB>\z@\relax #2\fi%
-\newpage%
-\fi\endgroup}
-
-
-
-% IEEEbiography ENVIRONMENT
-% Allows user to enter biography leaving place for picture (adapts to font size)
-% As of V1.5, a new optional argument allows you to have a real graphic!
-% V1.5 and later also fixes the "colliding biographies" which could happen when a
-% biography's text was shorter than the space for the photo.
-% MDS 7/2001
-% V1.6 prevent multiple biographies from making multiple TOC entries
-\newif\if@IEEEbiographyTOCentrynotmade
-\global\@IEEEbiographyTOCentrynotmadetrue
-
-% biography counter so hyperref can jump directly to the biographies
-% and not just the previous section
-\newcounter{IEEEbiography}
-\setcounter{IEEEbiography}{0}
-
-% photo area size
-\def\@IEEEBIOphotowidth{1.0in} % width of the biography photo area
-\def\@IEEEBIOphotodepth{1.25in} % depth (height) of the biography photo area
-% area cleared for photo
-\def\@IEEEBIOhangwidth{1.14in} % width cleared for the biography photo area
-\def\@IEEEBIOhangdepth{1.25in} % depth cleared for the biography photo area
- % actual depth will be a multiple of
- % \baselineskip, rounded up
-\def\@IEEEBIOskipN{4\baselineskip}% nominal value of the vskip above the biography
-
-\newenvironment{IEEEbiography}[2][]{\normalfont\@IEEEcompsoconly{\sffamily}\footnotesize%
-\unitlength 1in\parskip=0pt\par\parindent 1em\interlinepenalty500%
-% we need enough space to support the hanging indent
-% the nominal value of the spacer
-% and one extra line for good measure
-\@IEEEtrantmpdimenA=\@IEEEBIOhangdepth%
-\advance\@IEEEtrantmpdimenA by \@IEEEBIOskipN%
-\advance\@IEEEtrantmpdimenA by 1\baselineskip%
-% if this page does not have enough space, break it and lets start
-% with a new one
-\@IEEEtranneedspace{\@IEEEtrantmpdimenA}{\relax}%
-% nominal spacer can strech, not shrink use 1fil so user can out stretch with \vfill
-\vskip \@IEEEBIOskipN plus 1fil minus 0\baselineskip%
-% the default box for where the photo goes
-\def\@IEEEtempbiographybox{{\setlength{\fboxsep}{0pt}\framebox{%
-\begin{minipage}[b][\@IEEEBIOphotodepth][c]{\@IEEEBIOphotowidth}\centering PLACE\\ PHOTO\\ HERE \end{minipage}}}}%
-%
-% detect if the optional argument was supplied, this requires the
-% \@ifmtarg command as defined in the appendix section above
-% and if so, override the default box with what they want
-\@ifmtarg{#1}{\relax}{\def\@IEEEtempbiographybox{\mbox{\begin{minipage}[b][\@IEEEBIOphotodepth][c]{\@IEEEBIOphotowidth}%
-\centering%
-#1%
-\end{minipage}}}}% end if optional argument supplied
-% Make an entry into the table of contents only if we have not done so before
-\if@IEEEbiographyTOCentrynotmade%
-% link labels to the biography counter so hyperref will jump
-% to the biography, not the previous section
-\setcounter{IEEEbiography}{-1}%
-\refstepcounter{IEEEbiography}%
-\addcontentsline{toc}{section}{Biographies}%
-\global\@IEEEbiographyTOCentrynotmadefalse%
-\fi%
-% one more biography
-\refstepcounter{IEEEbiography}%
-% Make an entry for this name into the table of contents
-\addcontentsline{toc}{subsection}{#2}%
-% V1.6 properly handle if a new paragraph should occur while the
-% hanging indent is still active. Do this by redefining \par so
-% that it will not start a new paragraph. (But it will appear to the
-% user as if it did.) Also, strip any leading pars, newlines, or spaces.
-\let\@IEEEBIOORGparCMD=\par% save the original \par command
-\edef\par{\hfil\break\indent}% the new \par will not be a "real" \par
-\settoheight{\@IEEEtrantmpdimenA}{\@IEEEtempbiographybox}% get height of biography box
-\@IEEEtrantmpdimenB=\@IEEEBIOhangdepth%
-\@IEEEtrantmpcountA=\@IEEEtrantmpdimenB% countA has the hang depth
-\divide\@IEEEtrantmpcountA by \baselineskip% calculates lines needed to produce the hang depth
-\advance\@IEEEtrantmpcountA by 1% ensure we overestimate
-% set the hanging indent
-\hangindent\@IEEEBIOhangwidth%
-\hangafter-\@IEEEtrantmpcountA%
-% reference the top of the photo area to the top of a capital T
-\settoheight{\@IEEEtrantmpdimenB}{\mbox{T}}%
-% set the photo box, give it zero width and height so as not to disturb anything
-\noindent\makebox[0pt][l]{\hspace{-\@IEEEBIOhangwidth}\raisebox{\@IEEEtrantmpdimenB}[0pt][0pt]{%
-\raisebox{-\@IEEEBIOphotodepth}[0pt][0pt]{\@IEEEtempbiographybox}}}%
-% now place the author name and begin the bio text
-\noindent\textbf{#2\ }\@IEEEgobbleleadPARNLSP}{\relax\let\par=\@IEEEBIOORGparCMD\par%
-% 7/2001 V1.5 detect when the biography text is shorter than the photo area
-% and pad the unused area - preventing a collision from the next biography entry
-% MDS
-\ifnum \prevgraf <\@IEEEtrantmpcountA\relax% detect when the biography text is shorter than the photo
- \advance\@IEEEtrantmpcountA by -\prevgraf% calculate how many lines we need to pad
- \advance\@IEEEtrantmpcountA by -1\relax% we compensate for the fact that we indented an extra line
- \@IEEEtrantmpdimenA=\baselineskip% calculate the length of the padding
- \multiply\@IEEEtrantmpdimenA by \@IEEEtrantmpcountA%
- \noindent\rule{0pt}{\@IEEEtrantmpdimenA}% insert an invisible support strut
-\fi%
-\par\normalfont}
-
-
-
-% V1.6
-% added biography without a photo environment
-\newenvironment{IEEEbiographynophoto}[1]{%
-% Make an entry into the table of contents only if we have not done so before
-\if@IEEEbiographyTOCentrynotmade%
-% link labels to the biography counter so hyperref will jump
-% to the biography, not the previous section
-\setcounter{IEEEbiography}{-1}%
-\refstepcounter{IEEEbiography}%
-\addcontentsline{toc}{section}{Biographies}%
-\global\@IEEEbiographyTOCentrynotmadefalse%
-\fi%
-% one more biography
-\refstepcounter{IEEEbiography}%
-% Make an entry for this name into the table of contents
-\addcontentsline{toc}{subsection}{#1}%
-\normalfont\@IEEEcompsoconly{\sffamily}\footnotesize\interlinepenalty500%
-\vskip 4\baselineskip plus 1fil minus 0\baselineskip%
-\parskip=0pt\par%
-\noindent\textbf{#1\ }\@IEEEgobbleleadPARNLSP}{\relax\par\normalfont}
-
-
-% provide the user with some old font commands
-% got this from article.cls
-\DeclareOldFontCommand{\rm}{\normalfont\rmfamily}{\mathrm}
-\DeclareOldFontCommand{\sf}{\normalfont\sffamily}{\mathsf}
-\DeclareOldFontCommand{\tt}{\normalfont\ttfamily}{\mathtt}
-\DeclareOldFontCommand{\bf}{\normalfont\bfseries}{\mathbf}
-\DeclareOldFontCommand{\it}{\normalfont\itshape}{\mathit}
-\DeclareOldFontCommand{\sl}{\normalfont\slshape}{\@nomath\sl}
-\DeclareOldFontCommand{\sc}{\normalfont\scshape}{\@nomath\sc}
-\DeclareRobustCommand*\cal{\@fontswitch\relax\mathcal}
-\DeclareRobustCommand*\mit{\@fontswitch\relax\mathnormal}
-
-
-% SPECIAL PAPER NOTICE COMMANDS
-%
-% holds the special notice text
-\def\@IEEEspecialpapernotice{\relax}
-
-% for special papers, like invited papers, the user can do:
-% \IEEEspecialpapernotice{(Invited Paper)} before \maketitle
-\def\IEEEspecialpapernotice#1{\ifCLASSOPTIONconference%
-\def\@IEEEspecialpapernotice{{\sublargesize\textit{#1}\vspace*{1em}}}%
-\else%
-\def\@IEEEspecialpapernotice{{\\*[1.5ex]\sublargesize\textit{#1}}\vspace*{-2ex}}%
-\fi}
-
-
-
-
-% PUBLISHER ID COMMANDS
-% to insert a publisher's ID footer
-% V1.6 \IEEEpubid has been changed so that the change in page size and style
-% occurs in \maketitle. \IEEEpubid must now be issued prior to \maketitle
-% use \IEEEpubidadjcol as before - in the second column of the title page
-% These changes allow \maketitle to take the reduced page height into
-% consideration when dynamically setting the space between the author
-% names and the maintext.
-%
-% the amount the main text is pulled up to make room for the
-% publisher's ID footer
-% IEEE uses about 1.3\baselineskip for journals,
-% dynamic title spacing will clean up the fraction
-\def\@IEEEpubidpullup{1.3\baselineskip}
-\ifCLASSOPTIONtechnote
-% for technotes it must be an integer of baselineskip as there can be no
-% dynamic title spacing for two column mode technotes (the title is in the
-% in first column) and we should maintain an integer number of lines in the
-% second column
-% There are some examples (such as older issues of "Transactions on
-% Information Theory") in which IEEE really pulls the text off the ID for
-% technotes - about 0.55in (or 4\baselineskip). We'll use 2\baselineskip
-% and call it even.
-\def\@IEEEpubidpullup{2\baselineskip}
-\fi
-
-% V1.7 compsoc does not use a pullup
-\ifCLASSOPTIONcompsoc
-\def\@IEEEpubidpullup{0pt}
-\fi
-
-% holds the ID text
-\def\@IEEEpubid{\relax}
-
-% flag so \maketitle can tell if \IEEEpubid was called
-\newif\if@IEEEusingpubid
-\global\@IEEEusingpubidfalse
-% issue this command in the page to have the ID at the bottom
-% V1.6 use before \maketitle
-\def\IEEEpubid#1{\def\@IEEEpubid{#1}\global\@IEEEusingpubidtrue}
-
-
-% command which will pull up (shorten) the column it is executed in
-% to make room for the publisher ID. Place in the second column of
-% the title page when using \IEEEpubid
-% Is smart enough not to do anything when in single column text or
-% if the user hasn't called \IEEEpubid
-% currently needed in for the second column of a page with the
-% publisher ID. If not needed in future releases, please provide this
-% command and define it as \relax for backward compatibility
-% v1.6b do not allow command to operate if the peer review option has been
-% selected because \IEEEpubidadjcol will not be on the cover page.
-% V1.7 do nothing if compsoc
-\def\IEEEpubidadjcol{\ifCLASSOPTIONcompsoc\else\ifCLASSOPTIONpeerreview\else
-\if@twocolumn\if@IEEEusingpubid\enlargethispage{-\@IEEEpubidpullup}\fi\fi\fi\fi}
-
-% Special thanks to Peter Wilson, Daniel Luecking, and the other
-% gurus at comp.text.tex, for helping me to understand how best to
-% implement the IEEEpubid command in LaTeX.
-
-
-
-%% Lockout some commands under various conditions
-
-% general purpose bit bucket
-\newsavebox{\@IEEEtranrubishbin}
-
-% flags to prevent multiple warning messages
-\newif\if@IEEEWARNthanks
-\newif\if@IEEEWARNIEEEPARstart
-\newif\if@IEEEWARNIEEEbiography
-\newif\if@IEEEWARNIEEEbiographynophoto
-\newif\if@IEEEWARNIEEEpubid
-\newif\if@IEEEWARNIEEEpubidadjcol
-\newif\if@IEEEWARNIEEEmembership
-\newif\if@IEEEWARNIEEEaftertitletext
-\@IEEEWARNthankstrue
-\@IEEEWARNIEEEPARstarttrue
-\@IEEEWARNIEEEbiographytrue
-\@IEEEWARNIEEEbiographynophototrue
-\@IEEEWARNIEEEpubidtrue
-\@IEEEWARNIEEEpubidadjcoltrue
-\@IEEEWARNIEEEmembershiptrue
-\@IEEEWARNIEEEaftertitletexttrue
-
-
-%% Lockout some commands when in various modes, but allow them to be restored if needed
-%%
-% save commands which might be locked out
-% so that the user can later restore them if needed
-\let\@IEEESAVECMDthanks\thanks
-\let\@IEEESAVECMDIEEEPARstart\IEEEPARstart
-\let\@IEEESAVECMDIEEEbiography\IEEEbiography
-\let\@IEEESAVECMDendIEEEbiography\endIEEEbiography
-\let\@IEEESAVECMDIEEEbiographynophoto\IEEEbiographynophoto
-\let\@IEEESAVECMDendIEEEbiographynophoto\endIEEEbiographynophoto
-\let\@IEEESAVECMDIEEEpubid\IEEEpubid
-\let\@IEEESAVECMDIEEEpubidadjcol\IEEEpubidadjcol
-\let\@IEEESAVECMDIEEEmembership\IEEEmembership
-\let\@IEEESAVECMDIEEEaftertitletext\IEEEaftertitletext
-
-
-% disable \IEEEPARstart when in draft mode
-% This may have originally been done because the pre-V1.6 drop letter
-% algorithm had problems with a non-unity baselinestretch
-% At any rate, it seems too formal to have a drop letter in a draft
-% paper.
-\ifCLASSOPTIONdraftcls
-\def\IEEEPARstart#1#2{#1#2\if@IEEEWARNIEEEPARstart\typeout{** ATTENTION: \noexpand\IEEEPARstart
- is disabled in draft mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEPARstartfalse}
-\fi
-% and for technotes
-\ifCLASSOPTIONtechnote
-\def\IEEEPARstart#1#2{#1#2\if@IEEEWARNIEEEPARstart\typeout{** WARNING: \noexpand\IEEEPARstart
- is locked out for technotes (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEPARstartfalse}
-\fi
-
-
-% lockout unneeded commands when in conference mode
-\ifCLASSOPTIONconference
-% when locked out, \thanks, \IEEEbiography, \IEEEbiographynophoto, \IEEEpubid,
-% \IEEEmembership and \IEEEaftertitletext will all swallow their given text.
-% \IEEEPARstart will output a normal character instead
-% warn the user about these commands only once to prevent the console screen
-% from filling up with redundant messages
-\def\thanks#1{\if@IEEEWARNthanks\typeout{** WARNING: \noexpand\thanks
- is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNthanksfalse}
-\def\IEEEPARstart#1#2{#1#2\if@IEEEWARNIEEEPARstart\typeout{** WARNING: \noexpand\IEEEPARstart
- is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEPARstartfalse}
-
-
-% LaTeX treats environments and commands with optional arguments differently.
-% the actual ("internal") command is stored as \\commandname
-% (accessed via \csname\string\commandname\endcsname )
-% the "external" command \commandname is a macro with code to determine
-% whether or not the optional argument is presented and to provide the
-% default if it is absent. So, in order to save and restore such a command
-% we would have to save and restore \\commandname as well. But, if LaTeX
-% ever changes the way it names the internal names, the trick would break.
-% Instead let us just define a new environment so that the internal
-% name can be left undisturbed.
-\newenvironment{@IEEEbogusbiography}[2][]{\if@IEEEWARNIEEEbiography\typeout{** WARNING: \noexpand\IEEEbiography
- is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEbiographyfalse%
-\setbox\@IEEEtranrubishbin\vbox\bgroup}{\egroup\relax}
-% and make biography point to our bogus biography
-\let\IEEEbiography=\@IEEEbogusbiography
-\let\endIEEEbiography=\end@IEEEbogusbiography
-
-\renewenvironment{IEEEbiographynophoto}[1]{\if@IEEEWARNIEEEbiographynophoto\typeout{** WARNING: \noexpand\IEEEbiographynophoto
- is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEbiographynophotofalse%
-\setbox\@IEEEtranrubishbin\vbox\bgroup}{\egroup\relax}
-
-\def\IEEEpubid#1{\if@IEEEWARNIEEEpubid\typeout{** WARNING: \noexpand\IEEEpubid
- is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEpubidfalse}
-\def\IEEEpubidadjcol{\if@IEEEWARNIEEEpubidadjcol\typeout{** WARNING: \noexpand\IEEEpubidadjcol
- is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEpubidadjcolfalse}
-\def\IEEEmembership#1{\if@IEEEWARNIEEEmembership\typeout{** WARNING: \noexpand\IEEEmembership
- is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEmembershipfalse}
-\def\IEEEaftertitletext#1{\if@IEEEWARNIEEEaftertitletext\typeout{** WARNING: \noexpand\IEEEaftertitletext
- is locked out when in conference mode (line \the\inputlineno).}\fi\global\@IEEEWARNIEEEaftertitletextfalse}
-\fi
-
-
-% provide a way to restore the commands that are locked out
-\def\IEEEoverridecommandlockouts{%
-\typeout{** ATTENTION: Overriding command lockouts (line \the\inputlineno).}%
-\let\thanks\@IEEESAVECMDthanks%
-\let\IEEEPARstart\@IEEESAVECMDIEEEPARstart%
-\let\IEEEbiography\@IEEESAVECMDIEEEbiography%
-\let\endIEEEbiography\@IEEESAVECMDendIEEEbiography%
-\let\IEEEbiographynophoto\@IEEESAVECMDIEEEbiographynophoto%
-\let\endIEEEbiographynophoto\@IEEESAVECMDendIEEEbiographynophoto%
-\let\IEEEpubid\@IEEESAVECMDIEEEpubid%
-\let\IEEEpubidadjcol\@IEEESAVECMDIEEEpubidadjcol%
-\let\IEEEmembership\@IEEESAVECMDIEEEmembership%
-\let\IEEEaftertitletext\@IEEESAVECMDIEEEaftertitletext}
-
-
-
-% need a backslash character for typeout output
-{\catcode`\|=0 \catcode`\\=12
-|xdef|@IEEEbackslash{\}}
-
-
-% hook to allow easy disabling of all legacy warnings
-\def\@IEEElegacywarn#1#2{\typeout{** ATTENTION: \@IEEEbackslash #1 is deprecated (line \the\inputlineno).
-Use \@IEEEbackslash #2 instead.}}
-
-
-% provide for legacy commands
-\def\authorblockA{\@IEEElegacywarn{authorblockA}{IEEEauthorblockA}\IEEEauthorblockA}
-\def\authorblockN{\@IEEElegacywarn{authorblockN}{IEEEauthorblockN}\IEEEauthorblockN}
-\def\authorrefmark{\@IEEElegacywarn{authorrefmark}{IEEEauthorrefmark}\IEEEauthorrefmark}
-\def\PARstart{\@IEEElegacywarn{PARstart}{IEEEPARstart}\IEEEPARstart}
-\def\pubid{\@IEEElegacywarn{pubid}{IEEEpubid}\IEEEpubid}
-\def\pubidadjcol{\@IEEElegacywarn{pubidadjcol}{IEEEpubidadjcol}\IEEEpubidadjcol}
-\def\QED{\@IEEElegacywarn{QED}{IEEEQED}\IEEEQED}
-\def\QEDclosed{\@IEEElegacywarn{QEDclosed}{IEEEQEDclosed}\IEEEQEDclosed}
-\def\QEDopen{\@IEEElegacywarn{QEDopen}{IEEEQEDopen}\IEEEQEDopen}
-\def\specialpapernotice{\@IEEElegacywarn{specialpapernotice}{IEEEspecialpapernotice}\IEEEspecialpapernotice}
-
-
-
-% provide for legacy environments
-\def\biography{\@IEEElegacywarn{biography}{IEEEbiography}\IEEEbiography}
-\def\biographynophoto{\@IEEElegacywarn{biographynophoto}{IEEEbiographynophoto}\IEEEbiographynophoto}
-\def\keywords{\@IEEElegacywarn{keywords}{IEEEkeywords}\IEEEkeywords}
-\def\endbiography{\endIEEEbiography}
-\def\endbiographynophoto{\endIEEEbiographynophoto}
-\def\endkeywords{\endIEEEkeywords}
-
-
-% provide for legacy IED commands/lengths when possible
-\let\labelindent\IEEElabelindent
-\def\calcleftmargin{\@IEEElegacywarn{calcleftmargin}{IEEEcalcleftmargin}\IEEEcalcleftmargin}
-\def\setlabelwidth{\@IEEElegacywarn{setlabelwidth}{IEEEsetlabelwidth}\IEEEsetlabelwidth}
-\def\usemathlabelsep{\@IEEElegacywarn{usemathlabelsep}{IEEEusemathlabelsep}\IEEEusemathlabelsep}
-\def\iedlabeljustifyc{\@IEEElegacywarn{iedlabeljustifyc}{IEEEiedlabeljustifyc}\IEEEiedlabeljustifyc}
-\def\iedlabeljustifyl{\@IEEElegacywarn{iedlabeljustifyl}{IEEEiedlabeljustifyl}\IEEEiedlabeljustifyl}
-\def\iedlabeljustifyr{\@IEEElegacywarn{iedlabeljustifyr}{IEEEiedlabeljustifyr}\IEEEiedlabeljustifyr}
-
-
-
-% let \proof use the IEEEtran version even after amsthm is loaded
-% \proof is now deprecated in favor of \IEEEproof
-\AtBeginDocument{\def\proof{\@IEEElegacywarn{proof}{IEEEproof}\IEEEproof}\def\endproof{\endIEEEproof}}
-
-% V1.7 \overrideIEEEmargins is no longer supported.
-\def\overrideIEEEmargins{%
-\typeout{** WARNING: \string\overrideIEEEmargins \space no longer supported (line \the\inputlineno).}%
-\typeout{** Use the \string\CLASSINPUTinnersidemargin, \string\CLASSINPUToutersidemargin \space controls instead.}}
-
-
-\endinput
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%% End of IEEEtran.cls %%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% That's all folks!
-
--- a/document/llncs.cls Thu Jan 17 11:51:00 2013 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1189 +0,0 @@
-% LLNCS DOCUMENT CLASS -- version 2.13 (28-Jan-2002)
-% Springer Verlag LaTeX2e support for Lecture Notes in Computer Science
-%
-%%
-%% \CharacterTable
-%% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
-%% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
-%% Digits \0\1\2\3\4\5\6\7\8\9
-%% Exclamation \! Double quote \" Hash (number) \#
-%% Dollar \$ Percent \% Ampersand \&
-%% Acute accent \' Left paren \( Right paren \)
-%% Asterisk \* Plus \+ Comma \,
-%% Minus \- Point \. Solidus \/
-%% Colon \: Semicolon \; Less than \<
-%% Equals \= Greater than \> Question mark \?
-%% Commercial at \@ Left bracket \[ Backslash \\
-%% Right bracket \] Circumflex \^ Underscore \_
-%% Grave accent \` Left brace \{ Vertical bar \|
-%% Right brace \} Tilde \~}
-%%
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-\ProvidesClass{llncs}[2002/01/28 v2.13
-^^J LaTeX document class for Lecture Notes in Computer Science]
-% Options
-\let\if@envcntreset\iffalse
-\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue}
-\DeclareOption{citeauthoryear}{\let\citeauthoryear=Y}
-\DeclareOption{oribibl}{\let\oribibl=Y}
-\let\if@custvec\iftrue
-\DeclareOption{orivec}{\let\if@custvec\iffalse}
-\let\if@envcntsame\iffalse
-\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue}
-\let\if@envcntsect\iffalse
-\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue}
-\let\if@runhead\iffalse
-\DeclareOption{runningheads}{\let\if@runhead\iftrue}
-
-\let\if@openbib\iffalse
-\DeclareOption{openbib}{\let\if@openbib\iftrue}
-
-% languages
-\let\switcht@@therlang\relax
-\def\ds@deutsch{\def\switcht@@therlang{\switcht@deutsch}}
-\def\ds@francais{\def\switcht@@therlang{\switcht@francais}}
-
-\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}}
-
-\ProcessOptions
-
-\LoadClass[twoside]{article}
-\RequirePackage{multicol} % needed for the list of participants, index
-
-\setlength{\textwidth}{12.2cm}
-\setlength{\textheight}{19.3cm}
-\renewcommand\@pnumwidth{2em}
-\renewcommand\@tocrmarg{3.5em}
-%
-\def\@dottedtocline#1#2#3#4#5{%
- \ifnum #1>\c@tocdepth \else
- \vskip \z@ \@plus.2\p@
- {\leftskip #2\relax \rightskip \@tocrmarg \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \parindent #2\relax\@afterindenttrue
- \interlinepenalty\@M
- \leavevmode
- \@tempdima #3\relax
- \advance\leftskip \@tempdima \null\nobreak\hskip -\leftskip
- {#4}\nobreak
- \leaders\hbox{$\m@th
- \mkern \@dotsep mu\hbox{.}\mkern \@dotsep
- mu$}\hfill
- \nobreak
- \hb@xt@\@pnumwidth{\hfil\normalfont \normalcolor #5}%
- \par}%
- \fi}
-%
-\def\switcht@albion{%
-\def\abstractname{Abstract.}
-\def\ackname{Acknowledgement.}
-\def\andname{and}
-\def\lastandname{\unskip, and}
-\def\appendixname{Appendix}
-\def\chaptername{Chapter}
-\def\claimname{Claim}
-\def\conjecturename{Conjecture}
-\def\contentsname{Table of Contents}
-\def\corollaryname{Corollary}
-\def\definitionname{Definition}
-\def\examplename{Example}
-\def\exercisename{Exercise}
-\def\figurename{Fig.}
-\def\keywordname{{\bf Key words:}}
-\def\indexname{Index}
-\def\lemmaname{Lemma}
-\def\contriblistname{List of Contributors}
-\def\listfigurename{List of Figures}
-\def\listtablename{List of Tables}
-\def\mailname{{\it Correspondence to\/}:}
-\def\noteaddname{Note added in proof}
-\def\notename{Note}
-\def\partname{Part}
-\def\problemname{Problem}
-\def\proofname{Proof}
-\def\propertyname{Property}
-\def\propositionname{Proposition}
-\def\questionname{Question}
-\def\remarkname{Remark}
-\def\seename{see}
-\def\solutionname{Solution}
-\def\subclassname{{\it Subject Classifications\/}:}
-\def\tablename{Table}
-\def\theoremname{Theorem}}
-\switcht@albion
-% Names of theorem like environments are already defined
-% but must be translated if another language is chosen
-%
-% French section
-\def\switcht@francais{%\typeout{On parle francais.}%
- \def\abstractname{R\'esum\'e.}%
- \def\ackname{Remerciements.}%
- \def\andname{et}%
- \def\lastandname{ et}%
- \def\appendixname{Appendice}
- \def\chaptername{Chapitre}%
- \def\claimname{Pr\'etention}%
- \def\conjecturename{Hypoth\`ese}%
- \def\contentsname{Table des mati\`eres}%
- \def\corollaryname{Corollaire}%
- \def\definitionname{D\'efinition}%
- \def\examplename{Exemple}%
- \def\exercisename{Exercice}%
- \def\figurename{Fig.}%
- \def\keywordname{{\bf Mots-cl\'e:}}
- \def\indexname{Index}
- \def\lemmaname{Lemme}%
- \def\contriblistname{Liste des contributeurs}
- \def\listfigurename{Liste des figures}%
- \def\listtablename{Liste des tables}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}%
- \def\notename{Remarque}%
- \def\partname{Partie}%
- \def\problemname{Probl\`eme}%
- \def\proofname{Preuve}%
- \def\propertyname{Caract\'eristique}%
-%\def\propositionname{Proposition}%
- \def\questionname{Question}%
- \def\remarkname{Remarque}%
- \def\seename{voir}
- \def\solutionname{Solution}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tableau}%
- \def\theoremname{Th\'eor\`eme}%
-}
-%
-% German section
-\def\switcht@deutsch{%\typeout{Man spricht deutsch.}%
- \def\abstractname{Zusammenfassung.}%
- \def\ackname{Danksagung.}%
- \def\andname{und}%
- \def\lastandname{ und}%
- \def\appendixname{Anhang}%
- \def\chaptername{Kapitel}%
- \def\claimname{Behauptung}%
- \def\conjecturename{Hypothese}%
- \def\contentsname{Inhaltsverzeichnis}%
- \def\corollaryname{Korollar}%
-%\def\definitionname{Definition}%
- \def\examplename{Beispiel}%
- \def\exercisename{\"Ubung}%
- \def\figurename{Abb.}%
- \def\keywordname{{\bf Schl\"usselw\"orter:}}
- \def\indexname{Index}
-%\def\lemmaname{Lemma}%
- \def\contriblistname{Mitarbeiter}
- \def\listfigurename{Abbildungsverzeichnis}%
- \def\listtablename{Tabellenverzeichnis}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Nachtrag}%
- \def\notename{Anmerkung}%
- \def\partname{Teil}%
-%\def\problemname{Problem}%
- \def\proofname{Beweis}%
- \def\propertyname{Eigenschaft}%
-%\def\propositionname{Proposition}%
- \def\questionname{Frage}%
- \def\remarkname{Anmerkung}%
- \def\seename{siehe}
- \def\solutionname{L\"osung}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tabelle}%
-%\def\theoremname{Theorem}%
-}
-
-% Ragged bottom for the actual page
-\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil
-\global\let\@textbottom\relax}}
-
-\renewcommand\small{%
- \@setfontsize\small\@ixpt{11}%
- \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
- \abovedisplayshortskip \z@ \@plus2\p@
- \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
- \def\@listi{\leftmargin\leftmargini
- \parsep 0\p@ \@plus1\p@ \@minus\p@
- \topsep 8\p@ \@plus2\p@ \@minus4\p@
- \itemsep0\p@}%
- \belowdisplayskip \abovedisplayskip
-}
-
-\frenchspacing
-\widowpenalty=10000
-\clubpenalty=10000
-
-\setlength\oddsidemargin {63\p@}
-\setlength\evensidemargin {63\p@}
-\setlength\marginparwidth {90\p@}
-
-\setlength\headsep {16\p@}
-
-\setlength\footnotesep{7.7\p@}
-\setlength\textfloatsep{8mm\@plus 2\p@ \@minus 4\p@}
-\setlength\intextsep {8mm\@plus 2\p@ \@minus 2\p@}
-
-\setcounter{secnumdepth}{2}
-
-\newcounter {chapter}
-\renewcommand\thechapter {\@arabic\c@chapter}
-
-\newif\if@mainmatter \@mainmattertrue
-\newcommand\frontmatter{\cleardoublepage
- \@mainmatterfalse\pagenumbering{Roman}}
-\newcommand\mainmatter{\cleardoublepage
- \@mainmattertrue\pagenumbering{arabic}}
-\newcommand\backmatter{\if@openright\cleardoublepage\else\clearpage\fi
- \@mainmatterfalse}
-
-\renewcommand\part{\cleardoublepage
- \thispagestyle{empty}%
- \if@twocolumn
- \onecolumn
- \@tempswatrue
- \else
- \@tempswafalse
- \fi
- \null\vfil
- \secdef\@part\@spart}
-
-\def\@part[#1]#2{%
- \ifnum \c@secnumdepth >-2\relax
- \refstepcounter{part}%
- \addcontentsline{toc}{part}{\thepart\hspace{1em}#1}%
- \else
- \addcontentsline{toc}{part}{#1}%
- \fi
- \markboth{}{}%
- {\centering
- \interlinepenalty \@M
- \normalfont
- \ifnum \c@secnumdepth >-2\relax
- \huge\bfseries \partname~\thepart
- \par
- \vskip 20\p@
- \fi
- \Huge \bfseries #2\par}%
- \@endpart}
-\def\@spart#1{%
- {\centering
- \interlinepenalty \@M
- \normalfont
- \Huge \bfseries #1\par}%
- \@endpart}
-\def\@endpart{\vfil\newpage
- \if@twoside
- \null
- \thispagestyle{empty}%
- \newpage
- \fi
- \if@tempswa
- \twocolumn
- \fi}
-
-\newcommand\chapter{\clearpage
- \thispagestyle{empty}%
- \global\@topnum\z@
- \@afterindentfalse
- \secdef\@chapter\@schapter}
-\def\@chapter[#1]#2{\ifnum \c@secnumdepth >\m@ne
- \if@mainmatter
- \refstepcounter{chapter}%
- \typeout{\@chapapp\space\thechapter.}%
- \addcontentsline{toc}{chapter}%
- {\protect\numberline{\thechapter}#1}%
- \else
- \addcontentsline{toc}{chapter}{#1}%
- \fi
- \else
- \addcontentsline{toc}{chapter}{#1}%
- \fi
- \chaptermark{#1}%
- \addtocontents{lof}{\protect\addvspace{10\p@}}%
- \addtocontents{lot}{\protect\addvspace{10\p@}}%
- \if@twocolumn
- \@topnewpage[\@makechapterhead{#2}]%
- \else
- \@makechapterhead{#2}%
- \@afterheading
- \fi}
-\def\@makechapterhead#1{%
-% \vspace*{50\p@}%
- {\centering
- \ifnum \c@secnumdepth >\m@ne
- \if@mainmatter
- \large\bfseries \@chapapp{} \thechapter
- \par\nobreak
- \vskip 20\p@
- \fi
- \fi
- \interlinepenalty\@M
- \Large \bfseries #1\par\nobreak
- \vskip 40\p@
- }}
-\def\@schapter#1{\if@twocolumn
- \@topnewpage[\@makeschapterhead{#1}]%
- \else
- \@makeschapterhead{#1}%
- \@afterheading
- \fi}
-\def\@makeschapterhead#1{%
-% \vspace*{50\p@}%
- {\centering
- \normalfont
- \interlinepenalty\@M
- \Large \bfseries #1\par\nobreak
- \vskip 40\p@
- }}
-
-\renewcommand\section{\@startsection{section}{1}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {12\p@ \@plus 4\p@ \@minus 4\p@}%
- {\normalfont\large\bfseries\boldmath
- \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsection{\@startsection{subsection}{2}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {8\p@ \@plus 4\p@ \@minus 4\p@}%
- {\normalfont\normalsize\bfseries\boldmath
- \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}%
- {-18\p@ \@plus -4\p@ \@minus -4\p@}%
- {-0.5em \@plus -0.22em \@minus -0.1em}%
- {\normalfont\normalsize\bfseries\boldmath}}
-\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}%
- {-12\p@ \@plus -4\p@ \@minus -4\p@}%
- {-0.5em \@plus -0.22em \@minus -0.1em}%
- {\normalfont\normalsize\itshape}}
-\renewcommand\subparagraph[1]{\typeout{LLNCS warning: You should not use
- \string\subparagraph\space with this class}\vskip0.5cm
-You should not use \verb|\subparagraph| with this class.\vskip0.5cm}
-
-\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00}
-\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01}
-\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02}
-\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03}
-\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04}
-\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05}
-\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06}
-\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07}
-\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08}
-\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09}
-\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A}
-
-\let\footnotesize\small
-
-\if@custvec
-\def\vec#1{\mathchoice{\mbox{\boldmath$\displaystyle#1$}}
-{\mbox{\boldmath$\textstyle#1$}}
-{\mbox{\boldmath$\scriptstyle#1$}}
-{\mbox{\boldmath$\scriptscriptstyle#1$}}}
-\fi
-
-\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}}
-\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil
-\penalty50\hskip1em\null\nobreak\hfil\squareforqed
-\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi}
-
-\def\getsto{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr\gets\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-\gets\cr\to\cr}}}}}
-\def\lid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr<\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr<\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr<\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-<\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\gid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr>\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr>\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\grole{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr
->\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.8pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.3pt}<\cr}}}}}
-\def\bbbr{{\rm I\!R}} %reelle Zahlen
-\def\bbbm{{\rm I\!M}}
-\def\bbbn{{\rm I\!N}} %natuerliche Zahlen
-\def\bbbf{{\rm I\!F}}
-\def\bbbh{{\rm I\!H}}
-\def\bbbk{{\rm I\!K}}
-\def\bbbp{{\rm I\!P}}
-\def\bbbone{{\mathchoice {\rm 1\mskip-4mu l} {\rm 1\mskip-4mu l}
-{\rm 1\mskip-4.5mu l} {\rm 1\mskip-5mu l}}}
-\def\bbbc{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbq{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}}}
-\def\bbbt{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-T$}\hbox{\hbox to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbs{{\mathchoice
-{\setbox0=\hbox{$\displaystyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.5\wd0\vrule height0.45\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.4\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.55\wd0\vrule height0.45\ht0\hss}\box0}}}}
-\def\bbbz{{\mathchoice {\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\scriptstyle Z\kern-0.3em Z$}}
-{\hbox{$\mathsf\scriptscriptstyle Z\kern-0.2em Z$}}}}
-
-\let\ts\,
-
-\setlength\leftmargini {17\p@}
-\setlength\leftmargin {\leftmargini}
-\setlength\leftmarginii {\leftmargini}
-\setlength\leftmarginiii {\leftmargini}
-\setlength\leftmarginiv {\leftmargini}
-\setlength \labelsep {.5em}
-\setlength \labelwidth{\leftmargini}
-\addtolength\labelwidth{-\labelsep}
-
-\def\@listI{\leftmargin\leftmargini
- \parsep 0\p@ \@plus1\p@ \@minus\p@
- \topsep 8\p@ \@plus2\p@ \@minus4\p@
- \itemsep0\p@}
-\let\@listi\@listI
-\@listi
-\def\@listii {\leftmargin\leftmarginii
- \labelwidth\leftmarginii
- \advance\labelwidth-\labelsep
- \topsep 0\p@ \@plus2\p@ \@minus\p@}
-\def\@listiii{\leftmargin\leftmarginiii
- \labelwidth\leftmarginiii
- \advance\labelwidth-\labelsep
- \topsep 0\p@ \@plus\p@\@minus\p@
- \parsep \z@
- \partopsep \p@ \@plus\z@ \@minus\p@}
-
-\renewcommand\labelitemi{\normalfont\bfseries --}
-\renewcommand\labelitemii{$\m@th\bullet$}
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\def\tableofcontents{\chapter*{\contentsname\@mkboth{{\contentsname}}%
- {{\contentsname}}}
- \def\authcount##1{\setcounter{auco}{##1}\setcounter{@auth}{1}}
- \def\lastand{\ifnum\value{auco}=2\relax
- \unskip{} \andname\
- \else
- \unskip \lastandname\
- \fi}%
- \def\and{\stepcounter{@auth}\relax
- \ifnum\value{@auth}=\value{auco}%
- \lastand
- \else
- \unskip,
- \fi}%
- \@starttoc{toc}\if@restonecol\twocolumn\fi}
-
-\def\l@part#1#2{\addpenalty{\@secpenalty}%
- \addvspace{2em plus\p@}% % space above part line
- \begingroup
- \parindent \z@
- \rightskip \z@ plus 5em
- \hrule\vskip5pt
- \large % same size as for a contribution heading
- \bfseries\boldmath % set line in boldface
- \leavevmode % TeX command to enter horizontal mode.
- #1\par
- \vskip5pt
- \hrule
- \vskip1pt
- \nobreak % Never break after part entry
- \endgroup}
-
-\def\@dotsep{2}
-
-\def\hyperhrefextend{\ifx\hyper@anchor\@undefined\else
-{chapter.\thechapter}\fi}
-
-\def\addnumcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{\protect\numberline
- {\thechapter}#3}{\thepage}\hyperhrefextend}}
-\def\addcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{\thepage}\hyperhrefextend}}
-\def\addcontentsmarkwop#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{0}\hyperhrefextend}}
-
-\def\@adcmk[#1]{\ifcase #1 \or
-\def\@gtempa{\addnumcontentsmark}%
- \or \def\@gtempa{\addcontentsmark}%
- \or \def\@gtempa{\addcontentsmarkwop}%
- \fi\@gtempa{toc}{chapter}}
-\def\addtocmark{\@ifnextchar[{\@adcmk}{\@adcmk[3]}}
-
-\def\l@chapter#1#2{\addpenalty{-\@highpenalty}
- \vskip 1.0em plus 1pt \@tempdima 1.5em \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- {\large\bfseries\boldmath#1}\ifx0#2\hfil\null
- \else
- \nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}%
- \fi\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@title#1#2{\addpenalty{-\@highpenalty}
- \addvspace{8pt plus 1pt}
- \@tempdima \z@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- #1\nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@author#1#2{\addpenalty{\@highpenalty}
- \@tempdima=\z@ %15\p@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima %\hskip -\leftskip
- \textit{#1}\par
- \penalty\@highpenalty \endgroup}
-
-%\setcounter{tocdepth}{0}
-\newdimen\tocchpnum
-\newdimen\tocsecnum
-\newdimen\tocsectotal
-\newdimen\tocsubsecnum
-\newdimen\tocsubsectotal
-\newdimen\tocsubsubsecnum
-\newdimen\tocsubsubsectotal
-\newdimen\tocparanum
-\newdimen\tocparatotal
-\newdimen\tocsubparanum
-\tocchpnum=\z@ % no chapter numbers
-\tocsecnum=15\p@ % section 88. plus 2.222pt
-\tocsubsecnum=23\p@ % subsection 88.8 plus 2.222pt
-\tocsubsubsecnum=27\p@ % subsubsection 88.8.8 plus 1.444pt
-\tocparanum=35\p@ % paragraph 88.8.8.8 plus 1.666pt
-\tocsubparanum=43\p@ % subparagraph 88.8.8.8.8 plus 1.888pt
-\def\calctocindent{%
-\tocsectotal=\tocchpnum
-\advance\tocsectotal by\tocsecnum
-\tocsubsectotal=\tocsectotal
-\advance\tocsubsectotal by\tocsubsecnum
-\tocsubsubsectotal=\tocsubsectotal
-\advance\tocsubsubsectotal by\tocsubsubsecnum
-\tocparatotal=\tocsubsubsectotal
-\advance\tocparatotal by\tocparanum}
-\calctocindent
-
-\def\l@section{\@dottedtocline{1}{\tocchpnum}{\tocsecnum}}
-\def\l@subsection{\@dottedtocline{2}{\tocsectotal}{\tocsubsecnum}}
-\def\l@subsubsection{\@dottedtocline{3}{\tocsubsectotal}{\tocsubsubsecnum}}
-\def\l@paragraph{\@dottedtocline{4}{\tocsubsubsectotal}{\tocparanum}}
-\def\l@subparagraph{\@dottedtocline{5}{\tocparatotal}{\tocsubparanum}}
-
-\def\listoffigures{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}}
- \@starttoc{lof}\if@restonecol\twocolumn\fi}
-\def\l@figure{\@dottedtocline{1}{0em}{1.5em}}
-
-\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}}
- \@starttoc{lot}\if@restonecol\twocolumn\fi}
-\let\l@table\l@figure
-
-\renewcommand\listoffigures{%
- \section*{\listfigurename
- \@mkboth{\listfigurename}{\listfigurename}}%
- \@starttoc{lof}%
- }
-
-\renewcommand\listoftables{%
- \section*{\listtablename
- \@mkboth{\listtablename}{\listtablename}}%
- \@starttoc{lot}%
- }
-
-\ifx\oribibl\undefined
-\ifx\citeauthoryear\undefined
-\renewenvironment{thebibliography}[1]
- {\section*{\refname}
- \def\@biblabel##1{##1.}
- \small
- \list{\@biblabel{\@arabic\c@enumiv}}%
- {\settowidth\labelwidth{\@biblabel{#1}}%
- \leftmargin\labelwidth
- \advance\leftmargin\labelsep
- \if@openbib
- \advance\leftmargin\bibindent
- \itemindent -\bibindent
- \listparindent \itemindent
- \parsep \z@
- \fi
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{\@arabic\c@enumiv}}%
- \if@openbib
- \renewcommand\newblock{\par}%
- \else
- \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
- \fi
- \sloppy\clubpenalty4000\widowpenalty4000%
- \sfcode`\.=\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
-\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw
- {\let\protect\noexpand\immediate
- \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
-\newcount\@tempcntc
-\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi
- \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do
- {\@ifundefined
- {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries
- ?}\@warning
- {Citation `\@citeb' on page \thepage \space undefined}}%
- {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}%
- \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne
- \@citea\def\@citea{,}\hbox{\csname b@\@citeb\endcsname}%
- \else
- \advance\@tempcntb\@ne
- \ifnum\@tempcntb=\@tempcntc
- \else\advance\@tempcntb\m@ne\@citeo
- \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}}
-\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else
- \@citea\def\@citea{,\,\hskip\z@skip}%
- \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else
- {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else
- \def\@citea{--}\fi
- \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi}
-\else
-\renewenvironment{thebibliography}[1]
- {\section*{\refname}
- \small
- \list{}%
- {\settowidth\labelwidth{}%
- \leftmargin\parindent
- \itemindent=-\parindent
- \labelsep=\z@
- \if@openbib
- \advance\leftmargin\bibindent
- \itemindent -\bibindent
- \listparindent \itemindent
- \parsep \z@
- \fi
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{}}%
- \if@openbib
- \renewcommand\newblock{\par}%
- \else
- \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
- \fi
- \sloppy\clubpenalty4000\widowpenalty4000%
- \sfcode`\.=\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
- \def\@cite#1{#1}%
- \def\@lbibitem[#1]#2{\item[]\if@filesw
- {\def\protect##1{\string ##1\space}\immediate
- \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
- \fi
-\else
-\@cons\@openbib@code{\noexpand\small}
-\fi
-
-\def\idxquad{\hskip 10\p@}% space that divides entry from number
-
-\def\@idxitem{\par\hangindent 10\p@}
-
-\def\subitem{\par\setbox0=\hbox{--\enspace}% second order
- \noindent\hangindent\wd0\box0}% index entry
-
-\def\subsubitem{\par\setbox0=\hbox{--\,--\enspace}% third
- \noindent\hangindent\wd0\box0}% order index entry
-
-\def\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax}
-
-\renewenvironment{theindex}
- {\@mkboth{\indexname}{\indexname}%
- \thispagestyle{empty}\parindent\z@
- \parskip\z@ \@plus .3\p@\relax
- \let\item\par
- \def\,{\relax\ifmmode\mskip\thinmuskip
- \else\hskip0.2em\ignorespaces\fi}%
- \normalfont\small
- \begin{multicols}{2}[\@makeschapterhead{\indexname}]%
- }
- {\end{multicols}}
-
-\renewcommand\footnoterule{%
- \kern-3\p@
- \hrule\@width 2truecm
- \kern2.6\p@}
- \newdimen\fnindent
- \fnindent1em
-\long\def\@makefntext#1{%
- \parindent \fnindent%
- \leftskip \fnindent%
- \noindent
- \llap{\hb@xt@1em{\hss\@makefnmark\ }}\ignorespaces#1}
-
-\long\def\@makecaption#1#2{%
- \vskip\abovecaptionskip
- \sbox\@tempboxa{{\bfseries #1.} #2}%
- \ifdim \wd\@tempboxa >\hsize
- {\bfseries #1.} #2\par
- \else
- \global \@minipagefalse
- \hb@xt@\hsize{\hfil\box\@tempboxa\hfil}%
- \fi
- \vskip\belowcaptionskip}
-
-\def\fps@figure{htbp}
-\def\fnum@figure{\figurename\thinspace\thefigure}
-\def \@floatboxreset {%
- \reset@font
- \small
- \@setnobreak
- \@setminipage
-}
-\def\fps@table{htbp}
-\def\fnum@table{\tablename~\thetable}
-\renewenvironment{table}
- {\setlength\abovecaptionskip{0\p@}%
- \setlength\belowcaptionskip{10\p@}%
- \@float{table}}
- {\end@float}
-\renewenvironment{table*}
- {\setlength\abovecaptionskip{0\p@}%
- \setlength\belowcaptionskip{10\p@}%
- \@dblfloat{table}}
- {\end@dblfloat}
-
-\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname
- ext@#1\endcsname}{#1}{\protect\numberline{\csname
- the#1\endcsname}{\ignorespaces #2}}\begingroup
- \@parboxrestore
- \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
- \endgroup}
-
-% LaTeX does not provide a command to enter the authors institute
-% addresses. The \institute command is defined here.
-
-\newcounter{@inst}
-\newcounter{@auth}
-\newcounter{auco}
-\newdimen\instindent
-\newbox\authrun
-\newtoks\authorrunning
-\newtoks\tocauthor
-\newbox\titrun
-\newtoks\titlerunning
-\newtoks\toctitle
-
-\def\clearheadinfo{\gdef\@author{No Author Given}%
- \gdef\@title{No Title Given}%
- \gdef\@subtitle{}%
- \gdef\@institute{No Institute Given}%
- \gdef\@thanks{}%
- \global\titlerunning={}\global\authorrunning={}%
- \global\toctitle={}\global\tocauthor={}}
-
-\def\institute#1{\gdef\@institute{#1}}
-
-\def\institutename{\par
- \begingroup
- \parskip=\z@
- \parindent=\z@
- \setcounter{@inst}{1}%
- \def\and{\par\stepcounter{@inst}%
- \noindent$^{\the@inst}$\enspace\ignorespaces}%
- \setbox0=\vbox{\def\thanks##1{}\@institute}%
- \ifnum\c@@inst=1\relax
- \gdef\fnnstart{0}%
- \else
- \xdef\fnnstart{\c@@inst}%
- \setcounter{@inst}{1}%
- \noindent$^{\the@inst}$\enspace
- \fi
- \ignorespaces
- \@institute\par
- \endgroup}
-
-\def\@fnsymbol#1{\ensuremath{\ifcase#1\or\star\or{\star\star}\or
- {\star\star\star}\or \dagger\or \ddagger\or
- \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger
- \or \ddagger\ddagger \else\@ctrerr\fi}}
-
-\def\inst#1{\unskip$^{#1}$}
-\def\fnmsep{\unskip$^,$}
-\def\email#1{{\tt#1}}
-\AtBeginDocument{\@ifundefined{url}{\def\url#1{#1}}{}%
-\@ifpackageloaded{babel}{%
-\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}%
-\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}%
-\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}%
-}{\switcht@@therlang}%
-}
-\def\homedir{\~{ }}
-
-\def\subtitle#1{\gdef\@subtitle{#1}}
-\clearheadinfo
-
-\renewcommand\maketitle{\newpage
- \refstepcounter{chapter}%
- \stepcounter{section}%
- \setcounter{section}{0}%
- \setcounter{subsection}{0}%
- \setcounter{figure}{0}
- \setcounter{table}{0}
- \setcounter{equation}{0}
- \setcounter{footnote}{0}%
- \begingroup
- \parindent=\z@
- \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
- \if@twocolumn
- \ifnum \col@number=\@ne
- \@maketitle
- \else
- \twocolumn[\@maketitle]%
- \fi
- \else
- \newpage
- \global\@topnum\z@ % Prevents figures from going at top of page.
- \@maketitle
- \fi
- \thispagestyle{empty}\@thanks
-%
- \def\\{\unskip\ \ignorespaces}\def\inst##1{\unskip{}}%
- \def\thanks##1{\unskip{}}\def\fnmsep{\unskip}%
- \instindent=\hsize
- \advance\instindent by-\headlineindent
-% \if!\the\toctitle!\addcontentsline{toc}{title}{\@title}\else
-% \addcontentsline{toc}{title}{\the\toctitle}\fi
- \if@runhead
- \if!\the\titlerunning!\else
- \edef\@title{\the\titlerunning}%
- \fi
- \global\setbox\titrun=\hbox{\small\rm\unboldmath\ignorespaces\@title}%
- \ifdim\wd\titrun>\instindent
- \typeout{Title too long for running head. Please supply}%
- \typeout{a shorter form with \string\titlerunning\space prior to
- \string\maketitle}%
- \global\setbox\titrun=\hbox{\small\rm
- Title Suppressed Due to Excessive Length}%
- \fi
- \xdef\@title{\copy\titrun}%
- \fi
-%
- \if!\the\tocauthor!\relax
- {\def\and{\noexpand\protect\noexpand\and}%
- \protected@xdef\toc@uthor{\@author}}%
- \else
- \def\\{\noexpand\protect\noexpand\newline}%
- \protected@xdef\scratch{\the\tocauthor}%
- \protected@xdef\toc@uthor{\scratch}%
- \fi
-% \addcontentsline{toc}{author}{\toc@uthor}%
- \if@runhead
- \if!\the\authorrunning!
- \value{@inst}=\value{@auth}%
- \setcounter{@auth}{1}%
- \else
- \edef\@author{\the\authorrunning}%
- \fi
- \global\setbox\authrun=\hbox{\small\unboldmath\@author\unskip}%
- \ifdim\wd\authrun>\instindent
- \typeout{Names of authors too long for running head. Please supply}%
- \typeout{a shorter form with \string\authorrunning\space prior to
- \string\maketitle}%
- \global\setbox\authrun=\hbox{\small\rm
- Authors Suppressed Due to Excessive Length}%
- \fi
- \xdef\@author{\copy\authrun}%
- \markboth{\@author}{\@title}%
- \fi
- \endgroup
- \setcounter{footnote}{\fnnstart}%
- \clearheadinfo}
-%
-\def\@maketitle{\newpage
- \markboth{}{}%
- \def\lastand{\ifnum\value{@inst}=2\relax
- \unskip{} \andname\
- \else
- \unskip \lastandname\
- \fi}%
- \def\and{\stepcounter{@auth}\relax
- \ifnum\value{@auth}=\value{@inst}%
- \lastand
- \else
- \unskip,
- \fi}%
- \begin{center}%
- \let\newline\\
- {\Large \bfseries\boldmath
- \pretolerance=10000
- \@title \par}\vskip .8cm
-\if!\@subtitle!\else {\large \bfseries\boldmath
- \vskip -.65cm
- \pretolerance=10000
- \@subtitle \par}\vskip .8cm\fi
- \setbox0=\vbox{\setcounter{@auth}{1}\def\and{\stepcounter{@auth}}%
- \def\thanks##1{}\@author}%
- \global\value{@inst}=\value{@auth}%
- \global\value{auco}=\value{@auth}%
- \setcounter{@auth}{1}%
-{\lineskip .5em
-\noindent\ignorespaces
-\@author\vskip.35cm}
- {\small\institutename}
- \end{center}%
- }
-
-% definition of the "\spnewtheorem" command.
-%
-% Usage:
-%
-% \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font}
-% or \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font}
-% or \spnewtheorem*{env_nam}{caption}{cap_font}{body_font}
-%
-% New is "cap_font" and "body_font". It stands for
-% fontdefinition of the caption and the text itself.
-%
-% "\spnewtheorem*" gives a theorem without number.
-%
-% A defined spnewthoerem environment is used as described
-% by Lamport.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\def\@thmcountersep{}
-\def\@thmcounterend{.}
-
-\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}}
-
-% definition of \spnewtheorem with number
-
-\def\@spnthm#1#2{%
- \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}}
-\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}}
-
-\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}\@addtoreset{#1}{#3}%
- \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand
- \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}%
- \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spothm#1[#2]#3#4#5{%
- \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}%
- {\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{the#1}{\@nameuse{the#2}}%
- \expandafter\xdef\csname #1name\endcsname{#3}%
- \global\@namedef{#1}{\@spthm{#2}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}}
-
-\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\refstepcounter{#1}%
-\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}}
-
-\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}%
- \ignorespaces}
-
-\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname
- the#1\endcsname}{#5}{#3}{#4}\ignorespaces}
-
-\def\@spbegintheorem#1#2#3#4{\trivlist
- \item[\hskip\labelsep{#3#1\ #2\@thmcounterend}]#4}
-
-\def\@spopargbegintheorem#1#2#3#4#5{\trivlist
- \item[\hskip\labelsep{#4#1\ #2}]{#4(#3)\@thmcounterend\ }#5}
-
-% definition of \spnewtheorem* without number
-
-\def\@sthm#1#2{\@Ynthm{#1}{#2}}
-
-\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}}
-
-\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces}
-
-\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1}
- {#4}{#2}{#3}\ignorespaces}
-
-\def\@Begintheorem#1#2#3{#3\trivlist
- \item[\hskip\labelsep{#2#1\@thmcounterend}]}
-
-\def\@Opargbegintheorem#1#2#3#4{#4\trivlist
- \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }}
-
-\if@envcntsect
- \def\@thmcountersep{.}
- \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape}
-\else
- \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape}
- \if@envcntreset
- \@addtoreset{theorem}{section}
- \else
- \@addtoreset{theorem}{chapter}
- \fi
-\fi
-
-%definition of divers theorem environments
-\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily}
-\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily}
-\if@envcntsame % alle Umgebungen wie Theorem.
- \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}}
-\else % alle Umgebungen mit eigenem Zaehler
- \if@envcntsect % mit section numeriert
- \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}}
- \else % nicht mit section numeriert
- \if@envcntreset
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{section}}
- \else
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{chapter}}%
- \fi
- \fi
-\fi
-\spn@wtheorem{case}{Case}{\itshape}{\rmfamily}
-\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily}
-\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape}
-\spn@wtheorem{definition}{Definition}{\bfseries}{\itshape}
-\spn@wtheorem{example}{Example}{\itshape}{\rmfamily}
-\spn@wtheorem{exercise}{Exercise}{\itshape}{\rmfamily}
-\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape}
-\spn@wtheorem{note}{Note}{\itshape}{\rmfamily}
-\spn@wtheorem{problem}{Problem}{\itshape}{\rmfamily}
-\spn@wtheorem{property}{Property}{\itshape}{\rmfamily}
-\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape}
-\spn@wtheorem{question}{Question}{\itshape}{\rmfamily}
-\spn@wtheorem{solution}{Solution}{\itshape}{\rmfamily}
-\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily}
-
-\def\@takefromreset#1#2{%
- \def\@tempa{#1}%
- \let\@tempd\@elt
- \def\@elt##1{%
- \def\@tempb{##1}%
- \ifx\@tempa\@tempb\else
- \@addtoreset{##1}{#2}%
- \fi}%
- \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
- \expandafter\def\csname cl@#2\endcsname{}%
- \@tempc
- \let\@elt\@tempd}
-
-\def\theopargself{\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5}
- \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
- \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }}
- }
-
-\renewenvironment{abstract}{%
- \list{}{\advance\topsep by0.35cm\relax\small
- \leftmargin=1cm
- \labelwidth=\z@
- \listparindent=\z@
- \itemindent\listparindent
- \rightmargin\leftmargin}\item[\hskip\labelsep
- \bfseries\abstractname]}
- {\endlist}
-
-\newdimen\headlineindent % dimension for space between
-\headlineindent=1.166cm % number and text of headings.
-
-\def\ps@headings{\let\@mkboth\@gobbletwo
- \let\@oddfoot\@empty\let\@evenfoot\@empty
- \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
- \leftmark\hfil}
- \def\@oddhead{\normalfont\small\hfil\rightmark\hspace{\headlineindent}%
- \llap{\thepage}}
- \def\chaptermark##1{}%
- \def\sectionmark##1{}%
- \def\subsectionmark##1{}}
-
-\def\ps@titlepage{\let\@mkboth\@gobbletwo
- \let\@oddfoot\@empty\let\@evenfoot\@empty
- \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
- \hfil}
- \def\@oddhead{\normalfont\small\hfil\hspace{\headlineindent}%
- \llap{\thepage}}
- \def\chaptermark##1{}%
- \def\sectionmark##1{}%
- \def\subsectionmark##1{}}
-
-\if@runhead\ps@headings\else
-\ps@empty\fi
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\endinput
-%end of file llncs.cls
--- a/document/mathpartir.sty Thu Jan 17 11:51:00 2013 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,446 +0,0 @@
-% Mathpartir --- Math Paragraph for Typesetting Inference Rules
-%
-% Copyright (C) 2001, 2002, 2003, 2004, 2005 Didier Rémy
-%
-% Author : Didier Remy
-% Version : 1.2.0
-% Bug Reports : to author
-% Web Site : http://pauillac.inria.fr/~remy/latex/
-%
-% Mathpartir is free software; you can redistribute it and/or modify
-% it under the terms of the GNU General Public License as published by
-% the Free Software Foundation; either version 2, or (at your option)
-% any later version.
-%
-% Mathpartir is distributed in the hope that it will be useful,
-% but WITHOUT ANY WARRANTY; without even the implied warranty of
-% MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
-% GNU General Public License for more details
-% (http://pauillac.inria.fr/~remy/license/GPL).
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% File mathpartir.sty (LaTeX macros)
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\NeedsTeXFormat{LaTeX2e}
-\ProvidesPackage{mathpartir}
- [2005/12/20 version 1.2.0 Math Paragraph for Typesetting Inference Rules]
-
-%%
-
-%% Identification
-%% Preliminary declarations
-
-\RequirePackage {keyval}
-
-%% Options
-%% More declarations
-
-%% PART I: Typesetting maths in paragraphe mode
-
-%% \newdimen \mpr@tmpdim
-%% Dimens are a precious ressource. Uses seems to be local.
-\let \mpr@tmpdim \@tempdima
-
-% To ensure hevea \hva compatibility, \hva should expands to nothing
-% in mathpar or in inferrule
-\let \mpr@hva \empty
-
-%% normal paragraph parametters, should rather be taken dynamically
-\def \mpr@savepar {%
- \edef \MathparNormalpar
- {\noexpand \lineskiplimit \the\lineskiplimit
- \noexpand \lineskip \the\lineskip}%
- }
-
-\def \mpr@rulelineskip {\lineskiplimit=0.3em\lineskip=0.2em plus 0.1em}
-\def \mpr@lesslineskip {\lineskiplimit=0.6em\lineskip=0.5em plus 0.2em}
-\def \mpr@lineskip {\lineskiplimit=1.2em\lineskip=1.2em plus 0.2em}
-\let \MathparLineskip \mpr@lineskip
-\def \mpr@paroptions {\MathparLineskip}
-\let \mpr@prebindings \relax
-
-\newskip \mpr@andskip \mpr@andskip 2em plus 0.5fil minus 0.5em
-
-\def \mpr@goodbreakand
- {\hskip -\mpr@andskip \penalty -1000\hskip \mpr@andskip}
-\def \mpr@and {\hskip \mpr@andskip}
-\def \mpr@andcr {\penalty 50\mpr@and}
-\def \mpr@cr {\penalty -10000\mpr@and}
-\def \mpr@eqno #1{\mpr@andcr #1\hskip 0em plus -1fil \penalty 10}
-
-\def \mpr@bindings {%
- \let \and \mpr@andcr
- \let \par \mpr@andcr
- \let \\\mpr@cr
- \let \eqno \mpr@eqno
- \let \hva \mpr@hva
- }
-\let \MathparBindings \mpr@bindings
-
-% \@ifundefined {ignorespacesafterend}
-% {\def \ignorespacesafterend {\aftergroup \ignorespaces}
-
-\newenvironment{mathpar}[1][]
- {$$\mpr@savepar \parskip 0em \hsize \linewidth \centering
- \vbox \bgroup \mpr@prebindings \mpr@paroptions #1\ifmmode $\else
- \noindent $\displaystyle\fi
- \MathparBindings}
- {\unskip \ifmmode $\fi\egroup $$\ignorespacesafterend}
-
-\newenvironment{mathparpagebreakable}[1][]
- {\begingroup
- \par
- \mpr@savepar \parskip 0em \hsize \linewidth \centering
- \mpr@prebindings \mpr@paroptions #1%
- \vskip \abovedisplayskip \vskip -\lineskip%
- \ifmmode \else $\displaystyle\fi
- \MathparBindings
- }
- {\unskip
- \ifmmode $\fi \par\endgroup
- \vskip \belowdisplayskip
- \noindent
- \ignorespacesafterend}
-
-% \def \math@mathpar #1{\setbox0 \hbox {$\displaystyle #1$}\ifnum
-% \wd0 < \hsize $$\box0$$\else \bmathpar #1\emathpar \fi}
-
-%%% HOV BOXES
-
-\def \mathvbox@ #1{\hbox \bgroup \mpr@normallineskip
- \vbox \bgroup \tabskip 0em \let \\ \cr
- \halign \bgroup \hfil $##$\hfil\cr #1\crcr \egroup \egroup
- \egroup}
-
-\def \mathhvbox@ #1{\setbox0 \hbox {\let \\\qquad $#1$}\ifnum \wd0 < \hsize
- \box0\else \mathvbox {#1}\fi}
-
-
-%% Part II -- operations on lists
-
-\newtoks \mpr@lista
-\newtoks \mpr@listb
-
-\long \def\mpr@cons #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter
-{#2}\edef #2{\the \mpr@lista \the \mpr@listb}}
-
-\long \def\mpr@snoc #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter
-{#2}\edef #2{\the \mpr@listb\the\mpr@lista}}
-
-\long \def \mpr@concat#1=#2\mpr@to#3{\mpr@lista \expandafter {#2}\mpr@listb
-\expandafter {#3}\edef #1{\the \mpr@listb\the\mpr@lista}}
-
-\def \mpr@head #1\mpr@to #2{\expandafter \mpr@head@ #1\mpr@head@ #1#2}
-\long \def \mpr@head@ #1#2\mpr@head@ #3#4{\def #4{#1}\def#3{#2}}
-
-\def \mpr@flatten #1\mpr@to #2{\expandafter \mpr@flatten@ #1\mpr@flatten@ #1#2}
-\long \def \mpr@flatten@ \\#1\\#2\mpr@flatten@ #3#4{\def #4{#1}\def #3{\\#2}}
-
-\def \mpr@makelist #1\mpr@to #2{\def \mpr@all {#1}%
- \mpr@lista {\\}\mpr@listb \expandafter {\mpr@all}\edef \mpr@all {\the
- \mpr@lista \the \mpr@listb \the \mpr@lista}\let #2\empty
- \def \mpr@stripof ##1##2\mpr@stripend{\def \mpr@stripped{##2}}\loop
- \mpr@flatten \mpr@all \mpr@to \mpr@one
- \expandafter \mpr@snoc \mpr@one \mpr@to #2\expandafter \mpr@stripof
- \mpr@all \mpr@stripend
- \ifx \mpr@stripped \empty \let \mpr@isempty 0\else \let \mpr@isempty 1\fi
- \ifx 1\mpr@isempty
- \repeat
-}
-
-\def \mpr@rev #1\mpr@to #2{\let \mpr@tmp \empty
- \def \\##1{\mpr@cons ##1\mpr@to \mpr@tmp}#1\let #2\mpr@tmp}
-
-%% Part III -- Type inference rules
-
-\newif \if@premisse
-\newbox \mpr@hlist
-\newbox \mpr@vlist
-\newif \ifmpr@center \mpr@centertrue
-\def \mpr@htovlist {%
- \setbox \mpr@hlist
- \hbox {\strut
- \ifmpr@center \hskip -0.5\wd\mpr@hlist\fi
- \unhbox \mpr@hlist}%
- \setbox \mpr@vlist
- \vbox {\if@premisse \box \mpr@hlist \unvbox \mpr@vlist
- \else \unvbox \mpr@vlist \box \mpr@hlist
- \fi}%
-}
-% OLD version
-% \def \mpr@htovlist {%
-% \setbox \mpr@hlist
-% \hbox {\strut \hskip -0.5\wd\mpr@hlist \unhbox \mpr@hlist}%
-% \setbox \mpr@vlist
-% \vbox {\if@premisse \box \mpr@hlist \unvbox \mpr@vlist
-% \else \unvbox \mpr@vlist \box \mpr@hlist
-% \fi}%
-% }
-
-\def \mpr@item #1{$\displaystyle #1$}
-\def \mpr@sep{2em}
-\def \mpr@blank { }
-\def \mpr@hovbox #1#2{\hbox
- \bgroup
- \ifx #1T\@premissetrue
- \else \ifx #1B\@premissefalse
- \else
- \PackageError{mathpartir}
- {Premisse orientation should either be T or B}
- {Fatal error in Package}%
- \fi \fi
- \def \@test {#2}\ifx \@test \mpr@blank\else
- \setbox \mpr@hlist \hbox {}%
- \setbox \mpr@vlist \vbox {}%
- \if@premisse \let \snoc \mpr@cons \else \let \snoc \mpr@snoc \fi
- \let \@hvlist \empty \let \@rev \empty
- \mpr@tmpdim 0em
- \expandafter \mpr@makelist #2\mpr@to \mpr@flat
- \if@premisse \mpr@rev \mpr@flat \mpr@to \@rev \else \let \@rev \mpr@flat \fi
- \def \\##1{%
- \def \@test {##1}\ifx \@test \empty
- \mpr@htovlist
- \mpr@tmpdim 0em %%% last bug fix not extensively checked
- \else
- \setbox0 \hbox{\mpr@item {##1}}\relax
- \advance \mpr@tmpdim by \wd0
- %\mpr@tmpdim 1.02\mpr@tmpdim
- \ifnum \mpr@tmpdim < \hsize
- \ifnum \wd\mpr@hlist > 0
- \if@premisse
- \setbox \mpr@hlist
- \hbox {\unhbox0 \hskip \mpr@sep \unhbox \mpr@hlist}%
- \else
- \setbox \mpr@hlist
- \hbox {\unhbox \mpr@hlist \hskip \mpr@sep \unhbox0}%
- \fi
- \else
- \setbox \mpr@hlist \hbox {\unhbox0}%
- \fi
- \else
- \ifnum \wd \mpr@hlist > 0
- \mpr@htovlist
- \mpr@tmpdim \wd0
- \fi
- \setbox \mpr@hlist \hbox {\unhbox0}%
- \fi
- \advance \mpr@tmpdim by \mpr@sep
- \fi
- }%
- \@rev
- \mpr@htovlist
- \ifmpr@center \hskip \wd\mpr@vlist\fi \box \mpr@vlist
- \fi
- \egroup
-}
-
-%%% INFERENCE RULES
-
-\@ifundefined{@@over}{%
- \let\@@over\over % fallback if amsmath is not loaded
- \let\@@overwithdelims\overwithdelims
- \let\@@atop\atop \let\@@atopwithdelims\atopwithdelims
- \let\@@above\above \let\@@abovewithdelims\abovewithdelims
- }{}
-
-%% The default
-
-\def \mpr@@fraction #1#2{\hbox {\advance \hsize by -0.5em
- $\displaystyle {#1\mpr@over #2}$}}
-\def \mpr@@nofraction #1#2{\hbox {\advance \hsize by -0.5em
- $\displaystyle {#1\@@atop #2}$}}
-
-\let \mpr@fraction \mpr@@fraction
-
-%% A generic solution to arrow
-
-\def \mpr@make@fraction #1#2#3#4#5{\hbox {%
- \def \mpr@tail{#1}%
- \def \mpr@body{#2}%
- \def \mpr@head{#3}%
- \setbox1=\hbox{$#4$}\setbox2=\hbox{$#5$}%
- \setbox3=\hbox{$\mkern -3mu\mpr@body\mkern -3mu$}%
- \setbox3=\hbox{$\mkern -3mu \mpr@body\mkern -3mu$}%
- \dimen0=\dp1\advance\dimen0 by \ht3\relax\dp1\dimen0\relax
- \dimen0=\ht2\advance\dimen0 by \dp3\relax\ht2\dimen0\relax
- \setbox0=\hbox {$\box1 \@@atop \box2$}%
- \dimen0=\wd0\box0
- \box0 \hskip -\dimen0\relax
- \hbox to \dimen0 {$%
- \mathrel{\mpr@tail}\joinrel
- \xleaders\hbox{\copy3}\hfil\joinrel\mathrel{\mpr@head}%
- $}}}
-
-%% Old stuff should be removed in next version
-\def \mpr@@nothing #1#2
- {$\lower 0.01pt \mpr@@nofraction {#1}{#2}$}
-\def \mpr@@reduce #1#2{\hbox
- {$\lower 0.01pt \mpr@@fraction {#1}{#2}\mkern -15mu\rightarrow$}}
-\def \mpr@@rewrite #1#2#3{\hbox
- {$\lower 0.01pt \mpr@@fraction {#2}{#3}\mkern -8mu#1$}}
-\def \mpr@infercenter #1{\vcenter {\mpr@hovbox{T}{#1}}}
-
-\def \mpr@empty {}
-\def \mpr@inferrule
- {\bgroup
- \ifnum \linewidth<\hsize \hsize \linewidth\fi
- \mpr@rulelineskip
- \let \and \qquad
- \let \hva \mpr@hva
- \let \@rulename \mpr@empty
- \let \@rule@options \mpr@empty
- \let \mpr@over \@@over
- \mpr@inferrule@}
-\newcommand {\mpr@inferrule@}[3][]
- {\everymath={\displaystyle}%
- \def \@test {#2}\ifx \empty \@test
- \setbox0 \hbox {$\vcenter {\mpr@hovbox{B}{#3}}$}%
- \else
- \def \@test {#3}\ifx \empty \@test
- \setbox0 \hbox {$\vcenter {\mpr@hovbox{T}{#2}}$}%
- \else
- \setbox0 \mpr@fraction {\mpr@hovbox{T}{#2}}{\mpr@hovbox{B}{#3}}%
- \fi \fi
- \def \@test {#1}\ifx \@test\empty \box0
- \else \vbox
-%%% Suggestion de Francois pour les etiquettes longues
-%%% {\hbox to \wd0 {\RefTirName {#1}\hfil}\box0}\fi
- {\hbox {\RefTirName {#1}}\box0}\fi
- \egroup}
-
-\def \mpr@vdotfil #1{\vbox to #1{\leaders \hbox{$\cdot$} \vfil}}
-
-% They are two forms
-% \inferrule [label]{[premisses}{conclusions}
-% or
-% \inferrule* [options]{[premisses}{conclusions}
-%
-% Premisses and conclusions are lists of elements separated by \\
-% Each \\ produces a break, attempting horizontal breaks if possible,
-% and vertical breaks if needed.
-%
-% An empty element obtained by \\\\ produces a vertical break in all cases.
-%
-% The former rule is aligned on the fraction bar.
-% The optional label appears on top of the rule
-% The second form to be used in a derivation tree is aligned on the last
-% line of its conclusion
-%
-% The second form can be parameterized, using the key=val interface. The
-% folloiwng keys are recognized:
-%
-% width set the width of the rule to val
-% narrower set the width of the rule to val\hsize
-% before execute val at the beginning/left
-% lab put a label [Val] on top of the rule
-% lskip add negative skip on the right
-% left put a left label [Val]
-% Left put a left label [Val], ignoring its width
-% right put a right label [Val]
-% Right put a right label [Val], ignoring its width
-% leftskip skip negative space on the left-hand side
-% rightskip skip negative space on the right-hand side
-% vdots lift the rule by val and fill vertical space with dots
-% after execute val at the end/right
-%
-% Note that most options must come in this order to avoid strange
-% typesetting (in particular leftskip must preceed left and Left and
-% rightskip must follow Right or right; vdots must come last
-% or be only followed by rightskip.
-%
-
-%% Keys that make sence in all kinds of rules
-\def \mprset #1{\setkeys{mprset}{#1}}
-\define@key {mprset}{andskip}[]{\mpr@andskip=#1}
-\define@key {mprset}{lineskip}[]{\lineskip=#1}
-\define@key {mprset}{flushleft}[]{\mpr@centerfalse}
-\define@key {mprset}{center}[]{\mpr@centertrue}
-\define@key {mprset}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite}
-\define@key {mprset}{atop}[]{\let \mpr@fraction \mpr@@nofraction}
-\define@key {mprset}{myfraction}[]{\let \mpr@fraction #1}
-\define@key {mprset}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}}
-\define@key {mprset}{sep}{\def\mpr@sep{#1}}
-
-\newbox \mpr@right
-\define@key {mpr}{flushleft}[]{\mpr@centerfalse}
-\define@key {mpr}{center}[]{\mpr@centertrue}
-\define@key {mpr}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite}
-\define@key {mpr}{myfraction}[]{\let \mpr@fraction #1}
-\define@key {mpr}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}}
-\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax
- \advance \hsize by -\wd0\box0}
-\define@key {mpr}{width}{\hsize #1}
-\define@key {mpr}{sep}{\def\mpr@sep{#1}}
-\define@key {mpr}{before}{#1}
-\define@key {mpr}{lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}}
-\define@key {mpr}{Lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}}
-\define@key {mpr}{narrower}{\hsize #1\hsize}
-\define@key {mpr}{leftskip}{\hskip -#1}
-\define@key {mpr}{reduce}[]{\let \mpr@fraction \mpr@@reduce}
-\define@key {mpr}{rightskip}
- {\setbox \mpr@right \hbox {\unhbox \mpr@right \hskip -#1}}
-\define@key {mpr}{LEFT}{\setbox0 \hbox {$#1$}\relax
- \advance \hsize by -\wd0\box0}
-\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax
- \advance \hsize by -\wd0\box0}
-\define@key {mpr}{Left}{\llap{$\TirName {#1}\;$}}
-\define@key {mpr}{right}
- {\setbox0 \hbox {$\;\TirName {#1}$}\relax \advance \hsize by -\wd0
- \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}}
-\define@key {mpr}{RIGHT}
- {\setbox0 \hbox {$#1$}\relax \advance \hsize by -\wd0
- \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}}
-\define@key {mpr}{Right}
- {\setbox \mpr@right \hbox {\unhbox \mpr@right \rlap {$\;\TirName {#1}$}}}
-\define@key {mpr}{vdots}{\def \mpr@vdots {\@@atop \mpr@vdotfil{#1}}}
-\define@key {mpr}{after}{\edef \mpr@after {\mpr@after #1}}
-
-\newcommand \mpr@inferstar@ [3][]{\setbox0
- \hbox {\let \mpr@rulename \mpr@empty \let \mpr@vdots \relax
- \setbox \mpr@right \hbox{}%
- $\setkeys{mpr}{#1}%
- \ifx \mpr@rulename \mpr@empty \mpr@inferrule {#2}{#3}\else
- \mpr@inferrule [{\mpr@rulename}]{#2}{#3}\fi
- \box \mpr@right \mpr@vdots$}
- \setbox1 \hbox {\strut}
- \@tempdima \dp0 \advance \@tempdima by -\dp1
- \raise \@tempdima \box0}
-
-\def \mpr@infer {\@ifnextchar *{\mpr@inferstar}{\mpr@inferrule}}
-\newcommand \mpr@err@skipargs[3][]{}
-\def \mpr@inferstar*{\ifmmode
- \let \@do \mpr@inferstar@
- \else
- \let \@do \mpr@err@skipargs
- \PackageError {mathpartir}
- {\string\inferrule* can only be used in math mode}{}%
- \fi \@do}
-
-
-%%% Exports
-
-% Envirnonment mathpar
-
-\let \inferrule \mpr@infer
-
-% make a short name \infer is not already defined
-\@ifundefined {infer}{\let \infer \mpr@infer}{}
-
-\def \TirNameStyle #1{\small \textsc{#1}}
-\def \tir@name #1{\hbox {\small \TirNameStyle{#1}}}
-\let \TirName \tir@name
-\let \DefTirName \TirName
-\let \RefTirName \TirName
-
-%%% Other Exports
-
-% \let \listcons \mpr@cons
-% \let \listsnoc \mpr@snoc
-% \let \listhead \mpr@head
-% \let \listmake \mpr@makelist
-
-
-
-
-\endinput
--- a/document/root.bib Thu Jan 17 11:51:00 2013 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,68 +0,0 @@
-@article{UrbanCheneyBerghofer11,
- author = {C.~Urban and J.~Cheney and S.~Berghofer},
- title = {{M}echanizing the {M}etatheory of {LF}},
- journal = {ACM Transactions on Computational Logic},
- volume = {12},
- issue = {2},
- year = {2011},
- pages = {15:1--15:42}
-}
-
-@inproceedings{Norrish11,
- author = {M.~Norrish},
- title = {{M}echanised {C}omputability {T}heory},
- booktitle = {Proc.~of the 2nd Conference on Interactive Theorem Proving (ITP)},
- year = {2011},
- series = {LNCS},
- volume = {6898},
- pages = {297--311}
-}
-
-@inproceedings{AspertiRicciotti12,
- author = {A.~Asperti and W.~Ricciotti},
- title = {{F}ormalizing {T}uring {M}achines},
- booktitle = {Proc.~of the 19th International Workshop on Logic, Language,
- Information and Computation (WoLLIC)},
- year = {2012},
- pages = {1-25},
- series = {LNCS},
- volume = {7456}
-}
-
-
-@Unpublished{WuZhangUrban12,
- author = {C.~Wu and X.~Zhang and C.~Urban},
- title = {???},
- note = {Submitted},
- year = {2012}
-}
-
-@book{Boolos87,
- author = {G.~Boolos and J.~P.~Burgess and R.~C.~Jeffrey},
- title = {{C}omputability and {L}ogic (5th~ed.)},
- publisher = {Cambridge University Press},
- year = {2007}
-}
-
-@inproceedings{WuZhangUrban11,
- author = {C.~Wu and X.~Zhang and C.~Urban},
- title = {{A} {F}ormalisation of the {M}yhill-{N}erode {T}heorem based on {R}egular {E}xpressions
- ({P}roof {P}earl)},
- booktitle = {Proc.~of the 2nd Conference on Interactive Theorem Proving},
- year = {2011},
- pages = {341--356},
- series = {LNCS},
- volume = {6898}
-}
-
-
-@Article{Post36,
- author = {E.~Post},
- title = {{F}inite {C}ombinatory {P}rocesses-{F}ormulation 1},
- journal = {Journal of Symbolic Logic},
- year = {1936},
- volume = {1},
- number = {3},
- pages = {103--105}
-}
-
--- a/document/root.tex Thu Jan 17 11:51:00 2013 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,62 +0,0 @@
-\documentclass[runningheads]{llncs}
-%\documentclass[10pt, conference, compsocconf]{IEEEtran}
-\usepackage{isabelle}
-\usepackage{isabellesym}
-\usepackage{times}
-\usepackage{amssymb}
-\usepackage{amsmath}
-\usepackage{mathpartir}
-\usepackage{pdfsetup}
-\usepackage{tikz}
-\usepackage{pgf}
-
-% urls in roman style, theory text in math-similar italics
-\urlstyle{rm}
-\isabellestyle{it}
-
-% for uniform font size
-%\renewcommand{\isastyle}{\isastyleminor}
-
-\def\dn{\,\stackrel{\mbox{\scriptsize def}}{=}\,}
-\renewcommand{\isasymequiv}{$\dn$}
-\renewcommand{\isasymemptyset}{$\varnothing$}
-\renewcommand{\isacharunderscore}{\mbox{$\_$}}
-\renewcommand{\isasymiota}{}
-\newcommand{\isasymulcorner}{$\ulcorner$}
-\newcommand{\isasymurcorner}{$\urcorner$}
-\begin{document}
-
-
-\title{Formalising Computability Theory in Isabelle/HOL}
-\author{Jian Xu\inst{1} \and Xingyuan Zhang\inst{1} \and Christian Urban\inst{2}}
-\institute{PLA University of Science and Technology, China \and King's College London, UK}
-
-\maketitle
-
-
-\begin{abstract}
-We present a formalised theory of computability in the
-theorem prover Isabelle/HOL. This theorem prover is based on classical
-logic which precludes \emph{direct} reasoning about computability: every
-boolean predicate is either true or false because of the law of excluded
-middle. The only way to reason about computability in a classical theorem
-prover is to formalise a concrete model for computation.
-We formalise Turing machines and relate them to abacus machines and recursive
-functions. Our theory can be used to formalise other computability results:
-we give one example about the undecidability of Wang's tiling problem, whose proof uses
-the notion of a universal Turing machine.
-\end{abstract}
-
-% generated text of all theories
-\input{session}
-
-% optional bibliography
-\bibliographystyle{abbrv}
-\bibliography{root}
-
-\end{document}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
Binary file paper.pdf has changed
--- a/thys/abacus.thy Thu Jan 17 11:51:00 2013 +0000
+++ b/thys/abacus.thy Fri Jan 18 11:40:01 2013 +0000
@@ -381,7 +381,7 @@
lemma abc_step_red:
"abc_steps_l ac aprog stp = (as, am) \<Longrightarrow>
abc_steps_l ac aprog (Suc stp) = abc_step_l (as, am) (abc_fetch as aprog) "
-sorry
+oops
lemma tm_shift_fetch:
"\<lbrakk>fetch A s b = (ac, ns); ns \<noteq> 0 \<rbrakk>
@@ -445,12 +445,7 @@
(Suc s - start_of ly as) b = (ac, ns)"
and notfinal: "ns \<noteq> 0"
shows "fetch tp s b = (ac, ns)"
-proof -
- have "s > start_of ly as \<and> s < start_of ly (Suc as)"
- sorry
- thus "fetch tp s b = (ac, ns)"
- sorry
-qed
+oops
lemma step_eq_in:
assumes layout: "ly = layout_of ap"