tphols-2011/generated/Myhill.tex
author zhang
Tue, 25 Jan 2011 12:14:31 +0000
changeset 31 b6815473ee2e
parent 30 f5db9e08effc
permissions -rw-r--r--
1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML

%
\begin{isabellebody}%
\def\isabellecontext{Myhill}%
%
\isadelimtheory
%
\endisadelimtheory
%
\isatagtheory
\isacommand{theory}\isamarkupfalse%
\ Myhill\isanewline
\ \ \isakeyword{imports}\ Main\ List{\isacharunderscore}Prefix\isanewline
\isakeyword{begin}%
\endisatagtheory
{\isafoldtheory}%
%
\isadelimtheory
%
\endisadelimtheory
%
\isamarkupsection{Preliminary definitions%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
Sequential composition of two languages \isa{L{\isadigit{1}}} and \isa{L{\isadigit{2}}}%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ Seq\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}string\ set\ {\isasymRightarrow}\ string\ set\ {\isasymRightarrow}\ string\ set{\isachardoublequoteclose}\ {\isacharparenleft}{\isachardoublequoteopen}{\isacharunderscore}\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharunderscore}{\isachardoublequoteclose}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharcomma}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharparenright}\isanewline
\isakeyword{where}\ \isanewline
\ \ {\isachardoublequoteopen}L{\isadigit{1}}\ {\isacharsemicolon}{\isacharsemicolon}\ L{\isadigit{2}}\ {\isacharequal}\ {\isacharbraceleft}s{\isadigit{1}}\ {\isacharat}\ s{\isadigit{2}}\ {\isacharbar}\ s{\isadigit{1}}\ s{\isadigit{2}}{\isachardot}\ s{\isadigit{1}}\ {\isasymin}\ L{\isadigit{1}}\ {\isasymand}\ s{\isadigit{2}}\ {\isasymin}\ L{\isadigit{2}}{\isacharbraceright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
Transitive closure of language \isa{L}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{inductive{\isacharunderscore}set}\isamarkupfalse%
\isanewline
\ \ Star\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}string\ set\ {\isasymRightarrow}\ string\ set{\isachardoublequoteclose}\ {\isacharparenleft}{\isachardoublequoteopen}{\isacharunderscore}{\isasymstar}{\isachardoublequoteclose}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{1}}{\isacharbrackright}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{2}}{\isacharparenright}\isanewline
\ \ \isakeyword{for}\ L\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}string\ set{\isachardoublequoteclose}\isanewline
\isakeyword{where}\isanewline
\ \ start{\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}\ {\isasymin}\ L{\isasymstar}{\isachardoublequoteclose}\isanewline
{\isacharbar}\ step{\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ \ {\isachardoublequoteopen}{\isasymlbrakk}s{\isadigit{1}}\ {\isasymin}\ L{\isacharsemicolon}\ s{\isadigit{2}}\ {\isasymin}\ L{\isasymstar}{\isasymrbrakk}\ {\isasymLongrightarrow}\ s{\isadigit{1}}{\isacharat}s{\isadigit{2}}\ {\isasymin}\ L{\isasymstar}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
Some properties of operator \isa{{\isacharsemicolon}{\isacharsemicolon}}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{lemma}\isamarkupfalse%
\ seq{\isacharunderscore}union{\isacharunderscore}distrib{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}{\isacharparenleft}A\ {\isasymunion}\ B{\isacharparenright}\ {\isacharsemicolon}{\isacharsemicolon}\ C\ {\isacharequal}\ {\isacharparenleft}A\ {\isacharsemicolon}{\isacharsemicolon}\ C{\isacharparenright}\ {\isasymunion}\ {\isacharparenleft}B\ {\isacharsemicolon}{\isacharsemicolon}\ C{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Seq{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ seq{\isacharunderscore}intro{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}{\isasymlbrakk}x\ {\isasymin}\ A{\isacharsemicolon}\ y\ {\isasymin}\ B{\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isacharat}\ y\ {\isasymin}\ A\ {\isacharsemicolon}{\isacharsemicolon}\ B\ {\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Seq{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ seq{\isacharunderscore}assoc{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}{\isacharparenleft}A\ {\isacharsemicolon}{\isacharsemicolon}\ B{\isacharparenright}\ {\isacharsemicolon}{\isacharsemicolon}\ C\ {\isacharequal}\ A\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharparenleft}B\ {\isacharsemicolon}{\isacharsemicolon}\ C{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
{\isacharparenleft}auto\ simp{\isacharcolon}Seq{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{apply}\isamarkupfalse%
\ blast\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}metis\ append{\isacharunderscore}assoc{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ star{\isacharunderscore}intro{\isadigit{1}}{\isacharbrackleft}rule{\isacharunderscore}format{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymin}\ lang{\isasymstar}\ {\isasymLongrightarrow}\ {\isasymforall}\ y{\isachardot}\ y\ {\isasymin}\ lang{\isasymstar}\ {\isasymlongrightarrow}\ x\ {\isacharat}\ y\ {\isasymin}\ lang{\isasymstar}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}erule\ Star{\isachardot}induct{\isacharcomma}\ auto{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ star{\isacharunderscore}intro{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}y\ {\isasymin}\ lang\ {\isasymLongrightarrow}\ y\ {\isasymin}\ lang{\isasymstar}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}drule\ step{\isacharbrackleft}of\ y\ lang\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharbrackright}{\isacharcomma}\ auto\ simp{\isacharcolon}start{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ star{\isacharunderscore}intro{\isadigit{3}}{\isacharbrackleft}rule{\isacharunderscore}format{\isacharbrackright}{\isacharcolon}\ \isanewline
\ \ {\isachardoublequoteopen}x\ {\isasymin}\ lang{\isasymstar}\ {\isasymLongrightarrow}\ {\isasymforall}y\ {\isachardot}\ y\ {\isasymin}\ lang\ {\isasymlongrightarrow}\ x\ {\isacharat}\ y\ {\isasymin}\ lang{\isasymstar}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}erule\ Star{\isachardot}induct{\isacharcomma}\ auto\ intro{\isacharcolon}star{\isacharunderscore}intro{\isadigit{2}}{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ star{\isacharunderscore}decom{\isacharcolon}\ \isanewline
\ \ {\isachardoublequoteopen}{\isasymlbrakk}x\ {\isasymin}\ lang{\isasymstar}{\isacharsemicolon}\ x\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isasymrbrakk}\ {\isasymLongrightarrow}{\isacharparenleft}{\isasymexists}\ a\ b{\isachardot}\ x\ {\isacharequal}\ a\ {\isacharat}\ b\ {\isasymand}\ a\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymand}\ a\ {\isasymin}\ lang\ {\isasymand}\ b\ {\isasymin}\ lang{\isasymstar}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}induct\ x\ rule{\isacharcolon}\ Star{\isachardot}induct{\isacharcomma}\ simp{\isacharcomma}\ blast{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ star{\isacharunderscore}decom{\isacharprime}{\isacharcolon}\ \isanewline
\ \ {\isachardoublequoteopen}{\isasymlbrakk}x\ {\isasymin}\ lang{\isasymstar}{\isacharsemicolon}\ x\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isasymexists}a\ b{\isachardot}\ x\ {\isacharequal}\ a\ {\isacharat}\ b\ {\isasymand}\ a\ {\isasymin}\ lang{\isasymstar}\ {\isasymand}\ b\ {\isasymin}\ lang{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}induct\ x\ rule{\isacharcolon}Star{\isachardot}induct{\isacharcomma}\ simp{\isacharparenright}\isanewline
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}case{\isacharunderscore}tac\ {\isachardoublequoteopen}s{\isadigit{2}}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharparenright}\isanewline
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharcomma}\ rule{\isacharunderscore}tac\ x\ {\isacharequal}\ s{\isadigit{1}}\ \isakeyword{in}\ exI{\isacharcomma}\ simp\ add{\isacharcolon}start{\isacharparenright}\isanewline
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}simp{\isacharcomma}\ {\isacharparenleft}erule\ exE{\isacharbar}\ erule\ conjE{\isacharparenright}{\isacharplus}{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}s{\isadigit{1}}\ {\isacharat}\ a{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharcomma}\ rule{\isacharunderscore}tac\ x\ {\isacharequal}\ b\ \isakeyword{in}\ exI{\isacharcomma}\ simp\ add{\isacharcolon}step{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\begin{isamarkuptext}%
Ardens lemma expressed at the level of language, rather than the level of regular expression.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{theorem}\isamarkupfalse%
\ ardens{\isacharunderscore}revised{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ nemp{\isacharcolon}\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}\ {\isasymnotin}\ A{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isacharparenleft}X\ {\isacharequal}\ X\ {\isacharsemicolon}{\isacharsemicolon}\ A\ {\isasymunion}\ B{\isacharparenright}\ {\isasymlongleftrightarrow}\ {\isacharparenleft}X\ {\isacharequal}\ B\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isasymstar}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\isanewline
\ \ \isacommand{assume}\isamarkupfalse%
\ eq{\isacharcolon}\ {\isachardoublequoteopen}X\ {\isacharequal}\ B\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isasymstar}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}A{\isasymstar}\ {\isacharequal}\ \ {\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}\ {\isasymunion}\ A{\isasymstar}\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Seq{\isacharunderscore}def\ star{\isacharunderscore}intro{\isadigit{3}}\ star{\isacharunderscore}decom{\isacharprime}{\isacharparenright}\ \ \isanewline
\ \ \isacommand{then}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}B\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isasymstar}\ {\isacharequal}\ B\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharparenleft}{\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}\ {\isasymunion}\ A{\isasymstar}\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{unfolding}\isamarkupfalse%
\ Seq{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \isacommand{also}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ B\ {\isasymunion}\ B\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharparenleft}A{\isasymstar}\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isacharparenright}{\isachardoublequoteclose}\ \ \isanewline
\ \ \ \ \isacommand{unfolding}\isamarkupfalse%
\ Seq{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \isacommand{also}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ B\ {\isasymunion}\ {\isacharparenleft}B\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isasymstar}{\isacharparenright}\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ only{\isacharcolon}seq{\isacharunderscore}assoc{\isacharparenright}\isanewline
\ \ \isacommand{finally}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}X\ {\isacharequal}\ X\ {\isacharsemicolon}{\isacharsemicolon}\ A\ {\isasymunion}\ B{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ eq\ \isacommand{by}\isamarkupfalse%
\ blast\ \isanewline
\isacommand{next}\isamarkupfalse%
\isanewline
\ \ \isacommand{assume}\isamarkupfalse%
\ eq{\isacharprime}{\isacharcolon}\ {\isachardoublequoteopen}X\ {\isacharequal}\ X\ {\isacharsemicolon}{\isacharsemicolon}\ A\ {\isasymunion}\ B{\isachardoublequoteclose}\isanewline
\ \ \isacommand{hence}\isamarkupfalse%
\ c{\isadigit{1}}{\isacharprime}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymAnd}\ x{\isachardot}\ x\ {\isasymin}\ B\ {\isasymLongrightarrow}\ x\ {\isasymin}\ X{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isakeyword{and}\ c{\isadigit{2}}{\isacharprime}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymAnd}\ x\ y{\isachardot}\ {\isasymlbrakk}x\ {\isasymin}\ X{\isacharsemicolon}\ y\ {\isasymin}\ A{\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isacharat}\ y\ {\isasymin}\ X{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ Seq{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}X\ {\isacharequal}\ B\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isasymstar}{\isachardoublequoteclose}\ \isanewline
\ \ \isacommand{proof}\isamarkupfalse%
\isanewline
\ \ \ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}B\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isasymstar}\ {\isasymsubseteq}\ X{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
\ \isacommand{fix}\isamarkupfalse%
\ x\ y\isanewline
\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymlbrakk}y\ {\isasymin}\ A{\isasymstar}{\isacharsemicolon}\ x\ {\isasymin}\ X{\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isacharat}\ y\ {\isasymin}\ X\ {\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}induct\ arbitrary{\isacharcolon}x\ rule{\isacharcolon}Star{\isachardot}induct{\isacharcomma}\ simp{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ only{\isacharcolon}append{\isacharunderscore}assoc{\isacharbrackleft}THEN\ sym{\isacharbrackright}\ dest{\isacharcolon}c{\isadigit{2}}{\isacharprime}{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
\ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ c{\isadigit{1}}{\isacharprime}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Seq{\isacharunderscore}def{\isacharparenright}\ \isanewline
\ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{next}\isamarkupfalse%
\isanewline
\ \ \ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}X\ {\isasymsubseteq}\ B\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isasymstar}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
\ \isacommand{fix}\isamarkupfalse%
\ x\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}x\ {\isasymin}\ X\ {\isasymLongrightarrow}\ x\ {\isasymin}\ B\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isasymstar}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharparenleft}induct\ x\ taking{\isacharcolon}length\ rule{\isacharcolon}measure{\isacharunderscore}induct{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{fix}\isamarkupfalse%
\ z\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
\ hyps{\isacharcolon}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ {\isachardoublequoteopen}{\isasymforall}y{\isachardot}\ length\ y\ {\isacharless}\ length\ z\ {\isasymlongrightarrow}\ y\ {\isasymin}\ X\ {\isasymlongrightarrow}\ y\ {\isasymin}\ B\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isasymstar}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isakeyword{and}\ z{\isacharunderscore}in{\isacharcolon}\ {\isachardoublequoteopen}z\ {\isasymin}\ X{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}z\ {\isasymin}\ B\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isasymstar}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharparenleft}cases\ {\isachardoublequoteopen}z\ {\isasymin}\ B{\isachardoublequoteclose}{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{case}\isamarkupfalse%
\ True\ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Seq{\isacharunderscore}def\ start{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{next}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{case}\isamarkupfalse%
\ False\ \isacommand{hence}\isamarkupfalse%
\ {\isachardoublequoteopen}z\ {\isasymin}\ X\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ eq{\isacharprime}\ z{\isacharunderscore}in\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{then}\isamarkupfalse%
\ \isacommand{obtain}\isamarkupfalse%
\ za\ zb\ \isakeyword{where}\ za{\isacharunderscore}in{\isacharcolon}\ {\isachardoublequoteopen}za\ {\isasymin}\ X{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \isakeyword{and}\ zab{\isacharcolon}\ {\isachardoublequoteopen}z\ {\isacharequal}\ za\ {\isacharat}\ zb\ {\isasymand}\ zb\ {\isasymin}\ A{\isachardoublequoteclose}\ \isakeyword{and}\ zbne{\isacharcolon}\ {\isachardoublequoteopen}zb\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ nemp\ \isacommand{unfolding}\isamarkupfalse%
\ Seq{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse%
\ zbne\ zab\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}length\ za\ {\isacharless}\ length\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
\ za{\isacharunderscore}in\ hyps\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}za\ {\isasymin}\ B\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isasymstar}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{hence}\isamarkupfalse%
\ {\isachardoublequoteopen}za\ {\isacharat}\ zb\ {\isasymin}\ B\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isasymstar}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ zab\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}clarsimp\ simp{\isacharcolon}Seq{\isacharunderscore}def{\isacharcomma}\ blast\ dest{\isacharcolon}star{\isacharunderscore}intro{\isadigit{3}}{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ zab\ \isacommand{by}\isamarkupfalse%
\ simp\ \ \ \ \ \ \ \isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
\ \isanewline
\ \ \ \ \ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
\ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\begin{isamarkuptext}%
The syntax of regular expressions is defined by the datatype \isa{rexp}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{datatype}\isamarkupfalse%
\ rexp\ {\isacharequal}\isanewline
\ \ NULL\isanewline
{\isacharbar}\ EMPTY\isanewline
{\isacharbar}\ CHAR\ char\isanewline
{\isacharbar}\ SEQ\ rexp\ rexp\isanewline
{\isacharbar}\ ALT\ rexp\ rexp\isanewline
{\isacharbar}\ STAR\ rexp%
\begin{isamarkuptext}%
The following \isa{L} is an overloaded operator, where \isa{L{\isacharparenleft}x{\isacharparenright}} evaluates to 
  the language represented by the syntactic object \isa{x}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{consts}\isamarkupfalse%
\ L{\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ string\ set{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The \isa{L{\isacharparenleft}rexp{\isacharparenright}} for regular expression \isa{rexp} is defined by the 
  following overloading function \isa{L{\isacharunderscore}rexp}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{overloading}\isamarkupfalse%
\ L{\isacharunderscore}rexp\ {\isasymequiv}\ {\isachardoublequoteopen}L{\isacharcolon}{\isacharcolon}\ \ rexp\ {\isasymRightarrow}\ string\ set{\isachardoublequoteclose}\isanewline
\isakeyword{begin}\isanewline
\isacommand{fun}\isamarkupfalse%
\isanewline
\ \ L{\isacharunderscore}rexp\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}rexp\ {\isasymRightarrow}\ string\ set{\isachardoublequoteclose}\isanewline
\isakeyword{where}\isanewline
\ \ \ \ {\isachardoublequoteopen}L{\isacharunderscore}rexp\ {\isacharparenleft}NULL{\isacharparenright}\ {\isacharequal}\ {\isacharbraceleft}{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\ \ {\isacharbar}\ {\isachardoublequoteopen}L{\isacharunderscore}rexp\ {\isacharparenleft}EMPTY{\isacharparenright}\ {\isacharequal}\ {\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\ \ {\isacharbar}\ {\isachardoublequoteopen}L{\isacharunderscore}rexp\ {\isacharparenleft}CHAR\ c{\isacharparenright}\ {\isacharequal}\ {\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\ \ {\isacharbar}\ {\isachardoublequoteopen}L{\isacharunderscore}rexp\ {\isacharparenleft}SEQ\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}L{\isacharunderscore}rexp\ r{\isadigit{1}}{\isacharparenright}\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharparenleft}L{\isacharunderscore}rexp\ r{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ {\isacharbar}\ {\isachardoublequoteopen}L{\isacharunderscore}rexp\ {\isacharparenleft}ALT\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}L{\isacharunderscore}rexp\ r{\isadigit{1}}{\isacharparenright}\ {\isasymunion}\ {\isacharparenleft}L{\isacharunderscore}rexp\ r{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ {\isacharbar}\ {\isachardoublequoteopen}L{\isacharunderscore}rexp\ {\isacharparenleft}STAR\ r{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}L{\isacharunderscore}rexp\ r{\isacharparenright}{\isasymstar}{\isachardoublequoteclose}\isanewline
\isacommand{end}\isamarkupfalse%
%
\begin{isamarkuptext}%
To obtain equational system out of finite set of equivalent classes, a fold operation
  on finite set \isa{folds} is defined. The use of \isa{SOME} makes \isa{fold}
  more robust than the \isa{fold} in Isabelle library. The expression \isa{folds\ f}
  makes sense when \isa{f} is not \isa{associative} and \isa{commutitive},
  while \isa{fold\ f} does not.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ folds\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ {\isasymRightarrow}\ {\isacharprime}b{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}b\ {\isasymRightarrow}\ {\isacharprime}a\ set\ {\isasymRightarrow}\ {\isacharprime}b{\isachardoublequoteclose}\isanewline
\isakeyword{where}\isanewline
\ \ {\isachardoublequoteopen}folds\ f\ z\ S\ {\isasymequiv}\ SOME\ x{\isachardot}\ fold{\isacharunderscore}graph\ f\ z\ S\ x{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The following lemma assures that the arbitrary choice made by the \isa{SOME} in \isa{folds}
  does not affect the \isa{L}-value of the resultant regular expression.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{lemma}\isamarkupfalse%
\ folds{\isacharunderscore}alt{\isacharunderscore}simp\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}finite\ rs\ {\isasymLongrightarrow}\ L\ {\isacharparenleft}folds\ ALT\ NULL\ rs{\isacharparenright}\ {\isacharequal}\ {\isasymUnion}\ {\isacharparenleft}L\ {\isacharbackquote}\ rs{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}rule\ set{\isacharunderscore}ext{\isacharcomma}\ simp\ add{\isacharcolon}folds{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}rule\ someI{\isadigit{2}}{\isacharunderscore}ex{\isacharcomma}\ erule\ finite{\isacharunderscore}imp{\isacharunderscore}fold{\isacharunderscore}graph{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}erule\ fold{\isacharunderscore}graph{\isachardot}induct{\isacharcomma}\ auto{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isasymin}\ {\isacharbraceleft}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isachardot}\ P\ x\ y{\isacharbraceright}\ {\isasymlongleftrightarrow}\ P\ x\ y{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ simp%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\begin{isamarkuptext}%
\isa{{\isasymapprox}L} is an equivalent class defined by language \isa{Lang}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\isanewline
\ \ str{\isacharunderscore}eq{\isacharunderscore}rel\ {\isacharparenleft}{\isachardoublequoteopen}{\isasymapprox}{\isacharunderscore}{\isachardoublequoteclose}{\isacharparenright}\isanewline
\isakeyword{where}\isanewline
\ \ {\isachardoublequoteopen}{\isasymapprox}Lang\ {\isasymequiv}\ {\isacharbraceleft}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isachardot}\ \ {\isacharparenleft}{\isasymforall}z{\isachardot}\ x\ {\isacharat}\ z\ {\isasymin}\ Lang\ {\isasymlongleftrightarrow}\ y\ {\isacharat}\ z\ {\isasymin}\ Lang{\isacharparenright}{\isacharbraceright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
Among equivlant clases of \isa{{\isasymapprox}Lang}, the set \isa{finals{\isacharparenleft}Lang{\isacharparenright}} singles out 
  those which contains strings from \isa{Lang}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ \ {\isachardoublequoteopen}finals\ Lang\ {\isasymequiv}\ {\isacharbraceleft}{\isasymapprox}Lang\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}x{\isacharbraceright}\ {\isacharbar}\ x\ {\isachardot}\ x\ {\isasymin}\ Lang{\isacharbraceright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The following lemma show the relationshipt between \isa{finals{\isacharparenleft}Lang{\isacharparenright}} and \isa{Lang}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{lemma}\isamarkupfalse%
\ lang{\isacharunderscore}is{\isacharunderscore}union{\isacharunderscore}of{\isacharunderscore}finals{\isacharcolon}\ \isanewline
\ \ {\isachardoublequoteopen}Lang\ {\isacharequal}\ {\isasymUnion}\ finals{\isacharparenleft}Lang{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ \isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}Lang\ {\isasymsubseteq}\ {\isasymUnion}\ {\isacharparenleft}finals\ Lang{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
\isanewline
\ \ \ \ \isacommand{fix}\isamarkupfalse%
\ x\isanewline
\ \ \ \ \isacommand{assume}\isamarkupfalse%
\ {\isachardoublequoteopen}x\ {\isasymin}\ Lang{\isachardoublequoteclose}\ \ \ \isanewline
\ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isasymUnion}\ {\isacharparenleft}finals\ Lang{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}finals{\isacharunderscore}def{\isacharcomma}\ rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}x{\isacharbraceright}{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Image{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def{\isacharparenright}\ \ \ \ \isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\isacommand{next}\isamarkupfalse%
\isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymUnion}\ {\isacharparenleft}finals\ Lang{\isacharparenright}\ {\isasymsubseteq}\ Lang{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}clarsimp\ simp{\isacharcolon}finals{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}drule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{in}\ spec{\isacharcomma}\ auto{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\isamarkupsection{Direction \isa{finite\ partition\ {\isasymRightarrow}\ regular\ language}%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
The relationship between equivalent classes can be described by an
  equational system.
  For example, in equational system \eqref{example_eqns},  $X_0, X_1$ are equivalent 
  classes. The first equation says every string in $X_0$ is obtained either by
  appending one $b$ to a string in $X_0$ or by appending one $a$ to a string in
  $X_1$ or just be an empty string (represented by the regular expression $\lambda$). Similary,
  the second equation tells how the strings inside $X_1$ are composed.
  \begin{equation}\label{example_eqns}
    \begin{aligned}
      X_0 & = X_0 b + X_1 a + \lambda \\
      X_1 & = X_0 a + X_1 b
    \end{aligned}
  \end{equation}
  The summands on the right hand side is represented by the following data type
  \isa{rhs{\isacharunderscore}item}, mnemonic for 'right hand side item'.
  Generally, there are two kinds of right hand side items, one kind corresponds to
  pure regular expressions, like the $\lambda$ in \eqref{example_eqns}, the other kind corresponds to
  transitions from one one equivalent class to another, like the $X_0 b, X_1 a$ etc.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{datatype}\isamarkupfalse%
\ rhs{\isacharunderscore}item\ {\isacharequal}\ \isanewline
\ \ \ Lam\ {\isachardoublequoteopen}rexp{\isachardoublequoteclose}\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isanewline
\ {\isacharbar}\ Trn\ {\isachardoublequoteopen}{\isacharparenleft}string\ set{\isacharparenright}{\isachardoublequoteclose}\ {\isachardoublequoteopen}rexp{\isachardoublequoteclose}%
\begin{isamarkuptext}%
In this formalization, pure regular expressions like $\lambda$ is 
  repsented by \isa{Lam{\isacharparenleft}EMPTY{\isacharparenright}}, while transitions like $X_0 a$ is represented by $Trn~X_0~(CHAR~a)$.%
\end{isamarkuptext}%
\isamarkuptrue%
%
\begin{isamarkuptext}%
The functions \isa{the{\isacharunderscore}r} and \isa{the{\isacharunderscore}Trn} are used to extract
  subcomponents from right hand side items.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{fun}\isamarkupfalse%
\ the{\isacharunderscore}r\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}rhs{\isacharunderscore}item\ {\isasymRightarrow}\ rexp{\isachardoublequoteclose}\isanewline
\isakeyword{where}\ {\isachardoublequoteopen}the{\isacharunderscore}r\ {\isacharparenleft}Lam\ r{\isacharparenright}\ {\isacharequal}\ r{\isachardoublequoteclose}\isanewline
\isanewline
\isacommand{fun}\isamarkupfalse%
\ the{\isacharunderscore}Trn{\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}rhs{\isacharunderscore}item\ {\isasymRightarrow}\ {\isacharparenleft}string\ set\ {\isasymtimes}\ rexp{\isacharparenright}{\isachardoublequoteclose}\isanewline
\isakeyword{where}\ {\isachardoublequoteopen}the{\isacharunderscore}Trn\ {\isacharparenleft}Trn\ Y\ r{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Y{\isacharcomma}\ r{\isacharparenright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
Every right hand side item \isa{itm} defines a string set given 
  \isa{L{\isacharparenleft}itm{\isacharparenright}}, defined as:%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{overloading}\isamarkupfalse%
\ L{\isacharunderscore}rhs{\isacharunderscore}e\ {\isasymequiv}\ {\isachardoublequoteopen}L{\isacharcolon}{\isacharcolon}\ rhs{\isacharunderscore}item\ {\isasymRightarrow}\ string\ set{\isachardoublequoteclose}\isanewline
\isakeyword{begin}\isanewline
\ \ \isacommand{fun}\isamarkupfalse%
\ L{\isacharunderscore}rhs{\isacharunderscore}e{\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}rhs{\isacharunderscore}item\ {\isasymRightarrow}\ string\ set{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{where}\isanewline
\ \ \ \ \ {\isachardoublequoteopen}L{\isacharunderscore}rhs{\isacharunderscore}e\ {\isacharparenleft}Lam\ r{\isacharparenright}\ {\isacharequal}\ L\ r{\isachardoublequoteclose}\ {\isacharbar}\isanewline
\ \ \ \ \ {\isachardoublequoteopen}L{\isacharunderscore}rhs{\isacharunderscore}e\ {\isacharparenleft}Trn\ X\ r{\isacharparenright}\ {\isacharequal}\ X\ {\isacharsemicolon}{\isacharsemicolon}\ L\ r{\isachardoublequoteclose}\isanewline
\isacommand{end}\isamarkupfalse%
%
\begin{isamarkuptext}%
The right hand side of every equation is represented by a set of
  items. The string set defined by such a set \isa{itms} is given
  by \isa{L{\isacharparenleft}itms{\isacharparenright}}, defined as:%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{overloading}\isamarkupfalse%
\ L{\isacharunderscore}rhs\ {\isasymequiv}\ {\isachardoublequoteopen}L{\isacharcolon}{\isacharcolon}\ rhs{\isacharunderscore}item\ set\ {\isasymRightarrow}\ string\ set{\isachardoublequoteclose}\isanewline
\isakeyword{begin}\isanewline
\ \ \ \isacommand{fun}\isamarkupfalse%
\ L{\isacharunderscore}rhs{\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}rhs{\isacharunderscore}item\ set\ {\isasymRightarrow}\ string\ set{\isachardoublequoteclose}\isanewline
\ \ \ \isakeyword{where}\ {\isachardoublequoteopen}L{\isacharunderscore}rhs\ rhs\ {\isacharequal}\ {\isasymUnion}\ {\isacharparenleft}L\ {\isacharbackquote}\ rhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
\isacommand{end}\isamarkupfalse%
%
\begin{isamarkuptext}%
Given a set of equivalent classses \isa{CS} and one equivalent class \isa{X} among
  \isa{CS}, the term \isa{init{\isacharunderscore}rhs\ CS\ X} is used to extract the right hand side of
  the equation describing the formation of \isa{X}. The definition of \isa{init{\isacharunderscore}rhs}
  is:%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\isanewline
\ \ {\isachardoublequoteopen}init{\isacharunderscore}rhs\ CS\ X\ {\isasymequiv}\ \ \isanewline
\ \ \ \ \ \ if\ {\isacharparenleft}{\isacharbrackleft}{\isacharbrackright}\ {\isasymin}\ X{\isacharparenright}\ then\ \isanewline
\ \ \ \ \ \ \ \ \ \ {\isacharbraceleft}Lam{\isacharparenleft}EMPTY{\isacharparenright}{\isacharbraceright}\ {\isasymunion}\ {\isacharbraceleft}Trn\ Y\ {\isacharparenleft}CHAR\ c{\isacharparenright}\ {\isacharbar}\ Y\ c{\isachardot}\ Y\ {\isasymin}\ CS\ {\isasymand}\ Y\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}\ {\isasymsubseteq}\ X{\isacharbraceright}\isanewline
\ \ \ \ \ \ else\ \isanewline
\ \ \ \ \ \ \ \ \ \ {\isacharbraceleft}Trn\ Y\ {\isacharparenleft}CHAR\ c{\isacharparenright}{\isacharbar}\ Y\ c{\isachardot}\ Y\ {\isasymin}\ CS\ {\isasymand}\ Y\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}\ {\isasymsubseteq}\ X{\isacharbraceright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
In the definition of \isa{init{\isacharunderscore}rhs}, the term 
  \isa{{\isacharbraceleft}Trn\ Y\ {\isacharparenleft}CHAR\ c{\isacharparenright}{\isacharbar}\ Y\ c{\isachardot}\ Y\ {\isasymin}\ CS\ {\isasymand}\ Y\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}\ {\isasymsubseteq}\ X{\isacharbraceright}} appearing on both branches
  describes the formation of strings in \isa{X} out of transitions, while 
  the term \isa{{\isacharbraceleft}Lam{\isacharparenleft}EMPTY{\isacharparenright}{\isacharbraceright}} describes the empty string which is intrinsically contained in
  \isa{X} rather than by transition. This \isa{{\isacharbraceleft}Lam{\isacharparenleft}EMPTY{\isacharparenright}{\isacharbraceright}} corresponds to 
  the $\lambda$ in \eqref{example_eqns}.

  With the help of \isa{init{\isacharunderscore}rhs}, the equitional system descrbing the formation of every
  equivalent class inside \isa{CS} is given by the following \isa{eqs{\isacharparenleft}CS{\isacharparenright}}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ {\isachardoublequoteopen}eqs\ CS\ {\isasymequiv}\ {\isacharbraceleft}{\isacharparenleft}X{\isacharcomma}\ init{\isacharunderscore}rhs\ CS\ X{\isacharparenright}\ {\isacharbar}\ X{\isachardot}\ \ X\ {\isasymin}\ CS{\isacharbraceright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The following \isa{items{\isacharunderscore}of\ rhs\ X} returns all \isa{X}-items in \isa{rhs}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\isanewline
\ \ {\isachardoublequoteopen}items{\isacharunderscore}of\ rhs\ X\ {\isasymequiv}\ {\isacharbraceleft}Trn\ X\ r\ {\isacharbar}\ r{\isachardot}\ {\isacharparenleft}Trn\ X\ r{\isacharparenright}\ {\isasymin}\ rhs{\isacharbraceright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The following \isa{rexp{\isacharunderscore}of\ rhs\ X} combines all regular expressions in \isa{X}-items
  using \isa{ALT} to form a single regular expression. 
  It will be used later to implement \isa{arden{\isacharunderscore}variate} and \isa{rhs{\isacharunderscore}subst}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ {\isachardoublequoteopen}rexp{\isacharunderscore}of\ rhs\ X\ {\isasymequiv}\ folds\ ALT\ NULL\ {\isacharparenleft}{\isacharparenleft}snd\ o\ the{\isacharunderscore}Trn{\isacharparenright}\ {\isacharbackquote}\ items{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The following \isa{lam{\isacharunderscore}of\ rhs} returns all pure regular expression items in \isa{rhs}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\isanewline
\ \ {\isachardoublequoteopen}lam{\isacharunderscore}of\ rhs\ {\isasymequiv}\ {\isacharbraceleft}Lam\ r\ {\isacharbar}\ r{\isachardot}\ Lam\ r\ {\isasymin}\ rhs{\isacharbraceright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The following \isa{rexp{\isacharunderscore}of{\isacharunderscore}lam\ rhs} combines pure regular expression items in \isa{rhs}
  using \isa{ALT} to form a single regular expression. 
  When all variables inside \isa{rhs} are eliminated, \isa{rexp{\isacharunderscore}of{\isacharunderscore}lam\ rhs}
  is used to compute compute the regular expression corresponds to \isa{rhs}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\isanewline
\ \ {\isachardoublequoteopen}rexp{\isacharunderscore}of{\isacharunderscore}lam\ rhs\ {\isasymequiv}\ folds\ ALT\ NULL\ {\isacharparenleft}the{\isacharunderscore}r\ {\isacharbackquote}\ lam{\isacharunderscore}of\ rhs{\isacharparenright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The following \isa{attach{\isacharunderscore}rexp\ rexp{\isacharprime}\ itm} attach 
  the regular expression \isa{rexp{\isacharprime}} to
  the right of right hand side item \isa{itm}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{fun}\isamarkupfalse%
\ attach{\isacharunderscore}rexp\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}rexp\ {\isasymRightarrow}\ rhs{\isacharunderscore}item\ {\isasymRightarrow}\ rhs{\isacharunderscore}item{\isachardoublequoteclose}\isanewline
\isakeyword{where}\isanewline
\ \ {\isachardoublequoteopen}attach{\isacharunderscore}rexp\ rexp{\isacharprime}\ {\isacharparenleft}Lam\ rexp{\isacharparenright}\ \ \ {\isacharequal}\ Lam\ {\isacharparenleft}SEQ\ rexp\ rexp{\isacharprime}{\isacharparenright}{\isachardoublequoteclose}\isanewline
{\isacharbar}\ {\isachardoublequoteopen}attach{\isacharunderscore}rexp\ rexp{\isacharprime}\ {\isacharparenleft}Trn\ X\ rexp{\isacharparenright}\ {\isacharequal}\ Trn\ X\ {\isacharparenleft}SEQ\ rexp\ rexp{\isacharprime}{\isacharparenright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The following \isa{append{\isacharunderscore}rhs{\isacharunderscore}rexp\ rhs\ rexp} attaches 
  \isa{rexp} to every item in \isa{rhs}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\isanewline
\ \ {\isachardoublequoteopen}append{\isacharunderscore}rhs{\isacharunderscore}rexp\ rhs\ rexp\ {\isasymequiv}\ {\isacharparenleft}attach{\isacharunderscore}rexp\ rexp{\isacharparenright}\ {\isacharbackquote}\ rhs{\isachardoublequoteclose}%
\begin{isamarkuptext}%
With the help of the two functions immediately above, Ardens'
  transformation on right hand side \isa{rhs} is implemented
  by the following function \isa{arden{\isacharunderscore}variate\ X\ rhs}.
  After this transformation, the recursive occurent of \isa{X}
  in \isa{rhs} will be eliminated, while the 
  string set defined by \isa{rhs} is kept unchanged.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ {\isachardoublequoteopen}arden{\isacharunderscore}variate\ X\ rhs\ {\isasymequiv}\ \isanewline
\ \ \ \ \ \ \ \ append{\isacharunderscore}rhs{\isacharunderscore}rexp\ {\isacharparenleft}rhs\ {\isacharminus}\ items{\isacharunderscore}of\ rhs\ X{\isacharparenright}\ {\isacharparenleft}STAR\ {\isacharparenleft}rexp{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
Suppose the equation defining \isa{X} is $X = xrhs$,
  the purpose of \isa{rhs{\isacharunderscore}subst} is to substitute all occurences of \isa{X} in
  \isa{rhs} by \isa{xrhs}.
  A litte thought may reveal that the final result
  should be: first append $(a_1 | a_2 | \ldots | a_n)$ to every item of \isa{xrhs} and then
  union the result with all non-\isa{X}-items of \isa{rhs}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ {\isachardoublequoteopen}rhs{\isacharunderscore}subst\ rhs\ X\ xrhs\ {\isasymequiv}\ \isanewline
\ \ \ \ \ \ \ \ {\isacharparenleft}rhs\ {\isacharminus}\ {\isacharparenleft}items{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isacharparenright}\ {\isasymunion}\ {\isacharparenleft}append{\isacharunderscore}rhs{\isacharunderscore}rexp\ xrhs\ {\isacharparenleft}rexp{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
Suppose the equation defining \isa{X} is $X = xrhs$, the follwing
  \isa{eqs{\isacharunderscore}subst\ ES\ X\ xrhs} substitute \isa{xrhs} into every equation
  of the equational system \isa{ES}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\isanewline
\ \ {\isachardoublequoteopen}eqs{\isacharunderscore}subst\ ES\ X\ xrhs\ {\isasymequiv}\ {\isacharbraceleft}{\isacharparenleft}Y{\isacharcomma}\ rhs{\isacharunderscore}subst\ yrhs\ X\ xrhs{\isacharparenright}\ {\isacharbar}\ Y\ yrhs{\isachardot}\ {\isacharparenleft}Y{\isacharcomma}\ yrhs{\isacharparenright}\ {\isasymin}\ ES{\isacharbraceright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The computation of regular expressions for equivalent classes is accomplished
  using a iteration principle given by the following lemma.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{lemma}\isamarkupfalse%
\ wf{\isacharunderscore}iter\ {\isacharbrackleft}rule{\isacharunderscore}format{\isacharbrackright}{\isacharcolon}\ \isanewline
\ \ \isakeyword{fixes}\ f\isanewline
\ \ \isakeyword{assumes}\ step{\isacharcolon}\ {\isachardoublequoteopen}{\isasymAnd}\ e{\isachardot}\ {\isasymlbrakk}P\ e{\isacharsemicolon}\ {\isasymnot}\ Q\ e{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymexists}\ e{\isacharprime}{\isachardot}\ P\ e{\isacharprime}\ {\isasymand}\ \ {\isacharparenleft}f{\isacharparenleft}e{\isacharprime}{\isacharparenright}{\isacharcomma}\ f{\isacharparenleft}e{\isacharparenright}{\isacharparenright}\ {\isasymin}\ less{\isacharunderscore}than{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ pe{\isacharcolon}\ \ \ \ \ {\isachardoublequoteopen}P\ e\ {\isasymlongrightarrow}\ {\isacharparenleft}{\isasymexists}\ e{\isacharprime}{\isachardot}\ P\ e{\isacharprime}\ {\isasymand}\ \ Q\ e{\isacharprime}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharparenleft}induct\ e\ rule{\isacharcolon}\ wf{\isacharunderscore}induct\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ {\isacharbrackleft}OF\ wf{\isacharunderscore}inv{\isacharunderscore}image{\isacharbrackleft}OF\ wf{\isacharunderscore}less{\isacharunderscore}than{\isacharcomma}\ \isakeyword{where}\ f\ {\isacharequal}\ {\isachardoublequoteopen}f{\isachardoublequoteclose}{\isacharbrackright}{\isacharbrackright}{\isacharcomma}\ clarify{\isacharparenright}\isanewline
\ \ \isacommand{fix}\isamarkupfalse%
\ x\ \isanewline
\ \ \isacommand{assume}\isamarkupfalse%
\ h\ {\isacharbrackleft}rule{\isacharunderscore}format{\isacharbrackright}{\isacharcolon}\ \isanewline
\ \ \ \ {\isachardoublequoteopen}{\isasymforall}y{\isachardot}\ {\isacharparenleft}y{\isacharcomma}\ x{\isacharparenright}\ {\isasymin}\ inv{\isacharunderscore}image\ less{\isacharunderscore}than\ f\ {\isasymlongrightarrow}\ P\ y\ {\isasymlongrightarrow}\ {\isacharparenleft}{\isasymexists}e{\isacharprime}{\isachardot}\ P\ e{\isacharprime}\ {\isasymand}\ Q\ e{\isacharprime}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isakeyword{and}\ px{\isacharcolon}\ {\isachardoublequoteopen}P\ x{\isachardoublequoteclose}\isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymexists}e{\isacharprime}{\isachardot}\ P\ e{\isacharprime}\ {\isasymand}\ Q\ e{\isacharprime}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
{\isacharparenleft}cases\ {\isachardoublequoteopen}Q\ x{\isachardoublequoteclose}{\isacharparenright}\isanewline
\ \ \ \ \isacommand{assume}\isamarkupfalse%
\ {\isachardoublequoteopen}Q\ x{\isachardoublequoteclose}\ \isacommand{with}\isamarkupfalse%
\ px\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \isacommand{next}\isamarkupfalse%
\isanewline
\ \ \ \ \isacommand{assume}\isamarkupfalse%
\ nq{\isacharcolon}\ {\isachardoublequoteopen}{\isasymnot}\ Q\ x{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{from}\isamarkupfalse%
\ step\ {\isacharbrackleft}OF\ px\ nq{\isacharbrackright}\isanewline
\ \ \ \ \isacommand{obtain}\isamarkupfalse%
\ e{\isacharprime}\ \isakeyword{where}\ pe{\isacharprime}{\isacharcolon}\ {\isachardoublequoteopen}P\ e{\isacharprime}{\isachardoublequoteclose}\ \isakeyword{and}\ ltf{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}f\ e{\isacharprime}{\isacharcomma}\ f\ x{\isacharparenright}\ {\isasymin}\ less{\isacharunderscore}than{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \ \ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\isanewline
\ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharparenleft}rule\ h{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{from}\isamarkupfalse%
\ ltf\ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}e{\isacharprime}{\isacharcomma}\ x{\isacharparenright}\ {\isasymin}\ inv{\isacharunderscore}image\ less{\isacharunderscore}than\ f{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}inv{\isacharunderscore}image{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{next}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \isacommand{from}\isamarkupfalse%
\ pe{\isacharprime}\ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}P\ e{\isacharprime}{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
\isanewline
\ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\begin{isamarkuptext}%
The \isa{P} in lemma \isa{wf{\isacharunderscore}iter} is an invaiant kept throughout the iteration procedure.
  The particular invariant used to solve our problem is defined by function \isa{Inv{\isacharparenleft}ES{\isacharparenright}},
  an invariant over equal system \isa{ES}.
  Every definition starting next till \isa{Inv} stipulates a property to be satisfied by \isa{ES}.%
\end{isamarkuptext}%
\isamarkuptrue%
%
\begin{isamarkuptext}%
Every variable is defined at most onece in \isa{ES}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ {\isachardoublequoteopen}distinct{\isacharunderscore}equas\ ES\ {\isasymequiv}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ {\isasymforall}\ X\ rhs\ rhs{\isacharprime}{\isachardot}\ {\isacharparenleft}X{\isacharcomma}\ rhs{\isacharparenright}\ {\isasymin}\ ES\ {\isasymand}\ {\isacharparenleft}X{\isacharcomma}\ rhs{\isacharprime}{\isacharparenright}\ {\isasymin}\ ES\ {\isasymlongrightarrow}\ rhs\ {\isacharequal}\ rhs{\isacharprime}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
Every equation in \isa{ES} (represented by \isa{{\isacharparenleft}X{\isacharcomma}\ rhs{\isacharparenright}}) is valid, i.e. \isa{{\isacharparenleft}X\ {\isacharequal}\ L\ rhs{\isacharparenright}}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ {\isachardoublequoteopen}valid{\isacharunderscore}eqns\ ES\ {\isasymequiv}\ {\isasymforall}\ X\ rhs{\isachardot}\ {\isacharparenleft}X{\isacharcomma}\ rhs{\isacharparenright}\ {\isasymin}\ ES\ {\isasymlongrightarrow}\ {\isacharparenleft}X\ {\isacharequal}\ L\ rhs{\isacharparenright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
\isa{rhs{\isacharunderscore}nonempty\ rhs} requires regular expressions occuring in transitional 
  items of \isa{rhs} does not contain empty string. This is necessary for
  the application of Arden's transformation to \isa{rhs}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ {\isachardoublequoteopen}rhs{\isacharunderscore}nonempty\ rhs\ {\isasymequiv}\ {\isacharparenleft}{\isasymforall}\ Y\ r{\isachardot}\ Trn\ Y\ r\ {\isasymin}\ rhs\ {\isasymlongrightarrow}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymnotin}\ L\ r{\isacharparenright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
\isa{ardenable\ ES} requires that Arden's transformation is applicable
  to every equation of equational system \isa{ES}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ {\isachardoublequoteopen}ardenable\ ES\ {\isasymequiv}\ {\isasymforall}\ X\ rhs{\isachardot}\ {\isacharparenleft}X{\isacharcomma}\ rhs{\isacharparenright}\ {\isasymin}\ ES\ {\isasymlongrightarrow}\ rhs{\isacharunderscore}nonempty\ rhs{\isachardoublequoteclose}\isanewline
\isanewline
\isanewline
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ {\isachardoublequoteopen}non{\isacharunderscore}empty\ ES\ {\isasymequiv}\ {\isasymforall}\ X\ rhs{\isachardot}\ {\isacharparenleft}X{\isacharcomma}\ rhs{\isacharparenright}\ {\isasymin}\ ES\ {\isasymlongrightarrow}\ X\ {\isasymnoteq}\ {\isacharbraceleft}{\isacharbraceright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The following \isa{finite{\isacharunderscore}rhs\ ES} requires every equation in \isa{rhs} be finite.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\isanewline
\ \ {\isachardoublequoteopen}finite{\isacharunderscore}rhs\ ES\ {\isasymequiv}\ {\isasymforall}\ X\ rhs{\isachardot}\ {\isacharparenleft}X{\isacharcomma}\ rhs{\isacharparenright}\ {\isasymin}\ ES\ {\isasymlongrightarrow}\ finite\ rhs{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The following \isa{classes{\isacharunderscore}of\ rhs} returns all variables (or equivalent classes)
  occuring in \isa{rhs}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ {\isachardoublequoteopen}classes{\isacharunderscore}of\ rhs\ {\isasymequiv}\ {\isacharbraceleft}X{\isachardot}\ {\isasymexists}\ r{\isachardot}\ Trn\ X\ r\ {\isasymin}\ rhs{\isacharbraceright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The following \isa{lefts{\isacharunderscore}of\ ES} returns all variables 
  defined by equational system \isa{ES}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\isanewline
\ \ {\isachardoublequoteopen}lefts{\isacharunderscore}of\ ES\ {\isasymequiv}\ {\isacharbraceleft}Y\ {\isacharbar}\ Y\ yrhs{\isachardot}\ {\isacharparenleft}Y{\isacharcomma}\ yrhs{\isacharparenright}\ {\isasymin}\ ES{\isacharbraceright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The following \isa{self{\isacharunderscore}contained\ ES} requires that every
  variable occuring on the right hand side of equations is already defined by some
  equation in \isa{ES}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ {\isachardoublequoteopen}self{\isacharunderscore}contained\ ES\ {\isasymequiv}\ {\isasymforall}\ {\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharparenright}\ {\isasymin}\ ES{\isachardot}\ classes{\isacharunderscore}of\ xrhs\ {\isasymsubseteq}\ lefts{\isacharunderscore}of\ ES{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The invariant \isa{Inv{\isacharparenleft}ES{\isacharparenright}} is obtained by conjunctioning all the previous
  defined constaints on \isa{ES}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ {\isachardoublequoteopen}Inv\ ES\ {\isasymequiv}\ valid{\isacharunderscore}eqns\ ES\ {\isasymand}\ finite\ ES\ {\isasymand}\ distinct{\isacharunderscore}equas\ ES\ {\isasymand}\ ardenable\ ES\ {\isasymand}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ non{\isacharunderscore}empty\ ES\ {\isasymand}\ finite{\isacharunderscore}rhs\ ES\ {\isasymand}\ self{\isacharunderscore}contained\ ES{\isachardoublequoteclose}%
\isamarkupsubsection{Proof for this direction%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
The following are some basic properties of the above definitions.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{lemma}\isamarkupfalse%
\ L{\isacharunderscore}rhs{\isacharunderscore}union{\isacharunderscore}distrib{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}\ L\ {\isacharparenleft}A{\isacharcolon}{\isacharcolon}rhs{\isacharunderscore}item\ set{\isacharparenright}\ {\isasymunion}\ L\ B\ {\isacharequal}\ L\ {\isacharparenleft}A\ {\isasymunion}\ B{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ simp%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ finite{\isacharunderscore}snd{\isacharunderscore}Trn{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ finite{\isacharcolon}{\isachardoublequoteopen}finite\ rhs{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}finite\ {\isacharbraceleft}r\isactrlisub {\isadigit{2}}{\isachardot}\ Trn\ Y\ r\isactrlisub {\isadigit{2}}\ {\isasymin}\ rhs{\isacharbraceright}{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}finite\ {\isacharquery}B{\isachardoublequoteclose}{\isacharparenright}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \isacommand{def}\isamarkupfalse%
\ rhs{\isacharprime}\ {\isasymequiv}\ {\isachardoublequoteopen}{\isacharbraceleft}e\ {\isasymin}\ rhs{\isachardot}\ {\isasymexists}\ r{\isachardot}\ e\ {\isacharequal}\ Trn\ Y\ r{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharquery}B\ {\isacharequal}\ {\isacharparenleft}snd\ o\ the{\isacharunderscore}Trn{\isacharparenright}\ {\isacharbackquote}\ rhs{\isacharprime}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ rhs{\isacharprime}{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}image{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ rhs{\isacharprime}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ finite\ rhs{\isacharprime}{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ rexp{\isacharunderscore}of{\isacharunderscore}empty{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ finite{\isacharcolon}{\isachardoublequoteopen}finite\ rhs{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{and}\ nonempty{\isacharcolon}{\isachardoublequoteopen}rhs{\isacharunderscore}nonempty\ rhs{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}\ {\isasymnotin}\ L\ {\isacharparenleft}rexp{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{using}\isamarkupfalse%
\ finite\ nonempty\ rhs{\isacharunderscore}nonempty{\isacharunderscore}def\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}drule{\isacharunderscore}tac\ finite{\isacharunderscore}snd{\isacharunderscore}Trn{\isacharbrackleft}\isakeyword{where}\ Y\ {\isacharequal}\ X{\isacharbrackright}{\isacharcomma}\ auto\ simp{\isacharcolon}rexp{\isacharunderscore}of{\isacharunderscore}def\ items{\isacharunderscore}of{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ {\isacharbrackleft}intro{\isacharbang}{\isacharbrackright}{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}P\ {\isacharparenleft}Trn\ X\ r{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymexists}a{\isachardot}\ {\isacharparenleft}{\isasymexists}r{\isachardot}\ a\ {\isacharequal}\ Trn\ X\ r\ {\isasymand}\ P\ a{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}%
\isadelimproof
\ %
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ auto%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
\isanewline
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ finite{\isacharunderscore}items{\isacharunderscore}of{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}finite\ rhs\ {\isasymLongrightarrow}\ finite\ {\isacharparenleft}items{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}items{\isacharunderscore}of{\isacharunderscore}def\ intro{\isacharcolon}finite{\isacharunderscore}subset{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ lang{\isacharunderscore}of{\isacharunderscore}rexp{\isacharunderscore}of{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ finite{\isacharcolon}{\isachardoublequoteopen}finite\ rhs{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}L\ {\isacharparenleft}items{\isacharunderscore}of\ rhs\ X{\isacharparenright}\ {\isacharequal}\ X\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharparenleft}L\ {\isacharparenleft}rexp{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharparenleft}{\isacharparenleft}snd\ {\isasymcirc}\ the{\isacharunderscore}Trn{\isacharparenright}\ {\isacharbackquote}\ items{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ finite{\isacharunderscore}items{\isacharunderscore}of{\isacharbrackleft}OF\ finite{\isacharbrackright}\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}rexp{\isacharunderscore}of{\isacharunderscore}def\ Seq{\isacharunderscore}def\ items{\isacharunderscore}of{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ s{\isadigit{1}}\ \isakeyword{in}\ exI{\isacharcomma}\ rule{\isacharunderscore}tac\ x\ {\isacharequal}\ s{\isadigit{2}}\ \isakeyword{in}\ exI{\isacharcomma}\ auto{\isacharparenright}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x{\isacharequal}\ {\isachardoublequoteopen}Trn\ X\ r{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharcomma}\ auto\ simp{\isacharcolon}Seq{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ rexp{\isacharunderscore}of{\isacharunderscore}lam{\isacharunderscore}eq{\isacharunderscore}lam{\isacharunderscore}set{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ finite{\isacharcolon}\ {\isachardoublequoteopen}finite\ rhs{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}L\ {\isacharparenleft}rexp{\isacharunderscore}of{\isacharunderscore}lam\ rhs{\isacharparenright}\ {\isacharequal}\ L\ {\isacharparenleft}lam{\isacharunderscore}of\ rhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharparenleft}the{\isacharunderscore}r\ {\isacharbackquote}\ {\isacharbraceleft}Lam\ r\ {\isacharbar}r{\isachardot}\ Lam\ r\ {\isasymin}\ rhs{\isacharbraceright}{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ finite\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ finite{\isacharunderscore}imageI{\isacharcomma}\ auto\ intro{\isacharcolon}finite{\isacharunderscore}subset{\isacharparenright}\isanewline
\ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}rexp{\isacharunderscore}of{\isacharunderscore}lam{\isacharunderscore}def\ lam{\isacharunderscore}of{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}\ L\ {\isacharparenleft}attach{\isacharunderscore}rexp\ r\ xb{\isacharparenright}\ {\isacharequal}\ L\ xb\ {\isacharsemicolon}{\isacharsemicolon}\ L\ r{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}cases\ xb{\isacharcomma}\ auto\ simp{\isacharcolon}Seq{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}s{\isadigit{1}}\ {\isacharat}\ s{\isadigit{1}}a{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharcomma}\ rule{\isacharunderscore}tac\ x\ {\isacharequal}\ s{\isadigit{2}}a\ \isakeyword{in}\ exI{\isacharcomma}auto\ simp{\isacharcolon}Seq{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ lang{\isacharunderscore}of{\isacharunderscore}append{\isacharunderscore}rhs{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}L\ {\isacharparenleft}append{\isacharunderscore}rhs{\isacharunderscore}rexp\ rhs\ r{\isacharparenright}\ {\isacharequal}\ L\ rhs\ {\isacharsemicolon}{\isacharsemicolon}\ L\ r{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}append{\isacharunderscore}rhs{\isacharunderscore}rexp{\isacharunderscore}def\ image{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Seq{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}L\ xb\ {\isacharsemicolon}{\isacharsemicolon}\ L\ r{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharcomma}\ auto\ simp\ add{\isacharcolon}Seq{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}attach{\isacharunderscore}rexp\ r\ xb{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharcomma}\ auto\ simp{\isacharcolon}Seq{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ classes{\isacharunderscore}of{\isacharunderscore}union{\isacharunderscore}distrib{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}classes{\isacharunderscore}of\ A\ {\isasymunion}\ classes{\isacharunderscore}of\ B\ {\isacharequal}\ classes{\isacharunderscore}of\ {\isacharparenleft}A\ {\isasymunion}\ B{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}classes{\isacharunderscore}of{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ lefts{\isacharunderscore}of{\isacharunderscore}union{\isacharunderscore}distrib{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}lefts{\isacharunderscore}of\ A\ {\isasymunion}\ lefts{\isacharunderscore}of\ B\ {\isacharequal}\ lefts{\isacharunderscore}of\ {\isacharparenleft}A\ {\isasymunion}\ B{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}lefts{\isacharunderscore}of{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\begin{isamarkuptext}%
The following several lemmas until \isa{init{\isacharunderscore}ES{\isacharunderscore}satisfy{\isacharunderscore}Inv} are
  to prove that initial equational system satisfies invariant \isa{Inv}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{lemma}\isamarkupfalse%
\ defined{\isacharunderscore}by{\isacharunderscore}str{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}{\isasymlbrakk}s\ {\isasymin}\ X{\isacharsemicolon}\ X\ {\isasymin}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ X\ {\isacharequal}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}s{\isacharbraceright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}quotient{\isacharunderscore}def\ Image{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ every{\isacharunderscore}eqclass{\isacharunderscore}has{\isacharunderscore}transition{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ has{\isacharunderscore}str{\isacharcolon}\ {\isachardoublequoteopen}s\ {\isacharat}\ {\isacharbrackleft}c{\isacharbrackright}\ {\isasymin}\ X{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{and}\ \ \ \ \ in{\isacharunderscore}CS{\isacharcolon}\ \ \ {\isachardoublequoteopen}X\ {\isasymin}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{obtains}\ Y\ \isakeyword{where}\ {\isachardoublequoteopen}Y\ {\isasymin}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}Y\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}\ {\isasymsubseteq}\ X{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}s\ {\isasymin}\ Y{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \isacommand{def}\isamarkupfalse%
\ Y\ {\isasymequiv}\ {\isachardoublequoteopen}{\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}s{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}Y\ {\isasymin}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{unfolding}\isamarkupfalse%
\ Y{\isacharunderscore}def\ quotient{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}X\ {\isacharequal}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}s\ {\isacharat}\ {\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ has{\isacharunderscore}str\ in{\isacharunderscore}CS\ defined{\isacharunderscore}by{\isacharunderscore}str\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \isacommand{then}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}Y\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}\ {\isasymsubseteq}\ X{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{unfolding}\isamarkupfalse%
\ Y{\isacharunderscore}def\ Image{\isacharunderscore}def\ Seq{\isacharunderscore}def\isanewline
\ \ \ \ \isacommand{unfolding}\isamarkupfalse%
\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ clarsimp\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}s\ {\isasymin}\ Y{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse%
\ Y{\isacharunderscore}def\ \isanewline
\ \ \ \ \isacommand{unfolding}\isamarkupfalse%
\ Image{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}blast\ intro{\isacharcolon}\ that{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ l{\isacharunderscore}eq{\isacharunderscore}r{\isacharunderscore}in{\isacharunderscore}eqs{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ X{\isacharunderscore}in{\isacharunderscore}eqs{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharparenright}\ {\isasymin}\ {\isacharparenleft}eqs\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}X\ {\isacharequal}\ L\ xrhs{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ \isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}X\ {\isasymsubseteq}\ L\ xrhs{\isachardoublequoteclose}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
\isanewline
\ \ \ \ \isacommand{fix}\isamarkupfalse%
\ x\isanewline
\ \ \ \ \isacommand{assume}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}{\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymin}\ X{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}x\ {\isasymin}\ L\ xrhs{\isachardoublequoteclose}\ \ \ \ \ \ \ \ \ \ \isanewline
\ \ \ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharparenleft}cases\ {\isachardoublequoteopen}x\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
\ empty{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ X{\isacharunderscore}in{\isacharunderscore}eqs\ {\isachardoublequoteopen}{\isacharparenleft}{\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}eqs{\isacharunderscore}def\ init{\isacharunderscore}rhs{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{next}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
\ not{\isacharunderscore}empty{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isacommand{then}\isamarkupfalse%
\ \isacommand{obtain}\isamarkupfalse%
\ clist\ c\ \isakeyword{where}\ decom{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isacharequal}\ clist\ {\isacharat}\ {\isacharbrackleft}c{\isacharbrackright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}case{\isacharunderscore}tac\ x\ rule{\isacharcolon}rev{\isacharunderscore}cases{\isacharcomma}\ auto{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}X\ {\isasymin}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ X{\isacharunderscore}in{\isacharunderscore}eqs\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}eqs{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{then}\isamarkupfalse%
\ \isacommand{obtain}\isamarkupfalse%
\ Y\ \isanewline
\ \ \ \ \ \ \ \ \isakeyword{where}\ {\isachardoublequoteopen}Y\ {\isasymin}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}Y\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}\ {\isasymsubseteq}\ X{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}clist\ {\isasymin}\ Y{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ decom\ {\isachardoublequoteopen}{\isacharparenleft}{\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\ every{\isacharunderscore}eqclass{\isacharunderscore}has{\isacharunderscore}transition\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \ \ \ \ \isacommand{hence}\isamarkupfalse%
\ \isanewline
\ \ \ \ \ \ \ \ {\isachardoublequoteopen}x\ {\isasymin}\ L\ {\isacharbraceleft}Trn\ Y\ {\isacharparenleft}CHAR\ c{\isacharparenright}{\isacharbar}\ Y\ c{\isachardot}\ Y\ {\isasymin}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}\ {\isasymand}\ Y\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}\ {\isasymsubseteq}\ X{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}{\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\ decom\isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp{\isacharcomma}\ rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}Trn\ Y\ {\isacharparenleft}CHAR\ c{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharcomma}\ simp\ add{\isacharcolon}Seq{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ X{\isacharunderscore}in{\isacharunderscore}eqs\ {\isachardoublequoteopen}{\isacharparenleft}{\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}eqs{\isacharunderscore}def\ init{\isacharunderscore}rhs{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\isacommand{next}\isamarkupfalse%
\isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}L\ xrhs\ {\isasymsubseteq}\ X{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ X{\isacharunderscore}in{\isacharunderscore}eqs\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}eqs{\isacharunderscore}def\ init{\isacharunderscore}rhs{\isacharunderscore}def{\isacharparenright}\ \isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ finite{\isacharunderscore}init{\isacharunderscore}rhs{\isacharcolon}\ \isanewline
\ \ \isakeyword{assumes}\ finite{\isacharcolon}\ {\isachardoublequoteopen}finite\ CS{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}finite\ {\isacharparenleft}init{\isacharunderscore}rhs\ CS\ X{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharbraceleft}Trn\ Y\ {\isacharparenleft}CHAR\ c{\isacharparenright}\ {\isacharbar}Y\ c{\isachardot}\ Y\ {\isasymin}\ CS\ {\isasymand}\ Y\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}\ {\isasymsubseteq}\ X{\isacharbraceright}{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}finite\ {\isacharquery}A{\isachardoublequoteclose}{\isacharparenright}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \ \ \isacommand{def}\isamarkupfalse%
\ S\ {\isasymequiv}\ {\isachardoublequoteopen}{\isacharbraceleft}{\isacharparenleft}Y{\isacharcomma}\ c{\isacharparenright}{\isacharbar}\ Y\ c{\isachardot}\ Y\ {\isasymin}\ CS\ {\isasymand}\ Y\ {\isacharsemicolon}{\isacharsemicolon}\ {\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}\ {\isasymsubseteq}\ X{\isacharbraceright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{def}\isamarkupfalse%
\ h\ {\isasymequiv}\ {\isachardoublequoteopen}{\isasymlambda}\ {\isacharparenleft}Y{\isacharcomma}\ c{\isacharparenright}{\isachardot}\ Trn\ Y\ {\isacharparenleft}CHAR\ c{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharparenleft}CS\ {\isasymtimes}\ {\isacharparenleft}UNIV{\isacharcolon}{\isacharcolon}char\ set{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ finite\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \ \ \isacommand{hence}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ S{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ S{\isacharunderscore}def\ \isanewline
\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ B\ {\isacharequal}\ {\isachardoublequoteopen}CS\ {\isasymtimes}\ UNIV{\isachardoublequoteclose}\ \isakeyword{in}\ finite{\isacharunderscore}subset{\isacharcomma}\ auto{\isacharparenright}\isanewline
\ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharquery}A\ {\isacharequal}\ h\ {\isacharbackquote}\ S{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}\ S{\isacharunderscore}def\ h{\isacharunderscore}def\ image{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isanewline
\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}init{\isacharunderscore}rhs{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ init{\isacharunderscore}ES{\isacharunderscore}satisfy{\isacharunderscore}Inv{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ finite{\isacharunderscore}CS{\isacharcolon}\ {\isachardoublequoteopen}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}Inv\ {\isacharparenleft}eqs\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharparenleft}eqs\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ finite{\isacharunderscore}CS\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}eqs{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}distinct{\isacharunderscore}equas\ {\isacharparenleft}eqs\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\ \ \ \ \ \isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}distinct{\isacharunderscore}equas{\isacharunderscore}def\ eqs{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}ardenable\ {\isacharparenleft}eqs\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}ardenable{\isacharunderscore}def\ eqs{\isacharunderscore}def\ init{\isacharunderscore}rhs{\isacharunderscore}def\ rhs{\isacharunderscore}nonempty{\isacharunderscore}def\ del{\isacharcolon}L{\isacharunderscore}rhs{\isachardot}simps{\isacharparenright}\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}valid{\isacharunderscore}eqns\ {\isacharparenleft}eqs\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ l{\isacharunderscore}eq{\isacharunderscore}r{\isacharunderscore}in{\isacharunderscore}eqs\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}valid{\isacharunderscore}eqns{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}non{\isacharunderscore}empty\ {\isacharparenleft}eqs\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}non{\isacharunderscore}empty{\isacharunderscore}def\ eqs{\isacharunderscore}def\ quotient{\isacharunderscore}def\ Image{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite{\isacharunderscore}rhs\ {\isacharparenleft}eqs\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ finite{\isacharunderscore}init{\isacharunderscore}rhs{\isacharbrackleft}OF\ finite{\isacharunderscore}CS{\isacharbrackright}\ \isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}finite{\isacharunderscore}rhs{\isacharunderscore}def\ eqs{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}self{\isacharunderscore}contained\ {\isacharparenleft}eqs\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}self{\isacharunderscore}contained{\isacharunderscore}def\ eqs{\isacharunderscore}def\ init{\isacharunderscore}rhs{\isacharunderscore}def\ classes{\isacharunderscore}of{\isacharunderscore}def\ lefts{\isacharunderscore}of{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}Inv{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\begin{isamarkuptext}%
From this point until \isa{iteration{\isacharunderscore}step}, we are trying to prove 
  that there exists iteration steps which keep \isa{Inv{\isacharparenleft}ES{\isacharparenright}} while
  decreasing the size of \isa{ES} with every iteration.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{lemma}\isamarkupfalse%
\ arden{\isacharunderscore}variate{\isacharunderscore}keeps{\isacharunderscore}eq{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ l{\isacharunderscore}eq{\isacharunderscore}r{\isacharcolon}\ {\isachardoublequoteopen}X\ {\isacharequal}\ L\ rhs{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{and}\ not{\isacharunderscore}empty{\isacharcolon}\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}\ {\isasymnotin}\ L\ {\isacharparenleft}rexp{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{and}\ finite{\isacharcolon}\ {\isachardoublequoteopen}finite\ rhs{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}X\ {\isacharequal}\ L\ {\isacharparenleft}arden{\isacharunderscore}variate\ X\ rhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \isacommand{def}\isamarkupfalse%
\ A\ {\isasymequiv}\ {\isachardoublequoteopen}L\ {\isacharparenleft}rexp{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{def}\isamarkupfalse%
\ b\ {\isasymequiv}\ {\isachardoublequoteopen}rhs\ {\isacharminus}\ items{\isacharunderscore}of\ rhs\ X{\isachardoublequoteclose}\isanewline
\ \ \isacommand{def}\isamarkupfalse%
\ B\ {\isasymequiv}\ {\isachardoublequoteopen}L\ b{\isachardoublequoteclose}\ \isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}X\ {\isacharequal}\ B\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isasymstar}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}rhs\ {\isacharequal}\ items{\isacharunderscore}of\ rhs\ X\ {\isasymunion}\ b{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}b{\isacharunderscore}def\ items{\isacharunderscore}of{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{hence}\isamarkupfalse%
\ {\isachardoublequoteopen}L\ rhs\ {\isacharequal}\ L{\isacharparenleft}items{\isacharunderscore}of\ rhs\ X\ {\isasymunion}\ b{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \ \ \isacommand{hence}\isamarkupfalse%
\ {\isachardoublequoteopen}L\ rhs\ {\isacharequal}\ L{\isacharparenleft}items{\isacharunderscore}of\ rhs\ X{\isacharparenright}\ {\isasymunion}\ B{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ only{\isacharcolon}L{\isacharunderscore}rhs{\isacharunderscore}union{\isacharunderscore}distrib\ B{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{with}\isamarkupfalse%
\ lang{\isacharunderscore}of{\isacharunderscore}rexp{\isacharunderscore}of\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}L\ rhs\ {\isacharequal}\ X\ {\isacharsemicolon}{\isacharsemicolon}\ A\ {\isasymunion}\ B\ {\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ finite\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ only{\isacharcolon}B{\isacharunderscore}def\ b{\isacharunderscore}def\ A{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\isanewline
\ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ l{\isacharunderscore}eq{\isacharunderscore}r\ not{\isacharunderscore}empty\isanewline
\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}drule{\isacharunderscore}tac\ B\ {\isacharequal}\ B\ \isakeyword{and}\ X\ {\isacharequal}\ X\ \isakeyword{in}\ ardens{\isacharunderscore}revised{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}A{\isacharunderscore}def\ simp\ del{\isacharcolon}L{\isacharunderscore}rhs{\isachardot}simps{\isacharparenright}\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}L\ {\isacharparenleft}arden{\isacharunderscore}variate\ X\ rhs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}B\ {\isacharsemicolon}{\isacharsemicolon}\ A{\isasymstar}{\isacharparenright}{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}{\isacharquery}L\ {\isacharequal}\ {\isacharquery}R{\isachardoublequoteclose}{\isacharparenright}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ only{\isacharcolon}arden{\isacharunderscore}variate{\isacharunderscore}def\ L{\isacharunderscore}rhs{\isacharunderscore}union{\isacharunderscore}distrib\ lang{\isacharunderscore}of{\isacharunderscore}append{\isacharunderscore}rhs\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ B{\isacharunderscore}def\ A{\isacharunderscore}def\ b{\isacharunderscore}def\ L{\isacharunderscore}rexp{\isachardot}simps\ seq{\isacharunderscore}union{\isacharunderscore}distrib{\isacharparenright}\isanewline
\ \ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\ \isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ append{\isacharunderscore}keeps{\isacharunderscore}finite{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}finite\ rhs\ {\isasymLongrightarrow}\ finite\ {\isacharparenleft}append{\isacharunderscore}rhs{\isacharunderscore}rexp\ rhs\ r{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}append{\isacharunderscore}rhs{\isacharunderscore}rexp{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ arden{\isacharunderscore}variate{\isacharunderscore}keeps{\isacharunderscore}finite{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}finite\ rhs\ {\isasymLongrightarrow}\ finite\ {\isacharparenleft}arden{\isacharunderscore}variate\ X\ rhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}arden{\isacharunderscore}variate{\isacharunderscore}def\ append{\isacharunderscore}keeps{\isacharunderscore}finite{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ append{\isacharunderscore}keeps{\isacharunderscore}nonempty{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}rhs{\isacharunderscore}nonempty\ rhs\ {\isasymLongrightarrow}\ rhs{\isacharunderscore}nonempty\ {\isacharparenleft}append{\isacharunderscore}rhs{\isacharunderscore}rexp\ rhs\ r{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}rhs{\isacharunderscore}nonempty{\isacharunderscore}def\ append{\isacharunderscore}rhs{\isacharunderscore}rexp{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}case{\isacharunderscore}tac\ x{\isacharcomma}\ auto\ simp{\isacharcolon}Seq{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ nonempty{\isacharunderscore}set{\isacharunderscore}sub{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}rhs{\isacharunderscore}nonempty\ rhs\ {\isasymLongrightarrow}\ rhs{\isacharunderscore}nonempty\ {\isacharparenleft}rhs\ {\isacharminus}\ A{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}rhs{\isacharunderscore}nonempty{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ nonempty{\isacharunderscore}set{\isacharunderscore}union{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}{\isasymlbrakk}rhs{\isacharunderscore}nonempty\ rhs{\isacharsemicolon}\ rhs{\isacharunderscore}nonempty\ rhs{\isacharprime}{\isasymrbrakk}\ {\isasymLongrightarrow}\ rhs{\isacharunderscore}nonempty\ {\isacharparenleft}rhs\ {\isasymunion}\ rhs{\isacharprime}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}rhs{\isacharunderscore}nonempty{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ arden{\isacharunderscore}variate{\isacharunderscore}keeps{\isacharunderscore}nonempty{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}rhs{\isacharunderscore}nonempty\ rhs\ {\isasymLongrightarrow}\ rhs{\isacharunderscore}nonempty\ {\isacharparenleft}arden{\isacharunderscore}variate\ X\ rhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ only{\isacharcolon}arden{\isacharunderscore}variate{\isacharunderscore}def\ append{\isacharunderscore}keeps{\isacharunderscore}nonempty\ nonempty{\isacharunderscore}set{\isacharunderscore}sub{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ rhs{\isacharunderscore}subst{\isacharunderscore}keeps{\isacharunderscore}nonempty{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}{\isasymlbrakk}rhs{\isacharunderscore}nonempty\ rhs{\isacharsemicolon}\ rhs{\isacharunderscore}nonempty\ xrhs{\isasymrbrakk}\ {\isasymLongrightarrow}\ rhs{\isacharunderscore}nonempty\ {\isacharparenleft}rhs{\isacharunderscore}subst\ rhs\ X\ xrhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ only{\isacharcolon}rhs{\isacharunderscore}subst{\isacharunderscore}def\ append{\isacharunderscore}keeps{\isacharunderscore}nonempty\ \ nonempty{\isacharunderscore}set{\isacharunderscore}union\ nonempty{\isacharunderscore}set{\isacharunderscore}sub{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ rhs{\isacharunderscore}subst{\isacharunderscore}keeps{\isacharunderscore}eq{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ substor{\isacharcolon}\ {\isachardoublequoteopen}X\ {\isacharequal}\ L\ xrhs{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{and}\ finite{\isacharcolon}\ {\isachardoublequoteopen}finite\ rhs{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}L\ {\isacharparenleft}rhs{\isacharunderscore}subst\ rhs\ X\ xrhs{\isacharparenright}\ {\isacharequal}\ L\ rhs{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}{\isacharquery}Left\ {\isacharequal}\ {\isacharquery}Right{\isachardoublequoteclose}{\isacharparenright}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \isacommand{def}\isamarkupfalse%
\ A\ {\isasymequiv}\ {\isachardoublequoteopen}L\ {\isacharparenleft}rhs\ {\isacharminus}\ items{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharquery}Left\ {\isacharequal}\ A\ {\isasymunion}\ L\ {\isacharparenleft}append{\isacharunderscore}rhs{\isacharunderscore}rexp\ xrhs\ {\isacharparenleft}rexp{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ only{\isacharcolon}rhs{\isacharunderscore}subst{\isacharunderscore}def\ L{\isacharunderscore}rhs{\isacharunderscore}union{\isacharunderscore}distrib\ A{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharquery}Right\ {\isacharequal}\ A\ {\isasymunion}\ L\ {\isacharparenleft}items{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}rhs\ {\isacharequal}\ {\isacharparenleft}rhs\ {\isacharminus}\ items{\isacharunderscore}of\ rhs\ X{\isacharparenright}\ {\isasymunion}\ {\isacharparenleft}items{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}items{\isacharunderscore}of{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ only{\isacharcolon}L{\isacharunderscore}rhs{\isacharunderscore}union{\isacharunderscore}distrib\ A{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}L\ {\isacharparenleft}append{\isacharunderscore}rhs{\isacharunderscore}rexp\ xrhs\ {\isacharparenleft}rexp{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ L\ {\isacharparenleft}items{\isacharunderscore}of\ rhs\ X{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ finite\ substor\ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ only{\isacharcolon}lang{\isacharunderscore}of{\isacharunderscore}append{\isacharunderscore}rhs\ lang{\isacharunderscore}of{\isacharunderscore}rexp{\isacharunderscore}of{\isacharparenright}\isanewline
\ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ rhs{\isacharunderscore}subst{\isacharunderscore}keeps{\isacharunderscore}finite{\isacharunderscore}rhs{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}{\isasymlbrakk}finite\ rhs{\isacharsemicolon}\ finite\ yrhs{\isasymrbrakk}\ {\isasymLongrightarrow}\ finite\ {\isacharparenleft}rhs{\isacharunderscore}subst\ rhs\ Y\ yrhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}rhs{\isacharunderscore}subst{\isacharunderscore}def\ append{\isacharunderscore}keeps{\isacharunderscore}finite{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ eqs{\isacharunderscore}subst{\isacharunderscore}keeps{\isacharunderscore}finite{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ finite{\isacharcolon}{\isachardoublequoteopen}finite\ {\isacharparenleft}ES{\isacharcolon}{\isacharcolon}\ {\isacharparenleft}string\ set\ {\isasymtimes}\ rhs{\isacharunderscore}item\ set{\isacharparenright}\ set{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}finite\ {\isacharparenleft}eqs{\isacharunderscore}subst\ ES\ Y\ yrhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharbraceleft}{\isacharparenleft}Ya{\isacharcomma}\ rhs{\isacharunderscore}subst\ yrhsa\ Y\ yrhs{\isacharparenright}\ {\isacharbar}Ya\ yrhsa{\isachardot}\ {\isacharparenleft}Ya{\isacharcomma}\ yrhsa{\isacharparenright}\ {\isasymin}\ ES{\isacharbraceright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}finite\ {\isacharquery}A{\isachardoublequoteclose}{\isacharparenright}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \isacommand{def}\isamarkupfalse%
\ eqns{\isacharprime}\ {\isasymequiv}\ {\isachardoublequoteopen}{\isacharbraceleft}{\isacharparenleft}{\isacharparenleft}Ya{\isacharcolon}{\isacharcolon}string\ set{\isacharparenright}{\isacharcomma}\ yrhsa{\isacharparenright}{\isacharbar}\ Ya\ yrhsa{\isachardot}\ {\isacharparenleft}Ya{\isacharcomma}\ yrhsa{\isacharparenright}\ {\isasymin}\ ES{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{def}\isamarkupfalse%
\ h\ {\isasymequiv}\ {\isachardoublequoteopen}{\isasymlambda}\ {\isacharparenleft}{\isacharparenleft}Ya{\isacharcolon}{\isacharcolon}string\ set{\isacharparenright}{\isacharcomma}\ yrhsa{\isacharparenright}{\isachardot}\ {\isacharparenleft}Ya{\isacharcomma}\ rhs{\isacharunderscore}subst\ yrhsa\ Y\ yrhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharparenleft}h\ {\isacharbackquote}\ eqns{\isacharprime}{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ finite\ h{\isacharunderscore}def\ eqns{\isacharprime}{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharquery}A\ {\isacharequal}\ h\ {\isacharbackquote}\ eqns{\isacharprime}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}h{\isacharunderscore}def\ eqns{\isacharprime}{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ auto\ \ \ \ \ \ \isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}eqs{\isacharunderscore}subst{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ eqs{\isacharunderscore}subst{\isacharunderscore}keeps{\isacharunderscore}finite{\isacharunderscore}rhs{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}{\isasymlbrakk}finite{\isacharunderscore}rhs\ ES{\isacharsemicolon}\ finite\ yrhs{\isasymrbrakk}\ {\isasymLongrightarrow}\ finite{\isacharunderscore}rhs\ {\isacharparenleft}eqs{\isacharunderscore}subst\ ES\ Y\ yrhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ intro{\isacharcolon}rhs{\isacharunderscore}subst{\isacharunderscore}keeps{\isacharunderscore}finite{\isacharunderscore}rhs\ simp\ add{\isacharcolon}eqs{\isacharunderscore}subst{\isacharunderscore}def\ finite{\isacharunderscore}rhs{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ append{\isacharunderscore}rhs{\isacharunderscore}keeps{\isacharunderscore}cls{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}classes{\isacharunderscore}of\ {\isacharparenleft}append{\isacharunderscore}rhs{\isacharunderscore}rexp\ rhs\ r{\isacharparenright}\ {\isacharequal}\ classes{\isacharunderscore}of\ rhs{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}classes{\isacharunderscore}of{\isacharunderscore}def\ append{\isacharunderscore}rhs{\isacharunderscore}rexp{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}case{\isacharunderscore}tac\ xa{\isacharcomma}\ auto\ simp{\isacharcolon}image{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}SEQ\ ra\ r{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharcomma}\ rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}Trn\ x\ ra{\isachardoublequoteclose}\ \isakeyword{in}\ bexI{\isacharcomma}\ simp{\isacharplus}{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ arden{\isacharunderscore}variate{\isacharunderscore}removes{\isacharunderscore}cl{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}classes{\isacharunderscore}of\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}\ {\isacharequal}\ classes{\isacharunderscore}of\ yrhs\ {\isacharminus}\ {\isacharbraceleft}Y{\isacharbraceright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}arden{\isacharunderscore}variate{\isacharunderscore}def\ append{\isacharunderscore}rhs{\isacharunderscore}keeps{\isacharunderscore}cls\ items{\isacharunderscore}of{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}classes{\isacharunderscore}of{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ lefts{\isacharunderscore}of{\isacharunderscore}keeps{\isacharunderscore}cls{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}lefts{\isacharunderscore}of\ {\isacharparenleft}eqs{\isacharunderscore}subst\ ES\ Y\ yrhs{\isacharparenright}\ {\isacharequal}\ lefts{\isacharunderscore}of\ ES{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}lefts{\isacharunderscore}of{\isacharunderscore}def\ eqs{\isacharunderscore}subst{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ rhs{\isacharunderscore}subst{\isacharunderscore}updates{\isacharunderscore}cls{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}X\ {\isasymnotin}\ classes{\isacharunderscore}of\ xrhs\ {\isasymLongrightarrow}\ \isanewline
\ \ \ \ \ \ classes{\isacharunderscore}of\ {\isacharparenleft}rhs{\isacharunderscore}subst\ rhs\ X\ xrhs{\isacharparenright}\ {\isacharequal}\ classes{\isacharunderscore}of\ rhs\ {\isasymunion}\ classes{\isacharunderscore}of\ xrhs\ {\isacharminus}\ {\isacharbraceleft}X{\isacharbraceright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}simp\ only{\isacharcolon}rhs{\isacharunderscore}subst{\isacharunderscore}def\ append{\isacharunderscore}rhs{\isacharunderscore}keeps{\isacharunderscore}cls\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ classes{\isacharunderscore}of{\isacharunderscore}union{\isacharunderscore}distrib{\isacharbrackleft}THEN\ sym{\isacharbrackright}{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}classes{\isacharunderscore}of{\isacharunderscore}def\ items{\isacharunderscore}of{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ eqs{\isacharunderscore}subst{\isacharunderscore}keeps{\isacharunderscore}self{\isacharunderscore}contained{\isacharcolon}\isanewline
\ \ \isakeyword{fixes}\ Y\isanewline
\ \ \isakeyword{assumes}\ sc{\isacharcolon}\ {\isachardoublequoteopen}self{\isacharunderscore}contained\ {\isacharparenleft}ES\ {\isasymunion}\ {\isacharbraceleft}{\isacharparenleft}Y{\isacharcomma}\ yrhs{\isacharparenright}{\isacharbraceright}{\isacharparenright}{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}self{\isacharunderscore}contained\ {\isacharquery}A{\isachardoublequoteclose}{\isacharparenright}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}self{\isacharunderscore}contained\ {\isacharparenleft}eqs{\isacharunderscore}subst\ ES\ Y\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}self{\isacharunderscore}contained\ {\isacharquery}B{\isachardoublequoteclose}{\isacharparenright}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
\ \isacommand{fix}\isamarkupfalse%
\ X\ xrhs{\isacharprime}\isanewline
\ \ \ \ \isacommand{assume}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharprime}{\isacharparenright}\ {\isasymin}\ {\isacharquery}B{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{then}\isamarkupfalse%
\ \isacommand{obtain}\isamarkupfalse%
\ xrhs\ \isanewline
\ \ \ \ \ \ \isakeyword{where}\ xrhs{\isacharunderscore}xrhs{\isacharprime}{\isacharcolon}\ {\isachardoublequoteopen}xrhs{\isacharprime}\ {\isacharequal}\ rhs{\isacharunderscore}subst\ xrhs\ Y\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isakeyword{and}\ X{\isacharunderscore}in{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharparenright}\ {\isasymin}\ ES{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}eqs{\isacharunderscore}subst{\isacharunderscore}def{\isacharcomma}\ blast{\isacharparenright}\ \ \ \ \isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}classes{\isacharunderscore}of\ xrhs{\isacharprime}\ {\isasymsubseteq}\ lefts{\isacharunderscore}of\ {\isacharquery}B{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}lefts{\isacharunderscore}of\ {\isacharquery}B\ {\isacharequal}\ lefts{\isacharunderscore}of\ ES{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}lefts{\isacharunderscore}of{\isacharunderscore}def\ eqs{\isacharunderscore}subst{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}classes{\isacharunderscore}of\ xrhs{\isacharprime}\ {\isasymsubseteq}\ lefts{\isacharunderscore}of\ ES{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}classes{\isacharunderscore}of\ xrhs{\isacharprime}\ {\isasymsubseteq}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ classes{\isacharunderscore}of\ xrhs\ {\isasymunion}\ classes{\isacharunderscore}of\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}\ {\isacharminus}\ {\isacharbraceleft}Y{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}Y\ {\isasymnotin}\ classes{\isacharunderscore}of\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ arden{\isacharunderscore}variate{\isacharunderscore}removes{\isacharunderscore}cl\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ xrhs{\isacharunderscore}xrhs{\isacharprime}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}rhs{\isacharunderscore}subst{\isacharunderscore}updates{\isacharunderscore}cls{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}classes{\isacharunderscore}of\ xrhs\ {\isasymsubseteq}\ lefts{\isacharunderscore}of\ ES\ {\isasymunion}\ {\isacharbraceleft}Y{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ X{\isacharunderscore}in\ sc\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}simp\ only{\isacharcolon}self{\isacharunderscore}contained{\isacharunderscore}def\ lefts{\isacharunderscore}of{\isacharunderscore}union{\isacharunderscore}distrib{\isacharbrackleft}THEN\ sym{\isacharbrackright}{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}drule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{in}\ bspec{\isacharcomma}\ auto\ simp{\isacharcolon}lefts{\isacharunderscore}of{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}classes{\isacharunderscore}of\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}\ {\isasymsubseteq}\ lefts{\isacharunderscore}of\ ES\ {\isasymunion}\ {\isacharbraceleft}Y{\isacharbraceright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ sc\ \isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}arden{\isacharunderscore}variate{\isacharunderscore}removes{\isacharunderscore}cl\ self{\isacharunderscore}contained{\isacharunderscore}def\ lefts{\isacharunderscore}of{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
\ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ only{\isacharcolon}eqs{\isacharunderscore}subst{\isacharunderscore}def\ self{\isacharunderscore}contained{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ eqs{\isacharunderscore}subst{\isacharunderscore}satisfy{\isacharunderscore}Inv{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ Inv{\isacharunderscore}ES{\isacharcolon}\ {\isachardoublequoteopen}Inv\ {\isacharparenleft}ES\ {\isasymunion}\ {\isacharbraceleft}{\isacharparenleft}Y{\isacharcomma}\ yrhs{\isacharparenright}{\isacharbraceright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}Inv\ {\isacharparenleft}eqs{\isacharunderscore}subst\ ES\ Y\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\ \ \isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ finite{\isacharunderscore}yrhs{\isacharcolon}\ {\isachardoublequoteopen}finite\ yrhs{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Inv{\isacharunderscore}def\ finite{\isacharunderscore}rhs{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ nonempty{\isacharunderscore}yrhs{\isacharcolon}\ {\isachardoublequoteopen}rhs{\isacharunderscore}nonempty\ yrhs{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Inv{\isacharunderscore}def\ ardenable{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ Y{\isacharunderscore}eq{\isacharunderscore}yrhs{\isacharcolon}\ {\isachardoublequoteopen}Y\ {\isacharequal}\ L\ yrhs{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ only{\isacharcolon}Inv{\isacharunderscore}def\ valid{\isacharunderscore}eqns{\isacharunderscore}def{\isacharcomma}\ blast{\isacharparenright}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}distinct{\isacharunderscore}equas\ {\isacharparenleft}eqs{\isacharunderscore}subst\ ES\ Y\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}distinct{\isacharunderscore}equas{\isacharunderscore}def\ eqs{\isacharunderscore}subst{\isacharunderscore}def\ Inv{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharparenleft}eqs{\isacharunderscore}subst\ ES\ Y\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}Inv{\isacharunderscore}def\ eqs{\isacharunderscore}subst{\isacharunderscore}keeps{\isacharunderscore}finite{\isacharparenright}\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite{\isacharunderscore}rhs\ {\isacharparenleft}eqs{\isacharunderscore}subst\ ES\ Y\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite{\isacharunderscore}rhs\ ES{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ \isanewline
\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}Inv{\isacharunderscore}def\ finite{\isacharunderscore}rhs{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ yrhs{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Inv{\isacharunderscore}def\ finite{\isacharunderscore}rhs{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ arden{\isacharunderscore}variate{\isacharunderscore}keeps{\isacharunderscore}finite\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isanewline
\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}eqs{\isacharunderscore}subst{\isacharunderscore}keeps{\isacharunderscore}finite{\isacharunderscore}rhs{\isacharparenright}\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}ardenable\ {\isacharparenleft}eqs{\isacharunderscore}subst\ ES\ Y\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\ \isanewline
\ \ \ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
\ \isacommand{fix}\isamarkupfalse%
\ X\ rhs\isanewline
\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}X{\isacharcomma}\ rhs{\isacharparenright}\ {\isasymin}\ ES{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isacommand{hence}\isamarkupfalse%
\ {\isachardoublequoteopen}rhs{\isacharunderscore}nonempty\ rhs{\isachardoublequoteclose}\ \ \isacommand{using}\isamarkupfalse%
\ prems\ Inv{\isacharunderscore}ES\ \ \isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}Inv{\isacharunderscore}def\ ardenable{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
\ nonempty{\isacharunderscore}yrhs\ \isanewline
\ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}rhs{\isacharunderscore}nonempty\ {\isacharparenleft}rhs{\isacharunderscore}subst\ rhs\ Y\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}nonempty{\isacharunderscore}yrhs\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ rhs{\isacharunderscore}subst{\isacharunderscore}keeps{\isacharunderscore}nonempty\ arden{\isacharunderscore}variate{\isacharunderscore}keeps{\isacharunderscore}nonempty{\isacharparenright}\isanewline
\ \ \ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
\ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}ardenable{\isacharunderscore}def\ eqs{\isacharunderscore}subst{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}valid{\isacharunderscore}eqns\ {\isacharparenleft}eqs{\isacharunderscore}subst\ ES\ Y\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}Y\ {\isacharequal}\ L\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ Y{\isacharunderscore}eq{\isacharunderscore}yrhs\ Inv{\isacharunderscore}ES\ finite{\isacharunderscore}yrhs\ nonempty{\isacharunderscore}yrhs\ \ \ \ \ \ \isanewline
\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ arden{\isacharunderscore}variate{\isacharunderscore}keeps{\isacharunderscore}eq{\isacharcomma}\ {\isacharparenleft}simp\ add{\isacharcolon}rexp{\isacharunderscore}of{\isacharunderscore}empty{\isacharparenright}{\isacharplus}{\isacharparenright}\isanewline
\ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ \isanewline
\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}clarsimp\ simp\ add{\isacharcolon}valid{\isacharunderscore}eqns{\isacharunderscore}def\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ eqs{\isacharunderscore}subst{\isacharunderscore}def\ rhs{\isacharunderscore}subst{\isacharunderscore}keeps{\isacharunderscore}eq\ Inv{\isacharunderscore}def\ finite{\isacharunderscore}rhs{\isacharunderscore}def\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ simp\ del{\isacharcolon}L{\isacharunderscore}rhs{\isachardot}simps{\isacharparenright}\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ \isanewline
\ \ \ \ non{\isacharunderscore}empty{\isacharunderscore}subst{\isacharcolon}\ {\isachardoublequoteopen}non{\isacharunderscore}empty\ {\isacharparenleft}eqs{\isacharunderscore}subst\ ES\ Y\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Inv{\isacharunderscore}def\ non{\isacharunderscore}empty{\isacharunderscore}def\ eqs{\isacharunderscore}subst{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ self{\isacharunderscore}subst{\isacharcolon}\ {\isachardoublequoteopen}self{\isacharunderscore}contained\ {\isacharparenleft}eqs{\isacharunderscore}subst\ ES\ Y\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ eqs{\isacharunderscore}subst{\isacharunderscore}keeps{\isacharunderscore}self{\isacharunderscore}contained\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}Inv{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}Inv{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ eqs{\isacharunderscore}subst{\isacharunderscore}card{\isacharunderscore}le{\isacharcolon}\ \isanewline
\ \ \isakeyword{assumes}\ finite{\isacharcolon}\ {\isachardoublequoteopen}finite\ {\isacharparenleft}ES{\isacharcolon}{\isacharcolon}{\isacharparenleft}string\ set\ {\isasymtimes}\ rhs{\isacharunderscore}item\ set{\isacharparenright}\ set{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}card\ {\isacharparenleft}eqs{\isacharunderscore}subst\ ES\ Y\ yrhs{\isacharparenright}\ {\isacharless}{\isacharequal}\ card\ ES{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \isacommand{def}\isamarkupfalse%
\ f\ {\isasymequiv}\ {\isachardoublequoteopen}{\isasymlambda}\ x{\isachardot}\ {\isacharparenleft}{\isacharparenleft}fst\ x{\isacharparenright}{\isacharcolon}{\isacharcolon}string\ set{\isacharcomma}\ rhs{\isacharunderscore}subst\ {\isacharparenleft}snd\ x{\isacharparenright}\ Y\ yrhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}eqs{\isacharunderscore}subst\ ES\ Y\ yrhs\ {\isacharequal}\ f\ {\isacharbackquote}\ ES{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}eqs{\isacharunderscore}subst{\isacharunderscore}def\ f{\isacharunderscore}def\ image{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharparenleft}Ya{\isacharcomma}\ yrhsa{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{in}\ bexI{\isacharcomma}\ simp{\isacharplus}{\isacharparenright}\isanewline
\ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ finite\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ intro{\isacharcolon}card{\isacharunderscore}image{\isacharunderscore}le{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ eqs{\isacharunderscore}subst{\isacharunderscore}cls{\isacharunderscore}remains{\isacharcolon}\ \isanewline
\ \ {\isachardoublequoteopen}{\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharparenright}\ {\isasymin}\ ES\ {\isasymLongrightarrow}\ {\isasymexists}\ xrhs{\isacharprime}{\isachardot}\ {\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharprime}{\isacharparenright}\ {\isasymin}\ {\isacharparenleft}eqs{\isacharunderscore}subst\ ES\ Y\ yrhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}eqs{\isacharunderscore}subst{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ card{\isacharunderscore}noteq{\isacharunderscore}{\isadigit{1}}{\isacharunderscore}has{\isacharunderscore}more{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ card{\isacharcolon}{\isachardoublequoteopen}card\ S\ {\isasymnoteq}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{and}\ e{\isacharunderscore}in{\isacharcolon}\ {\isachardoublequoteopen}e\ {\isasymin}\ S{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{and}\ finite{\isacharcolon}\ {\isachardoublequoteopen}finite\ S{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{obtains}\ e{\isacharprime}\ \isakeyword{where}\ {\isachardoublequoteopen}e{\isacharprime}\ {\isasymin}\ S\ {\isasymand}\ e\ {\isasymnoteq}\ e{\isacharprime}{\isachardoublequoteclose}\ \isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}card\ {\isacharparenleft}S\ {\isacharminus}\ {\isacharbraceleft}e{\isacharbraceright}{\isacharparenright}\ {\isachargreater}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}card\ S\ {\isachargreater}\ {\isadigit{1}}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ card\ e{\isacharunderscore}in\ finite\ \ \isanewline
\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}case{\isacharunderscore}tac\ {\isachardoublequoteopen}card\ S{\isachardoublequoteclose}{\isacharcomma}\ auto{\isacharparenright}\ \isanewline
\ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ finite\ e{\isacharunderscore}in\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{hence}\isamarkupfalse%
\ {\isachardoublequoteopen}S\ {\isacharminus}\ {\isacharbraceleft}e{\isacharbraceright}\ {\isasymnoteq}\ {\isacharbraceleft}{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ finite\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ notI{\isacharcomma}\ simp{\isacharparenright}\isanewline
\ \ \isacommand{thus}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}{\isasymAnd}e{\isacharprime}{\isachardot}\ e{\isacharprime}\ {\isasymin}\ S\ {\isasymand}\ e\ {\isasymnoteq}\ e{\isacharprime}\ {\isasymLongrightarrow}\ thesis{\isacharparenright}\ {\isasymLongrightarrow}\ thesis{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ iteration{\isacharunderscore}step{\isacharcolon}\ \isanewline
\ \ \isakeyword{assumes}\ Inv{\isacharunderscore}ES{\isacharcolon}\ {\isachardoublequoteopen}Inv\ ES{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{and}\ \ \ \ X{\isacharunderscore}in{\isacharunderscore}ES{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharparenright}\ {\isasymin}\ ES{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{and}\ \ \ \ not{\isacharunderscore}T{\isacharcolon}\ {\isachardoublequoteopen}card\ ES\ {\isasymnoteq}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isasymexists}\ ES{\isacharprime}{\isachardot}\ {\isacharparenleft}Inv\ ES{\isacharprime}\ {\isasymand}\ {\isacharparenleft}{\isasymexists}\ xrhs{\isacharprime}{\isachardot}{\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharprime}{\isacharparenright}\ {\isasymin}\ ES{\isacharprime}{\isacharparenright}{\isacharparenright}\ {\isasymand}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}card\ ES{\isacharprime}{\isacharcomma}\ card\ ES{\isacharparenright}\ {\isasymin}\ less{\isacharunderscore}than{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}{\isasymexists}\ ES{\isacharprime}{\isachardot}\ {\isacharquery}P\ ES{\isacharprime}{\isachardoublequoteclose}{\isacharparenright}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ finite{\isacharunderscore}ES{\isacharcolon}\ {\isachardoublequoteopen}finite\ ES{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}Inv{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{then}\isamarkupfalse%
\ \isacommand{obtain}\isamarkupfalse%
\ Y\ yrhs\ \isanewline
\ \ \ \ \isakeyword{where}\ Y{\isacharunderscore}in{\isacharunderscore}ES{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}Y{\isacharcomma}\ yrhs{\isacharparenright}\ {\isasymin}\ ES{\isachardoublequoteclose}\ \isakeyword{and}\ not{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharparenright}\ {\isasymnoteq}\ {\isacharparenleft}Y{\isacharcomma}\ yrhs{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ not{\isacharunderscore}T\ X{\isacharunderscore}in{\isacharunderscore}ES\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}drule{\isacharunderscore}tac\ card{\isacharunderscore}noteq{\isacharunderscore}{\isadigit{1}}{\isacharunderscore}has{\isacharunderscore}more{\isacharcomma}\ auto{\isacharparenright}\isanewline
\ \ \isacommand{def}\isamarkupfalse%
\ ES{\isacharprime}\ {\isacharequal}{\isacharequal}\ {\isachardoublequoteopen}ES\ {\isacharminus}\ {\isacharbraceleft}{\isacharparenleft}Y{\isacharcomma}\ yrhs{\isacharparenright}{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{let}\isamarkupfalse%
\ {\isacharquery}ES{\isacharprime}{\isacharprime}\ {\isacharequal}\ {\isachardoublequoteopen}eqs{\isacharunderscore}subst\ ES{\isacharprime}\ Y\ {\isacharparenleft}arden{\isacharunderscore}variate\ Y\ yrhs{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharquery}P\ {\isacharquery}ES{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}Inv\ {\isacharquery}ES{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ Y{\isacharunderscore}in{\isacharunderscore}ES\ Inv{\isacharunderscore}ES\isanewline
\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ eqs{\isacharunderscore}subst{\isacharunderscore}satisfy{\isacharunderscore}Inv{\isacharcomma}\ simp\ add{\isacharcolon}ES{\isacharprime}{\isacharunderscore}def\ insert{\isacharunderscore}absorb{\isacharparenright}\isanewline
\ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymexists}xrhs{\isacharprime}{\isachardot}\ {\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharprime}{\isacharparenright}\ {\isasymin}\ {\isacharquery}ES{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ \ \isacommand{using}\isamarkupfalse%
\ not{\isacharunderscore}eq\ X{\isacharunderscore}in{\isacharunderscore}ES\isanewline
\ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ ES\ {\isacharequal}\ ES{\isacharprime}\ \isakeyword{in}\ eqs{\isacharunderscore}subst{\isacharunderscore}cls{\isacharunderscore}remains{\isacharcomma}\ auto\ simp\ add{\isacharcolon}ES{\isacharprime}{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}card\ {\isacharquery}ES{\isacharprime}{\isacharprime}{\isacharcomma}\ card\ ES{\isacharparenright}\ {\isasymin}\ less{\isacharunderscore}than{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ ES{\isacharprime}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ finite{\isacharunderscore}ES\ ES{\isacharprime}{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}card\ ES{\isacharprime}\ {\isacharless}\ card\ ES{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ finite{\isacharunderscore}ES\ Y{\isacharunderscore}in{\isacharunderscore}ES\isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}ES{\isacharprime}{\isacharunderscore}def\ card{\isacharunderscore}gt{\isacharunderscore}{\isadigit{0}}{\isacharunderscore}iff\ intro{\isacharcolon}diff{\isacharunderscore}Suc{\isacharunderscore}less{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ dest{\isacharcolon}eqs{\isacharunderscore}subst{\isacharunderscore}card{\isacharunderscore}le\ elim{\isacharcolon}le{\isacharunderscore}less{\isacharunderscore}trans{\isacharparenright}\isanewline
\ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\begin{isamarkuptext}%
From this point until \isa{hard{\isacharunderscore}direction}, the hard direction is proved
  through a simple application of the iteration principle.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{lemma}\isamarkupfalse%
\ iteration{\isacharunderscore}conc{\isacharcolon}\ \isanewline
\ \ \isakeyword{assumes}\ history{\isacharcolon}\ {\isachardoublequoteopen}Inv\ ES{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{and}\ \ \ \ X{\isacharunderscore}in{\isacharunderscore}ES{\isacharcolon}\ {\isachardoublequoteopen}{\isasymexists}\ xrhs{\isachardot}\ {\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharparenright}\ {\isasymin}\ ES{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ \isanewline
\ \ {\isachardoublequoteopen}{\isasymexists}\ ES{\isacharprime}{\isachardot}\ {\isacharparenleft}Inv\ ES{\isacharprime}\ {\isasymand}\ {\isacharparenleft}{\isasymexists}\ xrhs{\isacharprime}{\isachardot}\ {\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharprime}{\isacharparenright}\ {\isasymin}\ ES{\isacharprime}{\isacharparenright}{\isacharparenright}\ {\isasymand}\ card\ ES{\isacharprime}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}{\isasymexists}\ ES{\isacharprime}{\isachardot}\ {\isacharquery}P\ ES{\isacharprime}{\isachardoublequoteclose}{\isacharparenright}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ {\isacharparenleft}cases\ {\isachardoublequoteopen}card\ ES\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}{\isacharparenright}\isanewline
\ \ \isacommand{case}\isamarkupfalse%
\ True\isanewline
\ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ history\ X{\isacharunderscore}in{\isacharunderscore}ES\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\isacommand{next}\isamarkupfalse%
\isanewline
\ \ \isacommand{case}\isamarkupfalse%
\ False\ \ \isanewline
\ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ history\ iteration{\isacharunderscore}step\ X{\isacharunderscore}in{\isacharunderscore}ES\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ f\ {\isacharequal}\ card\ \isakeyword{in}\ wf{\isacharunderscore}iter{\isacharcomma}\ auto{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\ \ \isanewline
\isacommand{lemma}\isamarkupfalse%
\ last{\isacharunderscore}cl{\isacharunderscore}exists{\isacharunderscore}rexp{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ ES{\isacharunderscore}single{\isacharcolon}\ {\isachardoublequoteopen}ES\ {\isacharequal}\ {\isacharbraceleft}{\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharparenright}{\isacharbraceright}{\isachardoublequoteclose}\ \isanewline
\ \ \isakeyword{and}\ Inv{\isacharunderscore}ES{\isacharcolon}\ {\isachardoublequoteopen}Inv\ ES{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isasymexists}\ {\isacharparenleft}r{\isacharcolon}{\isacharcolon}rexp{\isacharparenright}{\isachardot}\ L\ r\ {\isacharequal}\ X{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}{\isasymexists}\ r{\isachardot}\ {\isacharquery}P\ r{\isachardoublequoteclose}{\isacharparenright}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \isacommand{let}\isamarkupfalse%
\ {\isacharquery}A\ {\isacharequal}\ {\isachardoublequoteopen}arden{\isacharunderscore}variate\ X\ xrhs{\isachardoublequoteclose}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharquery}P\ {\isacharparenleft}rexp{\isacharunderscore}of{\isacharunderscore}lam\ {\isacharquery}A{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}L\ {\isacharparenleft}rexp{\isacharunderscore}of{\isacharunderscore}lam\ {\isacharquery}A{\isacharparenright}\ {\isacharequal}\ L\ {\isacharparenleft}lam{\isacharunderscore}of\ {\isacharquery}A{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharparenleft}rule\ rexp{\isacharunderscore}of{\isacharunderscore}lam{\isacharunderscore}eq{\isacharunderscore}lam{\isacharunderscore}set{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharparenleft}arden{\isacharunderscore}variate\ X\ xrhs{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ ES{\isacharunderscore}single\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ arden{\isacharunderscore}variate{\isacharunderscore}keeps{\isacharunderscore}finite{\isacharcomma}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ auto\ simp\ add{\isacharcolon}Inv{\isacharunderscore}def\ finite{\isacharunderscore}rhs{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \isacommand{also}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ L\ {\isacharquery}A{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}lam{\isacharunderscore}of\ {\isacharquery}A\ {\isacharequal}\ {\isacharquery}A{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}classes{\isacharunderscore}of\ {\isacharquery}A\ {\isacharequal}\ {\isacharbraceleft}{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ ES{\isacharunderscore}single\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}arden{\isacharunderscore}variate{\isacharunderscore}removes{\isacharunderscore}cl\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ self{\isacharunderscore}contained{\isacharunderscore}def\ Inv{\isacharunderscore}def\ lefts{\isacharunderscore}of{\isacharunderscore}def{\isacharparenright}\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ only{\isacharcolon}lam{\isacharunderscore}of{\isacharunderscore}def\ classes{\isacharunderscore}of{\isacharunderscore}def{\isacharcomma}\ case{\isacharunderscore}tac\ x{\isacharcomma}\ auto{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \isacommand{also}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ X{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharparenleft}rule\ arden{\isacharunderscore}variate{\isacharunderscore}keeps{\isacharunderscore}eq\ {\isacharbrackleft}THEN\ sym{\isacharbrackright}{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}X\ {\isacharequal}\ L\ xrhs{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ ES{\isacharunderscore}single\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ only{\isacharcolon}Inv{\isacharunderscore}def\ valid{\isacharunderscore}eqns{\isacharunderscore}def{\isacharparenright}\ \ \isanewline
\ \ \ \ \isacommand{next}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \isacommand{from}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ ES{\isacharunderscore}single\ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}\ {\isasymnotin}\ L\ {\isacharparenleft}rexp{\isacharunderscore}of\ xrhs\ X{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
{\isacharparenleft}simp\ add{\isacharcolon}Inv{\isacharunderscore}def\ ardenable{\isacharunderscore}def\ rexp{\isacharunderscore}of{\isacharunderscore}empty\ finite{\isacharunderscore}rhs{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{next}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \isacommand{from}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\ ES{\isacharunderscore}single\ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ xrhs{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}Inv{\isacharunderscore}def\ finite{\isacharunderscore}rhs{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \isacommand{finally}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\ \ \ \isanewline
\isacommand{lemma}\isamarkupfalse%
\ every{\isacharunderscore}eqcl{\isacharunderscore}has{\isacharunderscore}reg{\isacharcolon}\ \isanewline
\ \ \isakeyword{assumes}\ finite{\isacharunderscore}CS{\isacharcolon}\ {\isachardoublequoteopen}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{and}\ X{\isacharunderscore}in{\isacharunderscore}CS{\isacharcolon}\ {\isachardoublequoteopen}X\ {\isasymin}\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isasymexists}\ {\isacharparenleft}reg{\isacharcolon}{\isacharcolon}rexp{\isacharparenright}{\isachardot}\ L\ reg\ {\isacharequal}\ X{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}{\isasymexists}\ r{\isachardot}\ {\isacharquery}E\ r{\isachardoublequoteclose}{\isacharparenright}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \isacommand{from}\isamarkupfalse%
\ X{\isacharunderscore}in{\isacharunderscore}CS\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymexists}\ xrhs{\isachardot}\ {\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharparenright}\ {\isasymin}\ {\isacharparenleft}eqs\ {\isacharparenleft}UNIV\ \ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}eqs{\isacharunderscore}def\ init{\isacharunderscore}rhs{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{then}\isamarkupfalse%
\ \isacommand{obtain}\isamarkupfalse%
\ ES\ xrhs\ \isakeyword{where}\ Inv{\isacharunderscore}ES{\isacharcolon}\ {\isachardoublequoteopen}Inv\ ES{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isakeyword{and}\ X{\isacharunderscore}in{\isacharunderscore}ES{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharparenright}\ {\isasymin}\ ES{\isachardoublequoteclose}\isanewline
\ \ \ \ \isakeyword{and}\ card{\isacharunderscore}ES{\isacharcolon}\ {\isachardoublequoteopen}card\ ES\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ finite{\isacharunderscore}CS\ X{\isacharunderscore}in{\isacharunderscore}CS\ init{\isacharunderscore}ES{\isacharunderscore}satisfy{\isacharunderscore}Inv\ iteration{\isacharunderscore}conc\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \isacommand{hence}\isamarkupfalse%
\ ES{\isacharunderscore}single{\isacharunderscore}equa{\isacharcolon}\ {\isachardoublequoteopen}ES\ {\isacharequal}\ {\isacharbraceleft}{\isacharparenleft}X{\isacharcomma}\ xrhs{\isacharparenright}{\isacharbraceright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Inv{\isacharunderscore}def\ dest{\isacharbang}{\isacharcolon}card{\isacharunderscore}Suc{\isacharunderscore}Diff{\isadigit{1}}\ simp{\isacharcolon}card{\isacharunderscore}eq{\isacharunderscore}{\isadigit{0}}{\isacharunderscore}iff{\isacharparenright}\ \isanewline
\ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ Inv{\isacharunderscore}ES\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule\ last{\isacharunderscore}cl{\isacharunderscore}exists{\isacharunderscore}rexp{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ finals{\isacharunderscore}in{\isacharunderscore}partitions{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}finals\ Lang\ {\isasymsubseteq}\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
\ \ %
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}finals{\isacharunderscore}def\ quotient{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\ \ \ \isanewline
%
\endisadelimproof
\isanewline
\isacommand{theorem}\isamarkupfalse%
\ hard{\isacharunderscore}direction{\isacharcolon}\ \isanewline
\ \ \isakeyword{assumes}\ finite{\isacharunderscore}CS{\isacharcolon}\ {\isachardoublequoteopen}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ \ \ {\isachardoublequoteopen}{\isasymexists}\ {\isacharparenleft}reg{\isacharcolon}{\isacharcolon}rexp{\isacharparenright}{\isachardot}\ Lang\ {\isacharequal}\ L\ reg{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymforall}\ X\ {\isasymin}\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isachardot}\ {\isasymexists}\ {\isacharparenleft}reg{\isacharcolon}{\isacharcolon}rexp{\isacharparenright}{\isachardot}\ X\ {\isacharequal}\ L\ reg{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ finite{\isacharunderscore}CS\ every{\isacharunderscore}eqcl{\isacharunderscore}has{\isacharunderscore}reg\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \isacommand{then}\isamarkupfalse%
\ \isacommand{obtain}\isamarkupfalse%
\ f\ \isanewline
\ \ \ \ \isakeyword{where}\ f{\isacharunderscore}prop{\isacharcolon}\ {\isachardoublequoteopen}{\isasymforall}\ X\ {\isasymin}\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isachardot}\ X\ {\isacharequal}\ L\ {\isacharparenleft}{\isacharparenleft}f\ X{\isacharparenright}{\isacharcolon}{\isacharcolon}rexp{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ dest{\isacharcolon}bchoice{\isacharparenright}\isanewline
\ \ \isacommand{def}\isamarkupfalse%
\ rs\ {\isasymequiv}\ {\isachardoublequoteopen}f\ {\isacharbackquote}\ {\isacharparenleft}finals\ Lang{\isacharparenright}{\isachardoublequoteclose}\ \ \isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}Lang\ {\isacharequal}\ {\isasymUnion}\ {\isacharparenleft}finals\ Lang{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ lang{\isacharunderscore}is{\isacharunderscore}union{\isacharunderscore}of{\isacharunderscore}finals\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \isacommand{also}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymdots}\ {\isacharequal}\ L\ {\isacharparenleft}folds\ ALT\ NULL\ rs{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ rs{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharparenleft}finals\ Lang{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ finite{\isacharunderscore}CS\ finals{\isacharunderscore}in{\isacharunderscore}partitions{\isacharbrackleft}of\ {\isachardoublequoteopen}Lang{\isachardoublequoteclose}{\isacharbrackright}\ \ \ \isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}erule{\isacharunderscore}tac\ finite{\isacharunderscore}subset{\isacharcomma}\ simp{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ rs{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isanewline
\ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ f{\isacharunderscore}prop\ rs{\isacharunderscore}def\ finals{\isacharunderscore}in{\isacharunderscore}partitions{\isacharbrackleft}of\ {\isachardoublequoteopen}Lang{\isachardoublequoteclose}{\isacharbrackright}\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{finally}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\isamarkupsection{Direction: \isa{regular\ language\ {\isasymRightarrow}finite\ partition}%
}
\isamarkuptrue%
%
\isamarkupsubsection{The scheme for this direction%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
The following convenient notation \isa{x\ {\isasymapprox}Lang\ y} means:
  string \isa{x} and \isa{y} are equivalent with respect to 
  language \isa{Lang}.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\isanewline
\ \ str{\isacharunderscore}eq\ {\isacharparenleft}{\isachardoublequoteopen}{\isacharunderscore}\ {\isasymapprox}{\isacharunderscore}\ {\isacharunderscore}{\isachardoublequoteclose}{\isacharparenright}\isanewline
\isakeyword{where}\isanewline
\ \ {\isachardoublequoteopen}x\ {\isasymapprox}Lang\ y\ {\isasymequiv}\ {\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isasymin}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isachardoublequoteclose}%
\begin{isamarkuptext}%
The very basic scheme to show the finiteness of the partion generated by a language \isa{Lang}
  is by attaching tags to every string. The set of tags are carfully choosen to make it finite.
  If it can be proved that strings with the same tag are equivlent with respect \isa{Lang},
  then the partition given rise by \isa{Lang} must be finite. The reason for this is a lemma 
  in standard library (\isa{finite{\isacharunderscore}imageD}), which says: if the image of an injective 
  function on a set \isa{A} is finite, then \isa{A} is finite. It can be shown that
  the function obtained by llifting \isa{tag}
  to the level of equalent classes (i.e. \isa{{\isacharparenleft}{\isacharparenleft}op\ {\isacharbackquote}{\isacharparenright}\ tag{\isacharparenright}}) is injective 
  (by lemma \isa{tag{\isacharunderscore}image{\isacharunderscore}injI}) and the image of this function is finite 
  (with the help of lemma \isa{finite{\isacharunderscore}tag{\isacharunderscore}imageI}).

  BUT, I think this argument can be encapsulated by one lemma instead of the current presentation.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{lemma}\isamarkupfalse%
\ eq{\isacharunderscore}class{\isacharunderscore}equalI{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}{\isasymlbrakk}X\ {\isasymin}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}lang{\isacharsemicolon}\ Y\ {\isasymin}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}lang{\isacharsemicolon}\ x\ {\isasymin}\ X{\isacharsemicolon}\ y\ {\isasymin}\ Y{\isacharsemicolon}\ x\ {\isasymapprox}lang\ y{\isasymrbrakk}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ X\ {\isacharequal}\ Y{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}quotient{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ tag{\isacharunderscore}image{\isacharunderscore}injI{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ str{\isacharunderscore}inj{\isacharcolon}\ {\isachardoublequoteopen}{\isasymAnd}\ x\ y{\isachardot}\ tag\ x\ {\isacharequal}\ tag\ {\isacharparenleft}y{\isacharcolon}{\isacharcolon}string{\isacharparenright}\ {\isasymLongrightarrow}\ x\ {\isasymapprox}lang\ y{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}inj{\isacharunderscore}on\ {\isacharparenleft}{\isacharparenleft}op\ {\isacharbackquote}{\isacharparenright}\ tag{\isacharparenright}\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}lang{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
\ \isacommand{fix}\isamarkupfalse%
\ X\ Y\isanewline
\ \ \ \ \isacommand{assume}\isamarkupfalse%
\ X{\isacharunderscore}in{\isacharcolon}\ {\isachardoublequoteopen}X\ {\isasymin}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}lang{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isakeyword{and}\ \ Y{\isacharunderscore}in{\isacharcolon}\ {\isachardoublequoteopen}Y\ {\isasymin}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}lang{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isakeyword{and}\ \ tag{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}tag\ {\isacharbackquote}\ X\ {\isacharequal}\ tag\ {\isacharbackquote}\ Y{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{then}\isamarkupfalse%
\ \isacommand{obtain}\isamarkupfalse%
\ x\ y\ \isakeyword{where}\ {\isachardoublequoteopen}x\ {\isasymin}\ X{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}y\ {\isasymin}\ Y{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}tag\ x\ {\isacharequal}\ tag\ y{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isacommand{unfolding}\isamarkupfalse%
\ quotient{\isacharunderscore}def\ Image{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}def\ image{\isacharunderscore}def\isanewline
\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
\ simp\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \ \ \isacommand{with}\isamarkupfalse%
\ X{\isacharunderscore}in\ Y{\isacharunderscore}in\ str{\isacharunderscore}inj\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}X\ {\isacharequal}\ Y{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ eq{\isacharunderscore}class{\isacharunderscore}equalI{\isacharcomma}\ simp{\isacharplus}{\isacharparenright}\isanewline
\ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
\isanewline
\ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{unfolding}\isamarkupfalse%
\ inj{\isacharunderscore}on{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ finite{\isacharunderscore}tag{\isacharunderscore}imageI{\isacharcolon}\ \isanewline
\ \ {\isachardoublequoteopen}finite\ {\isacharparenleft}range\ tag{\isacharparenright}\ {\isasymLongrightarrow}\ finite\ {\isacharparenleft}{\isacharparenleft}{\isacharparenleft}op\ {\isacharbackquote}{\isacharparenright}\ tag{\isacharparenright}\ {\isacharbackquote}\ S{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ B\ {\isacharequal}\ {\isachardoublequoteopen}Pow\ {\isacharparenleft}tag\ {\isacharbackquote}\ UNIV{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{in}\ finite{\isacharunderscore}subset{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}image{\isacharunderscore}def\ Pow{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\isamarkupsubsection{A small theory for list difference%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
The notion of list diffrence is need to make proofs more readable.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{fun}\isamarkupfalse%
\ list{\isacharunderscore}diff\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{infix}\ {\isachardoublequoteopen}{\isacharminus}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{1}}{\isacharparenright}\isanewline
\isakeyword{where}\isanewline
\ \ {\isachardoublequoteopen}list{\isacharunderscore}diff\ {\isacharbrackleft}{\isacharbrackright}\ \ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ {\isacharbar}\isanewline
\ \ {\isachardoublequoteopen}list{\isacharunderscore}diff\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ x{\isacharhash}xs{\isachardoublequoteclose}\ {\isacharbar}\isanewline
\ \ {\isachardoublequoteopen}list{\isacharunderscore}diff\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}y{\isacharhash}ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ x\ {\isacharequal}\ y\ then\ list{\isacharunderscore}diff\ xs\ ys\ else\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isacharat}\ y{\isacharparenright}\ {\isacharminus}\ x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}induct\ x{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}case{\isacharunderscore}tac\ y{\isacharcomma}\ simp{\isacharplus}{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isacharminus}\ x\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}induct\ x{\isacharcomma}\ auto{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isacharequal}\ xa\ {\isacharat}\ y\ {\isasymLongrightarrow}\ x\ {\isacharminus}\ xa\ {\isacharequal}\ y\ {\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}induct\ x{\isacharcomma}\ auto{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isacharminus}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}induct\ x{\isacharcomma}\ auto{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isacharminus}\ y\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}x\ {\isasymle}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharminus}\ \ \ \isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymexists}xa{\isachardot}\ x\ {\isacharequal}\ xa\ {\isacharat}\ {\isacharparenleft}x\ {\isacharminus}\ y{\isacharparenright}\ {\isasymand}\ xa\ {\isasymle}\ y{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}rule\ list{\isacharunderscore}diff{\isachardot}induct{\isacharbrackleft}of\ {\isacharunderscore}\ x\ y{\isacharbrackright}{\isacharcomma}\ simp{\isacharplus}{\isacharparenright}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}clarsimp{\isacharcomma}\ rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}y\ {\isacharhash}\ xa{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharcomma}\ simp{\isacharplus}{\isacharparenright}\isanewline
\ \ \isacommand{thus}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isacharminus}\ y\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}x\ {\isasymle}\ y{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ diff{\isacharunderscore}prefix{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}{\isasymlbrakk}c\ {\isasymle}\ a\ {\isacharminus}\ b{\isacharsemicolon}\ b\ {\isasymle}\ a{\isasymrbrakk}\ {\isasymLongrightarrow}\ b\ {\isacharat}\ c\ {\isasymle}\ a{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ elim{\isacharcolon}prefixE{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ diff{\isacharunderscore}diff{\isacharunderscore}appd{\isacharcolon}\ \isanewline
\ \ {\isachardoublequoteopen}{\isasymlbrakk}c\ {\isacharless}\ a\ {\isacharminus}\ b{\isacharsemicolon}\ b\ {\isacharless}\ a{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}a\ {\isacharminus}\ b{\isacharparenright}\ {\isacharminus}\ c\ {\isacharequal}\ a\ {\isacharminus}\ {\isacharparenleft}b\ {\isacharat}\ c{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}clarsimp\ simp{\isacharcolon}strict{\isacharunderscore}prefix{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}drule\ diff{\isacharunderscore}prefix{\isacharcomma}\ auto\ elim{\isacharcolon}prefixE{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ app{\isacharunderscore}eq{\isacharunderscore}cases{\isacharbrackleft}rule{\isacharunderscore}format{\isacharbrackright}{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}{\isasymforall}\ x\ {\isachardot}\ x\ {\isacharat}\ y\ {\isacharequal}\ m\ {\isacharat}\ n\ {\isasymlongrightarrow}\ {\isacharparenleft}x\ {\isasymle}\ m\ {\isasymor}\ m\ {\isasymle}\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}induct\ y{\isacharcomma}\ simp{\isacharparenright}\isanewline
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}clarify{\isacharcomma}\ drule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}x\ {\isacharat}\ {\isacharbrackleft}a{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{in}\ spec{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}clarsimp{\isacharcomma}\ auto\ simp{\isacharcolon}prefix{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ app{\isacharunderscore}eq{\isacharunderscore}dest{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}x\ {\isacharat}\ y\ {\isacharequal}\ m\ {\isacharat}\ n\ {\isasymLongrightarrow}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}x\ {\isasymle}\ m\ {\isasymand}\ {\isacharparenleft}m\ {\isacharminus}\ x{\isacharparenright}\ {\isacharat}\ n\ {\isacharequal}\ y{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}m\ {\isasymle}\ x\ {\isasymand}\ {\isacharparenleft}x\ {\isacharminus}\ m{\isacharparenright}\ {\isacharat}\ y\ {\isacharequal}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}frule{\isacharunderscore}tac\ app{\isacharunderscore}eq{\isacharunderscore}cases{\isacharcomma}\ auto\ elim{\isacharcolon}prefixE{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\isamarkupsubsection{Lemmas for basic cases%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
The the final result of this direction is in \isa{easier{\isacharunderscore}direction}, which
  is an induction on the structure of regular expressions. There is one case 
  for each regular expression operator. For basic operators such as \isa{NULL{\isacharcomma}\ EMPTY{\isacharcomma}\ CHAR\ c},
  the finiteness of their language partition can be established directly with no need
  of taggiing. This section contains several technical lemma for these base cases.

  The inductive cases involve operators \isa{ALT{\isacharcomma}\ SEQ} and \isa{STAR}. 
  Tagging functions need to be defined individually for each of them. There will be one
  dedicated section for each of these cases, and each section goes virtually the same way:
  gives definition of the tagging function and prove that strings 
  with the same tag are equivalent.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{lemma}\isamarkupfalse%
\ quot{\isacharunderscore}empty{\isacharunderscore}subset{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}{\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isacharparenright}\ {\isasymsubseteq}\ {\isacharbraceleft}{\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isacharcomma}\ UNIV\ {\isacharminus}\ {\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isacharbraceright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\isanewline
\ \ \isacommand{fix}\isamarkupfalse%
\ x\isanewline
\ \ \isacommand{assume}\isamarkupfalse%
\ {\isachardoublequoteopen}x\ {\isasymin}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{then}\isamarkupfalse%
\ \isacommand{obtain}\isamarkupfalse%
\ y\ \isakeyword{where}\ h{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isacharequal}\ {\isacharbraceleft}z{\isachardot}\ {\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ {\isasymapprox}{\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isacharbraceright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{unfolding}\isamarkupfalse%
\ quotient{\isacharunderscore}def\ Image{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isacharbraceleft}{\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isacharcomma}\ UNIV\ {\isacharminus}\ {\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isacharbraceright}{\isachardoublequoteclose}\ \isanewline
\ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharparenleft}cases\ {\isachardoublequoteopen}y\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharparenright}\isanewline
\ \ \ \ \isacommand{case}\isamarkupfalse%
\ True\ \isacommand{with}\isamarkupfalse%
\ h\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}x\ {\isacharequal}\ {\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \isacommand{next}\isamarkupfalse%
\isanewline
\ \ \ \ \isacommand{case}\isamarkupfalse%
\ False\ \isacommand{with}\isamarkupfalse%
\ h\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}x\ {\isacharequal}\ UNIV\ {\isacharminus}\ {\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ quot{\isacharunderscore}char{\isacharunderscore}subset{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}{\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}{\isacharparenright}\ {\isasymsubseteq}\ {\isacharbraceleft}{\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isacharcomma}{\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}{\isacharcomma}\ UNIV\ {\isacharminus}\ {\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharcomma}\ {\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}{\isacharbraceright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ \isanewline
\ \ \isacommand{fix}\isamarkupfalse%
\ x\ \isanewline
\ \ \isacommand{assume}\isamarkupfalse%
\ {\isachardoublequoteopen}x\ {\isasymin}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{then}\isamarkupfalse%
\ \isacommand{obtain}\isamarkupfalse%
\ y\ \isakeyword{where}\ h{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isacharequal}\ {\isacharbraceleft}z{\isachardot}\ {\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ {\isasymapprox}{\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}{\isacharbraceright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{unfolding}\isamarkupfalse%
\ quotient{\isacharunderscore}def\ Image{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}x\ {\isasymin}\ {\isacharbraceleft}{\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isacharcomma}{\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}{\isacharcomma}\ UNIV\ {\isacharminus}\ {\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharcomma}\ {\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
\ \isacommand{assume}\isamarkupfalse%
\ {\isachardoublequoteopen}y\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isacommand{hence}\isamarkupfalse%
\ {\isachardoublequoteopen}x\ {\isacharequal}\ {\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ h\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
\ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
\ {\isachardoublequoteopen}y\ {\isacharequal}\ {\isacharbrackleft}c{\isacharbrackright}{\isachardoublequoteclose}\ \isacommand{hence}\isamarkupfalse%
\ {\isachardoublequoteopen}x\ {\isacharequal}\ {\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ h\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ dest{\isacharbang}{\isacharcolon}spec{\isacharbrackleft}\isakeyword{where}\ x\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharbrackright}\ simp{\isacharcolon}str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
\ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
\ {\isachardoublequoteopen}y\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}y\ {\isasymnoteq}\ {\isacharbrackleft}c{\isacharbrackright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isacommand{hence}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymforall}\ z{\isachardot}\ {\isacharparenleft}y\ {\isacharat}\ z{\isacharparenright}\ {\isasymnoteq}\ {\isacharbrackleft}c{\isacharbrackright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}case{\isacharunderscore}tac\ y{\isacharcomma}\ auto{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymAnd}\ p{\isachardot}\ {\isacharparenleft}p\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymand}\ p\ {\isasymnoteq}\ {\isacharbrackleft}c{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isasymforall}\ q{\isachardot}\ p\ {\isacharat}\ q\ {\isasymnoteq}\ {\isacharbrackleft}c{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}case{\isacharunderscore}tac\ p{\isacharcomma}\ auto{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}x\ {\isacharequal}\ UNIV\ {\isacharminus}\ {\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharcomma}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ h\isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
\ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\isamarkupsubsection{The case for \isa{SEQ}%
}
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ {\isachardoublequoteopen}tag{\isacharunderscore}str{\isacharunderscore}SEQ\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}\ x\ {\isasymequiv}\ \isanewline
\ \ \ \ \ \ \ {\isacharparenleft}{\isacharparenleft}{\isasymapprox}L\isactrlisub {\isadigit{1}}{\isacharparenright}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}x{\isacharbraceright}{\isacharcomma}\ {\isacharbraceleft}{\isacharparenleft}{\isasymapprox}L\isactrlisub {\isadigit{2}}{\isacharparenright}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}x\ {\isacharminus}\ xa{\isacharbraceright}{\isacharbar}\ xa{\isachardot}\ \ xa\ {\isasymle}\ x\ {\isasymand}\ xa\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isacharbraceright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ tag{\isacharunderscore}str{\isacharunderscore}seq{\isacharunderscore}range{\isacharunderscore}finite{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}{\isasymlbrakk}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{1}}{\isacharparenright}{\isacharsemicolon}\ finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isasymrbrakk}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ finite\ {\isacharparenleft}range\ {\isacharparenleft}tag{\isacharunderscore}str{\isacharunderscore}SEQ\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ B\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{1}}{\isacharparenright}\ {\isasymtimes}\ {\isacharparenleft}Pow\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{in}\ finite{\isacharunderscore}subset{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}tag{\isacharunderscore}str{\isacharunderscore}SEQ{\isacharunderscore}def\ Image{\isacharunderscore}def\ quotient{\isacharunderscore}def\ split{\isacharcolon}if{\isacharunderscore}splits{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ append{\isacharunderscore}seq{\isacharunderscore}elim{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}x\ {\isacharat}\ y\ {\isasymin}\ L\isactrlisub {\isadigit{1}}\ {\isacharsemicolon}{\isacharsemicolon}\ L\isactrlisub {\isadigit{2}}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isacharparenleft}{\isasymexists}\ xa\ {\isasymle}\ x{\isachardot}\ xa\ {\isasymin}\ L\isactrlisub {\isadigit{1}}\ {\isasymand}\ {\isacharparenleft}x\ {\isacharminus}\ xa{\isacharparenright}\ {\isacharat}\ y\ {\isasymin}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}\ {\isasymor}\ \isanewline
\ \ \ \ \ \ \ \ \ \ {\isacharparenleft}{\isasymexists}\ ya\ {\isasymle}\ y{\isachardot}\ {\isacharparenleft}x\ {\isacharat}\ ya{\isacharparenright}\ {\isasymin}\ L\isactrlisub {\isadigit{1}}\ {\isasymand}\ {\isacharparenleft}y\ {\isacharminus}\ ya{\isacharparenright}\ {\isasymin}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \isacommand{from}\isamarkupfalse%
\ assms\ \isacommand{obtain}\isamarkupfalse%
\ s\isactrlisub {\isadigit{1}}\ s\isactrlisub {\isadigit{2}}\ \isanewline
\ \ \ \ \isakeyword{where}\ {\isachardoublequoteopen}x\ {\isacharat}\ y\ {\isacharequal}\ s\isactrlisub {\isadigit{1}}\ {\isacharat}\ s\isactrlisub {\isadigit{2}}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isakeyword{and}\ in{\isacharunderscore}seq{\isacharcolon}\ {\isachardoublequoteopen}s\isactrlisub {\isadigit{1}}\ {\isasymin}\ L\isactrlisub {\isadigit{1}}\ {\isasymand}\ s\isactrlisub {\isadigit{2}}\ {\isasymin}\ L\isactrlisub {\isadigit{2}}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Seq{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{hence}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymle}\ s\isactrlisub {\isadigit{1}}\ {\isasymand}\ {\isacharparenleft}s\isactrlisub {\isadigit{1}}\ {\isacharminus}\ x{\isacharparenright}\ {\isacharat}\ s\isactrlisub {\isadigit{2}}\ {\isacharequal}\ y{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}s\isactrlisub {\isadigit{1}}\ {\isasymle}\ x\ {\isasymand}\ {\isacharparenleft}x\ {\isacharminus}\ s\isactrlisub {\isadigit{1}}{\isacharparenright}\ {\isacharat}\ y\ {\isacharequal}\ s\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ app{\isacharunderscore}eq{\isacharunderscore}dest\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymlbrakk}x\ {\isasymle}\ s\isactrlisub {\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}s\isactrlisub {\isadigit{1}}\ {\isacharminus}\ x{\isacharparenright}\ {\isacharat}\ s\isactrlisub {\isadigit{2}}\ {\isacharequal}\ y{\isasymrbrakk}\ {\isasymLongrightarrow}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isasymexists}\ ya\ {\isasymle}\ y{\isachardot}\ {\isacharparenleft}x\ {\isacharat}\ ya{\isacharparenright}\ {\isasymin}\ L\isactrlisub {\isadigit{1}}\ {\isasymand}\ {\isacharparenleft}y\ {\isacharminus}\ ya{\isacharparenright}\ {\isasymin}\ L\isactrlisub {\isadigit{2}}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ in{\isacharunderscore}seq\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}s\isactrlisub {\isadigit{1}}\ {\isacharminus}\ x{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharcomma}\ auto\ elim{\isacharcolon}prefixE{\isacharparenright}\isanewline
\ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymlbrakk}s\isactrlisub {\isadigit{1}}\ {\isasymle}\ x{\isacharsemicolon}\ {\isacharparenleft}x\ {\isacharminus}\ s\isactrlisub {\isadigit{1}}{\isacharparenright}\ {\isacharat}\ y\ {\isacharequal}\ s\isactrlisub {\isadigit{2}}{\isasymrbrakk}\ {\isasymLongrightarrow}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isasymexists}\ xa\ {\isasymle}\ x{\isachardot}\ xa\ {\isasymin}\ L\isactrlisub {\isadigit{1}}\ {\isasymand}\ {\isacharparenleft}x\ {\isacharminus}\ xa{\isacharparenright}\ {\isacharat}\ y\ {\isasymin}\ L\isactrlisub {\isadigit{2}}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ in{\isacharunderscore}seq\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ s\isactrlisub {\isadigit{1}}\ \isakeyword{in}\ exI{\isacharcomma}\ auto{\isacharparenright}\isanewline
\ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ tag{\isacharunderscore}str{\isacharunderscore}SEQ{\isacharunderscore}injI{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}tag{\isacharunderscore}str{\isacharunderscore}SEQ\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}\ m\ {\isacharequal}\ tag{\isacharunderscore}str{\isacharunderscore}SEQ\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}\ n\ {\isasymLongrightarrow}\ m\ {\isasymapprox}{\isacharparenleft}L\isactrlisub {\isadigit{1}}\ {\isacharsemicolon}{\isacharsemicolon}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}\ n{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
\ \isacommand{fix}\isamarkupfalse%
\ x\ y\ z\isanewline
\ \ \ \ \isacommand{assume}\isamarkupfalse%
\ xz{\isacharunderscore}in{\isacharunderscore}seq{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isacharat}\ z\ {\isasymin}\ L\isactrlisub {\isadigit{1}}\ {\isacharsemicolon}{\isacharsemicolon}\ L\isactrlisub {\isadigit{2}}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isakeyword{and}\ tag{\isacharunderscore}xy{\isacharcolon}\ {\isachardoublequoteopen}tag{\isacharunderscore}str{\isacharunderscore}SEQ\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}\ x\ {\isacharequal}\ tag{\isacharunderscore}str{\isacharunderscore}SEQ\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}\ y{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
{\isachardoublequoteopen}y\ {\isacharat}\ z\ {\isasymin}\ L\isactrlisub {\isadigit{1}}\ {\isacharsemicolon}{\isacharsemicolon}\ L\isactrlisub {\isadigit{2}}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}{\isasymexists}\ xa\ {\isasymle}\ x{\isachardot}\ xa\ {\isasymin}\ L\isactrlisub {\isadigit{1}}\ {\isasymand}\ {\isacharparenleft}x\ {\isacharminus}\ xa{\isacharparenright}\ {\isacharat}\ z\ {\isasymin}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}\ {\isasymor}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}{\isasymexists}\ za\ {\isasymle}\ z{\isachardot}\ {\isacharparenleft}x\ {\isacharat}\ za{\isacharparenright}\ {\isasymin}\ L\isactrlisub {\isadigit{1}}\ {\isasymand}\ {\isacharparenleft}z\ {\isacharminus}\ za{\isacharparenright}\ {\isasymin}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ xz{\isacharunderscore}in{\isacharunderscore}seq\ append{\isacharunderscore}seq{\isacharunderscore}elim\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \ \ \isacommand{fix}\isamarkupfalse%
\ xa\isanewline
\ \ \ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
\ h{\isadigit{1}}{\isacharcolon}\ {\isachardoublequoteopen}xa\ {\isasymle}\ x{\isachardoublequoteclose}\ \isakeyword{and}\ h{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}xa\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isachardoublequoteclose}\ \isakeyword{and}\ h{\isadigit{3}}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isacharminus}\ xa{\isacharparenright}\ {\isacharat}\ z\ {\isasymin}\ L\isactrlisub {\isadigit{2}}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isacommand{obtain}\isamarkupfalse%
\ ya\ \isakeyword{where}\ {\isachardoublequoteopen}ya\ {\isasymle}\ y{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}ya\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}{\isacharparenleft}y\ {\isacharminus}\ ya{\isacharparenright}\ {\isacharat}\ z\ {\isasymin}\ L\isactrlisub {\isadigit{2}}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymexists}\ ya{\isachardot}\ \ ya\ {\isasymle}\ y\ {\isasymand}\ ya\ {\isasymin}\ L\isactrlisub {\isadigit{1}}\ {\isasymand}\ {\isacharparenleft}x\ {\isacharminus}\ xa{\isacharparenright}\ {\isasymapprox}L\isactrlisub {\isadigit{2}}\ {\isacharparenleft}y\ {\isacharminus}\ ya{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharbraceleft}{\isasymapprox}L\isactrlisub {\isadigit{2}}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}x\ {\isacharminus}\ xa{\isacharbraceright}\ {\isacharbar}xa{\isachardot}\ xa\ {\isasymle}\ x\ {\isasymand}\ xa\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isacharbraceright}\ {\isacharequal}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbraceleft}{\isasymapprox}L\isactrlisub {\isadigit{2}}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}y\ {\isacharminus}\ xa{\isacharbraceright}\ {\isacharbar}xa{\isachardot}\ xa\ {\isasymle}\ y\ {\isasymand}\ xa\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isacharbraceright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}{\isacharquery}Left\ {\isacharequal}\ {\isacharquery}Right{\isachardoublequoteclose}{\isacharparenright}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ h{\isadigit{1}}\ tag{\isacharunderscore}xy\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}tag{\isacharunderscore}str{\isacharunderscore}SEQ{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymapprox}L\isactrlisub {\isadigit{2}}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}x\ {\isacharminus}\ xa{\isacharbraceright}\ {\isasymin}\ {\isacharquery}Left{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ h{\isadigit{1}}\ h{\isadigit{2}}\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymapprox}L\isactrlisub {\isadigit{2}}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}x\ {\isacharminus}\ xa{\isacharbraceright}\ {\isasymin}\ {\isacharquery}Right{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Image{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
\ prems\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \ \ \isacommand{hence}\isamarkupfalse%
\ {\isachardoublequoteopen}y\ {\isacharat}\ z\ {\isasymin}\ L\isactrlisub {\isadigit{1}}\ {\isacharsemicolon}{\isacharsemicolon}\ L\isactrlisub {\isadigit{2}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}erule{\isacharunderscore}tac\ prefixE{\isacharcomma}\ auto\ simp{\isacharcolon}Seq{\isacharunderscore}def{\isacharparenright}\ \ \ \ \ \ \ \ \ \ \isanewline
\ \ \ \ \ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
\ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \ \ \isacommand{fix}\isamarkupfalse%
\ za\isanewline
\ \ \ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
\ h{\isadigit{1}}{\isacharcolon}\ {\isachardoublequoteopen}za\ {\isasymle}\ z{\isachardoublequoteclose}\ \isakeyword{and}\ h{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isacharat}\ za{\isacharparenright}\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isachardoublequoteclose}\ \isakeyword{and}\ h{\isadigit{3}}{\isacharcolon}\ {\isachardoublequoteopen}z\ {\isacharminus}\ za\ {\isasymin}\ L\isactrlisub {\isadigit{2}}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isacommand{hence}\isamarkupfalse%
\ {\isachardoublequoteopen}y\ {\isacharat}\ za\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymapprox}L\isactrlisub {\isadigit{1}}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}x{\isacharbraceright}\ {\isacharequal}\ {\isasymapprox}L\isactrlisub {\isadigit{1}}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}y{\isacharbraceright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ h{\isadigit{1}}\ tag{\isacharunderscore}xy\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}tag{\isacharunderscore}str{\isacharunderscore}SEQ{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
\ h{\isadigit{2}}\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}Image{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}def{\isacharparenright}\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
\ h{\isadigit{1}}\ h{\isadigit{3}}\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}y\ {\isacharat}\ z\ {\isasymin}\ L\isactrlisub {\isadigit{1}}\ {\isacharsemicolon}{\isacharsemicolon}\ L\isactrlisub {\isadigit{2}}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}drule{\isacharunderscore}tac\ A\ {\isacharequal}\ L\isactrlisub {\isadigit{1}}\ \isakeyword{in}\ seq{\isacharunderscore}intro{\isacharcomma}\ auto\ elim{\isacharcolon}prefixE{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
\ \isacommand{thus}\isamarkupfalse%
\ {\isachardoublequoteopen}tag{\isacharunderscore}str{\isacharunderscore}SEQ\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}\ m\ {\isacharequal}\ tag{\isacharunderscore}str{\isacharunderscore}SEQ\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}\ n\ {\isasymLongrightarrow}\ m\ {\isasymapprox}{\isacharparenleft}L\isactrlisub {\isadigit{1}}\ {\isacharsemicolon}{\isacharsemicolon}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}\ n{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ str{\isacharunderscore}eq{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\ \isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ quot{\isacharunderscore}seq{\isacharunderscore}finiteI{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ finite{\isadigit{1}}{\isacharcolon}\ {\isachardoublequoteopen}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharparenleft}L\isactrlisub {\isadigit{1}}{\isacharcolon}{\isacharcolon}string\ set{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{and}\ finite{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharparenleft}L\isactrlisub {\isadigit{1}}\ {\isacharsemicolon}{\isacharsemicolon}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharparenleft}rule{\isacharunderscore}tac\ f\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharparenleft}op\ {\isacharbackquote}{\isacharparenright}\ {\isacharparenleft}tag{\isacharunderscore}str{\isacharunderscore}SEQ\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{in}\ finite{\isacharunderscore}imageD{\isacharparenright}\isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharparenleft}op\ {\isacharbackquote}\ {\isacharparenleft}tag{\isacharunderscore}str{\isacharunderscore}SEQ\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}\ {\isacharbackquote}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{1}}\ {\isacharsemicolon}{\isacharsemicolon}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ finite{\isadigit{1}}\ finite{\isadigit{2}}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ intro{\isacharcolon}finite{\isacharunderscore}tag{\isacharunderscore}imageI\ tag{\isacharunderscore}str{\isacharunderscore}seq{\isacharunderscore}range{\isacharunderscore}finite{\isacharparenright}\isanewline
\isacommand{next}\isamarkupfalse%
\isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ \ {\isachardoublequoteopen}inj{\isacharunderscore}on\ {\isacharparenleft}op\ {\isacharbackquote}\ {\isacharparenleft}tag{\isacharunderscore}str{\isacharunderscore}SEQ\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isacharparenright}\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{1}}\ {\isacharsemicolon}{\isacharsemicolon}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}rule\ tag{\isacharunderscore}image{\isacharunderscore}injI{\isacharparenright}\isanewline
\ \ \ \ \isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}rule\ tag{\isacharunderscore}str{\isacharunderscore}SEQ{\isacharunderscore}injI{\isacharparenright}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ intro{\isacharcolon}tag{\isacharunderscore}image{\isacharunderscore}injI\ tag{\isacharunderscore}str{\isacharunderscore}SEQ{\isacharunderscore}injI\ simp{\isacharcolon}{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\isamarkupsubsection{The case for \isa{ALT}%
}
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ {\isachardoublequoteopen}tag{\isacharunderscore}str{\isacharunderscore}ALT\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}string{\isacharparenright}\ {\isasymequiv}\ {\isacharparenleft}{\isacharparenleft}{\isasymapprox}L\isactrlisub {\isadigit{1}}{\isacharparenright}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}x{\isacharbraceright}{\isacharcomma}\ {\isacharparenleft}{\isasymapprox}L\isactrlisub {\isadigit{2}}{\isacharparenright}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}x{\isacharbraceright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ tag{\isacharunderscore}str{\isacharunderscore}alt{\isacharunderscore}range{\isacharunderscore}finite{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}{\isasymlbrakk}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{1}}{\isacharparenright}{\isacharsemicolon}\ finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isasymrbrakk}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ finite\ {\isacharparenleft}range\ {\isacharparenleft}tag{\isacharunderscore}str{\isacharunderscore}ALT\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ B\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{1}}{\isacharparenright}\ {\isasymtimes}\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{in}\ finite{\isacharunderscore}subset{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}tag{\isacharunderscore}str{\isacharunderscore}ALT{\isacharunderscore}def\ Image{\isacharunderscore}def\ quotient{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ quot{\isacharunderscore}union{\isacharunderscore}finiteI{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ finite{\isadigit{1}}{\isacharcolon}\ {\isachardoublequoteopen}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharparenleft}L\isactrlisub {\isadigit{1}}{\isacharcolon}{\isacharcolon}string\ set{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{and}\ finite{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharparenleft}L\isactrlisub {\isadigit{1}}\ {\isasymunion}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharparenleft}rule{\isacharunderscore}tac\ f\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharparenleft}op\ {\isacharbackquote}{\isacharparenright}\ {\isacharparenleft}tag{\isacharunderscore}str{\isacharunderscore}ALT\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{in}\ finite{\isacharunderscore}imageD{\isacharparenright}\isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharparenleft}op\ {\isacharbackquote}\ {\isacharparenleft}tag{\isacharunderscore}str{\isacharunderscore}ALT\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}\ {\isacharbackquote}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{1}}\ {\isasymunion}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{using}\isamarkupfalse%
\ finite{\isadigit{1}}\ finite{\isadigit{2}}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ intro{\isacharcolon}finite{\isacharunderscore}tag{\isacharunderscore}imageI\ tag{\isacharunderscore}str{\isacharunderscore}alt{\isacharunderscore}range{\isacharunderscore}finite{\isacharparenright}\isanewline
\isacommand{next}\isamarkupfalse%
\isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}inj{\isacharunderscore}on\ {\isacharparenleft}op\ {\isacharbackquote}\ {\isacharparenleft}tag{\isacharunderscore}str{\isacharunderscore}ALT\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isacharparenright}\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{1}}\ {\isasymunion}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymAnd}m\ n{\isachardot}\ tag{\isacharunderscore}str{\isacharunderscore}ALT\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}\ m\ {\isacharequal}\ tag{\isacharunderscore}str{\isacharunderscore}ALT\ L\isactrlisub {\isadigit{1}}\ L\isactrlisub {\isadigit{2}}\ n\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ m\ {\isasymapprox}{\isacharparenleft}L\isactrlisub {\isadigit{1}}\ {\isasymunion}\ L\isactrlisub {\isadigit{2}}{\isacharparenright}\ n{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isacommand{unfolding}\isamarkupfalse%
\ tag{\isacharunderscore}str{\isacharunderscore}ALT{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}def\ Image{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ intro{\isacharcolon}tag{\isacharunderscore}image{\isacharunderscore}injI{\isacharparenright}\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\isamarkupsubsection{The case for \isa{STAR}%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
This turned out to be the most tricky case.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{definition}\isamarkupfalse%
\ \isanewline
\ \ {\isachardoublequoteopen}tag{\isacharunderscore}str{\isacharunderscore}STAR\ L\isactrlisub {\isadigit{1}}\ x\ {\isasymequiv}\ {\isacharbraceleft}{\isacharparenleft}{\isasymapprox}L\isactrlisub {\isadigit{1}}{\isacharparenright}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}x\ {\isacharminus}\ xa{\isacharbraceright}\ {\isacharbar}\ xa{\isachardot}\ xa\ {\isacharless}\ x\ {\isasymand}\ xa\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ finite{\isacharunderscore}set{\isacharunderscore}has{\isacharunderscore}max{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}finite\ A{\isacharsemicolon}\ A\ {\isasymnoteq}\ {\isacharbraceleft}{\isacharbraceright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}{\isasymexists}\ max\ {\isasymin}\ A{\isachardot}\ {\isasymforall}\ a\ {\isasymin}\ A{\isachardot}\ f\ a\ {\isacharless}{\isacharequal}\ {\isacharparenleft}f\ max\ {\isacharcolon}{\isacharcolon}\ nat{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ {\isacharparenleft}induct\ rule{\isacharcolon}finite{\isachardot}induct{\isacharparenright}\isanewline
\ \ \isacommand{case}\isamarkupfalse%
\ emptyI\ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}case\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\isacommand{next}\isamarkupfalse%
\isanewline
\ \ \isacommand{case}\isamarkupfalse%
\ {\isacharparenleft}insertI\ A\ a{\isacharparenright}\isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}case\isanewline
\ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharparenleft}cases\ {\isachardoublequoteopen}A\ {\isacharequal}\ {\isacharbraceleft}{\isacharbraceright}{\isachardoublequoteclose}{\isacharparenright}\isanewline
\ \ \ \ \isacommand{case}\isamarkupfalse%
\ True\ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ a\ \isakeyword{in}\ bexI{\isacharcomma}\ auto{\isacharparenright}\isanewline
\ \ \isacommand{next}\isamarkupfalse%
\isanewline
\ \ \ \ \isacommand{case}\isamarkupfalse%
\ False\isanewline
\ \ \ \ \isacommand{with}\isamarkupfalse%
\ prems\ \isacommand{obtain}\isamarkupfalse%
\ max\ \isanewline
\ \ \ \ \ \ \isakeyword{where}\ h{\isadigit{1}}{\isacharcolon}\ {\isachardoublequoteopen}max\ {\isasymin}\ A{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \isakeyword{and}\ h{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymforall}a{\isasymin}A{\isachardot}\ f\ a\ {\isasymle}\ f\ max{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \ \ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\isanewline
\ \ \ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharparenleft}cases\ {\isachardoublequoteopen}f\ a\ {\isasymle}\ f\ max{\isachardoublequoteclose}{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
\ {\isachardoublequoteopen}f\ a\ {\isasymle}\ f\ max{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
\ h{\isadigit{1}}\ h{\isadigit{2}}\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ max\ \isakeyword{in}\ bexI{\isacharcomma}\ auto{\isacharparenright}\isanewline
\ \ \ \ \isacommand{next}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymnot}\ {\isacharparenleft}f\ a\ {\isasymle}\ f\ max{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ h{\isadigit{2}}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ a\ \isakeyword{in}\ bexI{\isacharcomma}\ auto{\isacharparenright}\isanewline
\ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ finite{\isacharunderscore}strict{\isacharunderscore}prefix{\isacharunderscore}set{\isacharcolon}\ {\isachardoublequoteopen}finite\ {\isacharbraceleft}xa{\isachardot}\ xa\ {\isacharless}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}string{\isacharparenright}{\isacharbraceright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}induct\ x\ rule{\isacharcolon}rev{\isacharunderscore}induct{\isacharcomma}\ simp{\isacharparenright}\isanewline
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}subgoal{\isacharunderscore}tac\ {\isachardoublequoteopen}{\isacharbraceleft}xa{\isachardot}\ xa\ {\isacharless}\ xs\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}{\isacharbraceright}\ {\isacharequal}\ {\isacharbraceleft}xa{\isachardot}\ xa\ {\isacharless}\ xs{\isacharbraceright}\ {\isasymunion}\ {\isacharbraceleft}xs{\isacharbraceright}{\isachardoublequoteclose}{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}strict{\isacharunderscore}prefix{\isacharunderscore}def{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ tag{\isacharunderscore}str{\isacharunderscore}star{\isacharunderscore}range{\isacharunderscore}finite{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{1}}{\isacharparenright}\ {\isasymLongrightarrow}\ finite\ {\isacharparenleft}range\ {\isacharparenleft}tag{\isacharunderscore}str{\isacharunderscore}STAR\ L\isactrlisub {\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{apply}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ B\ {\isacharequal}\ {\isachardoublequoteopen}Pow\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{in}\ finite{\isacharunderscore}subset{\isacharparenright}\isanewline
\isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}tag{\isacharunderscore}str{\isacharunderscore}STAR{\isacharunderscore}def\ Image{\isacharunderscore}def\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ quotient{\isacharunderscore}def\ split{\isacharcolon}if{\isacharunderscore}splits{\isacharparenright}%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ tag{\isacharunderscore}str{\isacharunderscore}STAR{\isacharunderscore}injI{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}tag{\isacharunderscore}str{\isacharunderscore}STAR\ L\isactrlisub {\isadigit{1}}\ m\ {\isacharequal}\ tag{\isacharunderscore}str{\isacharunderscore}STAR\ L\isactrlisub {\isadigit{1}}\ n\ {\isasymLongrightarrow}\ m\ {\isasymapprox}{\isacharparenleft}L\isactrlisub {\isadigit{1}}{\isasymstar}{\isacharparenright}\ n{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
\ \isacommand{fix}\isamarkupfalse%
\ x\ y\ z\isanewline
\ \ \ \ \isacommand{assume}\isamarkupfalse%
\ xz{\isacharunderscore}in{\isacharunderscore}star{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isacharat}\ z\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isakeyword{and}\ tag{\isacharunderscore}xy{\isacharcolon}\ {\isachardoublequoteopen}tag{\isacharunderscore}str{\isacharunderscore}STAR\ L\isactrlisub {\isadigit{1}}\ x\ {\isacharequal}\ tag{\isacharunderscore}str{\isacharunderscore}STAR\ L\isactrlisub {\isadigit{1}}\ y{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}y\ {\isacharat}\ z\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharparenleft}cases\ {\isachardoublequoteopen}x\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{case}\isamarkupfalse%
\ True\isanewline
\ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
\ tag{\isacharunderscore}xy\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}y\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}tag{\isacharunderscore}str{\isacharunderscore}STAR{\isacharunderscore}def\ strict{\isacharunderscore}prefix{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ xz{\isacharunderscore}in{\isacharunderscore}star\ True\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \ \ \isacommand{next}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \isacommand{case}\isamarkupfalse%
\ False\isanewline
\ \ \ \ \ \ \isacommand{obtain}\isamarkupfalse%
\ x{\isacharunderscore}max\ \isanewline
\ \ \ \ \ \ \ \ \isakeyword{where}\ h{\isadigit{1}}{\isacharcolon}\ {\isachardoublequoteopen}x{\isacharunderscore}max\ {\isacharless}\ x{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \isakeyword{and}\ h{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}x{\isacharunderscore}max\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \isakeyword{and}\ h{\isadigit{3}}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isacharminus}\ x{\isacharunderscore}max{\isacharparenright}\ {\isacharat}\ z\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \isakeyword{and}\ h{\isadigit{4}}{\isacharcolon}{\isachardoublequoteopen}{\isasymforall}\ xa\ {\isacharless}\ x{\isachardot}\ xa\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}\ {\isasymand}\ {\isacharparenleft}x\ {\isacharminus}\ xa{\isacharparenright}\ {\isacharat}\ z\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isasymlongrightarrow}\ length\ xa\ {\isasymle}\ length\ x{\isacharunderscore}max{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \ \ \ \ \isacommand{let}\isamarkupfalse%
\ {\isacharquery}S\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharbraceleft}xa{\isachardot}\ xa\ {\isacharless}\ x\ {\isasymand}\ xa\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}\ {\isasymand}\ {\isacharparenleft}x\ {\isacharminus}\ xa{\isacharparenright}\ {\isacharat}\ z\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharquery}S{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule{\isacharunderscore}tac\ B\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharbraceleft}xa{\isachardot}\ xa\ {\isacharless}\ x{\isacharbraceright}{\isachardoublequoteclose}\ \isakeyword{in}\ finite{\isacharunderscore}subset{\isacharcomma}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ auto\ simp{\isacharcolon}finite{\isacharunderscore}strict{\isacharunderscore}prefix{\isacharunderscore}set{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharquery}S\ {\isasymnoteq}\ {\isacharbraceleft}{\isacharbraceright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ False\ xz{\isacharunderscore}in{\isacharunderscore}star\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp{\isacharcomma}\ rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{in}\ exI{\isacharcomma}\ auto\ simp{\isacharcolon}strict{\isacharunderscore}prefix{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymexists}\ max\ {\isasymin}\ {\isacharquery}S{\isachardot}\ {\isasymforall}\ a\ {\isasymin}\ {\isacharquery}S{\isachardot}\ length\ a\ {\isasymle}\ length\ max{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ finite{\isacharunderscore}set{\isacharunderscore}has{\isacharunderscore}max\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
\ prems\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \isacommand{obtain}\isamarkupfalse%
\ ya\ \isanewline
\ \ \ \ \ \ \ \ \isakeyword{where}\ h{\isadigit{5}}{\isacharcolon}\ {\isachardoublequoteopen}ya\ {\isacharless}\ y{\isachardoublequoteclose}\ \isakeyword{and}\ h{\isadigit{6}}{\isacharcolon}\ {\isachardoublequoteopen}ya\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}{\isachardoublequoteclose}\ \isakeyword{and}\ h{\isadigit{7}}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isacharminus}\ x{\isacharunderscore}max{\isacharparenright}\ {\isasymapprox}L\isactrlisub {\isadigit{1}}\ {\isacharparenleft}y\ {\isacharminus}\ ya{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse%
\ tag{\isacharunderscore}xy\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharbraceleft}{\isasymapprox}L\isactrlisub {\isadigit{1}}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}x\ {\isacharminus}\ xa{\isacharbraceright}\ {\isacharbar}xa{\isachardot}\ xa\ {\isacharless}\ x\ {\isasymand}\ xa\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}{\isacharbraceright}\ {\isacharequal}\ \isanewline
\ \ \ \ \ \ \ \ \ \ {\isacharbraceleft}{\isasymapprox}L\isactrlisub {\isadigit{1}}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}y\ {\isacharminus}\ xa{\isacharbraceright}\ {\isacharbar}xa{\isachardot}\ xa\ {\isacharless}\ y\ {\isasymand}\ xa\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}{\isacharbraceright}{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}{\isacharquery}left\ {\isacharequal}\ {\isacharquery}right{\isachardoublequoteclose}{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}tag{\isacharunderscore}str{\isacharunderscore}STAR{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymapprox}L\isactrlisub {\isadigit{1}}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}x\ {\isacharminus}\ x{\isacharunderscore}max{\isacharbraceright}\ {\isasymin}\ {\isacharquery}left{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ h{\isadigit{1}}\ h{\isadigit{2}}\ \isacommand{by}\isamarkupfalse%
\ auto\isanewline
\ \ \ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymapprox}L\isactrlisub {\isadigit{1}}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharbraceleft}x\ {\isacharminus}\ x{\isacharunderscore}max{\isacharbraceright}\ {\isasymin}\ {\isacharquery}right{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
\ prems\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{apply}\isamarkupfalse%
\ \isanewline
\ \ \ \ \ \ \ \ \ \ {\isacharparenleft}simp\ add{\isacharcolon}Image{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}def{\isacharparenright}\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
\ \ \ \ \ \ \isanewline
\ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}y\ {\isacharminus}\ ya{\isacharparenright}\ {\isacharat}\ z\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
{\isacharminus}\isanewline
\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse%
\ h{\isadigit{3}}\ h{\isadigit{1}}\ \isacommand{obtain}\isamarkupfalse%
\ a\ b\ \isakeyword{where}\ a{\isacharunderscore}in{\isacharcolon}\ {\isachardoublequoteopen}a\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \isakeyword{and}\ a{\isacharunderscore}neq{\isacharcolon}\ {\isachardoublequoteopen}a\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{and}\ b{\isacharunderscore}in{\isacharcolon}\ {\isachardoublequoteopen}b\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \isakeyword{and}\ ab{\isacharunderscore}max{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isacharminus}\ x{\isacharunderscore}max{\isacharparenright}\ {\isacharat}\ z\ {\isacharequal}\ a\ {\isacharat}\ b{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}drule{\isacharunderscore}tac\ star{\isacharunderscore}decom{\isacharcomma}\ auto\ simp{\isacharcolon}strict{\isacharunderscore}prefix{\isacharunderscore}def\ elim{\isacharcolon}prefixE{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isacharminus}\ x{\isacharunderscore}max{\isacharparenright}\ {\isasymle}\ a\ {\isasymand}\ {\isacharparenleft}a\ {\isacharminus}\ {\isacharparenleft}x\ {\isacharminus}\ x{\isacharunderscore}max{\isacharparenright}{\isacharparenright}\ {\isacharat}\ b\ {\isacharequal}\ z{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharparenleft}x\ {\isacharminus}\ x{\isacharunderscore}max{\isacharparenright}\ {\isasymle}\ a\ {\isasymand}\ {\isacharparenleft}a\ {\isacharminus}\ {\isacharparenleft}x\ {\isacharminus}\ x{\isacharunderscore}max{\isacharparenright}{\isacharparenright}\ {\isacharat}\ b\ {\isacharequal}\ z{\isacharparenright}\ {\isasymor}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}a\ {\isacharless}\ {\isacharparenleft}x\ {\isacharminus}\ x{\isacharunderscore}max{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}{\isacharparenleft}x\ {\isacharminus}\ x{\isacharunderscore}max{\isacharparenright}\ {\isacharminus}\ a{\isacharparenright}\ {\isacharat}\ z\ {\isacharequal}\ b{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ app{\isacharunderscore}eq{\isacharunderscore}dest{\isacharbrackleft}OF\ ab{\isacharunderscore}max{\isacharbrackright}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}strict{\isacharunderscore}prefix{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
\ np{\isacharcolon}\ {\isachardoublequoteopen}a\ {\isacharless}\ {\isacharparenleft}x\ {\isacharminus}\ x{\isacharunderscore}max{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{and}\ b{\isacharunderscore}eqs{\isacharcolon}\ {\isachardoublequoteopen}\ {\isacharparenleft}{\isacharparenleft}x\ {\isacharminus}\ x{\isacharunderscore}max{\isacharparenright}\ {\isacharminus}\ a{\isacharparenright}\ {\isacharat}\ z\ {\isacharequal}\ b{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}False{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{let}\isamarkupfalse%
\ {\isacharquery}x{\isacharunderscore}max{\isacharprime}\ {\isacharequal}\ {\isachardoublequoteopen}x{\isacharunderscore}max\ {\isacharat}\ a{\isachardoublequoteclose}\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharquery}x{\isacharunderscore}max{\isacharprime}\ {\isacharless}\ x{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ np\ h{\isadigit{1}}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}clarsimp\ simp{\isacharcolon}strict{\isacharunderscore}prefix{\isacharunderscore}def\ diff{\isacharunderscore}prefix{\isacharparenright}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharquery}x{\isacharunderscore}max{\isacharprime}\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ a{\isacharunderscore}in\ h{\isadigit{2}}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}star{\isacharunderscore}intro{\isadigit{3}}{\isacharparenright}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isacharminus}\ {\isacharquery}x{\isacharunderscore}max{\isacharprime}{\isacharparenright}\ {\isacharat}\ z\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isasymstar}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ b{\isacharunderscore}eqs\ b{\isacharunderscore}in\ np\ h{\isadigit{1}}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}simp\ add{\isacharcolon}diff{\isacharunderscore}diff{\isacharunderscore}appd{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse%
\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymnot}\ {\isacharparenleft}length\ {\isacharquery}x{\isacharunderscore}max{\isacharprime}\ {\isasymle}\ length\ x{\isacharunderscore}max{\isacharparenright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ a{\isacharunderscore}neq\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse%
\ h{\isadigit{4}}\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
\ \isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
\ \isacommand{ultimately}\isamarkupfalse%
\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ blast\isanewline
\ \ \ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \ \ \isacommand{then}\isamarkupfalse%
\ \isacommand{obtain}\isamarkupfalse%
\ za\ \isakeyword{where}\ z{\isacharunderscore}decom{\isacharcolon}\ {\isachardoublequoteopen}z\ {\isacharequal}\ za\ {\isacharat}\ b{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \isakeyword{and}\ x{\isacharunderscore}za{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isacharminus}\ x{\isacharunderscore}max{\isacharparenright}\ {\isacharat}\ za\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
\ a{\isacharunderscore}in\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ elim{\isacharcolon}prefixE{\isacharparenright}\ \ \ \ \ \ \ \ \isanewline
\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse%
\ x{\isacharunderscore}za\ h{\isadigit{7}}\ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharparenleft}y\ {\isacharminus}\ ya{\isacharparenright}\ {\isacharat}\ za\ {\isasymin}\ L\isactrlisub {\isadigit{1}}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}str{\isacharunderscore}eq{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
\ z{\isacharunderscore}decom\ b{\isacharunderscore}in\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ dest{\isacharbang}{\isacharcolon}step{\isacharbrackleft}of\ {\isachardoublequoteopen}{\isacharparenleft}y\ {\isacharminus}\ ya{\isacharparenright}\ {\isacharat}\ za{\isachardoublequoteclose}{\isacharbrackright}{\isacharparenright}\isanewline
\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse%
\isanewline
\ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
\ h{\isadigit{5}}\ h{\isadigit{6}}\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}thesis\ \isanewline
\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}drule{\isacharunderscore}tac\ star{\isacharunderscore}intro{\isadigit{1}}{\isacharcomma}\ auto\ simp{\isacharcolon}strict{\isacharunderscore}prefix{\isacharunderscore}def\ elim{\isacharcolon}prefixE{\isacharparenright}\isanewline
\ \ \ \ \isacommand{qed}\isamarkupfalse%
\ \ \ \ \ \ \isanewline
\ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
\ \isacommand{thus}\isamarkupfalse%
\ {\isachardoublequoteopen}tag{\isacharunderscore}str{\isacharunderscore}STAR\ L\isactrlisub {\isadigit{1}}\ m\ {\isacharequal}\ tag{\isacharunderscore}str{\isacharunderscore}STAR\ L\isactrlisub {\isadigit{1}}\ n\ {\isasymLongrightarrow}\ m\ {\isasymapprox}{\isacharparenleft}L\isactrlisub {\isadigit{1}}{\isasymstar}{\isacharparenright}\ n{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}str{\isacharunderscore}eq{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\isanewline
%
\endisadelimproof
\isanewline
\isacommand{lemma}\isamarkupfalse%
\ quot{\isacharunderscore}star{\isacharunderscore}finiteI{\isacharcolon}\isanewline
\ \ \isakeyword{assumes}\ finite{\isacharcolon}\ {\isachardoublequoteopen}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharparenleft}L\isactrlisub {\isadigit{1}}{\isacharcolon}{\isacharcolon}string\ set{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \isakeyword{shows}\ {\isachardoublequoteopen}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharparenleft}L\isactrlisub {\isadigit{1}}{\isasymstar}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
{\isacharparenleft}rule{\isacharunderscore}tac\ f\ {\isacharequal}\ {\isachardoublequoteopen}{\isacharparenleft}op\ {\isacharbackquote}{\isacharparenright}\ {\isacharparenleft}tag{\isacharunderscore}str{\isacharunderscore}STAR\ L\isactrlisub {\isadigit{1}}{\isacharparenright}{\isachardoublequoteclose}\ \isakeyword{in}\ finite{\isacharunderscore}imageD{\isacharparenright}\isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharparenleft}op\ {\isacharbackquote}\ {\isacharparenleft}tag{\isacharunderscore}str{\isacharunderscore}STAR\ L\isactrlisub {\isadigit{1}}{\isacharparenright}\ {\isacharbackquote}\ UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{1}}{\isasymstar}{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse%
\ finite\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ intro{\isacharcolon}finite{\isacharunderscore}tag{\isacharunderscore}imageI\ tag{\isacharunderscore}str{\isacharunderscore}star{\isacharunderscore}range{\isacharunderscore}finite{\isacharparenright}\isanewline
\isacommand{next}\isamarkupfalse%
\isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}inj{\isacharunderscore}on\ {\isacharparenleft}op\ {\isacharbackquote}\ {\isacharparenleft}tag{\isacharunderscore}str{\isacharunderscore}STAR\ L\isactrlisub {\isadigit{1}}{\isacharparenright}{\isacharparenright}\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}L\isactrlisub {\isadigit{1}}{\isasymstar}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ intro{\isacharcolon}tag{\isacharunderscore}image{\isacharunderscore}injI\ tag{\isacharunderscore}str{\isacharunderscore}STAR{\isacharunderscore}injI{\isacharparenright}\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\isamarkupsubsection{The main lemma%
}
\isamarkuptrue%
\isacommand{lemma}\isamarkupfalse%
\ easier{\isacharunderscore}directio{\isasymnu}{\isacharcolon}\isanewline
\ \ {\isachardoublequoteopen}Lang\ {\isacharequal}\ L\ {\isacharparenleft}r{\isacharcolon}{\isacharcolon}rexp{\isacharparenright}\ {\isasymLongrightarrow}\ finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}Lang{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ {\isacharparenleft}induct\ arbitrary{\isacharcolon}Lang\ rule{\isacharcolon}rexp{\isachardot}induct{\isacharparenright}\isanewline
\ \ \isacommand{case}\isamarkupfalse%
\ NULL\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}{\isacharbraceleft}{\isacharbraceright}{\isacharparenright}\ {\isasymsubseteq}\ {\isacharbraceleft}UNIV{\isacharbraceright}\ {\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ simp{\isacharcolon}quotient{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}rel{\isacharunderscore}def\ str{\isacharunderscore}eq{\isacharunderscore}def{\isacharparenright}\isanewline
\ \ \isacommand{with}\isamarkupfalse%
\ prems\ \isacommand{show}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isacharquery}case{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ intro{\isacharcolon}finite{\isacharunderscore}subset{\isacharparenright}\isanewline
\isacommand{next}\isamarkupfalse%
\isanewline
\ \ \isacommand{case}\isamarkupfalse%
\ EMPTY\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}{\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isacharparenright}\ {\isasymsubseteq}\ {\isacharbraceleft}{\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isacharcomma}\ UNIV\ {\isacharminus}\ {\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isacharbraceright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule\ quot{\isacharunderscore}empty{\isacharunderscore}subset{\isacharparenright}\isanewline
\ \ \isacommand{with}\isamarkupfalse%
\ prems\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}case\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ intro{\isacharcolon}finite{\isacharunderscore}subset{\isacharparenright}\isanewline
\isacommand{next}\isamarkupfalse%
\isanewline
\ \ \isacommand{case}\isamarkupfalse%
\ {\isacharparenleft}CHAR\ c{\isacharparenright}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}UNIV\ {\isacharslash}{\isacharslash}\ {\isacharparenleft}{\isasymapprox}{\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}{\isacharparenright}\ {\isasymsubseteq}\ {\isacharbraceleft}{\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharbraceright}{\isacharcomma}{\isacharbraceleft}{\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}{\isacharcomma}\ UNIV\ {\isacharminus}\ {\isacharbraceleft}{\isacharbrackleft}{\isacharbrackright}{\isacharcomma}\ {\isacharbrackleft}c{\isacharbrackright}{\isacharbraceright}{\isacharbraceright}{\isachardoublequoteclose}\ \isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}rule\ quot{\isacharunderscore}char{\isacharunderscore}subset{\isacharparenright}\isanewline
\ \ \isacommand{with}\isamarkupfalse%
\ prems\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}case\ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}auto\ intro{\isacharcolon}finite{\isacharunderscore}subset{\isacharparenright}\isanewline
\isacommand{next}\isamarkupfalse%
\isanewline
\ \ \isacommand{case}\isamarkupfalse%
\ {\isacharparenleft}SEQ\ r\isactrlisub {\isadigit{1}}\ r\isactrlisub {\isadigit{2}}{\isacharparenright}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymlbrakk}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharparenleft}L\ r\isactrlisub {\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isacharsemicolon}\ finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharparenleft}L\ r\isactrlisub {\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isasymrbrakk}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharparenleft}L\ r\isactrlisub {\isadigit{1}}\ {\isacharsemicolon}{\isacharsemicolon}\ L\ r\isactrlisub {\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}erule\ quot{\isacharunderscore}seq{\isacharunderscore}finiteI{\isacharcomma}\ simp{\isacharparenright}\isanewline
\ \ \isacommand{with}\isamarkupfalse%
\ prems\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}case\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\isacommand{next}\isamarkupfalse%
\isanewline
\ \ \isacommand{case}\isamarkupfalse%
\ {\isacharparenleft}ALT\ r\isactrlisub {\isadigit{1}}\ r\isactrlisub {\isadigit{2}}{\isacharparenright}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}{\isasymlbrakk}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharparenleft}L\ r\isactrlisub {\isadigit{1}}{\isacharparenright}{\isacharparenright}{\isacharsemicolon}\ finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharparenleft}L\ r\isactrlisub {\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isasymrbrakk}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharparenleft}L\ r\isactrlisub {\isadigit{1}}\ {\isasymunion}\ L\ r\isactrlisub {\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}erule\ quot{\isacharunderscore}union{\isacharunderscore}finiteI{\isacharcomma}\ simp{\isacharparenright}\isanewline
\ \ \isacommand{with}\isamarkupfalse%
\ prems\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}case\ \isacommand{by}\isamarkupfalse%
\ simp\ \ \isanewline
\isacommand{next}\isamarkupfalse%
\isanewline
\ \ \isacommand{case}\isamarkupfalse%
\ {\isacharparenleft}STAR\ r{\isacharparenright}\isanewline
\ \ \isacommand{have}\isamarkupfalse%
\ {\isachardoublequoteopen}finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharparenleft}L\ r{\isacharparenright}{\isacharparenright}\ \isanewline
\ \ \ \ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ finite\ {\isacharparenleft}UNIV\ {\isacharslash}{\isacharslash}\ {\isasymapprox}{\isacharparenleft}{\isacharparenleft}L\ r{\isacharparenright}{\isasymstar}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
\ \ \ \ \isacommand{by}\isamarkupfalse%
\ {\isacharparenleft}erule\ quot{\isacharunderscore}star{\isacharunderscore}finiteI{\isacharparenright}\isanewline
\ \ \isacommand{with}\isamarkupfalse%
\ prems\ \isacommand{show}\isamarkupfalse%
\ {\isacharquery}case\ \isacommand{by}\isamarkupfalse%
\ simp\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
\ \isanewline
%
\endisadelimproof
%
\isadelimtheory
\isanewline
%
\endisadelimtheory
%
\isatagtheory
\isacommand{end}\isamarkupfalse%
%
\endisatagtheory
{\isafoldtheory}%
%
\isadelimtheory
\isanewline
%
\endisadelimtheory
\isanewline
\end{isabellebody}%
%%% Local Variables:
%%% mode: latex
%%% TeX-master: "root"
%%% End: