thys2/Journal/Paper.tex
changeset 371 f65444d29e74
parent 369 e00950ba4514
child 372 78cc255e286f
--- a/thys2/Journal/Paper.tex	Mon Nov 01 10:52:44 2021 +0000
+++ b/thys2/Journal/Paper.tex	Tue Nov 02 13:57:59 2021 +0000
@@ -47,7 +47,7 @@
 %
 \isatagdocument
 %
-\isamarkupsection{Introduction%
+\isamarkupsection{Core of the proof%
 }
 \isamarkuptrue%
 %
@@ -69,12 +69,93 @@
 
 In the previous work, Ausaf and Urban established the below equality:
 \begin{lemma}
-\isa{{\normalsize{}If\,}\ v\ {\isacharcolon}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}{\isacharbackslash}{\kern0pt}c\ {\normalsize \,then\,}\ retrieve\ {\isacharparenleft}{\kern0pt}r\mbox{$\bbslash$}c{\isacharparenright}{\kern0pt}\ v\ {\isacharequal}{\kern0pt}\ retrieve\ r\ {\isacharparenleft}{\kern0pt}inj\ {\isacharparenleft}{\kern0pt}r\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}\ c\ v{\isacharparenright}{\kern0pt}{\isachardot}{\kern0pt}}
+\isa{{\normalsize{}If\,}\ v\ {\isacharcolon}{\kern0pt}\ r{\isacharbackslash}{\kern0pt}s\ {\normalsize \,then\,}\ Some\ {\isacharparenleft}{\kern0pt}flex\ r\ id\ s\ v{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ decode\ {\isacharparenleft}{\kern0pt}retrieve\ {\isacharparenleft}{\kern0pt}r\mbox{$^\uparrow$}\mbox{$\bbslash$}s{\isacharparenright}{\kern0pt}\ v{\isacharparenright}{\kern0pt}\ r{\isachardot}{\kern0pt}}
+\end{lemma}
+
+This lemma establishes a link with the lexer without bit-codes.
+
+With it we get the correctness of bit-coded algorithm.
+\begin{lemma}
+\isa{lexer\mbox{$_b$}\ r\ s\ {\isacharequal}{\kern0pt}\ lexer\ r\ s}
 \end{lemma}
 
-This lemma links the derivative of a bit-coded regular expression with
-the regular expression itself before the derivative. 
+However what is not certain is whether we can add simplification
+to the bit-coded algorithm, without breaking the correct lexing output.
+This might sound trivial in the case of producing a YES/NO answer,
+but once we require a lexing output to be produced (which is required
+in applications like compiler front-end, malicious attack domain extraction, 
+etc.), it is not straightforward if we still extract what is needed according
+to the POSIX standard.
+
+By simplification, we mean specifically the following rules:
+
+\begin{center}
+  \begin{tabular}{lcl}
+  \isa{\mbox{}\inferrule{\mbox{}}{\mbox{ASEQ\ bs\ AZERO\ r\isactrlsub {\isadigit{2}}\ {\isasymleadsto}\ AZERO}}}\\
+  \isa{\mbox{}\inferrule{\mbox{}}{\mbox{ASEQ\ bs\ r\isactrlsub {\isadigit{1}}\ AZERO\ {\isasymleadsto}\ AZERO}}}\\
+  \isa{\mbox{}\inferrule{\mbox{}}{\mbox{ASEQ\ bs\ {\isacharparenleft}{\kern0pt}AONE\ bs\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isasymleadsto}\ fuse\ {\isacharparenleft}{\kern0pt}bs\ {\isacharat}{\kern0pt}\ bs\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ r\isactrlsub {\isadigit{1}}}}}\\
+  \isa{\mbox{}\inferrule{\mbox{bs\ {\isasymleadsto}\ r\isactrlsub {\isadigit{1}}}}{\mbox{ASEQ\ bs\ bs\ r{\isadigit{3}}{\isachardot}{\kern0pt}{\isadigit{0}}\ {\isasymleadsto}\ ASEQ\ bs\ r\isactrlsub {\isadigit{1}}\ r{\isadigit{3}}{\isachardot}{\kern0pt}{\isadigit{0}}}}}\\
+  \isa{\mbox{}\inferrule{\mbox{bs\ {\isasymleadsto}\ r\isactrlsub {\isadigit{2}}}}{\mbox{ASEQ\ bs\ r{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ bs\ {\isasymleadsto}\ ASEQ\ bs\ r{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ r\isactrlsub {\isadigit{2}}}}}\\
+  \isa{\mbox{}\inferrule{\mbox{bs\ {\isasymleadsto}\ r\isactrlsub {\isadigit{1}}}}{\mbox{AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}bs{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rs{\isadigit{2}}{\isachardot}{\kern0pt}{\isadigit{0}}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rs{\isadigit{2}}{\isachardot}{\kern0pt}{\isadigit{0}}{\isacharparenright}{\kern0pt}}}}\\
+  \isa{\mbox{}\inferrule{\mbox{}}{\mbox{AALTs\ bs\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}\ {\isacharat}{\kern0pt}\ AZERO\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}rsb{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}\ {\isacharat}{\kern0pt}\ rsb{\isacharparenright}{\kern0pt}}}}\\
+  \isa{\mbox{}\inferrule{\mbox{}}{\mbox{AALTs\ bs\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ AALTs\ bs{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ rs{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}rsb{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ map\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}{\isacharparenright}{\kern0pt}\ rs{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ {\isacharat}{\kern0pt}\ rsb{\isacharparenright}{\kern0pt}}}}\\
+  \isa{\mbox{}\inferrule{\mbox{}}{\mbox{AALTs\ bs\ {\isacharparenleft}{\kern0pt}map\ {\isacharparenleft}{\kern0pt}fuse\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ {\isacharparenleft}{\kern0pt}bs\ {\isacharat}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ rs}}}\\
+  \isa{\mbox{}\inferrule{\mbox{}}{\mbox{AALTs\ {\isacharparenleft}{\kern0pt}bs\ {\isacharat}{\kern0pt}\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}map\ {\isacharparenleft}{\kern0pt}fuse\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}}}}\\
+  \isa{\mbox{}\inferrule{\mbox{}}{\mbox{AALTs\ bs\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isasymleadsto}\ AZERO}}}\\
+  \isa{\mbox{}\inferrule{\mbox{}}{\mbox{AALTs\ bs\ {\isacharbrackleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharbrackright}{\kern0pt}\ {\isasymleadsto}\ fuse\ bs\ r\isactrlsub {\isadigit{1}}}}}\\
+  \isa{\mbox{}\inferrule{\mbox{bs\mbox{$^\downarrow$}\ {\isacharequal}{\kern0pt}\ r\isactrlsub {\isadigit{2}}\mbox{$^\downarrow$}}}{\mbox{AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}bs{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rsb\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rsc{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}bs{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rsb\ {\isacharat}{\kern0pt}\ rsc{\isacharparenright}{\kern0pt}}}}\\
+
+
+  \isa{\mbox{}\inferrule{\mbox{}}{\mbox{Empty\ {\isacharcolon}{\kern0pt}\ \isactrlbold {\isadigit{1}}}}} & 
+  \isa{\mbox{}\inferrule{\mbox{}}{\mbox{Char\ c\ {\isacharcolon}{\kern0pt}\ c}}}\\[4mm]
+  \isa{\mbox{}\inferrule{\mbox{v\isactrlsub {\isadigit{1}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{1}}}}{\mbox{Left\ v\isactrlsub {\isadigit{1}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}}}} &
+  \isa{\mbox{}\inferrule{\mbox{v\isactrlsub {\isadigit{2}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{1}}}}{\mbox{Right\ v\isactrlsub {\isadigit{2}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{2}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{1}}}}}\\[4mm]
+  \isa{\mbox{}\inferrule{\mbox{v\isactrlsub {\isadigit{1}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{1}}}\\\ \mbox{v\isactrlsub {\isadigit{2}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{2}}}}{\mbox{Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}}}}  &
+  \isa{\mbox{}\inferrule{\mbox{{\isasymforall}v{\isasymin}vs{\isachardot}{\kern0pt}\ v\ {\isacharcolon}{\kern0pt}\ r\ {\isasymand}\ {\isacharbar}{\kern0pt}v{\isacharbar}{\kern0pt}\ {\isasymnoteq}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}}}{\mbox{Stars\ vs\ {\isacharcolon}{\kern0pt}\ r\isactrlsup {\isasymstar}}}}
+
+
+  \end{tabular}
+\end{center}
+
 
+And these can be made compact by the following simplification function:
+
+\begin{center}
+  \begin{tabular}{lcl}
+  \isa{bsimp\ {\isacharparenleft}{\kern0pt}ASEQ\ bs\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{bsimp{\isacharunderscore}{\kern0pt}ASEQ\ bs\ {\isacharparenleft}{\kern0pt}bsimp\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bsimp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
+  \isa{bsimp\ {\isacharparenleft}{\kern0pt}AALTs\ bs{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ rs{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ {\isacharparenleft}{\kern0pt}distinctBy\ {\isacharparenleft}{\kern0pt}flts\ {\isacharparenleft}{\kern0pt}map\ bsimp\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ erase\ {\isasymemptyset}{\isacharparenright}{\kern0pt}}\\
+  \isa{bsimp\ AZERO} & $\dn$ & \isa{AZERO}\\
+  \isa{{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\mbox{$^\uparrow$}} & $\dn$ & \isa{AALT\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}fuse\ {\isacharbrackleft}{\kern0pt}Z{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\mbox{$^\uparrow$}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}fuse\ {\isacharbrackleft}{\kern0pt}S{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}\mbox{$^\uparrow$}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}}\\
+\end{tabular}
+\end{center}
+
+The core idea of the proof is that two regular expressions,
+if "isomorphic" up to a finite number of rewrite steps, will
+remain so when we take derivative on both of them.
+This can be expressed by the following rewrite relation lemma:
+\begin{lemma}
+\isa{{\isacharparenleft}{\kern0pt}r\mbox{$\bbslash$}s{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ bders{\isacharunderscore}{\kern0pt}simp\ r\ s}
+\end{lemma}%
+\end{isamarkuptext}\isamarkuptrue%
+%
+\isadelimdocument
+%
+\endisadelimdocument
+%
+\isatagdocument
+%
+\isamarkupsection{Introduction%
+}
+\isamarkuptrue%
+%
+\endisatagdocument
+{\isafolddocument}%
+%
+\isadelimdocument
+%
+\endisadelimdocument
+%
+\begin{isamarkuptext}%
 Brzozowski \cite{Brzozowski1964} introduced the notion of the {\em
 derivative} \isa{r{\isacharbackslash}{\kern0pt}c} of a regular expression \isa{r} w.r.t.\
 a character~\isa{c}, and showed that it gave a simple solution to the
@@ -1586,1311 +1667,1311 @@
 \end{isabellebody}%
 \endinput
 %:%file=~/Dropbox/Workspace/journalpaper/lexing/thys2/Journal/Paper.thy%:%
-%:%50=136%:%
-%:%62=141%:%
-%:%63=142%:%
-%:%64=143%:%
-%:%65=144%:%
-%:%66=145%:%
-%:%67=146%:%
-%:%68=147%:%
-%:%69=148%:%
-%:%70=149%:%
-%:%71=150%:%
-%:%72=151%:%
-%:%73=152%:%
-%:%74=153%:%
-%:%75=154%:%
-%:%76=155%:%
-%:%77=156%:%
-%:%78=157%:%
-%:%79=158%:%
-%:%80=159%:%
-%:%81=160%:%
-%:%81=161%:%
-%:%82=162%:%
-%:%83=163%:%
-%:%84=164%:%
-%:%85=165%:%
-%:%86=166%:%
-%:%87=167%:%
-%:%88=168%:%
-%:%89=169%:%
-%:%90=170%:%
-%:%91=171%:%
-%:%92=172%:%
-%:%93=173%:%
-%:%94=174%:%
-%:%95=175%:%
-%:%96=176%:%
-%:%97=177%:%
-%:%98=178%:%
-%:%99=179%:%
-%:%100=180%:%
-%:%101=181%:%
-%:%102=182%:%
-%:%103=183%:%
-%:%104=184%:%
-%:%105=185%:%
-%:%106=186%:%
-%:%107=187%:%
-%:%108=188%:%
-%:%109=189%:%
-%:%110=190%:%
-%:%111=191%:%
-%:%112=192%:%
-%:%113=193%:%
-%:%114=194%:%
-%:%115=195%:%
-%:%116=196%:%
-%:%117=197%:%
-%:%118=198%:%
-%:%119=199%:%
-%:%120=200%:%
-%:%121=201%:%
-%:%122=202%:%
-%:%123=203%:%
-%:%124=204%:%
-%:%125=205%:%
-%:%126=206%:%
-%:%127=207%:%
-%:%128=208%:%
-%:%129=209%:%
-%:%130=210%:%
-%:%131=211%:%
-%:%132=212%:%
-%:%133=213%:%
-%:%134=214%:%
-%:%135=215%:%
-%:%136=216%:%
-%:%137=217%:%
-%:%138=218%:%
-%:%139=219%:%
-%:%140=220%:%
-%:%141=221%:%
-%:%142=222%:%
-%:%143=223%:%
-%:%144=224%:%
-%:%145=225%:%
-%:%145=226%:%
-%:%146=227%:%
-%:%147=228%:%
-%:%148=229%:%
-%:%149=230%:%
-%:%150=231%:%
-%:%151=232%:%
-%:%152=233%:%
-%:%153=234%:%
-%:%154=235%:%
-%:%155=236%:%
-%:%156=237%:%
-%:%157=238%:%
-%:%158=239%:%
-%:%159=240%:%
-%:%160=241%:%
-%:%161=242%:%
-%:%162=243%:%
-%:%163=244%:%
-%:%164=245%:%
-%:%165=246%:%
-%:%166=247%:%
-%:%167=248%:%
-%:%168=249%:%
-%:%169=250%:%
-%:%170=251%:%
-%:%171=252%:%
-%:%172=253%:%
-%:%173=254%:%
-%:%174=255%:%
-%:%175=256%:%
-%:%176=257%:%
-%:%177=258%:%
-%:%178=259%:%
-%:%179=260%:%
-%:%180=261%:%
-%:%181=262%:%
-%:%182=263%:%
-%:%183=264%:%
-%:%184=265%:%
-%:%185=266%:%
-%:%186=267%:%
-%:%187=268%:%
-%:%188=269%:%
-%:%189=270%:%
-%:%190=271%:%
-%:%191=272%:%
-%:%192=273%:%
-%:%193=274%:%
-%:%194=275%:%
-%:%195=276%:%
-%:%196=277%:%
-%:%197=278%:%
-%:%198=279%:%
-%:%199=280%:%
-%:%200=281%:%
-%:%201=282%:%
-%:%202=283%:%
-%:%203=284%:%
-%:%204=285%:%
-%:%205=286%:%
-%:%206=287%:%
-%:%207=288%:%
-%:%208=289%:%
-%:%209=290%:%
-%:%210=291%:%
-%:%211=292%:%
-%:%212=293%:%
-%:%213=294%:%
-%:%214=295%:%
-%:%215=296%:%
-%:%216=297%:%
-%:%217=298%:%
-%:%218=299%:%
-%:%219=300%:%
-%:%220=301%:%
-%:%221=302%:%
-%:%222=303%:%
-%:%223=304%:%
-%:%224=305%:%
-%:%225=306%:%
-%:%226=307%:%
-%:%227=308%:%
-%:%228=309%:%
-%:%229=310%:%
-%:%230=311%:%
-%:%231=312%:%
-%:%232=313%:%
-%:%233=314%:%
-%:%234=315%:%
-%:%235=316%:%
-%:%236=317%:%
-%:%237=318%:%
-%:%238=319%:%
-%:%239=320%:%
-%:%240=321%:%
-%:%241=322%:%
-%:%242=323%:%
-%:%243=324%:%
-%:%244=325%:%
-%:%245=326%:%
-%:%246=327%:%
-%:%247=328%:%
-%:%248=329%:%
-%:%249=330%:%
-%:%250=331%:%
-%:%259=335%:%
+%:%50=134%:%
+%:%62=136%:%
+%:%63=137%:%
+%:%64=138%:%
+%:%65=139%:%
+%:%66=140%:%
+%:%67=141%:%
+%:%68=142%:%
+%:%69=143%:%
+%:%70=144%:%
+%:%71=145%:%
+%:%72=146%:%
+%:%73=147%:%
+%:%74=148%:%
+%:%75=149%:%
+%:%76=150%:%
+%:%77=151%:%
+%:%78=152%:%
+%:%79=153%:%
+%:%80=154%:%
+%:%81=155%:%
+%:%82=156%:%
+%:%83=157%:%
+%:%84=158%:%
+%:%85=159%:%
+%:%86=160%:%
+%:%87=161%:%
+%:%88=162%:%
+%:%89=163%:%
+%:%90=164%:%
+%:%91=165%:%
+%:%92=166%:%
+%:%93=167%:%
+%:%94=168%:%
+%:%95=169%:%
+%:%96=170%:%
+%:%97=171%:%
+%:%98=172%:%
+%:%99=173%:%
+%:%100=174%:%
+%:%101=175%:%
+%:%102=176%:%
+%:%103=177%:%
+%:%104=178%:%
+%:%105=179%:%
+%:%106=180%:%
+%:%107=181%:%
+%:%108=182%:%
+%:%109=183%:%
+%:%110=184%:%
+%:%111=185%:%
+%:%112=186%:%
+%:%113=187%:%
+%:%114=188%:%
+%:%115=189%:%
+%:%116=190%:%
+%:%117=191%:%
+%:%118=192%:%
+%:%119=193%:%
+%:%120=194%:%
+%:%121=195%:%
+%:%122=196%:%
+%:%123=197%:%
+%:%124=198%:%
+%:%125=199%:%
+%:%126=200%:%
+%:%127=201%:%
+%:%128=202%:%
+%:%129=203%:%
+%:%130=204%:%
+%:%131=205%:%
+%:%132=206%:%
+%:%133=207%:%
+%:%134=208%:%
+%:%135=209%:%
+%:%136=210%:%
+%:%137=211%:%
+%:%138=212%:%
+%:%147=217%:%
+%:%159=223%:%
+%:%160=224%:%
+%:%161=225%:%
+%:%162=226%:%
+%:%162=227%:%
+%:%163=228%:%
+%:%164=229%:%
+%:%165=230%:%
+%:%166=231%:%
+%:%167=232%:%
+%:%168=233%:%
+%:%169=234%:%
+%:%170=235%:%
+%:%171=236%:%
+%:%172=237%:%
+%:%173=238%:%
+%:%174=239%:%
+%:%175=240%:%
+%:%176=241%:%
+%:%177=242%:%
+%:%178=243%:%
+%:%179=244%:%
+%:%180=245%:%
+%:%181=246%:%
+%:%182=247%:%
+%:%183=248%:%
+%:%184=249%:%
+%:%185=250%:%
+%:%186=251%:%
+%:%187=252%:%
+%:%188=253%:%
+%:%189=254%:%
+%:%190=255%:%
+%:%191=256%:%
+%:%192=257%:%
+%:%193=258%:%
+%:%194=259%:%
+%:%195=260%:%
+%:%196=261%:%
+%:%197=262%:%
+%:%198=263%:%
+%:%199=264%:%
+%:%200=265%:%
+%:%201=266%:%
+%:%202=267%:%
+%:%203=268%:%
+%:%204=269%:%
+%:%205=270%:%
+%:%206=271%:%
+%:%207=272%:%
+%:%208=273%:%
+%:%209=274%:%
+%:%210=275%:%
+%:%211=276%:%
+%:%212=277%:%
+%:%213=278%:%
+%:%214=279%:%
+%:%215=280%:%
+%:%216=281%:%
+%:%217=282%:%
+%:%218=283%:%
+%:%219=284%:%
+%:%220=285%:%
+%:%221=286%:%
+%:%222=287%:%
+%:%223=288%:%
+%:%224=289%:%
+%:%225=290%:%
+%:%226=291%:%
+%:%226=292%:%
+%:%227=293%:%
+%:%228=294%:%
+%:%229=295%:%
+%:%230=296%:%
+%:%231=297%:%
+%:%232=298%:%
+%:%233=299%:%
+%:%234=300%:%
+%:%235=301%:%
+%:%236=302%:%
+%:%237=303%:%
+%:%238=304%:%
+%:%239=305%:%
+%:%240=306%:%
+%:%241=307%:%
+%:%242=308%:%
+%:%243=309%:%
+%:%244=310%:%
+%:%245=311%:%
+%:%246=312%:%
+%:%247=313%:%
+%:%248=314%:%
+%:%249=315%:%
+%:%250=316%:%
+%:%251=317%:%
+%:%252=318%:%
+%:%253=319%:%
+%:%254=320%:%
+%:%255=321%:%
+%:%256=322%:%
+%:%257=323%:%
+%:%258=324%:%
+%:%259=325%:%
+%:%260=326%:%
+%:%261=327%:%
+%:%262=328%:%
+%:%263=329%:%
+%:%264=330%:%
+%:%265=331%:%
+%:%266=332%:%
+%:%267=333%:%
+%:%268=334%:%
+%:%269=335%:%
+%:%270=336%:%
 %:%271=337%:%
 %:%272=338%:%
-%:%272=339%:%
-%:%273=340%:%
-%:%274=341%:%
-%:%275=342%:%
-%:%276=343%:%
-%:%277=344%:%
-%:%278=345%:%
-%:%279=346%:%
-%:%280=347%:%
-%:%281=348%:%
-%:%282=349%:%
-%:%283=350%:%
-%:%284=351%:%
-%:%285=352%:%
-%:%286=353%:%
-%:%287=354%:%
-%:%288=355%:%
-%:%289=356%:%
-%:%290=357%:%
-%:%291=358%:%
-%:%292=359%:%
-%:%293=360%:%
-%:%294=361%:%
-%:%295=362%:%
-%:%296=363%:%
-%:%297=364%:%
-%:%298=365%:%
-%:%299=366%:%
-%:%300=367%:%
-%:%301=368%:%
-%:%302=369%:%
-%:%303=370%:%
-%:%304=371%:%
-%:%305=372%:%
-%:%306=373%:%
-%:%307=374%:%
-%:%308=375%:%
-%:%309=376%:%
-%:%310=377%:%
-%:%311=378%:%
-%:%311=379%:%
-%:%312=380%:%
-%:%313=381%:%
-%:%314=382%:%
-%:%315=383%:%
-%:%316=384%:%
-%:%316=385%:%
-%:%317=386%:%
-%:%318=387%:%
-%:%319=388%:%
-%:%320=389%:%
-%:%321=390%:%
-%:%322=391%:%
-%:%323=392%:%
-%:%324=393%:%
-%:%325=394%:%
-%:%326=395%:%
-%:%327=396%:%
-%:%328=397%:%
-%:%329=398%:%
-%:%330=399%:%
-%:%331=400%:%
-%:%332=401%:%
-%:%333=402%:%
-%:%334=403%:%
-%:%335=404%:%
-%:%336=405%:%
-%:%337=406%:%
-%:%338=407%:%
-%:%339=408%:%
-%:%340=409%:%
-%:%341=410%:%
-%:%342=411%:%
-%:%343=412%:%
-%:%344=413%:%
-%:%345=414%:%
-%:%346=415%:%
-%:%347=416%:%
-%:%348=417%:%
-%:%349=418%:%
-%:%350=419%:%
-%:%351=420%:%
-%:%352=421%:%
-%:%353=422%:%
-%:%354=423%:%
-%:%355=424%:%
-%:%356=425%:%
-%:%357=426%:%
-%:%358=427%:%
-%:%359=428%:%
-%:%360=429%:%
-%:%361=430%:%
-%:%362=431%:%
-%:%363=432%:%
-%:%364=433%:%
-%:%365=434%:%
-%:%366=435%:%
-%:%367=436%:%
-%:%368=437%:%
-%:%369=438%:%
-%:%370=439%:%
-%:%371=440%:%
-%:%372=441%:%
-%:%373=442%:%
-%:%374=443%:%
-%:%375=444%:%
-%:%376=445%:%
-%:%377=446%:%
-%:%378=447%:%
-%:%379=448%:%
-%:%380=449%:%
-%:%381=450%:%
-%:%382=451%:%
-%:%383=452%:%
-%:%384=453%:%
-%:%385=454%:%
-%:%386=455%:%
-%:%387=456%:%
-%:%388=457%:%
-%:%389=458%:%
-%:%390=459%:%
-%:%391=460%:%
-%:%392=461%:%
-%:%393=462%:%
-%:%394=463%:%
-%:%395=464%:%
-%:%396=465%:%
-%:%397=466%:%
-%:%398=467%:%
-%:%399=468%:%
-%:%400=469%:%
-%:%401=470%:%
-%:%402=471%:%
-%:%403=472%:%
-%:%404=473%:%
-%:%405=474%:%
-%:%414=478%:%
-%:%426=482%:%
-%:%427=483%:%
-%:%428=484%:%
-%:%429=485%:%
-%:%430=486%:%
-%:%431=487%:%
-%:%432=488%:%
-%:%433=489%:%
-%:%434=490%:%
-%:%435=491%:%
-%:%436=492%:%
-%:%437=493%:%
-%:%438=494%:%
-%:%439=495%:%
-%:%440=496%:%
-%:%441=497%:%
-%:%442=498%:%
-%:%443=499%:%
-%:%444=500%:%
-%:%445=501%:%
-%:%446=502%:%
-%:%447=503%:%
-%:%448=504%:%
-%:%449=505%:%
-%:%450=506%:%
-%:%451=507%:%
-%:%452=508%:%
-%:%453=509%:%
-%:%454=510%:%
-%:%455=511%:%
-%:%456=512%:%
-%:%457=513%:%
-%:%458=514%:%
-%:%459=515%:%
-%:%460=516%:%
-%:%461=517%:%
-%:%462=518%:%
-%:%463=519%:%
-%:%464=520%:%
-%:%465=521%:%
-%:%466=522%:%
-%:%467=523%:%
-%:%468=524%:%
-%:%469=525%:%
-%:%470=526%:%
-%:%471=527%:%
-%:%472=528%:%
-%:%473=529%:%
-%:%474=530%:%
-%:%475=531%:%
-%:%476=532%:%
-%:%477=533%:%
-%:%478=534%:%
-%:%479=535%:%
-%:%480=536%:%
-%:%481=537%:%
-%:%482=538%:%
-%:%483=539%:%
-%:%484=540%:%
-%:%485=541%:%
-%:%486=542%:%
-%:%487=543%:%
-%:%488=544%:%
-%:%489=545%:%
-%:%490=546%:%
-%:%491=547%:%
-%:%492=548%:%
-%:%493=549%:%
-%:%494=550%:%
-%:%495=551%:%
-%:%496=552%:%
-%:%497=553%:%
-%:%498=554%:%
-%:%499=555%:%
-%:%500=556%:%
-%:%501=557%:%
-%:%502=558%:%
-%:%503=559%:%
-%:%504=560%:%
-%:%505=561%:%
-%:%506=562%:%
-%:%507=563%:%
-%:%508=564%:%
-%:%509=565%:%
-%:%510=566%:%
-%:%511=567%:%
-%:%512=568%:%
-%:%513=569%:%
-%:%514=570%:%
-%:%515=571%:%
-%:%516=572%:%
-%:%517=573%:%
-%:%518=574%:%
-%:%519=575%:%
-%:%520=576%:%
-%:%521=577%:%
-%:%522=578%:%
-%:%523=579%:%
-%:%524=580%:%
-%:%525=581%:%
-%:%526=582%:%
-%:%527=583%:%
-%:%528=584%:%
-%:%529=585%:%
-%:%530=586%:%
-%:%531=587%:%
-%:%532=588%:%
-%:%533=589%:%
-%:%534=590%:%
-%:%535=591%:%
-%:%536=592%:%
-%:%537=593%:%
-%:%538=594%:%
-%:%539=595%:%
-%:%540=596%:%
-%:%540=597%:%
-%:%541=598%:%
-%:%541=599%:%
-%:%542=600%:%
-%:%543=601%:%
-%:%543=602%:%
-%:%544=603%:%
-%:%545=604%:%
-%:%546=605%:%
-%:%547=606%:%
-%:%548=607%:%
-%:%549=608%:%
-%:%550=609%:%
-%:%551=610%:%
-%:%552=611%:%
-%:%553=612%:%
-%:%554=613%:%
-%:%555=614%:%
-%:%556=615%:%
-%:%557=616%:%
-%:%558=617%:%
-%:%559=618%:%
-%:%560=619%:%
-%:%561=620%:%
-%:%562=621%:%
-%:%563=622%:%
-%:%564=623%:%
-%:%565=624%:%
-%:%566=625%:%
-%:%567=626%:%
-%:%568=627%:%
-%:%569=628%:%
-%:%570=629%:%
-%:%571=630%:%
-%:%572=631%:%
-%:%573=632%:%
-%:%574=633%:%
-%:%575=634%:%
-%:%576=635%:%
-%:%577=636%:%
-%:%578=637%:%
-%:%579=638%:%
-%:%580=639%:%
-%:%581=640%:%
-%:%582=641%:%
-%:%583=642%:%
-%:%584=643%:%
-%:%585=644%:%
-%:%586=645%:%
-%:%587=646%:%
-%:%588=647%:%
-%:%589=648%:%
-%:%590=649%:%
-%:%591=650%:%
-%:%592=651%:%
-%:%593=652%:%
-%:%594=653%:%
-%:%595=654%:%
-%:%596=655%:%
-%:%597=656%:%
-%:%598=657%:%
-%:%599=658%:%
-%:%599=659%:%
-%:%599=660%:%
-%:%600=661%:%
-%:%601=662%:%
-%:%602=663%:%
-%:%603=664%:%
-%:%604=665%:%
-%:%605=666%:%
-%:%606=667%:%
-%:%607=668%:%
-%:%608=669%:%
-%:%608=670%:%
-%:%609=671%:%
-%:%610=672%:%
-%:%611=673%:%
-%:%612=674%:%
-%:%613=675%:%
-%:%614=676%:%
-%:%615=677%:%
-%:%616=678%:%
-%:%617=679%:%
-%:%618=680%:%
-%:%619=681%:%
-%:%620=682%:%
-%:%621=683%:%
-%:%622=684%:%
-%:%623=685%:%
-%:%624=686%:%
-%:%625=687%:%
-%:%626=688%:%
-%:%627=689%:%
-%:%628=690%:%
-%:%629=691%:%
-%:%630=692%:%
-%:%631=693%:%
-%:%632=694%:%
-%:%633=695%:%
-%:%634=696%:%
-%:%635=697%:%
-%:%636=698%:%
-%:%637=699%:%
-%:%638=700%:%
-%:%639=701%:%
-%:%640=702%:%
-%:%641=703%:%
-%:%642=704%:%
-%:%643=705%:%
-%:%643=706%:%
-%:%644=707%:%
-%:%645=708%:%
-%:%646=709%:%
-%:%647=710%:%
-%:%648=711%:%
-%:%649=712%:%
-%:%650=713%:%
-%:%650=714%:%
-%:%651=715%:%
-%:%652=716%:%
-%:%653=717%:%
-%:%653=718%:%
-%:%654=719%:%
-%:%655=720%:%
-%:%656=721%:%
-%:%657=722%:%
-%:%658=723%:%
-%:%659=724%:%
-%:%660=725%:%
-%:%661=726%:%
-%:%662=727%:%
-%:%663=728%:%
-%:%664=729%:%
-%:%664=730%:%
-%:%665=731%:%
-%:%666=732%:%
-%:%667=733%:%
-%:%668=734%:%
-%:%668=735%:%
-%:%669=736%:%
-%:%670=737%:%
-%:%671=738%:%
-%:%672=739%:%
-%:%673=740%:%
-%:%674=741%:%
-%:%675=742%:%
-%:%676=743%:%
-%:%677=744%:%
-%:%678=745%:%
-%:%679=746%:%
-%:%680=747%:%
-%:%681=748%:%
-%:%682=749%:%
-%:%683=750%:%
-%:%684=751%:%
-%:%685=752%:%
-%:%686=753%:%
-%:%687=754%:%
-%:%688=755%:%
-%:%689=756%:%
-%:%690=757%:%
-%:%691=758%:%
-%:%692=759%:%
-%:%693=760%:%
-%:%694=761%:%
-%:%695=762%:%
-%:%696=763%:%
-%:%697=764%:%
-%:%698=765%:%
-%:%699=766%:%
-%:%700=767%:%
-%:%701=768%:%
-%:%702=769%:%
-%:%703=770%:%
-%:%704=771%:%
-%:%705=772%:%
-%:%706=773%:%
-%:%707=774%:%
-%:%708=775%:%
-%:%709=776%:%
-%:%710=777%:%
-%:%711=778%:%
-%:%712=779%:%
-%:%713=780%:%
-%:%714=781%:%
-%:%715=782%:%
-%:%716=783%:%
-%:%717=784%:%
-%:%718=785%:%
-%:%719=786%:%
-%:%720=787%:%
-%:%721=788%:%
-%:%722=789%:%
-%:%723=790%:%
-%:%724=791%:%
-%:%725=792%:%
-%:%726=793%:%
-%:%726=794%:%
-%:%727=795%:%
-%:%728=796%:%
-%:%729=797%:%
-%:%730=798%:%
-%:%731=799%:%
-%:%732=800%:%
-%:%733=801%:%
-%:%734=802%:%
-%:%735=803%:%
-%:%736=804%:%
-%:%737=805%:%
-%:%738=806%:%
-%:%739=807%:%
-%:%740=808%:%
-%:%741=809%:%
-%:%742=810%:%
-%:%743=811%:%
-%:%744=812%:%
-%:%745=813%:%
-%:%746=814%:%
-%:%747=815%:%
-%:%748=816%:%
-%:%749=817%:%
-%:%750=818%:%
-%:%751=819%:%
-%:%752=820%:%
-%:%753=821%:%
-%:%754=822%:%
-%:%755=823%:%
-%:%756=824%:%
-%:%757=825%:%
-%:%758=826%:%
-%:%759=827%:%
-%:%759=828%:%
-%:%760=829%:%
-%:%761=830%:%
-%:%762=831%:%
-%:%763=832%:%
-%:%764=833%:%
-%:%765=834%:%
-%:%766=835%:%
-%:%767=836%:%
-%:%768=837%:%
-%:%769=838%:%
-%:%770=839%:%
-%:%771=840%:%
-%:%772=841%:%
-%:%773=842%:%
-%:%774=843%:%
-%:%775=844%:%
-%:%776=845%:%
-%:%777=846%:%
-%:%778=847%:%
-%:%779=848%:%
-%:%780=849%:%
-%:%781=850%:%
-%:%782=851%:%
-%:%783=852%:%
-%:%784=853%:%
-%:%785=854%:%
-%:%786=855%:%
-%:%787=856%:%
-%:%788=857%:%
-%:%789=858%:%
-%:%789=859%:%
-%:%790=860%:%
-%:%791=861%:%
-%:%792=862%:%
-%:%793=863%:%
-%:%794=864%:%
-%:%795=865%:%
-%:%796=866%:%
-%:%797=867%:%
-%:%797=868%:%
-%:%798=869%:%
-%:%799=870%:%
-%:%800=871%:%
-%:%801=872%:%
-%:%802=873%:%
-%:%803=874%:%
-%:%804=875%:%
-%:%805=876%:%
-%:%806=877%:%
-%:%807=878%:%
-%:%808=879%:%
-%:%809=880%:%
-%:%810=881%:%
-%:%811=882%:%
-%:%812=883%:%
-%:%813=884%:%
-%:%814=885%:%
-%:%815=886%:%
-%:%816=887%:%
-%:%817=888%:%
-%:%818=889%:%
-%:%819=890%:%
-%:%820=891%:%
-%:%821=892%:%
-%:%822=893%:%
-%:%823=894%:%
-%:%824=895%:%
-%:%825=896%:%
-%:%826=897%:%
-%:%827=898%:%
-%:%828=899%:%
-%:%829=900%:%
-%:%830=901%:%
-%:%831=902%:%
-%:%832=903%:%
-%:%833=904%:%
-%:%834=905%:%
-%:%834=906%:%
-%:%834=907%:%
-%:%835=908%:%
-%:%835=909%:%
-%:%835=910%:%
-%:%836=911%:%
-%:%837=912%:%
-%:%837=913%:%
-%:%838=914%:%
-%:%839=915%:%
-%:%840=916%:%
-%:%841=917%:%
-%:%842=918%:%
-%:%843=919%:%
-%:%844=920%:%
-%:%845=921%:%
-%:%846=922%:%
-%:%847=923%:%
-%:%848=924%:%
-%:%849=925%:%
-%:%850=926%:%
-%:%851=927%:%
-%:%852=928%:%
-%:%853=929%:%
-%:%854=930%:%
-%:%855=931%:%
-%:%856=932%:%
-%:%857=933%:%
-%:%858=934%:%
-%:%859=935%:%
-%:%860=936%:%
-%:%861=937%:%
-%:%862=938%:%
-%:%863=939%:%
-%:%864=940%:%
-%:%865=941%:%
-%:%866=942%:%
-%:%867=943%:%
-%:%868=944%:%
-%:%869=945%:%
-%:%870=946%:%
-%:%871=947%:%
-%:%871=948%:%
-%:%872=949%:%
-%:%873=950%:%
-%:%874=951%:%
-%:%874=952%:%
-%:%874=953%:%
-%:%874=954%:%
-%:%875=955%:%
-%:%876=956%:%
-%:%877=957%:%
-%:%878=958%:%
-%:%879=959%:%
-%:%880=960%:%
-%:%881=961%:%
-%:%882=962%:%
-%:%883=963%:%
-%:%884=964%:%
-%:%885=965%:%
-%:%886=966%:%
-%:%887=967%:%
-%:%888=968%:%
-%:%889=969%:%
-%:%890=970%:%
-%:%891=971%:%
-%:%892=972%:%
-%:%893=973%:%
-%:%894=974%:%
-%:%895=975%:%
-%:%896=976%:%
-%:%897=977%:%
-%:%898=978%:%
-%:%899=979%:%
-%:%900=980%:%
-%:%901=981%:%
-%:%902=982%:%
-%:%903=983%:%
-%:%904=984%:%
-%:%905=985%:%
-%:%906=986%:%
-%:%907=987%:%
-%:%908=988%:%
-%:%909=989%:%
-%:%910=990%:%
-%:%911=991%:%
-%:%920=995%:%
-%:%932=999%:%
-%:%933=1000%:%
-%:%934=1001%:%
-%:%935=1002%:%
-%:%936=1003%:%
-%:%937=1004%:%
-%:%938=1005%:%
-%:%939=1006%:%
-%:%940=1007%:%
-%:%941=1008%:%
-%:%942=1009%:%
-%:%943=1010%:%
-%:%944=1011%:%
-%:%945=1012%:%
-%:%946=1013%:%
-%:%947=1014%:%
-%:%948=1015%:%
-%:%949=1016%:%
-%:%950=1017%:%
-%:%951=1018%:%
-%:%952=1019%:%
-%:%953=1020%:%
-%:%954=1021%:%
+%:%273=339%:%
+%:%274=340%:%
+%:%275=341%:%
+%:%276=342%:%
+%:%277=343%:%
+%:%278=344%:%
+%:%279=345%:%
+%:%280=346%:%
+%:%281=347%:%
+%:%282=348%:%
+%:%283=349%:%
+%:%284=350%:%
+%:%285=351%:%
+%:%286=352%:%
+%:%287=353%:%
+%:%288=354%:%
+%:%289=355%:%
+%:%290=356%:%
+%:%291=357%:%
+%:%292=358%:%
+%:%293=359%:%
+%:%294=360%:%
+%:%295=361%:%
+%:%296=362%:%
+%:%297=363%:%
+%:%298=364%:%
+%:%299=365%:%
+%:%300=366%:%
+%:%301=367%:%
+%:%302=368%:%
+%:%303=369%:%
+%:%304=370%:%
+%:%305=371%:%
+%:%306=372%:%
+%:%307=373%:%
+%:%308=374%:%
+%:%309=375%:%
+%:%310=376%:%
+%:%311=377%:%
+%:%312=378%:%
+%:%313=379%:%
+%:%314=380%:%
+%:%315=381%:%
+%:%316=382%:%
+%:%317=383%:%
+%:%318=384%:%
+%:%319=385%:%
+%:%320=386%:%
+%:%321=387%:%
+%:%322=388%:%
+%:%323=389%:%
+%:%324=390%:%
+%:%325=391%:%
+%:%326=392%:%
+%:%327=393%:%
+%:%328=394%:%
+%:%329=395%:%
+%:%330=396%:%
+%:%331=397%:%
+%:%340=404%:%
+%:%352=406%:%
+%:%353=407%:%
+%:%353=408%:%
+%:%354=409%:%
+%:%355=410%:%
+%:%356=411%:%
+%:%357=412%:%
+%:%358=413%:%
+%:%359=414%:%
+%:%360=415%:%
+%:%361=416%:%
+%:%362=417%:%
+%:%363=418%:%
+%:%364=419%:%
+%:%365=420%:%
+%:%366=421%:%
+%:%367=422%:%
+%:%368=423%:%
+%:%369=424%:%
+%:%370=425%:%
+%:%371=426%:%
+%:%372=427%:%
+%:%373=428%:%
+%:%374=429%:%
+%:%375=430%:%
+%:%376=431%:%
+%:%377=432%:%
+%:%378=433%:%
+%:%379=434%:%
+%:%380=435%:%
+%:%381=436%:%
+%:%382=437%:%
+%:%383=438%:%
+%:%384=439%:%
+%:%385=440%:%
+%:%386=441%:%
+%:%387=442%:%
+%:%388=443%:%
+%:%389=444%:%
+%:%390=445%:%
+%:%391=446%:%
+%:%392=447%:%
+%:%392=448%:%
+%:%393=449%:%
+%:%394=450%:%
+%:%395=451%:%
+%:%396=452%:%
+%:%397=453%:%
+%:%397=454%:%
+%:%398=455%:%
+%:%399=456%:%
+%:%400=457%:%
+%:%401=458%:%
+%:%402=459%:%
+%:%403=460%:%
+%:%404=461%:%
+%:%405=462%:%
+%:%406=463%:%
+%:%407=464%:%
+%:%408=465%:%
+%:%409=466%:%
+%:%410=467%:%
+%:%411=468%:%
+%:%412=469%:%
+%:%413=470%:%
+%:%414=471%:%
+%:%415=472%:%
+%:%416=473%:%
+%:%417=474%:%
+%:%418=475%:%
+%:%419=476%:%
+%:%420=477%:%
+%:%421=478%:%
+%:%422=479%:%
+%:%423=480%:%
+%:%424=481%:%
+%:%425=482%:%
+%:%426=483%:%
+%:%427=484%:%
+%:%428=485%:%
+%:%429=486%:%
+%:%430=487%:%
+%:%431=488%:%
+%:%432=489%:%
+%:%433=490%:%
+%:%434=491%:%
+%:%435=492%:%
+%:%436=493%:%
+%:%437=494%:%
+%:%438=495%:%
+%:%439=496%:%
+%:%440=497%:%
+%:%441=498%:%
+%:%442=499%:%
+%:%443=500%:%
+%:%444=501%:%
+%:%445=502%:%
+%:%446=503%:%
+%:%447=504%:%
+%:%448=505%:%
+%:%449=506%:%
+%:%450=507%:%
+%:%451=508%:%
+%:%452=509%:%
+%:%453=510%:%
+%:%454=511%:%
+%:%455=512%:%
+%:%456=513%:%
+%:%457=514%:%
+%:%458=515%:%
+%:%459=516%:%
+%:%460=517%:%
+%:%461=518%:%
+%:%462=519%:%
+%:%463=520%:%
+%:%464=521%:%
+%:%465=522%:%
+%:%466=523%:%
+%:%467=524%:%
+%:%468=525%:%
+%:%469=526%:%
+%:%470=527%:%
+%:%471=528%:%
+%:%472=529%:%
+%:%473=530%:%
+%:%474=531%:%
+%:%475=532%:%
+%:%476=533%:%
+%:%477=534%:%
+%:%478=535%:%
+%:%479=536%:%
+%:%480=537%:%
+%:%481=538%:%
+%:%482=539%:%
+%:%483=540%:%
+%:%484=541%:%
+%:%485=542%:%
+%:%486=543%:%
+%:%495=547%:%
+%:%507=551%:%
+%:%508=552%:%
+%:%509=553%:%
+%:%510=554%:%
+%:%511=555%:%
+%:%512=556%:%
+%:%513=557%:%
+%:%514=558%:%
+%:%515=559%:%
+%:%516=560%:%
+%:%517=561%:%
+%:%518=562%:%
+%:%519=563%:%
+%:%520=564%:%
+%:%521=565%:%
+%:%522=566%:%
+%:%523=567%:%
+%:%524=568%:%
+%:%525=569%:%
+%:%526=570%:%
+%:%527=571%:%
+%:%528=572%:%
+%:%529=573%:%
+%:%530=574%:%
+%:%531=575%:%
+%:%532=576%:%
+%:%533=577%:%
+%:%534=578%:%
+%:%535=579%:%
+%:%536=580%:%
+%:%537=581%:%
+%:%538=582%:%
+%:%539=583%:%
+%:%540=584%:%
+%:%541=585%:%
+%:%542=586%:%
+%:%543=587%:%
+%:%544=588%:%
+%:%545=589%:%
+%:%546=590%:%
+%:%547=591%:%
+%:%548=592%:%
+%:%549=593%:%
+%:%550=594%:%
+%:%551=595%:%
+%:%552=596%:%
+%:%553=597%:%
+%:%554=598%:%
+%:%555=599%:%
+%:%556=600%:%
+%:%557=601%:%
+%:%558=602%:%
+%:%559=603%:%
+%:%560=604%:%
+%:%561=605%:%
+%:%562=606%:%
+%:%563=607%:%
+%:%564=608%:%
+%:%565=609%:%
+%:%566=610%:%
+%:%567=611%:%
+%:%568=612%:%
+%:%569=613%:%
+%:%570=614%:%
+%:%571=615%:%
+%:%572=616%:%
+%:%573=617%:%
+%:%574=618%:%
+%:%575=619%:%
+%:%576=620%:%
+%:%577=621%:%
+%:%578=622%:%
+%:%579=623%:%
+%:%580=624%:%
+%:%581=625%:%
+%:%582=626%:%
+%:%583=627%:%
+%:%584=628%:%
+%:%585=629%:%
+%:%586=630%:%
+%:%587=631%:%
+%:%588=632%:%
+%:%589=633%:%
+%:%590=634%:%
+%:%591=635%:%
+%:%592=636%:%
+%:%593=637%:%
+%:%594=638%:%
+%:%595=639%:%
+%:%596=640%:%
+%:%597=641%:%
+%:%598=642%:%
+%:%599=643%:%
+%:%600=644%:%
+%:%601=645%:%
+%:%602=646%:%
+%:%603=647%:%
+%:%604=648%:%
+%:%605=649%:%
+%:%606=650%:%
+%:%607=651%:%
+%:%608=652%:%
+%:%609=653%:%
+%:%610=654%:%
+%:%611=655%:%
+%:%612=656%:%
+%:%613=657%:%
+%:%614=658%:%
+%:%615=659%:%
+%:%616=660%:%
+%:%617=661%:%
+%:%618=662%:%
+%:%619=663%:%
+%:%620=664%:%
+%:%621=665%:%
+%:%621=666%:%
+%:%622=667%:%
+%:%622=668%:%
+%:%623=669%:%
+%:%624=670%:%
+%:%624=671%:%
+%:%625=672%:%
+%:%626=673%:%
+%:%627=674%:%
+%:%628=675%:%
+%:%629=676%:%
+%:%630=677%:%
+%:%631=678%:%
+%:%632=679%:%
+%:%633=680%:%
+%:%634=681%:%
+%:%635=682%:%
+%:%636=683%:%
+%:%637=684%:%
+%:%638=685%:%
+%:%639=686%:%
+%:%640=687%:%
+%:%641=688%:%
+%:%642=689%:%
+%:%643=690%:%
+%:%644=691%:%
+%:%645=692%:%
+%:%646=693%:%
+%:%647=694%:%
+%:%648=695%:%
+%:%649=696%:%
+%:%650=697%:%
+%:%651=698%:%
+%:%652=699%:%
+%:%653=700%:%
+%:%654=701%:%
+%:%655=702%:%
+%:%656=703%:%
+%:%657=704%:%
+%:%658=705%:%
+%:%659=706%:%
+%:%660=707%:%
+%:%661=708%:%
+%:%662=709%:%
+%:%663=710%:%
+%:%664=711%:%
+%:%665=712%:%
+%:%666=713%:%
+%:%667=714%:%
+%:%668=715%:%
+%:%669=716%:%
+%:%670=717%:%
+%:%671=718%:%
+%:%672=719%:%
+%:%673=720%:%
+%:%674=721%:%
+%:%675=722%:%
+%:%676=723%:%
+%:%677=724%:%
+%:%678=725%:%
+%:%679=726%:%
+%:%680=727%:%
+%:%680=728%:%
+%:%680=729%:%
+%:%681=730%:%
+%:%682=731%:%
+%:%683=732%:%
+%:%684=733%:%
+%:%685=734%:%
+%:%686=735%:%
+%:%687=736%:%
+%:%688=737%:%
+%:%689=738%:%
+%:%689=739%:%
+%:%690=740%:%
+%:%691=741%:%
+%:%692=742%:%
+%:%693=743%:%
+%:%694=744%:%
+%:%695=745%:%
+%:%696=746%:%
+%:%697=747%:%
+%:%698=748%:%
+%:%699=749%:%
+%:%700=750%:%
+%:%701=751%:%
+%:%702=752%:%
+%:%703=753%:%
+%:%704=754%:%
+%:%705=755%:%
+%:%706=756%:%
+%:%707=757%:%
+%:%708=758%:%
+%:%709=759%:%
+%:%710=760%:%
+%:%711=761%:%
+%:%712=762%:%
+%:%713=763%:%
+%:%714=764%:%
+%:%715=765%:%
+%:%716=766%:%
+%:%717=767%:%
+%:%718=768%:%
+%:%719=769%:%
+%:%720=770%:%
+%:%721=771%:%
+%:%722=772%:%
+%:%723=773%:%
+%:%724=774%:%
+%:%724=775%:%
+%:%725=776%:%
+%:%726=777%:%
+%:%727=778%:%
+%:%728=779%:%
+%:%729=780%:%
+%:%730=781%:%
+%:%731=782%:%
+%:%731=783%:%
+%:%732=784%:%
+%:%733=785%:%
+%:%734=786%:%
+%:%734=787%:%
+%:%735=788%:%
+%:%736=789%:%
+%:%737=790%:%
+%:%738=791%:%
+%:%739=792%:%
+%:%740=793%:%
+%:%741=794%:%
+%:%742=795%:%
+%:%743=796%:%
+%:%744=797%:%
+%:%745=798%:%
+%:%745=799%:%
+%:%746=800%:%
+%:%747=801%:%
+%:%748=802%:%
+%:%749=803%:%
+%:%749=804%:%
+%:%750=805%:%
+%:%751=806%:%
+%:%752=807%:%
+%:%753=808%:%
+%:%754=809%:%
+%:%755=810%:%
+%:%756=811%:%
+%:%757=812%:%
+%:%758=813%:%
+%:%759=814%:%
+%:%760=815%:%
+%:%761=816%:%
+%:%762=817%:%
+%:%763=818%:%
+%:%764=819%:%
+%:%765=820%:%
+%:%766=821%:%
+%:%767=822%:%
+%:%768=823%:%
+%:%769=824%:%
+%:%770=825%:%
+%:%771=826%:%
+%:%772=827%:%
+%:%773=828%:%
+%:%774=829%:%
+%:%775=830%:%
+%:%776=831%:%
+%:%777=832%:%
+%:%778=833%:%
+%:%779=834%:%
+%:%780=835%:%
+%:%781=836%:%
+%:%782=837%:%
+%:%783=838%:%
+%:%784=839%:%
+%:%785=840%:%
+%:%786=841%:%
+%:%787=842%:%
+%:%788=843%:%
+%:%789=844%:%
+%:%790=845%:%
+%:%791=846%:%
+%:%792=847%:%
+%:%793=848%:%
+%:%794=849%:%
+%:%795=850%:%
+%:%796=851%:%
+%:%797=852%:%
+%:%798=853%:%
+%:%799=854%:%
+%:%800=855%:%
+%:%801=856%:%
+%:%802=857%:%
+%:%803=858%:%
+%:%804=859%:%
+%:%805=860%:%
+%:%806=861%:%
+%:%807=862%:%
+%:%807=863%:%
+%:%808=864%:%
+%:%809=865%:%
+%:%810=866%:%
+%:%811=867%:%
+%:%812=868%:%
+%:%813=869%:%
+%:%814=870%:%
+%:%815=871%:%
+%:%816=872%:%
+%:%817=873%:%
+%:%818=874%:%
+%:%819=875%:%
+%:%820=876%:%
+%:%821=877%:%
+%:%822=878%:%
+%:%823=879%:%
+%:%824=880%:%
+%:%825=881%:%
+%:%826=882%:%
+%:%827=883%:%
+%:%828=884%:%
+%:%829=885%:%
+%:%830=886%:%
+%:%831=887%:%
+%:%832=888%:%
+%:%833=889%:%
+%:%834=890%:%
+%:%835=891%:%
+%:%836=892%:%
+%:%837=893%:%
+%:%838=894%:%
+%:%839=895%:%
+%:%840=896%:%
+%:%840=897%:%
+%:%841=898%:%
+%:%842=899%:%
+%:%843=900%:%
+%:%844=901%:%
+%:%845=902%:%
+%:%846=903%:%
+%:%847=904%:%
+%:%848=905%:%
+%:%849=906%:%
+%:%850=907%:%
+%:%851=908%:%
+%:%852=909%:%
+%:%853=910%:%
+%:%854=911%:%
+%:%855=912%:%
+%:%856=913%:%
+%:%857=914%:%
+%:%858=915%:%
+%:%859=916%:%
+%:%860=917%:%
+%:%861=918%:%
+%:%862=919%:%
+%:%863=920%:%
+%:%864=921%:%
+%:%865=922%:%
+%:%866=923%:%
+%:%867=924%:%
+%:%868=925%:%
+%:%869=926%:%
+%:%870=927%:%
+%:%870=928%:%
+%:%871=929%:%
+%:%872=930%:%
+%:%873=931%:%
+%:%874=932%:%
+%:%875=933%:%
+%:%876=934%:%
+%:%877=935%:%
+%:%878=936%:%
+%:%878=937%:%
+%:%879=938%:%
+%:%880=939%:%
+%:%881=940%:%
+%:%882=941%:%
+%:%883=942%:%
+%:%884=943%:%
+%:%885=944%:%
+%:%886=945%:%
+%:%887=946%:%
+%:%888=947%:%
+%:%889=948%:%
+%:%890=949%:%
+%:%891=950%:%
+%:%892=951%:%
+%:%893=952%:%
+%:%894=953%:%
+%:%895=954%:%
+%:%896=955%:%
+%:%897=956%:%
+%:%898=957%:%
+%:%899=958%:%
+%:%900=959%:%
+%:%901=960%:%
+%:%902=961%:%
+%:%903=962%:%
+%:%904=963%:%
+%:%905=964%:%
+%:%906=965%:%
+%:%907=966%:%
+%:%908=967%:%
+%:%909=968%:%
+%:%910=969%:%
+%:%911=970%:%
+%:%912=971%:%
+%:%913=972%:%
+%:%914=973%:%
+%:%915=974%:%
+%:%915=975%:%
+%:%915=976%:%
+%:%916=977%:%
+%:%916=978%:%
+%:%916=979%:%
+%:%917=980%:%
+%:%918=981%:%
+%:%918=982%:%
+%:%919=983%:%
+%:%920=984%:%
+%:%921=985%:%
+%:%922=986%:%
+%:%923=987%:%
+%:%924=988%:%
+%:%925=989%:%
+%:%926=990%:%
+%:%927=991%:%
+%:%928=992%:%
+%:%929=993%:%
+%:%930=994%:%
+%:%931=995%:%
+%:%932=996%:%
+%:%933=997%:%
+%:%934=998%:%
+%:%935=999%:%
+%:%936=1000%:%
+%:%937=1001%:%
+%:%938=1002%:%
+%:%939=1003%:%
+%:%940=1004%:%
+%:%941=1005%:%
+%:%942=1006%:%
+%:%943=1007%:%
+%:%944=1008%:%
+%:%945=1009%:%
+%:%946=1010%:%
+%:%947=1011%:%
+%:%948=1012%:%
+%:%949=1013%:%
+%:%950=1014%:%
+%:%951=1015%:%
+%:%952=1016%:%
+%:%952=1017%:%
+%:%953=1018%:%
+%:%954=1019%:%
+%:%955=1020%:%
+%:%955=1021%:%
 %:%955=1022%:%
-%:%956=1023%:%
-%:%957=1024%:%
-%:%958=1025%:%
-%:%959=1026%:%
-%:%960=1027%:%
-%:%961=1028%:%
-%:%962=1029%:%
-%:%963=1030%:%
-%:%964=1031%:%
-%:%965=1032%:%
-%:%966=1033%:%
-%:%967=1034%:%
-%:%968=1035%:%
-%:%969=1036%:%
-%:%970=1037%:%
-%:%971=1038%:%
-%:%972=1039%:%
-%:%973=1040%:%
-%:%974=1041%:%
-%:%975=1042%:%
-%:%976=1043%:%
-%:%977=1044%:%
-%:%978=1045%:%
-%:%979=1046%:%
-%:%980=1047%:%
-%:%981=1048%:%
-%:%982=1049%:%
-%:%983=1050%:%
-%:%984=1051%:%
-%:%985=1052%:%
-%:%986=1053%:%
-%:%987=1054%:%
-%:%988=1055%:%
-%:%989=1056%:%
-%:%990=1057%:%
-%:%991=1058%:%
-%:%992=1059%:%
-%:%993=1060%:%
-%:%994=1061%:%
-%:%995=1062%:%
-%:%996=1063%:%
-%:%997=1064%:%
-%:%998=1065%:%
-%:%999=1066%:%
-%:%1000=1067%:%
-%:%1001=1068%:%
-%:%1002=1069%:%
-%:%1003=1070%:%
-%:%1004=1071%:%
-%:%1005=1072%:%
-%:%1006=1073%:%
-%:%1007=1074%:%
-%:%1008=1075%:%
-%:%1009=1076%:%
-%:%1010=1077%:%
-%:%1011=1078%:%
-%:%1012=1079%:%
-%:%1013=1080%:%
-%:%1014=1081%:%
-%:%1015=1082%:%
-%:%1016=1083%:%
-%:%1017=1084%:%
-%:%1018=1085%:%
-%:%1019=1086%:%
-%:%1020=1087%:%
-%:%1021=1088%:%
-%:%1022=1089%:%
-%:%1023=1090%:%
-%:%1024=1091%:%
-%:%1025=1092%:%
-%:%1026=1093%:%
-%:%1027=1094%:%
-%:%1028=1095%:%
-%:%1029=1096%:%
-%:%1030=1097%:%
-%:%1031=1098%:%
-%:%1032=1099%:%
-%:%1033=1100%:%
-%:%1034=1101%:%
-%:%1035=1102%:%
-%:%1036=1103%:%
-%:%1037=1104%:%
-%:%1038=1105%:%
-%:%1039=1106%:%
-%:%1040=1107%:%
-%:%1041=1108%:%
-%:%1042=1109%:%
-%:%1043=1110%:%
-%:%1044=1111%:%
-%:%1045=1112%:%
-%:%1046=1113%:%
-%:%1047=1114%:%
-%:%1047=1115%:%
-%:%1048=1116%:%
-%:%1049=1117%:%
-%:%1050=1118%:%
-%:%1051=1119%:%
-%:%1052=1120%:%
-%:%1053=1121%:%
-%:%1054=1122%:%
-%:%1055=1123%:%
-%:%1056=1124%:%
-%:%1057=1125%:%
-%:%1058=1126%:%
-%:%1058=1127%:%
-%:%1059=1128%:%
-%:%1060=1129%:%
-%:%1061=1130%:%
-%:%1062=1131%:%
-%:%1063=1132%:%
-%:%1064=1133%:%
-%:%1065=1134%:%
-%:%1066=1135%:%
-%:%1067=1136%:%
-%:%1068=1137%:%
-%:%1069=1138%:%
-%:%1070=1139%:%
-%:%1071=1140%:%
-%:%1072=1141%:%
-%:%1073=1142%:%
-%:%1074=1143%:%
-%:%1075=1144%:%
-%:%1076=1145%:%
-%:%1077=1146%:%
-%:%1078=1147%:%
-%:%1079=1148%:%
-%:%1080=1149%:%
-%:%1081=1150%:%
-%:%1082=1151%:%
-%:%1083=1152%:%
-%:%1084=1153%:%
-%:%1085=1154%:%
-%:%1086=1155%:%
-%:%1087=1156%:%
-%:%1088=1157%:%
-%:%1089=1158%:%
-%:%1090=1159%:%
-%:%1091=1160%:%
-%:%1092=1161%:%
-%:%1093=1162%:%
-%:%1094=1163%:%
-%:%1095=1164%:%
-%:%1096=1165%:%
-%:%1097=1166%:%
-%:%1098=1167%:%
-%:%1099=1168%:%
-%:%1100=1169%:%
-%:%1101=1170%:%
-%:%1102=1171%:%
-%:%1103=1172%:%
-%:%1104=1173%:%
-%:%1105=1174%:%
-%:%1106=1175%:%
-%:%1107=1176%:%
-%:%1108=1177%:%
-%:%1109=1178%:%
-%:%1110=1179%:%
-%:%1111=1180%:%
-%:%1112=1181%:%
-%:%1112=1182%:%
-%:%1112=1183%:%
-%:%1113=1184%:%
-%:%1114=1185%:%
-%:%1114=1186%:%
-%:%1114=1187%:%
-%:%1114=1188%:%
-%:%1115=1189%:%
-%:%1116=1190%:%
-%:%1117=1191%:%
-%:%1118=1192%:%
-%:%1119=1193%:%
-%:%1119=1194%:%
-%:%1120=1195%:%
-%:%1121=1196%:%
-%:%1122=1197%:%
-%:%1123=1198%:%
-%:%1124=1199%:%
-%:%1125=1200%:%
-%:%1126=1201%:%
-%:%1127=1202%:%
-%:%1128=1203%:%
-%:%1129=1204%:%
-%:%1130=1205%:%
-%:%1131=1206%:%
-%:%1132=1207%:%
-%:%1133=1208%:%
-%:%1134=1209%:%
-%:%1135=1210%:%
-%:%1136=1211%:%
-%:%1137=1212%:%
-%:%1138=1213%:%
-%:%1139=1214%:%
-%:%1140=1215%:%
-%:%1141=1216%:%
-%:%1142=1217%:%
-%:%1143=1218%:%
-%:%1144=1219%:%
-%:%1145=1220%:%
-%:%1146=1221%:%
-%:%1147=1222%:%
-%:%1148=1223%:%
-%:%1149=1224%:%
-%:%1150=1225%:%
-%:%1151=1226%:%
-%:%1152=1227%:%
-%:%1153=1228%:%
-%:%1154=1229%:%
-%:%1155=1230%:%
-%:%1156=1231%:%
-%:%1157=1232%:%
-%:%1158=1233%:%
-%:%1159=1234%:%
-%:%1160=1235%:%
-%:%1161=1236%:%
-%:%1162=1237%:%
-%:%1163=1238%:%
-%:%1164=1239%:%
-%:%1165=1240%:%
-%:%1166=1241%:%
-%:%1167=1242%:%
-%:%1167=1243%:%
-%:%1168=1244%:%
-%:%1168=1245%:%
-%:%1169=1246%:%
-%:%1169=1247%:%
-%:%1170=1248%:%
-%:%1171=1249%:%
-%:%1172=1250%:%
-%:%1173=1251%:%
-%:%1174=1252%:%
-%:%1175=1253%:%
-%:%1176=1254%:%
-%:%1176=1255%:%
-%:%1177=1256%:%
-%:%1177=1257%:%
-%:%1178=1258%:%
-%:%1178=1259%:%
-%:%1179=1260%:%
-%:%1180=1261%:%
-%:%1181=1262%:%
-%:%1182=1263%:%
-%:%1183=1264%:%
-%:%1184=1265%:%
-%:%1185=1266%:%
-%:%1186=1267%:%
-%:%1187=1268%:%
-%:%1188=1269%:%
-%:%1189=1270%:%
-%:%1190=1271%:%
-%:%1191=1272%:%
-%:%1192=1273%:%
-%:%1193=1274%:%
-%:%1194=1275%:%
-%:%1195=1276%:%
-%:%1196=1277%:%
-%:%1197=1278%:%
-%:%1198=1279%:%
-%:%1199=1280%:%
-%:%1200=1281%:%
-%:%1201=1282%:%
-%:%1202=1283%:%
-%:%1203=1284%:%
-%:%1204=1285%:%
-%:%1205=1286%:%
-%:%1206=1287%:%
-%:%1207=1288%:%
-%:%1208=1289%:%
-%:%1209=1290%:%
-%:%1209=1291%:%
-%:%1210=1292%:%
-%:%1211=1293%:%
-%:%1212=1294%:%
-%:%1213=1295%:%
-%:%1214=1296%:%
-%:%1215=1297%:%
-%:%1216=1298%:%
-%:%1216=1299%:%
-%:%1217=1300%:%
-%:%1218=1301%:%
-%:%1219=1302%:%
-%:%1219=1303%:%
-%:%1220=1304%:%
-%:%1221=1305%:%
-%:%1222=1306%:%
-%:%1222=1307%:%
-%:%1223=1308%:%
-%:%1223=1309%:%
-%:%1224=1310%:%
-%:%1224=1311%:%
-%:%1225=1312%:%
-%:%1226=1313%:%
-%:%1227=1314%:%
-%:%1227=1315%:%
-%:%1228=1316%:%
-%:%1228=1317%:%
-%:%1229=1318%:%
-%:%1229=1319%:%
-%:%1229=1320%:%
-%:%1229=1321%:%
-%:%1230=1322%:%
-%:%1230=1323%:%
-%:%1231=1324%:%
-%:%1232=1325%:%
-%:%1232=1327%:%
-%:%1232=1328%:%
-%:%1232=1329%:%
-%:%1232=1330%:%
-%:%1232=1331%:%
-%:%1233=1332%:%
-%:%1233=1333%:%
-%:%1234=1334%:%
-%:%1234=1335%:%
-%:%1235=1336%:%
-%:%1235=1337%:%
-%:%1236=1338%:%
-%:%1237=1339%:%
-%:%1238=1340%:%
-%:%1239=1341%:%
-%:%1239=1342%:%
-%:%1240=1343%:%
-%:%1240=1344%:%
-%:%1241=1345%:%
-%:%1242=1346%:%
-%:%1242=1347%:%
-%:%1243=1348%:%
-%:%1244=1349%:%
-%:%1245=1350%:%
-%:%1246=1351%:%
-%:%1247=1352%:%
-%:%1248=1353%:%
-%:%1249=1354%:%
-%:%1250=1355%:%
-%:%1251=1356%:%
-%:%1252=1357%:%
-%:%1253=1358%:%
-%:%1254=1359%:%
-%:%1255=1360%:%
-%:%1256=1361%:%
-%:%1257=1362%:%
-%:%1258=1363%:%
-%:%1259=1364%:%
-%:%1260=1365%:%
-%:%1261=1366%:%
-%:%1262=1367%:%
-%:%1263=1368%:%
-%:%1264=1369%:%
-%:%1265=1370%:%
-%:%1266=1371%:%
-%:%1267=1372%:%
-%:%1268=1373%:%
-%:%1269=1374%:%
-%:%1270=1375%:%
-%:%1271=1376%:%
-%:%1272=1377%:%
-%:%1273=1378%:%
-%:%1274=1379%:%
-%:%1275=1380%:%
-%:%1276=1381%:%
-%:%1277=1382%:%
-%:%1278=1383%:%
-%:%1279=1384%:%
-%:%1280=1385%:%
-%:%1281=1386%:%
-%:%1282=1387%:%
-%:%1283=1388%:%
-%:%1284=1389%:%
-%:%1285=1390%:%
-%:%1286=1391%:%
-%:%1287=1392%:%
-%:%1288=1393%:%
-%:%1289=1394%:%
-%:%1290=1395%:%
-%:%1291=1396%:%
-%:%1292=1397%:%
-%:%1293=1398%:%
-%:%1302=1402%:%
-%:%1314=1409%:%
-%:%1315=1410%:%
-%:%1316=1411%:%
-%:%1317=1412%:%
-%:%1318=1413%:%
-%:%1319=1414%:%
-%:%1320=1415%:%
-%:%1329=1420%:%
-%:%1341=1424%:%
-%:%1342=1425%:%
-%:%1343=1426%:%
-%:%1344=1427%:%
-%:%1345=1428%:%
-%:%1346=1429%:%
-%:%1347=1430%:%
-%:%1348=1431%:%
-%:%1349=1432%:%
-%:%1350=1433%:%
-%:%1351=1434%:%
-%:%1352=1435%:%
-%:%1353=1436%:%
-%:%1354=1437%:%
-%:%1355=1438%:%
-%:%1356=1439%:%
-%:%1357=1440%:%
-%:%1358=1441%:%
-%:%1359=1442%:%
-%:%1360=1443%:%
-%:%1361=1444%:%
-%:%1362=1445%:%
-%:%1363=1446%:%
-%:%1364=1447%:%
-%:%1365=1448%:%
-%:%1366=1449%:%
-%:%1367=1450%:%
-%:%1368=1451%:%
-%:%1369=1452%:%
-%:%1370=1453%:%
-%:%1371=1454%:%
-%:%1372=1455%:%
-%:%1373=1456%:%
-%:%1374=1457%:%
-%:%1375=1458%:%
-%:%1376=1459%:%
-%:%1377=1460%:%
-%:%1378=1461%:%
-%:%1379=1462%:%
-%:%1380=1463%:%
-%:%1381=1464%:%
-%:%1382=1465%:%
-%:%1383=1466%:%
-%:%1384=1467%:%
-%:%1385=1468%:%
-%:%1386=1469%:%
-%:%1387=1470%:%
-%:%1388=1471%:%
-%:%1389=1472%:%
-%:%1390=1473%:%
-%:%1391=1474%:%
-%:%1392=1475%:%
-%:%1393=1476%:%
-%:%1394=1477%:%
+%:%955=1023%:%
+%:%956=1024%:%
+%:%957=1025%:%
+%:%958=1026%:%
+%:%959=1027%:%
+%:%960=1028%:%
+%:%961=1029%:%
+%:%962=1030%:%
+%:%963=1031%:%
+%:%964=1032%:%
+%:%965=1033%:%
+%:%966=1034%:%
+%:%967=1035%:%
+%:%968=1036%:%
+%:%969=1037%:%
+%:%970=1038%:%
+%:%971=1039%:%
+%:%972=1040%:%
+%:%973=1041%:%
+%:%974=1042%:%
+%:%975=1043%:%
+%:%976=1044%:%
+%:%977=1045%:%
+%:%978=1046%:%
+%:%979=1047%:%
+%:%980=1048%:%
+%:%981=1049%:%
+%:%982=1050%:%
+%:%983=1051%:%
+%:%984=1052%:%
+%:%985=1053%:%
+%:%986=1054%:%
+%:%987=1055%:%
+%:%988=1056%:%
+%:%989=1057%:%
+%:%990=1058%:%
+%:%991=1059%:%
+%:%992=1060%:%
+%:%1001=1064%:%
+%:%1013=1068%:%
+%:%1014=1069%:%
+%:%1015=1070%:%
+%:%1016=1071%:%
+%:%1017=1072%:%
+%:%1018=1073%:%
+%:%1019=1074%:%
+%:%1020=1075%:%
+%:%1021=1076%:%
+%:%1022=1077%:%
+%:%1023=1078%:%
+%:%1024=1079%:%
+%:%1025=1080%:%
+%:%1026=1081%:%
+%:%1027=1082%:%
+%:%1028=1083%:%
+%:%1029=1084%:%
+%:%1030=1085%:%
+%:%1031=1086%:%
+%:%1032=1087%:%
+%:%1033=1088%:%
+%:%1034=1089%:%
+%:%1035=1090%:%
+%:%1036=1091%:%
+%:%1037=1092%:%
+%:%1038=1093%:%
+%:%1039=1094%:%
+%:%1040=1095%:%
+%:%1041=1096%:%
+%:%1042=1097%:%
+%:%1043=1098%:%
+%:%1044=1099%:%
+%:%1045=1100%:%
+%:%1046=1101%:%
+%:%1047=1102%:%
+%:%1048=1103%:%
+%:%1049=1104%:%
+%:%1050=1105%:%
+%:%1051=1106%:%
+%:%1052=1107%:%
+%:%1053=1108%:%
+%:%1054=1109%:%
+%:%1055=1110%:%
+%:%1056=1111%:%
+%:%1057=1112%:%
+%:%1058=1113%:%
+%:%1059=1114%:%
+%:%1060=1115%:%
+%:%1061=1116%:%
+%:%1062=1117%:%
+%:%1063=1118%:%
+%:%1064=1119%:%
+%:%1065=1120%:%
+%:%1066=1121%:%
+%:%1067=1122%:%
+%:%1068=1123%:%
+%:%1069=1124%:%
+%:%1070=1125%:%
+%:%1071=1126%:%
+%:%1072=1127%:%
+%:%1073=1128%:%
+%:%1074=1129%:%
+%:%1075=1130%:%
+%:%1076=1131%:%
+%:%1077=1132%:%
+%:%1078=1133%:%
+%:%1079=1134%:%
+%:%1080=1135%:%
+%:%1081=1136%:%
+%:%1082=1137%:%
+%:%1083=1138%:%
+%:%1084=1139%:%
+%:%1085=1140%:%
+%:%1086=1141%:%
+%:%1087=1142%:%
+%:%1088=1143%:%
+%:%1089=1144%:%
+%:%1090=1145%:%
+%:%1091=1146%:%
+%:%1092=1147%:%
+%:%1093=1148%:%
+%:%1094=1149%:%
+%:%1095=1150%:%
+%:%1096=1151%:%
+%:%1097=1152%:%
+%:%1098=1153%:%
+%:%1099=1154%:%
+%:%1100=1155%:%
+%:%1101=1156%:%
+%:%1102=1157%:%
+%:%1103=1158%:%
+%:%1104=1159%:%
+%:%1105=1160%:%
+%:%1106=1161%:%
+%:%1107=1162%:%
+%:%1108=1163%:%
+%:%1109=1164%:%
+%:%1110=1165%:%
+%:%1111=1166%:%
+%:%1112=1167%:%
+%:%1113=1168%:%
+%:%1114=1169%:%
+%:%1115=1170%:%
+%:%1116=1171%:%
+%:%1117=1172%:%
+%:%1118=1173%:%
+%:%1119=1174%:%
+%:%1120=1175%:%
+%:%1121=1176%:%
+%:%1122=1177%:%
+%:%1123=1178%:%
+%:%1124=1179%:%
+%:%1125=1180%:%
+%:%1126=1181%:%
+%:%1127=1182%:%
+%:%1128=1183%:%
+%:%1128=1184%:%
+%:%1129=1185%:%
+%:%1130=1186%:%
+%:%1131=1187%:%
+%:%1132=1188%:%
+%:%1133=1189%:%
+%:%1134=1190%:%
+%:%1135=1191%:%
+%:%1136=1192%:%
+%:%1137=1193%:%
+%:%1138=1194%:%
+%:%1139=1195%:%
+%:%1139=1196%:%
+%:%1140=1197%:%
+%:%1141=1198%:%
+%:%1142=1199%:%
+%:%1143=1200%:%
+%:%1144=1201%:%
+%:%1145=1202%:%
+%:%1146=1203%:%
+%:%1147=1204%:%
+%:%1148=1205%:%
+%:%1149=1206%:%
+%:%1150=1207%:%
+%:%1151=1208%:%
+%:%1152=1209%:%
+%:%1153=1210%:%
+%:%1154=1211%:%
+%:%1155=1212%:%
+%:%1156=1213%:%
+%:%1157=1214%:%
+%:%1158=1215%:%
+%:%1159=1216%:%
+%:%1160=1217%:%
+%:%1161=1218%:%
+%:%1162=1219%:%
+%:%1163=1220%:%
+%:%1164=1221%:%
+%:%1165=1222%:%
+%:%1166=1223%:%
+%:%1167=1224%:%
+%:%1168=1225%:%
+%:%1169=1226%:%
+%:%1170=1227%:%
+%:%1171=1228%:%
+%:%1172=1229%:%
+%:%1173=1230%:%
+%:%1174=1231%:%
+%:%1175=1232%:%
+%:%1176=1233%:%
+%:%1177=1234%:%
+%:%1178=1235%:%
+%:%1179=1236%:%
+%:%1180=1237%:%
+%:%1181=1238%:%
+%:%1182=1239%:%
+%:%1183=1240%:%
+%:%1184=1241%:%
+%:%1185=1242%:%
+%:%1186=1243%:%
+%:%1187=1244%:%
+%:%1188=1245%:%
+%:%1189=1246%:%
+%:%1190=1247%:%
+%:%1191=1248%:%
+%:%1192=1249%:%
+%:%1193=1250%:%
+%:%1193=1251%:%
+%:%1193=1252%:%
+%:%1194=1253%:%
+%:%1195=1254%:%
+%:%1195=1255%:%
+%:%1195=1256%:%
+%:%1195=1257%:%
+%:%1196=1258%:%
+%:%1197=1259%:%
+%:%1198=1260%:%
+%:%1199=1261%:%
+%:%1200=1262%:%
+%:%1200=1263%:%
+%:%1201=1264%:%
+%:%1202=1265%:%
+%:%1203=1266%:%
+%:%1204=1267%:%
+%:%1205=1268%:%
+%:%1206=1269%:%
+%:%1207=1270%:%
+%:%1208=1271%:%
+%:%1209=1272%:%
+%:%1210=1273%:%
+%:%1211=1274%:%
+%:%1212=1275%:%
+%:%1213=1276%:%
+%:%1214=1277%:%
+%:%1215=1278%:%
+%:%1216=1279%:%
+%:%1217=1280%:%
+%:%1218=1281%:%
+%:%1219=1282%:%
+%:%1220=1283%:%
+%:%1221=1284%:%
+%:%1222=1285%:%
+%:%1223=1286%:%
+%:%1224=1287%:%
+%:%1225=1288%:%
+%:%1226=1289%:%
+%:%1227=1290%:%
+%:%1228=1291%:%
+%:%1229=1292%:%
+%:%1230=1293%:%
+%:%1231=1294%:%
+%:%1232=1295%:%
+%:%1233=1296%:%
+%:%1234=1297%:%
+%:%1235=1298%:%
+%:%1236=1299%:%
+%:%1237=1300%:%
+%:%1238=1301%:%
+%:%1239=1302%:%
+%:%1240=1303%:%
+%:%1241=1304%:%
+%:%1242=1305%:%
+%:%1243=1306%:%
+%:%1244=1307%:%
+%:%1245=1308%:%
+%:%1246=1309%:%
+%:%1247=1310%:%
+%:%1248=1311%:%
+%:%1248=1312%:%
+%:%1249=1313%:%
+%:%1249=1314%:%
+%:%1250=1315%:%
+%:%1250=1316%:%
+%:%1251=1317%:%
+%:%1252=1318%:%
+%:%1253=1319%:%
+%:%1254=1320%:%
+%:%1255=1321%:%
+%:%1256=1322%:%
+%:%1257=1323%:%
+%:%1257=1324%:%
+%:%1258=1325%:%
+%:%1258=1326%:%
+%:%1259=1327%:%
+%:%1259=1328%:%
+%:%1260=1329%:%
+%:%1261=1330%:%
+%:%1262=1331%:%
+%:%1263=1332%:%
+%:%1264=1333%:%
+%:%1265=1334%:%
+%:%1266=1335%:%
+%:%1267=1336%:%
+%:%1268=1337%:%
+%:%1269=1338%:%
+%:%1270=1339%:%
+%:%1271=1340%:%
+%:%1272=1341%:%
+%:%1273=1342%:%
+%:%1274=1343%:%
+%:%1275=1344%:%
+%:%1276=1345%:%
+%:%1277=1346%:%
+%:%1278=1347%:%
+%:%1279=1348%:%
+%:%1280=1349%:%
+%:%1281=1350%:%
+%:%1282=1351%:%
+%:%1283=1352%:%
+%:%1284=1353%:%
+%:%1285=1354%:%
+%:%1286=1355%:%
+%:%1287=1356%:%
+%:%1288=1357%:%
+%:%1289=1358%:%
+%:%1290=1359%:%
+%:%1290=1360%:%
+%:%1291=1361%:%
+%:%1292=1362%:%
+%:%1293=1363%:%
+%:%1294=1364%:%
+%:%1295=1365%:%
+%:%1296=1366%:%
+%:%1297=1367%:%
+%:%1297=1368%:%
+%:%1298=1369%:%
+%:%1299=1370%:%
+%:%1300=1371%:%
+%:%1300=1372%:%
+%:%1301=1373%:%
+%:%1302=1374%:%
+%:%1303=1375%:%
+%:%1303=1376%:%
+%:%1304=1377%:%
+%:%1304=1378%:%
+%:%1305=1379%:%
+%:%1305=1380%:%
+%:%1306=1381%:%
+%:%1307=1382%:%
+%:%1308=1383%:%
+%:%1308=1384%:%
+%:%1309=1385%:%
+%:%1309=1386%:%
+%:%1310=1387%:%
+%:%1310=1388%:%
+%:%1310=1389%:%
+%:%1310=1390%:%
+%:%1311=1391%:%
+%:%1311=1392%:%
+%:%1312=1393%:%
+%:%1313=1394%:%
+%:%1313=1396%:%
+%:%1313=1397%:%
+%:%1313=1398%:%
+%:%1313=1399%:%
+%:%1313=1400%:%
+%:%1314=1401%:%
+%:%1314=1402%:%
+%:%1315=1403%:%
+%:%1315=1404%:%
+%:%1316=1405%:%
+%:%1316=1406%:%
+%:%1317=1407%:%
+%:%1318=1408%:%
+%:%1319=1409%:%
+%:%1320=1410%:%
+%:%1320=1411%:%
+%:%1321=1412%:%
+%:%1321=1413%:%
+%:%1322=1414%:%
+%:%1323=1415%:%
+%:%1323=1416%:%
+%:%1324=1417%:%
+%:%1325=1418%:%
+%:%1326=1419%:%
+%:%1327=1420%:%
+%:%1328=1421%:%
+%:%1329=1422%:%
+%:%1330=1423%:%
+%:%1331=1424%:%
+%:%1332=1425%:%
+%:%1333=1426%:%
+%:%1334=1427%:%
+%:%1335=1428%:%
+%:%1336=1429%:%
+%:%1337=1430%:%
+%:%1338=1431%:%
+%:%1339=1432%:%
+%:%1340=1433%:%
+%:%1341=1434%:%
+%:%1342=1435%:%
+%:%1343=1436%:%
+%:%1344=1437%:%
+%:%1345=1438%:%
+%:%1346=1439%:%
+%:%1347=1440%:%
+%:%1348=1441%:%
+%:%1349=1442%:%
+%:%1350=1443%:%
+%:%1351=1444%:%
+%:%1352=1445%:%
+%:%1353=1446%:%
+%:%1354=1447%:%
+%:%1355=1448%:%
+%:%1356=1449%:%
+%:%1357=1450%:%
+%:%1358=1451%:%
+%:%1359=1452%:%
+%:%1360=1453%:%
+%:%1361=1454%:%
+%:%1362=1455%:%
+%:%1363=1456%:%
+%:%1364=1457%:%
+%:%1365=1458%:%
+%:%1366=1459%:%
+%:%1367=1460%:%
+%:%1368=1461%:%
+%:%1369=1462%:%
+%:%1370=1463%:%
+%:%1371=1464%:%
+%:%1372=1465%:%
+%:%1373=1466%:%
+%:%1374=1467%:%
+%:%1383=1471%:%
 %:%1395=1478%:%
 %:%1396=1479%:%
 %:%1397=1480%:%
@@ -2898,156 +2979,156 @@
 %:%1399=1482%:%
 %:%1400=1483%:%
 %:%1401=1484%:%
-%:%1402=1485%:%
-%:%1403=1486%:%
-%:%1404=1487%:%
-%:%1405=1488%:%
-%:%1406=1489%:%
-%:%1407=1490%:%
-%:%1408=1491%:%
-%:%1409=1492%:%
-%:%1410=1493%:%
-%:%1411=1494%:%
-%:%1412=1495%:%
-%:%1413=1496%:%
-%:%1414=1497%:%
-%:%1415=1498%:%
-%:%1416=1499%:%
-%:%1417=1500%:%
-%:%1418=1501%:%
-%:%1419=1502%:%
-%:%1420=1503%:%
-%:%1421=1504%:%
-%:%1422=1505%:%
-%:%1423=1506%:%
-%:%1424=1507%:%
-%:%1425=1508%:%
-%:%1426=1509%:%
-%:%1427=1510%:%
-%:%1428=1511%:%
-%:%1429=1512%:%
-%:%1430=1513%:%
-%:%1430=1629%:%
-%:%1431=1630%:%
-%:%1432=1631%:%
-%:%1433=1632%:%
-%:%1434=1633%:%
-%:%1434=1794%:%
-%:%1435=1795%:%
-%:%1436=1796%:%
-%:%1437=1797%:%
-%:%1438=1798%:%
-%:%1439=1799%:%
-%:%1440=1800%:%
-%:%1441=1801%:%
-%:%1442=1802%:%
-%:%1443=1803%:%
-%:%1444=1804%:%
-%:%1445=1805%:%
-%:%1446=1806%:%
-%:%1447=1807%:%
-%:%1448=1808%:%
-%:%1449=1809%:%
-%:%1450=1810%:%
-%:%1451=1811%:%
-%:%1452=1812%:%
-%:%1453=1813%:%
-%:%1453=1814%:%
-%:%1454=1815%:%
-%:%1454=1816%:%
-%:%1454=1817%:%
-%:%1455=1818%:%
-%:%1455=1819%:%
-%:%1456=1820%:%
-%:%1457=1821%:%
-%:%1458=1822%:%
-%:%1459=1823%:%
-%:%1460=1824%:%
-%:%1461=1825%:%
-%:%1462=1826%:%
-%:%1463=1827%:%
-%:%1464=1828%:%
-%:%1465=1829%:%
-%:%1466=1830%:%
-%:%1467=1831%:%
-%:%1468=1832%:%
-%:%1469=1833%:%
-%:%1470=1834%:%
-%:%1471=1835%:%
-%:%1472=1836%:%
-%:%1472=1837%:%
-%:%1473=1838%:%
-%:%1474=1839%:%
-%:%1475=1840%:%
-%:%1475=1841%:%
-%:%1475=1842%:%
-%:%1476=1843%:%
-%:%1477=1844%:%
-%:%1478=1845%:%
-%:%1478=1846%:%
-%:%1479=1847%:%
-%:%1479=1848%:%
-%:%1480=1849%:%
-%:%1481=1850%:%
-%:%1482=1851%:%
-%:%1483=1852%:%
-%:%1484=1853%:%
-%:%1485=1854%:%
-%:%1486=1855%:%
-%:%1487=1856%:%
-%:%1488=1857%:%
-%:%1489=1858%:%
-%:%1490=1859%:%
-%:%1491=1860%:%
-%:%1492=1861%:%
-%:%1493=1862%:%
-%:%1494=1863%:%
-%:%1495=1864%:%
-%:%1496=1865%:%
-%:%1505=1870%:%
-%:%1517=1874%:%
-%:%1518=1875%:%
-%:%1519=1876%:%
-%:%1520=1877%:%
-%:%1521=1878%:%
-%:%1522=1879%:%
-%:%1523=1880%:%
-%:%1524=1881%:%
-%:%1525=1882%:%
-%:%1526=1883%:%
-%:%1527=1884%:%
-%:%1528=1885%:%
-%:%1529=1886%:%
-%:%1530=1887%:%
-%:%1531=1888%:%
-%:%1532=1889%:%
-%:%1533=1890%:%
-%:%1534=1891%:%
-%:%1535=1892%:%
-%:%1536=1893%:%
-%:%1537=1894%:%
-%:%1538=1895%:%
-%:%1539=1896%:%
-%:%1540=1897%:%
-%:%1541=1898%:%
-%:%1542=1899%:%
-%:%1543=1900%:%
-%:%1544=1901%:%
-%:%1545=1902%:%
-%:%1546=1903%:%
-%:%1547=1904%:%
-%:%1548=1905%:%
-%:%1549=1906%:%
-%:%1550=1907%:%
-%:%1551=1908%:%
-%:%1552=1909%:%
-%:%1553=1910%:%
-%:%1554=1911%:%
-%:%1555=1912%:%
-%:%1556=1913%:%
-%:%1557=1914%:%
-%:%1558=1915%:%
-%:%1559=1916%:%
+%:%1410=1489%:%
+%:%1422=1493%:%
+%:%1423=1494%:%
+%:%1424=1495%:%
+%:%1425=1496%:%
+%:%1426=1497%:%
+%:%1427=1498%:%
+%:%1428=1499%:%
+%:%1429=1500%:%
+%:%1430=1501%:%
+%:%1431=1502%:%
+%:%1432=1503%:%
+%:%1433=1504%:%
+%:%1434=1505%:%
+%:%1435=1506%:%
+%:%1436=1507%:%
+%:%1437=1508%:%
+%:%1438=1509%:%
+%:%1439=1510%:%
+%:%1440=1511%:%
+%:%1441=1512%:%
+%:%1442=1513%:%
+%:%1443=1514%:%
+%:%1444=1515%:%
+%:%1445=1516%:%
+%:%1446=1517%:%
+%:%1447=1518%:%
+%:%1448=1519%:%
+%:%1449=1520%:%
+%:%1450=1521%:%
+%:%1451=1522%:%
+%:%1452=1523%:%
+%:%1453=1524%:%
+%:%1454=1525%:%
+%:%1455=1526%:%
+%:%1456=1527%:%
+%:%1457=1528%:%
+%:%1458=1529%:%
+%:%1459=1530%:%
+%:%1460=1531%:%
+%:%1461=1532%:%
+%:%1462=1533%:%
+%:%1463=1534%:%
+%:%1464=1535%:%
+%:%1465=1536%:%
+%:%1466=1537%:%
+%:%1467=1538%:%
+%:%1468=1539%:%
+%:%1469=1540%:%
+%:%1470=1541%:%
+%:%1471=1542%:%
+%:%1472=1543%:%
+%:%1473=1544%:%
+%:%1474=1545%:%
+%:%1475=1546%:%
+%:%1476=1547%:%
+%:%1477=1548%:%
+%:%1478=1549%:%
+%:%1479=1550%:%
+%:%1480=1551%:%
+%:%1481=1552%:%
+%:%1482=1553%:%
+%:%1483=1554%:%
+%:%1484=1555%:%
+%:%1485=1556%:%
+%:%1486=1557%:%
+%:%1487=1558%:%
+%:%1488=1559%:%
+%:%1489=1560%:%
+%:%1490=1561%:%
+%:%1491=1562%:%
+%:%1492=1563%:%
+%:%1493=1564%:%
+%:%1494=1565%:%
+%:%1495=1566%:%
+%:%1496=1567%:%
+%:%1497=1568%:%
+%:%1498=1569%:%
+%:%1499=1570%:%
+%:%1500=1571%:%
+%:%1501=1572%:%
+%:%1502=1573%:%
+%:%1503=1574%:%
+%:%1504=1575%:%
+%:%1505=1576%:%
+%:%1506=1577%:%
+%:%1507=1578%:%
+%:%1508=1579%:%
+%:%1509=1580%:%
+%:%1510=1581%:%
+%:%1511=1582%:%
+%:%1511=1698%:%
+%:%1512=1699%:%
+%:%1513=1700%:%
+%:%1514=1701%:%
+%:%1515=1702%:%
+%:%1515=1863%:%
+%:%1516=1864%:%
+%:%1517=1865%:%
+%:%1518=1866%:%
+%:%1519=1867%:%
+%:%1520=1868%:%
+%:%1521=1869%:%
+%:%1522=1870%:%
+%:%1523=1871%:%
+%:%1524=1872%:%
+%:%1525=1873%:%
+%:%1526=1874%:%
+%:%1527=1875%:%
+%:%1528=1876%:%
+%:%1529=1877%:%
+%:%1530=1878%:%
+%:%1531=1879%:%
+%:%1532=1880%:%
+%:%1533=1881%:%
+%:%1534=1882%:%
+%:%1534=1883%:%
+%:%1535=1884%:%
+%:%1535=1885%:%
+%:%1535=1886%:%
+%:%1536=1887%:%
+%:%1536=1888%:%
+%:%1537=1889%:%
+%:%1538=1890%:%
+%:%1539=1891%:%
+%:%1540=1892%:%
+%:%1541=1893%:%
+%:%1542=1894%:%
+%:%1543=1895%:%
+%:%1544=1896%:%
+%:%1545=1897%:%
+%:%1546=1898%:%
+%:%1547=1899%:%
+%:%1548=1900%:%
+%:%1549=1901%:%
+%:%1550=1902%:%
+%:%1551=1903%:%
+%:%1552=1904%:%
+%:%1553=1905%:%
+%:%1553=1906%:%
+%:%1554=1907%:%
+%:%1555=1908%:%
+%:%1556=1909%:%
+%:%1556=1910%:%
+%:%1556=1911%:%
+%:%1557=1912%:%
+%:%1558=1913%:%
+%:%1559=1914%:%
+%:%1559=1915%:%
+%:%1560=1916%:%
 %:%1560=1917%:%
 %:%1561=1918%:%
 %:%1562=1919%:%
@@ -3058,4 +3139,66 @@
 %:%1567=1924%:%
 %:%1568=1925%:%
 %:%1569=1926%:%
-%:%1582=1932%:%
\ No newline at end of file
+%:%1570=1927%:%
+%:%1571=1928%:%
+%:%1572=1929%:%
+%:%1573=1930%:%
+%:%1574=1931%:%
+%:%1575=1932%:%
+%:%1576=1933%:%
+%:%1577=1934%:%
+%:%1586=1939%:%
+%:%1598=1943%:%
+%:%1599=1944%:%
+%:%1600=1945%:%
+%:%1601=1946%:%
+%:%1602=1947%:%
+%:%1603=1948%:%
+%:%1604=1949%:%
+%:%1605=1950%:%
+%:%1606=1951%:%
+%:%1607=1952%:%
+%:%1608=1953%:%
+%:%1609=1954%:%
+%:%1610=1955%:%
+%:%1611=1956%:%
+%:%1612=1957%:%
+%:%1613=1958%:%
+%:%1614=1959%:%
+%:%1615=1960%:%
+%:%1616=1961%:%
+%:%1617=1962%:%
+%:%1618=1963%:%
+%:%1619=1964%:%
+%:%1620=1965%:%
+%:%1621=1966%:%
+%:%1622=1967%:%
+%:%1623=1968%:%
+%:%1624=1969%:%
+%:%1625=1970%:%
+%:%1626=1971%:%
+%:%1627=1972%:%
+%:%1628=1973%:%
+%:%1629=1974%:%
+%:%1630=1975%:%
+%:%1631=1976%:%
+%:%1632=1977%:%
+%:%1633=1978%:%
+%:%1634=1979%:%
+%:%1635=1980%:%
+%:%1636=1981%:%
+%:%1637=1982%:%
+%:%1638=1983%:%
+%:%1639=1984%:%
+%:%1640=1985%:%
+%:%1641=1986%:%
+%:%1642=1987%:%
+%:%1643=1988%:%
+%:%1644=1989%:%
+%:%1645=1990%:%
+%:%1646=1991%:%
+%:%1647=1992%:%
+%:%1648=1993%:%
+%:%1649=1994%:%
+%:%1650=1995%:%
+%:%1663=2001%:%
\ No newline at end of file