# HG changeset patch # User zhang # Date 1295868595 0 # Node ID f5db9e08effc6e15a349a4a794d12544af9057ca # Parent c64241fa4dff33fce774ccb01ba7db8334e52ef7 Directory [tphols-2011] is added to host the commented version of [Myhill.thy]. Directory [pres] is created to host a build session for logical image [ListP] which is essentially [Main] extended with [List_prefix.thy]. diff -r c64241fa4dff -r f5db9e08effc IsaMakefile --- a/IsaMakefile Fri Jan 07 14:25:23 2011 +0000 +++ b/IsaMakefile Mon Jan 24 11:29:55 2011 +0000 @@ -19,7 +19,7 @@ session1: Slides/ROOT.ML \ Slides/document/root* \ Slides/Slides.thy - @$(USEDIR) -D generated -f ROOT.ML HOL Slides + @$(USEDIR) -D generated -f ROOT.ML ListP tphols-2011 slides: session1 rm -f Slides/generated/*.aux # otherwise latex will fall over @@ -29,19 +29,19 @@ ## Paper -session2: Paper/ROOT.ML \ - Paper/document/root* \ - Paper/Paper.thy - @$(USEDIR) -D generated -f ROOT.ML HOL Paper +session2: tphols-2011/ROOT.ML \ + tphols-2011/document/root* \ + ../*.thy + @$(USEDIR) -D generated -f ROOT.ML ListP tphols-2011 paper: session2 - rm -f Paper/generated/*.aux # otherwise latex will fall over - cd Paper/generated ; $(ISABELLE_TOOL) latex -o pdf root.tex - cp Paper/generated/root.pdf Paper/paper.pdf + rm -f tphols-2011/generated/*.aux # otherwise latex will fall over + cd tphols-2011/generated ; $(ISABELLE_TOOL) latex -o pdf root.tex + cp tphols-2011/generated/root.pdf tphols-2011/myhill.pdf ## clean clean: rm -rf Slides/generated/* - rm -rf Paper/generated/* + rm -rf tphols-2011/generated/* diff -r c64241fa4dff -r f5db9e08effc Myhill.thy --- a/Myhill.thy Fri Jan 07 14:25:23 2011 +0000 +++ b/Myhill.thy Mon Jan 24 11:29:55 2011 +0000 @@ -1,12 +1,15 @@ -theory MyhillNerode - imports "Main" "List_Prefix" +theory Myhill + imports Main List_Prefix begin -text {* sequential composition of languages *} +section {* Preliminary definitions *} + +text {* Sequential composition of two languages @{text "L1"} and @{text "L2"} *} definition Seq :: "string set \ string set \ string set" ("_ ;; _" [100,100] 100) where "L1 ;; L2 = {s1 @ s2 | s1 s2. s1 \ L1 \ s2 \ L2}" +text {* Transitive closure of language @{text "L"}. *} inductive_set Star :: "string set \ string set" ("_\" [101] 102) for L :: "string set" @@ -14,6 +17,8 @@ start[intro]: "[] \ L\" | step[intro]: "\s1 \ L; s2 \ L\\ \ s1@s2 \ L\" +text {* Some properties of operator @{text ";;"}.*} + lemma seq_union_distrib: "(A \ B) ;; C = (A ;; C) \ (B ;; C)" by (auto simp:Seq_def) @@ -28,23 +33,92 @@ apply blast by (metis append_assoc) +lemma star_intro1[rule_format]: "x \ lang\ \ \ y. y \ lang\ \ x @ y \ lang\" +by (erule Star.induct, auto) + +lemma star_intro2: "y \ lang \ y \ lang\" +by (drule step[of y lang "[]"], auto simp:start) + +lemma star_intro3[rule_format]: + "x \ lang\ \ \y . y \ lang \ x @ y \ lang\" +by (erule Star.induct, auto intro:star_intro2) + +lemma star_decom: + "\x \ lang\; x \ []\ \(\ a b. x = a @ b \ a \ [] \ a \ lang \ b \ lang\)" +by (induct x rule: Star.induct, simp, blast) + +lemma star_decom': + "\x \ lang\; x \ []\ \ \a b. x = a @ b \ a \ lang\ \ b \ lang" +apply (induct x rule:Star.induct, simp) +apply (case_tac "s2 = []") +apply (rule_tac x = "[]" in exI, rule_tac x = s1 in exI, simp add:start) +apply (simp, (erule exE| erule conjE)+) +by (rule_tac x = "s1 @ a" in exI, rule_tac x = b in exI, simp add:step) + +text {* Ardens lemma expressed at the level of language, rather than the level of regular expression. *} + theorem ardens_revised: assumes nemp: "[] \ A" shows "(X = X ;; A \ B) \ (X = B ;; A\)" proof assume eq: "X = B ;; A\" - have "A\ = {[]} \ A\ ;; A" sorry - then have "B ;; A\ = B ;; ({[]} \ A\ ;; A)" unfolding Seq_def by simp - also have "\ = B \ B ;; (A\ ;; A)" unfolding Seq_def by auto - also have "\ = B \ (B ;; A\) ;; A" unfolding Seq_def - by (auto) (metis append_assoc)+ - finally show "X = X ;; A \ B" using eq by auto + have "A\ = {[]} \ A\ ;; A" + by (auto simp:Seq_def star_intro3 star_decom') + then have "B ;; A\ = B ;; ({[]} \ A\ ;; A)" + unfolding Seq_def by simp + also have "\ = B \ B ;; (A\ ;; A)" + unfolding Seq_def by auto + also have "\ = B \ (B ;; A\) ;; A" + by (simp only:seq_assoc) + finally show "X = X ;; A \ B" + using eq by blast next - assume "X = X ;; A \ B" - then have "B \ X" "X ;; A \ X" by auto - thus "X = B ;; A\" sorry + assume eq': "X = X ;; A \ B" + hence c1': "\ x. x \ B \ x \ X" + and c2': "\ x y. \x \ X; y \ A\ \ x @ y \ X" + using Seq_def by auto + show "X = B ;; A\" + proof + show "B ;; A\ \ X" + proof- + { fix x y + have "\y \ A\; x \ X\ \ x @ y \ X " + apply (induct arbitrary:x rule:Star.induct, simp) + by (auto simp only:append_assoc[THEN sym] dest:c2') + } thus ?thesis using c1' by (auto simp:Seq_def) + qed + next + show "X \ B ;; A\" + proof- + { fix x + have "x \ X \ x \ B ;; A\" + proof (induct x taking:length rule:measure_induct) + fix z + assume hyps: + "\y. length y < length z \ y \ X \ y \ B ;; A\" + and z_in: "z \ X" + show "z \ B ;; A\" + proof (cases "z \ B") + case True thus ?thesis by (auto simp:Seq_def start) + next + case False hence "z \ X ;; A" using eq' z_in by auto + then obtain za zb where za_in: "za \ X" + and zab: "z = za @ zb \ zb \ A" and zbne: "zb \ []" + using nemp unfolding Seq_def by blast + from zbne zab have "length za < length z" by auto + with za_in hyps have "za \ B ;; A\" by blast + hence "za @ zb \ B ;; A\" using zab + by (clarsimp simp:Seq_def, blast dest:star_intro3) + thus ?thesis using zab by simp + qed + qed + } thus ?thesis by blast + qed + qed qed + +text {* The syntax of regular expressions is defined by the datatype @{text "rexp"}. *} datatype rexp = NULL | EMPTY @@ -53,11 +127,20 @@ | ALT rexp rexp | STAR rexp + +text {* + The following @{text "L"} is an overloaded operator, where @{text "L(x)"} evaluates to + the language represented by the syntactic object @{text "x"}. +*} consts L:: "'a \ string set" + +text {* + The @{text "L(rexp)"} for regular expression @{text "rexp"} is defined by the + following overloading function @{text "L_rexp"}. +*} overloading L_rexp \ "L:: rexp \ string set" begin - fun L_rexp :: "rexp \ string set" where @@ -69,157 +152,258 @@ | "L_rexp (STAR r) = (L_rexp r)\" end +text {* + To obtain equational system out of finite set of equivalent classes, a fold operation + on finite set @{text "folds"} is defined. The use of @{text "SOME"} makes @{text "fold"} + more robust than the @{text "fold"} in Isabelle library. The expression @{text "folds f"} + makes sense when @{text "f"} is not @{text "associative"} and @{text "commutitive"}, + while @{text "fold f"} does not. +*} + definition folds :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b" where "folds f z S \ SOME x. fold_graph f z S x" +text {* + The following lemma assures that the arbitrary choice made by the @{text "SOME"} in @{text "folds"} + does not affect the @{text "L"}-value of the resultant regular expression. + *} lemma folds_alt_simp [simp]: "finite rs \ L (folds ALT NULL rs) = \ (L ` rs)" apply (rule set_ext, simp add:folds_def) apply (rule someI2_ex, erule finite_imp_fold_graph) by (erule fold_graph.induct, auto) +(* Just a technical lemma. *) lemma [simp]: shows "(x, y) \ {(x, y). P x y} \ P x y" by simp -definition - str_eq ("_ \_ _") -where - "x \Lang y \ (\z. x @ z \ Lang \ y @ z \ Lang)" +text {* + @{text "\L"} is an equivalent class defined by language @{text "Lang"}. +*} definition str_eq_rel ("\_") where - "\Lang \ {(x, y). x \Lang y}" + "\Lang \ {(x, y). (\z. x @ z \ Lang \ y @ z \ Lang)}" + +text {* + Among equivlant clases of @{text "\Lang"}, the set @{text "finals(Lang)"} singles out + those which contains strings from @{text "Lang"}. +*} -definition - final :: "string set \ string set \ bool" -where - "final X Lang \ (X \ UNIV // \Lang) \ (\s \ X. s \ Lang)" +definition + "finals Lang \ {\Lang `` {x} | x . x \ Lang}" +text {* + The following lemma show the relationshipt between @{text "finals(Lang)"} and @{text "Lang"}. +*} lemma lang_is_union_of_finals: - "Lang = \ {X. final X Lang}" + "Lang = \ finals(Lang)" proof - show "Lang \ \ {X. final X Lang}" + show "Lang \ \ (finals Lang)" proof fix x assume "x \ Lang" - thus "x \ \ {X. final X Lang}" - apply (simp, rule_tac x = "(\Lang) `` {x}" in exI) - apply (auto simp:final_def quotient_def Image_def str_eq_rel_def str_eq_def) - by (drule_tac x = "[]" in spec, simp) + thus "x \ \ (finals Lang)" + apply (simp add:finals_def, rule_tac x = "(\Lang) `` {x}" in exI) + by (auto simp:Image_def str_eq_rel_def) qed next - show "\{X. final X Lang} \ Lang" - by (auto simp:final_def) + show "\ (finals Lang) \ Lang" + apply (clarsimp simp:finals_def str_eq_rel_def) + by (drule_tac x = "[]" in spec, auto) qed -section {* finite \ regular *} +section {* Direction @{text "finite partition \ regular language"}*} + +text {* + 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 + @{text "rhs_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. + *} datatype rhs_item = Lam "rexp" (* Lambda *) - | Trn "string set" "rexp" (* Transition *) + | Trn "(string set)" "rexp" (* Transition *) -fun the_Trn:: "rhs_item \ (string set \ rexp)" -where "the_Trn (Trn Y r) = (Y, r)" +text {* + In this formalization, pure regular expressions like $\lambda$ is + repsented by @{text "Lam(EMPTY)"}, while transitions like $X_0 a$ is represented by $Trn~X_0~(CHAR~a)$. + *} + +text {* + The functions @{text "the_r"} and @{text "the_Trn"} are used to extract + subcomponents from right hand side items. + *} fun the_r :: "rhs_item \ rexp" where "the_r (Lam r) = r" +fun the_Trn:: "rhs_item \ (string set \ rexp)" +where "the_Trn (Trn Y r) = (Y, r)" + +text {* + Every right hand side item @{text "itm"} defines a string set given + @{text "L(itm)"}, defined as: +*} overloading L_rhs_e \ "L:: rhs_item \ string set" begin -fun L_rhs_e:: "rhs_item \ string set" -where - "L_rhs_e (Lam r) = L r" | - "L_rhs_e (Trn X r) = X ;; L r" + fun L_rhs_e:: "rhs_item \ string set" + where + "L_rhs_e (Lam r) = L r" | + "L_rhs_e (Trn X r) = X ;; L r" end +text {* + The right hand side of every equation is represented by a set of + items. The string set defined by such a set @{text "itms"} is given + by @{text "L(itms)"}, defined as: +*} + overloading L_rhs \ "L:: rhs_item set \ string set" begin -fun L_rhs:: "rhs_item set \ string set" -where - "L_rhs rhs = \ (L ` rhs)" + fun L_rhs:: "rhs_item set \ string set" + where "L_rhs rhs = \ (L ` rhs)" end +text {* + Given a set of equivalent classses @{text "CS"} and one equivalent class @{text "X"} among + @{text "CS"}, the term @{text "init_rhs CS X"} is used to extract the right hand side of + the equation describing the formation of @{text "X"}. The definition of @{text "init_rhs"} + is: + *} + definition - "init_rhs CS X \ if ([] \ X) - then {Lam EMPTY} \ {Trn Y (CHAR c)| Y c. Y \ CS \ Y ;; {[c]} \ X} - else {Trn Y (CHAR c)| Y c. Y \ CS \ Y ;; {[c]} \ X}" + "init_rhs CS X \ + if ([] \ X) then + {Lam(EMPTY)} \ {Trn Y (CHAR c) | Y c. Y \ CS \ Y ;; {[c]} \ X} + else + {Trn Y (CHAR c)| Y c. Y \ CS \ Y ;; {[c]} \ X}" -definition - "eqs CS \ {(X, init_rhs CS X)|X. X \ CS}" +text {* + In the definition of @{text "init_rhs"}, the term + @{text "{Trn Y (CHAR c)| Y c. Y \ CS \ Y ;; {[c]} \ X}"} appearing on both branches + describes the formation of strings in @{text "X"} out of transitions, while + the term @{text "{Lam(EMPTY)}"} describes the empty string which is intrinsically contained in + @{text "X"} rather than by transition. This @{text "{Lam(EMPTY)}"} corresponds to + the $\lambda$ in \eqref{example_eqns}. + With the help of @{text "init_rhs"}, the equitional system descrbing the formation of every + equivalent class inside @{text "CS"} is given by the following @{text "eqs(CS)"}. + *} + +definition "eqs CS \ {(X, init_rhs CS X) | X. X \ CS}" (************ arden's lemma variation ********************) +text {* + The following @{text "items_of rhs X"} returns all @{text "X"}-items in @{text "rhs"}. + *} definition "items_of rhs X \ {Trn X r | r. (Trn X r) \ rhs}" +text {* + The following @{text "rexp_of rhs X"} combines all regular expressions in @{text "X"}-items + using @{text "ALT"} to form a single regular expression. + It will be used later to implement @{text "arden_variate"} and @{text "rhs_subst"}. + *} + +definition + "rexp_of rhs X \ folds ALT NULL ((snd o the_Trn) ` items_of rhs X)" + +text {* + The following @{text "lam_of rhs"} returns all pure regular expression items in @{text "rhs"}. + *} + definition "lam_of rhs \ {Lam r | r. Lam r \ rhs}" -definition - "rexp_of rhs X \ folds ALT NULL ((snd o the_Trn) ` items_of rhs X)" +text {* + The following @{text "rexp_of_lam rhs"} combines pure regular expression items in @{text "rhs"} + using @{text "ALT"} to form a single regular expression. + When all variables inside @{text "rhs"} are eliminated, @{text "rexp_of_lam rhs"} + is used to compute compute the regular expression corresponds to @{text "rhs"}. + *} definition "rexp_of_lam rhs \ folds ALT NULL (the_r ` lam_of rhs)" +text {* + The following @{text "attach_rexp rexp' itm"} attach + the regular expression @{text "rexp'"} to + the right of right hand side item @{text "itm"}. + *} + fun attach_rexp :: "rexp \ rhs_item \ rhs_item" where - "attach_rexp r' (Lam r) = Lam (SEQ r r')" -| "attach_rexp r' (Trn X r) = Trn X (SEQ r r')" + "attach_rexp rexp' (Lam rexp) = Lam (SEQ rexp rexp')" +| "attach_rexp rexp' (Trn X rexp) = Trn X (SEQ rexp rexp')" + +text {* + The following @{text "append_rhs_rexp rhs rexp"} attaches + @{text "rexp"} to every item in @{text "rhs"}. + *} definition - "append_rhs_rexp rhs r \ (attach_rexp r) ` rhs" + "append_rhs_rexp rhs rexp \ (attach_rexp rexp) ` rhs" +text {* + With the help of the two functions immediately above, Ardens' + transformation on right hand side @{text "rhs"} is implemented + by the following function @{text "arden_variate X rhs"}. + After this transformation, the recursive occurent of @{text "X"} + in @{text "rhs"} will be eliminated, while the + string set defined by @{text "rhs"} is kept unchanged. + *} definition - "arden_variate X rhs \ append_rhs_rexp (rhs - items_of rhs X) (STAR (rexp_of rhs X))" + "arden_variate X rhs \ + append_rhs_rexp (rhs - items_of rhs X) (STAR (rexp_of rhs X))" (*********** substitution of ES *************) -text {* rhs_subst rhs X xrhs: substitude all occurence of X in rhs with xrhs *} +text {* + Suppose the equation defining @{text "X"} is $X = xrhs$, + the purpose of @{text "rhs_subst"} is to substitute all occurences of @{text "X"} in + @{text "rhs"} by @{text "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 @{text "xrhs"} and then + union the result with all non-@{text "X"}-items of @{text "rhs"}. + *} definition - "rhs_subst rhs X xrhs \ (rhs - (items_of rhs X)) \ (append_rhs_rexp xrhs (rexp_of rhs X))" + "rhs_subst rhs X xrhs \ + (rhs - (items_of rhs X)) \ (append_rhs_rexp xrhs (rexp_of rhs X))" + +text {* + Suppose the equation defining @{text "X"} is $X = xrhs$, the follwing + @{text "eqs_subst ES X xrhs"} substitute @{text "xrhs"} into every equation + of the equational system @{text "ES"}. + *} definition "eqs_subst ES X xrhs \ {(Y, rhs_subst yrhs X xrhs) | Y yrhs. (Y, yrhs) \ ES}" text {* - Inv: Invairance of the equation-system, during the decrease of the equation-system, Inv holds. -*} - -definition - "distinct_equas ES \ \ X rhs rhs'. (X, rhs) \ ES \ (X, rhs') \ ES \ rhs = rhs'" - -definition - "valid_eqns ES \ \ X rhs. (X, rhs) \ ES \ (X = L rhs)" - -definition - "rhs_nonempty rhs \ (\ Y r. Trn Y r \ rhs \ [] \ L r)" - -definition - "ardenable ES \ \ X rhs. (X, rhs) \ ES \ rhs_nonempty rhs" - -definition - "non_empty ES \ \ X rhs. (X, rhs) \ ES \ X \ {}" - -definition - "finite_rhs ES \ \ X rhs. (X, rhs) \ ES \ finite rhs" - -definition - "classes_of rhs \ {X. \ r. Trn X r \ rhs}" - -definition - "lefts_of ES \ {Y | Y yrhs. (Y, yrhs) \ ES}" - -definition - "self_contained ES \ \ (X, xrhs) \ ES. classes_of xrhs \ lefts_of ES" - -definition - "Inv ES \ valid_eqns ES \ finite ES \ distinct_equas ES \ ardenable ES \ - non_empty ES \ finite_rhs ES \ self_contained ES" + The computation of regular expressions for equivalent classes is accomplished + using a iteration principle given by the following lemma. + *} lemma wf_iter [rule_format]: fixes f @@ -248,7 +432,88 @@ qed qed -text {* ************* basic properties of definitions above ************************ *} +text {* + The @{text "P"} in lemma @{text "wf_iter"} is an invaiant kept throughout the iteration procedure. + The particular invariant used to solve our problem is defined by function @{text "Inv(ES)"}, + an invariant over equal system @{text "ES"}. + Every definition starting next till @{text "Inv"} stipulates a property to be satisfied by @{text "ES"}. +*} + +text {* + Every variable is defined at most onece in @{text "ES"}. + *} +definition + "distinct_equas ES \ + \ X rhs rhs'. (X, rhs) \ ES \ (X, rhs') \ ES \ rhs = rhs'" +text {* + Every equation in @{text "ES"} (represented by @{text "(X, rhs)"}) is valid, i.e. @{text "(X = L rhs)"}. + *} +definition + "valid_eqns ES \ \ X rhs. (X, rhs) \ ES \ (X = L rhs)" + +text {* + @{text "rhs_nonempty rhs"} requires regular expressions occuring in transitional + items of @{text "rhs"} does not contain empty string. This is necessary for + the application of Arden's transformation to @{text "rhs"}. + *} +definition + "rhs_nonempty rhs \ (\ Y r. Trn Y r \ rhs \ [] \ L r)" + +text {* + @{text "ardenable ES"} requires that Arden's transformation is applicable + to every equation of equational system @{text "ES"}. + *} +definition + "ardenable ES \ \ X rhs. (X, rhs) \ ES \ rhs_nonempty rhs" + +(* The following non_empty seems useless. *) +definition + "non_empty ES \ \ X rhs. (X, rhs) \ ES \ X \ {}" + +text {* + The following @{text "finite_rhs ES"} requires every equation in @{text "rhs"} be finite. + *} +definition + "finite_rhs ES \ \ X rhs. (X, rhs) \ ES \ finite rhs" + +text {* + The following @{text "classes_of rhs"} returns all variables (or equivalent classes) + occuring in @{text "rhs"}. + *} +definition + "classes_of rhs \ {X. \ r. Trn X r \ rhs}" + +text {* + The following @{text "lefts_of ES"} returns all variables + defined by equational system @{text "ES"}. + *} +definition + "lefts_of ES \ {Y | Y yrhs. (Y, yrhs) \ ES}" + +text {* + The following @{text "self_contained ES"} requires that every + variable occuring on the right hand side of equations is already defined by some + equation in @{text "ES"}. + *} +definition + "self_contained ES \ \ (X, xrhs) \ ES. classes_of xrhs \ lefts_of ES" + + +text {* + The invariant @{text "Inv(ES)"} is obtained by conjunctioning all the previous + defined constaints on @{text "ES"}. + *} +definition + "Inv ES \ valid_eqns ES \ finite ES \ distinct_equas ES \ ardenable ES \ + non_empty ES \ finite_rhs ES \ self_contained ES" + +subsection {* Proof for this direction *} + + + +text {* + The following are some basic properties of the above definitions. +*} lemma L_rhs_union_distrib: " L (A::rhs_item set) \ L B = L (A \ B)" @@ -319,11 +584,14 @@ by (auto simp:lefts_of_def) -text {* ******BEGIN: proving the initial equation-system satisfies Inv ****** *} +text {* + The following several lemmas until @{text "init_ES_satisfy_Inv"} are + to prove that initial equational system satisfies invariant @{text "Inv"}. + *} lemma defined_by_str: "\s \ X; X \ UNIV // (\Lang)\ \ X = (\Lang) `` {s}" -by (auto simp:quotient_def Image_def str_eq_rel_def str_eq_def) +by (auto simp:quotient_def Image_def str_eq_rel_def) lemma every_eqclass_has_transition: assumes has_str: "s @ [c] \ X" @@ -339,10 +607,10 @@ then have "Y ;; {[c]} \ X" unfolding Y_def Image_def Seq_def unfolding str_eq_rel_def - by (auto) (simp add: str_eq_def) + by clarsimp moreover have "s \ Y" unfolding Y_def - unfolding Image_def str_eq_rel_def str_eq_def by simp + unfolding Image_def str_eq_rel_def by simp ultimately show thesis by (blast intro: that) qed @@ -369,7 +637,8 @@ and "Y ;; {[c]} \ X" and "clist \ Y" using decom "(1)" every_eqclass_has_transition by blast - hence "x \ L {Trn Y (CHAR c)| Y c. Y \ UNIV // (\Lang) \ Y ;; {[c]} \ X}" + hence + "x \ L {Trn Y (CHAR c)| Y c. Y \ UNIV // (\Lang) \ Y ;; {[c]} \ X}" using "(1)" decom by (simp, rule_tac x = "Trn Y (CHAR c)" in exI, simp add:Seq_def) thus ?thesis using X_in_eqs "(1)" @@ -412,7 +681,7 @@ moreover have "valid_eqns (eqs (UNIV // (\Lang)))" using l_eq_r_in_eqs by (simp add:valid_eqns_def) moreover have "non_empty (eqs (UNIV // (\Lang)))" - by (auto simp:non_empty_def eqs_def quotient_def Image_def str_eq_rel_def str_eq_def) + by (auto simp:non_empty_def eqs_def quotient_def Image_def str_eq_rel_def) moreover have "finite_rhs (eqs (UNIV // (\Lang)))" using finite_init_rhs[OF finite_CS] by (auto simp:finite_rhs_def eqs_def) @@ -421,8 +690,11 @@ ultimately show ?thesis by (simp add:Inv_def) qed -text {* ****** BEGIN: proving every equation-system's iteration step satisfies Inv ***** *} - +text {* + From this point until @{text "iteration_step"}, we are trying to prove + that there exists iteration steps which keep @{text "Inv(ES)"} while + decreasing the size of @{text "ES"} with every iteration. + *} lemma arden_variate_keeps_eq: assumes l_eq_r: "X = L rhs" and not_empty: "[] \ L (rexp_of rhs X)" @@ -506,7 +778,8 @@ assumes finite:"finite (ES:: (string set \ rhs_item set) set)" shows "finite (eqs_subst ES Y yrhs)" proof - - have "finite {(Ya, rhs_subst yrhsa Y yrhs) |Ya yrhsa. (Ya, yrhsa) \ ES}" (is "finite ?A") + have "finite {(Ya, rhs_subst yrhsa Y yrhs) |Ya yrhsa. (Ya, yrhsa) \ ES}" + (is "finite ?A") proof- def eqns' \ "{((Ya::string set), yrhsa)| Ya yrhsa. (Ya, yrhsa) \ ES}" def h \ "\ ((Ya::string set), yrhsa). (Ya, rhs_subst yrhsa Y yrhs)" @@ -537,14 +810,17 @@ by (auto simp:lefts_of_def eqs_subst_def) lemma rhs_subst_updates_cls: - "X \ classes_of xrhs \ classes_of (rhs_subst rhs X xrhs) = classes_of rhs \ classes_of xrhs - {X}" -apply (simp only:rhs_subst_def append_rhs_keeps_cls classes_of_union_distrib[THEN sym]) + "X \ classes_of xrhs \ + classes_of (rhs_subst rhs X xrhs) = classes_of rhs \ classes_of xrhs - {X}" +apply (simp only:rhs_subst_def append_rhs_keeps_cls + classes_of_union_distrib[THEN sym]) by (auto simp:classes_of_def items_of_def) lemma eqs_subst_keeps_self_contained: fixes Y assumes sc: "self_contained (ES \ {(Y, yrhs)})" (is "self_contained ?A") - shows "self_contained (eqs_subst ES Y (arden_variate Y yrhs))" (is "self_contained ?B") + shows "self_contained (eqs_subst ES Y (arden_variate Y yrhs))" + (is "self_contained ?B") proof- { fix X xrhs' assume "(X, xrhs') \ ?B" @@ -556,15 +832,18 @@ have "lefts_of ?B = lefts_of ES" by (auto simp add:lefts_of_def eqs_subst_def) moreover have "classes_of xrhs' \ lefts_of ES" proof- - have "classes_of xrhs' \ classes_of xrhs \ classes_of (arden_variate Y yrhs) - {Y}" + have "classes_of xrhs' \ + classes_of xrhs \ classes_of (arden_variate Y yrhs) - {Y}" proof- - have "Y \ classes_of (arden_variate Y yrhs)" using arden_variate_removes_cl by simp + have "Y \ classes_of (arden_variate Y yrhs)" + using arden_variate_removes_cl by simp thus ?thesis using xrhs_xrhs' by (auto simp:rhs_subst_updates_cls) qed moreover have "classes_of xrhs \ lefts_of ES \ {Y}" using X_in sc apply (simp only:self_contained_def lefts_of_union_distrib[THEN sym]) by (drule_tac x = "(X, xrhs)" in bspec, auto simp:lefts_of_def) - moreover have "classes_of (arden_variate Y yrhs) \ lefts_of ES \ {Y}" using sc + moreover have "classes_of (arden_variate Y yrhs) \ lefts_of ES \ {Y}" + using sc by (auto simp add:arden_variate_removes_cl self_contained_def lefts_of_def) ultimately show ?thesis by auto qed @@ -577,44 +856,57 @@ assumes Inv_ES: "Inv (ES \ {(Y, yrhs)})" shows "Inv (eqs_subst ES Y (arden_variate Y yrhs))" proof - - have finite_yrhs: "finite yrhs" using Inv_ES by (auto simp:Inv_def finite_rhs_def) - have nonempty_yrhs: "rhs_nonempty yrhs" using Inv_ES by (auto simp:Inv_def ardenable_def) - have Y_eq_yrhs: "Y = L yrhs" using Inv_ES by (simp only:Inv_def valid_eqns_def, blast) - - have "distinct_equas (eqs_subst ES Y (arden_variate Y yrhs))" using Inv_ES + have finite_yrhs: "finite yrhs" + using Inv_ES by (auto simp:Inv_def finite_rhs_def) + have nonempty_yrhs: "rhs_nonempty yrhs" + using Inv_ES by (auto simp:Inv_def ardenable_def) + have Y_eq_yrhs: "Y = L yrhs" + using Inv_ES by (simp only:Inv_def valid_eqns_def, blast) + have "distinct_equas (eqs_subst ES Y (arden_variate Y yrhs))" + using Inv_ES by (auto simp:distinct_equas_def eqs_subst_def Inv_def) - moreover have "finite (eqs_subst ES Y (arden_variate Y yrhs))" using Inv_ES - by (simp add:Inv_def eqs_subst_keeps_finite) + moreover have "finite (eqs_subst ES Y (arden_variate Y yrhs))" + using Inv_ES by (simp add:Inv_def eqs_subst_keeps_finite) moreover have "finite_rhs (eqs_subst ES Y (arden_variate Y yrhs))" proof- - have "finite_rhs ES" using Inv_ES by (simp add:Inv_def finite_rhs_def) + have "finite_rhs ES" using Inv_ES + by (simp add:Inv_def finite_rhs_def) moreover have "finite (arden_variate Y yrhs)" proof - - have "finite yrhs" using Inv_ES by (auto simp:Inv_def finite_rhs_def) + have "finite yrhs" using Inv_ES + by (auto simp:Inv_def finite_rhs_def) thus ?thesis using arden_variate_keeps_finite by simp qed - ultimately show ?thesis by (simp add:eqs_subst_keeps_finite_rhs) + ultimately show ?thesis + by (simp add:eqs_subst_keeps_finite_rhs) qed moreover have "ardenable (eqs_subst ES Y (arden_variate Y yrhs))" proof - { fix X rhs assume "(X, rhs) \ ES" - hence "rhs_nonempty rhs" using prems Inv_ES by (simp add:Inv_def ardenable_def) - with nonempty_yrhs have "rhs_nonempty (rhs_subst rhs Y (arden_variate Y yrhs))" - by (simp add:nonempty_yrhs rhs_subst_keeps_nonempty arden_variate_keeps_nonempty) + hence "rhs_nonempty rhs" using prems Inv_ES + by (simp add:Inv_def ardenable_def) + with nonempty_yrhs + have "rhs_nonempty (rhs_subst rhs Y (arden_variate Y yrhs))" + by (simp add:nonempty_yrhs + rhs_subst_keeps_nonempty arden_variate_keeps_nonempty) } thus ?thesis by (auto simp add:ardenable_def eqs_subst_def) qed moreover have "valid_eqns (eqs_subst ES Y (arden_variate Y yrhs))" proof- - have "Y = L (arden_variate Y yrhs)" using Y_eq_yrhs Inv_ES finite_yrhs nonempty_yrhs - by (rule_tac arden_variate_keeps_eq, (simp add:rexp_of_empty)+) + have "Y = L (arden_variate Y yrhs)" + using Y_eq_yrhs Inv_ES finite_yrhs nonempty_yrhs + by (rule_tac arden_variate_keeps_eq, (simp add:rexp_of_empty)+) thus ?thesis using Inv_ES - by (clarsimp simp add:valid_eqns_def eqs_subst_def rhs_subst_keeps_eq Inv_def finite_rhs_def + by (clarsimp simp add:valid_eqns_def + eqs_subst_def rhs_subst_keeps_eq Inv_def finite_rhs_def simp del:L_rhs.simps) qed - moreover have non_empty_subst: "non_empty (eqs_subst ES Y (arden_variate Y yrhs))" + moreover have + non_empty_subst: "non_empty (eqs_subst ES Y (arden_variate Y yrhs))" using Inv_ES by (auto simp:Inv_def non_empty_def eqs_subst_def) - moreover have self_subst: "self_contained (eqs_subst ES Y (arden_variate Y yrhs))" + moreover + have self_subst: "self_contained (eqs_subst ES Y (arden_variate Y yrhs))" using Inv_ES eqs_subst_keeps_self_contained by (simp add:Inv_def) ultimately show ?thesis using Inv_ES by (simp add:Inv_def) qed @@ -642,7 +934,8 @@ proof- have "card (S - {e}) > 0" proof - - have "card S > 1" using card e_in finite by (case_tac "card S", auto) + have "card S > 1" using card e_in finite + by (case_tac "card S", auto) thus ?thesis using finite e_in by auto qed hence "S - {e} \ {}" using finite by (rule_tac notI, simp) @@ -653,10 +946,12 @@ assumes Inv_ES: "Inv ES" and X_in_ES: "(X, xrhs) \ ES" and not_T: "card ES \ 1" - shows "\ ES'. (Inv ES' \ (\ xrhs'.(X, xrhs') \ ES')) \ (card ES', card ES) \ less_than" (is "\ ES'. ?P ES'") + shows "\ ES'. (Inv ES' \ (\ xrhs'.(X, xrhs') \ ES')) \ + (card ES', card ES) \ less_than" (is "\ ES'. ?P ES'") proof - have finite_ES: "finite ES" using Inv_ES by (simp add:Inv_def) - then obtain Y yrhs where Y_in_ES: "(Y, yrhs) \ ES" and not_eq: "(X, xrhs) \ (Y, yrhs)" + then obtain Y yrhs + where Y_in_ES: "(Y, yrhs) \ ES" and not_eq: "(X, xrhs) \ (Y, yrhs)" using not_T X_in_ES by (drule_tac card_noteq_1_has_more, auto) def ES' == "ES - {(Y, yrhs)}" let ?ES'' = "eqs_subst ES' Y (arden_variate Y yrhs)" @@ -679,12 +974,17 @@ thus ?thesis by blast qed -text {* ***** END: proving every equation-system's iteration step satisfies Inv ************** *} +text {* + From this point until @{text "hard_direction"}, the hard direction is proved + through a simple application of the iteration principle. +*} lemma iteration_conc: assumes history: "Inv ES" and X_in_ES: "\ xrhs. (X, xrhs) \ ES" - shows "\ ES'. (Inv ES' \ (\ xrhs'. (X, xrhs') \ ES')) \ card ES' = 1" (is "\ ES'. ?P ES'") + shows + "\ ES'. (Inv ES' \ (\ xrhs'. (X, xrhs') \ ES')) \ card ES' = 1" + (is "\ ES'. ?P ES'") proof (cases "card ES = 1") case True thus ?thesis using history X_in_ES @@ -706,26 +1006,31 @@ have "L (rexp_of_lam ?A) = L (lam_of ?A)" proof(rule rexp_of_lam_eq_lam_set) show "finite (arden_variate X xrhs)" using Inv_ES ES_single - by (rule_tac arden_variate_keeps_finite, auto simp add:Inv_def finite_rhs_def) + by (rule_tac arden_variate_keeps_finite, + auto simp add:Inv_def finite_rhs_def) qed also have "\ = L ?A" proof- have "lam_of ?A = ?A" proof- have "classes_of ?A = {}" using Inv_ES ES_single - by (simp add:arden_variate_removes_cl self_contained_def Inv_def lefts_of_def) - thus ?thesis by (auto simp only:lam_of_def classes_of_def, case_tac x, auto) + by (simp add:arden_variate_removes_cl + self_contained_def Inv_def lefts_of_def) + thus ?thesis + by (auto simp only:lam_of_def classes_of_def, case_tac x, auto) qed thus ?thesis by simp qed also have "\ = X" proof(rule arden_variate_keeps_eq [THEN sym]) - show "X = L xrhs" using Inv_ES ES_single by (auto simp only:Inv_def valid_eqns_def) + show "X = L xrhs" using Inv_ES ES_single + by (auto simp only:Inv_def valid_eqns_def) next from Inv_ES ES_single show "[] \ L (rexp_of xrhs X)" by(simp add:Inv_def ardenable_def rexp_of_empty finite_rhs_def) next - from Inv_ES ES_single show "finite xrhs" by (simp add:Inv_def finite_rhs_def) + from Inv_ES ES_single show "finite xrhs" + by (simp add:Inv_def finite_rhs_def) qed finally show ?thesis by simp qed @@ -750,79 +1055,73 @@ by (rule last_cl_exists_rexp) qed +lemma finals_in_partitions: + "finals Lang \ (UNIV // (\Lang))" + by (auto simp:finals_def quotient_def) + theorem hard_direction: assumes finite_CS: "finite (UNIV // (\Lang))" shows "\ (reg::rexp). Lang = L reg" proof - have "\ X \ (UNIV // (\Lang)). \ (reg::rexp). X = L reg" using finite_CS every_eqcl_has_reg by blast - then obtain f where f_prop: "\ X \ (UNIV // (\Lang)). X = L ((f X)::rexp)" + then obtain f + where f_prop: "\ X \ (UNIV // (\Lang)). X = L ((f X)::rexp)" by (auto dest:bchoice) - def rs \ "f ` {X. final X Lang}" - have "Lang = \ {X. final X Lang}" using lang_is_union_of_finals by simp + def rs \ "f ` (finals Lang)" + have "Lang = \ (finals Lang)" using lang_is_union_of_finals by auto also have "\ = L (folds ALT NULL rs)" proof - - have "finite {X. final X Lang}" using finite_CS by (auto simp:final_def) - thus ?thesis using f_prop by (auto simp:rs_def final_def) + have "finite rs" + proof - + have "finite (finals Lang)" + using finite_CS finals_in_partitions[of "Lang"] + by (erule_tac finite_subset, simp) + thus ?thesis using rs_def by auto + qed + thus ?thesis + using f_prop rs_def finals_in_partitions[of "Lang"] by auto qed finally show ?thesis by blast qed -section {* regular \ finite*} +section {* Direction: @{text "regular language \finite partition"} *} + +subsection {* The scheme for this direction *} -lemma quot_empty_subset: - "UNIV // (\{[]}) \ {{[]}, UNIV - {[]}}" -proof - fix x - assume "x \ UNIV // \{[]}" - then obtain y where h: "x = {z. (y, z) \ \{[]}}" unfolding quotient_def Image_def by blast - show "x \ {{[]}, UNIV - {[]}}" - proof (cases "y = []") - case True with h - have "x = {[]}" by (auto simp:str_eq_rel_def str_eq_def) - thus ?thesis by simp - next - case False with h - have "x = UNIV - {[]}" by (auto simp:str_eq_rel_def str_eq_def) - thus ?thesis by simp - qed -qed +text {* + The following convenient notation @{text "x \Lang y"} means: + string @{text "x"} and @{text "y"} are equivalent with respect to + language @{text "Lang"}. + *} + +definition + str_eq ("_ \_ _") +where + "x \Lang y \ (x, y) \ (\Lang)" -lemma quot_char_subset: - "UNIV // (\{[c]}) \ {{[]},{[c]}, UNIV - {[], [c]}}" -proof - fix x - assume "x \ UNIV // \{[c]}" - then obtain y where h: "x = {z. (y, z) \ \{[c]}}" unfolding quotient_def Image_def by blast - show "x \ {{[]},{[c]}, UNIV - {[], [c]}}" - proof - - { assume "y = []" hence "x = {[]}" using h by (auto simp:str_eq_rel_def str_eq_def) - } moreover { - assume "y = [c]" hence "x = {[c]}" using h - by (auto dest!:spec[where x = "[]"] simp:str_eq_rel_def str_eq_def) - } moreover { - assume "y \ []" and "y \ [c]" - hence "\ z. (y @ z) \ [c]" by (case_tac y, auto) - moreover have "\ p. (p \ [] \ p \ [c]) = (\ q. p @ q \ [c])" by (case_tac p, auto) - ultimately have "x = UNIV - {[],[c]}" using h - by (auto simp add:str_eq_rel_def str_eq_def) - } ultimately show ?thesis by blast - qed -qed +text {* + The very basic scheme to show the finiteness of the partion generated by a language @{text "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 @{text "Lang"}, + then the partition given rise by @{text "Lang"} must be finite. The reason for this is a lemma + in standard library (@{text "finite_imageD"}), which says: if the image of an injective + function on a set @{text "A"} is finite, then @{text "A"} is finite. It can be shown that + the function obtained by llifting @{text "tag"} + to the level of equalent classes (i.e. @{text "((op `) tag)"}) is injective + (by lemma @{text "tag_image_injI"}) and the image of this function is finite + (with the help of lemma @{text "finite_tag_imageI"}). -text {* *************** Some common lemmas for following ALT, SEQ & STAR cases ******************* *} - -lemma finite_tag_imageI: - "finite (range tag) \ finite (((op `) tag) ` S)" -apply (rule_tac B = "Pow (tag ` UNIV)" in finite_subset) -by (auto simp add:image_def Pow_def) + BUT, I think this argument can be encapsulated by one lemma instead of the current presentation. + *} lemma eq_class_equalI: - "\X \ UNIV // \lang; Y \ UNIV // \lang; x \ X; y \ Y; x \lang y\ \ X = Y" + "\X \ UNIV // \lang; Y \ UNIV // \lang; x \ X; y \ Y; x \lang y\ + \ X = Y" by (auto simp:quotient_def str_eq_rel_def str_eq_def) lemma tag_image_injI: - assumes str_inj: "\ m n. tag m = tag (n::string) \ m \lang n" + assumes str_inj: "\ x y. tag x = tag (y::string) \ x \lang y" shows "inj_on ((op `) tag) (UNIV // \lang)" proof- { fix X Y @@ -838,7 +1137,17 @@ thus ?thesis unfolding inj_on_def by auto qed -text {* **************** the SEQ case ************************ *} +lemma finite_tag_imageI: + "finite (range tag) \ finite (((op `) tag) ` S)" +apply (rule_tac B = "Pow (tag ` UNIV)" in finite_subset) +by (auto simp add:image_def Pow_def) + + +subsection {* A small theory for list difference *} + +text {* + The notion of list diffrence is need to make proofs more readable. + *} (* list_diff:: list substract, once different return tailer *) fun list_diff :: "'a list \ 'a list \ 'a list" (infix "-" 51) @@ -884,29 +1193,99 @@ by (clarsimp, auto simp:prefix_def) lemma app_eq_dest: - "x @ y = m @ n \ (x \ m \ (m - x) @ n = y) \ (m \ x \ (x - m) @ y = n)" + "x @ y = m @ n \ + (x \ m \ (m - x) @ n = y) \ (m \ x \ (x - m) @ y = n)" by (frule_tac app_eq_cases, auto elim:prefixE) +subsection {* Lemmas for basic cases *} + +text {* + The the final result of this direction is in @{text "easier_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 @{text "NULL, EMPTY, 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 @{text "ALT, SEQ"} and @{text "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. + *} + +lemma quot_empty_subset: + "UNIV // (\{[]}) \ {{[]}, UNIV - {[]}}" +proof + fix x + assume "x \ UNIV // \{[]}" + then obtain y where h: "x = {z. (y, z) \ \{[]}}" + unfolding quotient_def Image_def by blast + show "x \ {{[]}, UNIV - {[]}}" + proof (cases "y = []") + case True with h + have "x = {[]}" by (auto simp:str_eq_rel_def) + thus ?thesis by simp + next + case False with h + have "x = UNIV - {[]}" by (auto simp:str_eq_rel_def) + thus ?thesis by simp + qed +qed + +lemma quot_char_subset: + "UNIV // (\{[c]}) \ {{[]},{[c]}, UNIV - {[], [c]}}" +proof + fix x + assume "x \ UNIV // \{[c]}" + then obtain y where h: "x = {z. (y, z) \ \{[c]}}" + unfolding quotient_def Image_def by blast + show "x \ {{[]},{[c]}, UNIV - {[], [c]}}" + proof - + { assume "y = []" hence "x = {[]}" using h + by (auto simp:str_eq_rel_def) + } moreover { + assume "y = [c]" hence "x = {[c]}" using h + by (auto dest!:spec[where x = "[]"] simp:str_eq_rel_def) + } moreover { + assume "y \ []" and "y \ [c]" + hence "\ z. (y @ z) \ [c]" by (case_tac y, auto) + moreover have "\ p. (p \ [] \ p \ [c]) = (\ q. p @ q \ [c])" + by (case_tac p, auto) + ultimately have "x = UNIV - {[],[c]}" using h + by (auto simp add:str_eq_rel_def) + } ultimately show ?thesis by blast + qed +qed + +subsection {* The case for @{text "SEQ"}*} + definition - "tag_str_SEQ L\<^isub>1 L\<^isub>2 x \ ((\L\<^isub>1) `` {x}, {(\L\<^isub>2) `` {x - xa}| xa. xa \ x \ xa \ L\<^isub>1})" + "tag_str_SEQ L\<^isub>1 L\<^isub>2 x \ + ((\L\<^isub>1) `` {x}, {(\L\<^isub>2) `` {x - xa}| xa. xa \ x \ xa \ L\<^isub>1})" lemma tag_str_seq_range_finite: - "\finite (UNIV // \L\<^isub>1); finite (UNIV // \L\<^isub>2)\ \ finite (range (tag_str_SEQ L\<^isub>1 L\<^isub>2))" + "\finite (UNIV // \L\<^isub>1); finite (UNIV // \L\<^isub>2)\ + \ finite (range (tag_str_SEQ L\<^isub>1 L\<^isub>2))" apply (rule_tac B = "(UNIV // \L\<^isub>1) \ (Pow (UNIV // \L\<^isub>2))" in finite_subset) by (auto simp:tag_str_SEQ_def Image_def quotient_def split:if_splits) lemma append_seq_elim: assumes "x @ y \ L\<^isub>1 ;; L\<^isub>2" - shows "(\ xa \ x. xa \ L\<^isub>1 \ (x - xa) @ y \ L\<^isub>2) \ (\ ya \ y. (x @ ya) \ L\<^isub>1 \ (y - ya) \ L\<^isub>2)" + shows "(\ xa \ x. xa \ L\<^isub>1 \ (x - xa) @ y \ L\<^isub>2) \ + (\ ya \ y. (x @ ya) \ L\<^isub>1 \ (y - ya) \ L\<^isub>2)" proof- - from assms obtain s\<^isub>1 s\<^isub>2 where "x @ y = s\<^isub>1 @ s\<^isub>2" and in_seq: "s\<^isub>1 \ L\<^isub>1 \ s\<^isub>2 \ L\<^isub>2" + from assms obtain s\<^isub>1 s\<^isub>2 + where "x @ y = s\<^isub>1 @ s\<^isub>2" + and in_seq: "s\<^isub>1 \ L\<^isub>1 \ s\<^isub>2 \ L\<^isub>2" by (auto simp:Seq_def) hence "(x \ s\<^isub>1 \ (s\<^isub>1 - x) @ s\<^isub>2 = y) \ (s\<^isub>1 \ x \ (x - s\<^isub>1) @ y = s\<^isub>2)" using app_eq_dest by auto - moreover have "\x \ s\<^isub>1; (s\<^isub>1 - x) @ s\<^isub>2 = y\ \ \ ya \ y. (x @ ya) \ L\<^isub>1 \ (y - ya) \ L\<^isub>2" using in_seq - by (rule_tac x = "s\<^isub>1 - x" in exI, auto elim:prefixE) - moreover have "\s\<^isub>1 \ x; (x - s\<^isub>1) @ y = s\<^isub>2\ \ \ xa \ x. xa \ L\<^isub>1 \ (x - xa) @ y \ L\<^isub>2" using in_seq - by (rule_tac x = s\<^isub>1 in exI, auto) + moreover have "\x \ s\<^isub>1; (s\<^isub>1 - x) @ s\<^isub>2 = y\ \ + \ ya \ y. (x @ ya) \ L\<^isub>1 \ (y - ya) \ L\<^isub>2" + using in_seq by (rule_tac x = "s\<^isub>1 - x" in exI, auto elim:prefixE) + moreover have "\s\<^isub>1 \ x; (x - s\<^isub>1) @ y = s\<^isub>2\ \ + \ xa \ x. xa \ L\<^isub>1 \ (x - xa) @ y \ L\<^isub>2" + using in_seq by (rule_tac x = s\<^isub>1 in exI, auto) ultimately show ?thesis by blast qed @@ -918,7 +1297,8 @@ and tag_xy: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y" have"y @ z \ L\<^isub>1 ;; L\<^isub>2" proof- - have "(\ xa \ x. xa \ L\<^isub>1 \ (x - xa) @ z \ L\<^isub>2) \ (\ za \ z. (x @ za) \ L\<^isub>1 \ (z - za) \ L\<^isub>2)" + have "(\ xa \ x. xa \ L\<^isub>1 \ (x - xa) @ z \ L\<^isub>2) \ + (\ za \ z. (x @ za) \ L\<^isub>1 \ (z - za) \ L\<^isub>2)" using xz_in_seq append_seq_elim by simp moreover { fix xa @@ -928,7 +1308,8 @@ have "\ ya. ya \ y \ ya \ L\<^isub>1 \ (x - xa) \L\<^isub>2 (y - ya)" proof - have "{\L\<^isub>2 `` {x - xa} |xa. xa \ x \ xa \ L\<^isub>1} = - {\L\<^isub>2 `` {y - xa} |xa. xa \ y \ xa \ L\<^isub>1}" (is "?Left = ?Right") + {\L\<^isub>2 `` {y - xa} |xa. xa \ y \ xa \ L\<^isub>1}" + (is "?Left = ?Right") using h1 tag_xy by (auto simp:tag_str_SEQ_def) moreover have "\L\<^isub>2 `` {x - xa} \ ?Left" using h1 h2 by auto ultimately have "\L\<^isub>2 `` {x - xa} \ ?Right" by simp @@ -942,10 +1323,13 @@ assume h1: "za \ z" and h2: "(x @ za) \ L\<^isub>1" and h3: "z - za \ L\<^isub>2" hence "y @ za \ L\<^isub>1" proof- - have "\L\<^isub>1 `` {x} = \L\<^isub>1 `` {y}" using h1 tag_xy by (auto simp:tag_str_SEQ_def) - with h2 show ?thesis by (auto simp:Image_def str_eq_rel_def str_eq_def) + have "\L\<^isub>1 `` {x} = \L\<^isub>1 `` {y}" + using h1 tag_xy by (auto simp:tag_str_SEQ_def) + with h2 show ?thesis + by (auto simp:Image_def str_eq_rel_def str_eq_def) qed - with h1 h3 have "y @ z \ L\<^isub>1 ;; L\<^isub>2" by (drule_tac A = L\<^isub>1 in seq_intro, auto elim:prefixE) + with h1 h3 have "y @ z \ L\<^isub>1 ;; L\<^isub>2" + by (drule_tac A = L\<^isub>1 in seq_intro, auto elim:prefixE) } ultimately show ?thesis by blast qed @@ -958,7 +1342,8 @@ and finite2: "finite (UNIV // \L\<^isub>2)" shows "finite (UNIV // \(L\<^isub>1 ;; L\<^isub>2))" proof(rule_tac f = "(op `) (tag_str_SEQ L\<^isub>1 L\<^isub>2)" in finite_imageD) - show "finite (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2) ` UNIV // \L\<^isub>1 ;; L\<^isub>2)" using finite1 finite2 + show "finite (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2) ` UNIV // \L\<^isub>1 ;; L\<^isub>2)" + using finite1 finite2 by (auto intro:finite_tag_imageI tag_str_seq_range_finite) next show "inj_on (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2)) (UNIV // \L\<^isub>1 ;; L\<^isub>2)" @@ -967,13 +1352,14 @@ by (auto intro:tag_image_injI tag_str_SEQ_injI simp:) qed -text {* **************** the ALT case ************************ *} +subsection {* The case for @{text "ALT"} *} definition "tag_str_ALT L\<^isub>1 L\<^isub>2 (x::string) \ ((\L\<^isub>1) `` {x}, (\L\<^isub>2) `` {x})" lemma tag_str_alt_range_finite: - "\finite (UNIV // \L\<^isub>1); finite (UNIV // \L\<^isub>2)\ \ finite (range (tag_str_ALT L\<^isub>1 L\<^isub>2))" + "\finite (UNIV // \L\<^isub>1); finite (UNIV // \L\<^isub>2)\ + \ finite (range (tag_str_ALT L\<^isub>1 L\<^isub>2))" apply (rule_tac B = "(UNIV // \L\<^isub>1) \ (UNIV // \L\<^isub>2)" in finite_subset) by (auto simp:tag_str_ALT_def Image_def quotient_def) @@ -982,20 +1368,33 @@ and finite2: "finite (UNIV // \L\<^isub>2)" shows "finite (UNIV // \(L\<^isub>1 \ L\<^isub>2))" proof(rule_tac f = "(op `) (tag_str_ALT L\<^isub>1 L\<^isub>2)" in finite_imageD) - show "finite (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2) ` UNIV // \L\<^isub>1 \ L\<^isub>2)" using finite1 finite2 + show "finite (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2) ` UNIV // \L\<^isub>1 \ L\<^isub>2)" + using finite1 finite2 by (auto intro:finite_tag_imageI tag_str_alt_range_finite) next show "inj_on (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2)) (UNIV // \L\<^isub>1 \ L\<^isub>2)" proof- - have "\m n. tag_str_ALT L\<^isub>1 L\<^isub>2 m = tag_str_ALT L\<^isub>1 L\<^isub>2 n \ m \(L\<^isub>1 \ L\<^isub>2) n" + have "\m n. tag_str_ALT L\<^isub>1 L\<^isub>2 m = tag_str_ALT L\<^isub>1 L\<^isub>2 n + \ m \(L\<^isub>1 \ L\<^isub>2) n" unfolding tag_str_ALT_def str_eq_def Image_def str_eq_rel_def by auto thus ?thesis by (auto intro:tag_image_injI) qed qed -text {* **************** the Star case ****************** *} + +subsection {* + The case for @{text "STAR"} + *} -lemma finite_set_has_max: "\finite A; A \ {}\ \ (\ max \ A. \ a \ A. f a <= (f max :: nat))" +text {* + This turned out to be the most tricky case. + *} (* I will make some illustrations for it. *) + +definition + "tag_str_STAR L\<^isub>1 x \ {(\L\<^isub>1) `` {x - xa} | xa. xa < x \ xa \ L\<^isub>1\}" + +lemma finite_set_has_max: "\finite A; A \ {}\ \ + (\ max \ A. \ a \ A. f a <= (f max :: nat))" proof (induct rule:finite.induct) case emptyI thus ?case by simp next @@ -1005,7 +1404,9 @@ case True thus ?thesis by (rule_tac x = a in bexI, auto) next case False - with prems obtain max where h1: "max \ A" and h2: "\a\A. f a \ f max" by blast + with prems obtain max + where h1: "max \ A" + and h2: "\a\A. f a \ f max" by blast show ?thesis proof (cases "f a \ f max") assume "f a \ f max" @@ -1017,30 +1418,17 @@ qed qed -lemma star_intro1[rule_format]: "x \ lang\ \ \ y. y \ lang\ \ x @ y \ lang\" -by (erule Star.induct, auto) - -lemma star_intro2: "y \ lang \ y \ lang\" -by (drule step[of y lang "[]"], auto simp:start) - -lemma star_intro3[rule_format]: "x \ lang\ \ \y . y \ lang \ x @ y \ lang\" -by (erule Star.induct, auto intro:star_intro2) - -lemma star_decom: "\x \ lang\; x \ []\ \(\ a b. x = a @ b \ a \ [] \ a \ lang \ b \ lang\)" -by (induct x rule: Star.induct, simp, blast) - lemma finite_strict_prefix_set: "finite {xa. xa < (x::string)}" apply (induct x rule:rev_induct, simp) apply (subgoal_tac "{xa. xa < xs @ [x]} = {xa. xa < xs} \ {xs}") by (auto simp:strict_prefix_def) -definition - "tag_str_STAR L\<^isub>1 x \ {(\L\<^isub>1) `` {x - xa} | xa. xa < x \ xa \ L\<^isub>1\}" lemma tag_str_star_range_finite: "finite (UNIV // \L\<^isub>1) \ finite (range (tag_str_STAR L\<^isub>1))" apply (rule_tac B = "Pow (UNIV // \L\<^isub>1)" in finite_subset) -by (auto simp:tag_str_STAR_def Image_def quotient_def split:if_splits) +by (auto simp:tag_str_STAR_def Image_def + quotient_def split:if_splits) lemma tag_str_STAR_injI: "tag_str_STAR L\<^isub>1 m = tag_str_STAR L\<^isub>1 n \ m \(L\<^isub>1\) n" @@ -1051,58 +1439,76 @@ have "y @ z \ L\<^isub>1\" proof(cases "x = []") case True - with tag_xy have "y = []" by (auto simp:tag_str_STAR_def strict_prefix_def) + with tag_xy have "y = []" + by (auto simp:tag_str_STAR_def strict_prefix_def) thus ?thesis using xz_in_star True by simp next case False - obtain x_max where h1: "x_max < x" and h2: "x_max \ L\<^isub>1\" and h3: "(x - x_max) @ z \ L\<^isub>1\" - and h4:"\ xa < x. xa \ L\<^isub>1\ \ (x - xa) @ z \ L\<^isub>1\ \ length xa \ length x_max" + obtain x_max + where h1: "x_max < x" + and h2: "x_max \ L\<^isub>1\" + and h3: "(x - x_max) @ z \ L\<^isub>1\" + and h4:"\ xa < x. xa \ L\<^isub>1\ \ (x - xa) @ z \ L\<^isub>1\ + \ length xa \ length x_max" proof- let ?S = "{xa. xa < x \ xa \ L\<^isub>1\ \ (x - xa) @ z \ L\<^isub>1\}" have "finite ?S" - by (rule_tac B = "{xa. xa < x}" in finite_subset, auto simp:finite_strict_prefix_set) + by (rule_tac B = "{xa. xa < x}" in finite_subset, + auto simp:finite_strict_prefix_set) moreover have "?S \ {}" using False xz_in_star by (simp, rule_tac x = "[]" in exI, auto simp:strict_prefix_def) - ultimately have "\ max \ ?S. \ a \ ?S. length a \ length max" using finite_set_has_max by blast + ultimately have "\ max \ ?S. \ a \ ?S. length a \ length max" + using finite_set_has_max by blast with prems show ?thesis by blast qed - obtain ya where h5: "ya < y" and h6: "ya \ L\<^isub>1\" and h7: "(x - x_max) \L\<^isub>1 (y - ya)" + obtain ya + where h5: "ya < y" and h6: "ya \ L\<^isub>1\" and h7: "(x - x_max) \L\<^isub>1 (y - ya)" proof- from tag_xy have "{\L\<^isub>1 `` {x - xa} |xa. xa < x \ xa \ L\<^isub>1\} = {\L\<^isub>1 `` {y - xa} |xa. xa < y \ xa \ L\<^isub>1\}" (is "?left = ?right") by (auto simp:tag_str_STAR_def) moreover have "\L\<^isub>1 `` {x - x_max} \ ?left" using h1 h2 by auto ultimately have "\L\<^isub>1 `` {x - x_max} \ ?right" by simp - with prems show ?thesis apply (simp add:Image_def str_eq_rel_def str_eq_def) by blast + with prems show ?thesis apply + (simp add:Image_def str_eq_rel_def str_eq_def) by blast qed have "(y - ya) @ z \ L\<^isub>1\" proof- - from h3 h1 obtain a b where a_in: "a \ L\<^isub>1" and a_neq: "a \ []" and b_in: "b \ L\<^isub>1\" + from h3 h1 obtain a b where a_in: "a \ L\<^isub>1" + and a_neq: "a \ []" and b_in: "b \ L\<^isub>1\" and ab_max: "(x - x_max) @ z = a @ b" by (drule_tac star_decom, auto simp:strict_prefix_def elim:prefixE) have "(x - x_max) \ a \ (a - (x - x_max)) @ b = z" proof - - have "((x - x_max) \ a \ (a - (x - x_max)) @ b = z) \ (a < (x - x_max) \ ((x - x_max) - a) @ z = b)" + have "((x - x_max) \ a \ (a - (x - x_max)) @ b = z) \ + (a < (x - x_max) \ ((x - x_max) - a) @ z = b)" using app_eq_dest[OF ab_max] by (auto simp:strict_prefix_def) moreover { assume np: "a < (x - x_max)" and b_eqs: " ((x - x_max) - a) @ z = b" have "False" proof - let ?x_max' = "x_max @ a" - have "?x_max' < x" using np h1 by (clarsimp simp:strict_prefix_def diff_prefix) - moreover have "?x_max' \ L\<^isub>1\" using a_in h2 by (simp add:star_intro3) - moreover have "(x - ?x_max') @ z \ L\<^isub>1\" using b_eqs b_in np h1 by (simp add:diff_diff_appd) - moreover have "\ (length ?x_max' \ length x_max)" using a_neq by simp + have "?x_max' < x" + using np h1 by (clarsimp simp:strict_prefix_def diff_prefix) + moreover have "?x_max' \ L\<^isub>1\" + using a_in h2 by (simp add:star_intro3) + moreover have "(x - ?x_max') @ z \ L\<^isub>1\" + using b_eqs b_in np h1 by (simp add:diff_diff_appd) + moreover have "\ (length ?x_max' \ length x_max)" + using a_neq by simp ultimately show ?thesis using h4 by blast qed } ultimately show ?thesis by blast qed - then obtain za where z_decom: "z = za @ b" and x_za: "(x - x_max) @ za \ L\<^isub>1" + then obtain za where z_decom: "z = za @ b" + and x_za: "(x - x_max) @ za \ L\<^isub>1" using a_in by (auto elim:prefixE) - from x_za h7 have "(y - ya) @ za \ L\<^isub>1" by (auto simp:str_eq_def) + from x_za h7 have "(y - ya) @ za \ L\<^isub>1" + by (auto simp:str_eq_def str_eq_rel_def) with z_decom b_in show ?thesis by (auto dest!:step[of "(y - ya) @ za"]) qed - with h5 h6 show ?thesis by (drule_tac star_intro1, auto simp:strict_prefix_def elim:prefixE) + with h5 h6 show ?thesis + by (drule_tac star_intro1, auto simp:strict_prefix_def elim:prefixE) qed } thus "tag_str_STAR L\<^isub>1 m = tag_str_STAR L\<^isub>1 n \ m \(L\<^isub>1\) n" by (auto simp add:str_eq_def str_eq_rel_def) @@ -1119,9 +1525,11 @@ by (auto intro:tag_image_injI tag_str_STAR_injI) qed -text {* **************** the Other Direction ************ *} +subsection {* + The main lemma + *} -lemma other_direction: +lemma easier_directio\: "Lang = L (r::rexp) \ finite (UNIV // (\Lang))" proof (induct arbitrary:Lang rule:rexp.induct) case NULL @@ -1130,27 +1538,33 @@ with prems show "?case" by (auto intro:finite_subset) next case EMPTY - have "UNIV // (\{[]}) \ {{[]}, UNIV - {[]}}" by (rule quot_empty_subset) + have "UNIV // (\{[]}) \ {{[]}, UNIV - {[]}}" + by (rule quot_empty_subset) with prems show ?case by (auto intro:finite_subset) next case (CHAR c) - have "UNIV // (\{[c]}) \ {{[]},{[c]}, UNIV - {[], [c]}}" by (rule quot_char_subset) + have "UNIV // (\{[c]}) \ {{[]},{[c]}, UNIV - {[], [c]}}" + by (rule quot_char_subset) with prems show ?case by (auto intro:finite_subset) next case (SEQ r\<^isub>1 r\<^isub>2) - have "\finite (UNIV // \(L r\<^isub>1)); finite (UNIV // \(L r\<^isub>2))\ \ finite (UNIV // \(L r\<^isub>1 ;; L r\<^isub>2))" + have "\finite (UNIV // \(L r\<^isub>1)); finite (UNIV // \(L r\<^isub>2))\ + \ finite (UNIV // \(L r\<^isub>1 ;; L r\<^isub>2))" by (erule quot_seq_finiteI, simp) with prems show ?case by simp next case (ALT r\<^isub>1 r\<^isub>2) - have "\finite (UNIV // \(L r\<^isub>1)); finite (UNIV // \(L r\<^isub>2))\ \ finite (UNIV // \(L r\<^isub>1 \ L r\<^isub>2))" + have "\finite (UNIV // \(L r\<^isub>1)); finite (UNIV // \(L r\<^isub>2))\ + \ finite (UNIV // \(L r\<^isub>1 \ L r\<^isub>2))" by (erule quot_union_finiteI, simp) with prems show ?case by simp next case (STAR r) - have "finite (UNIV // \(L r)) \ finite (UNIV // \((L r)\))" + have "finite (UNIV // \(L r)) + \ finite (UNIV // \((L r)\))" by (erule quot_star_finiteI) with prems show ?case by simp qed -end \ No newline at end of file +end + diff -r c64241fa4dff -r f5db9e08effc pres/IsaMakefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/pres/IsaMakefile Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,31 @@ + +## targets + +default: ListP +images: ListP +test: + +all: images test + + +## global settings + +SRC = $(ISABELLE_HOME)/src +OUT = $(ISABELLE_OUTPUT) +LOG = $(OUT)/log + +USEDIR = $(ISABELLE_TOOL) usedir -v true -i true -d pdf ## -D generated + + +## ListP + +ListP: $(OUT)/ListP + +$(OUT)/ListP: ROOT.ML document/root.tex *.thy + @$(USEDIR) -b HOL ListP + + +## clean + +clean: + @rm -f $(OUT)/ListP diff -r c64241fa4dff -r f5db9e08effc pres/ListP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/pres/ListP.thy Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,4 @@ +theory ListP +imports Main List_Prefix +begin +end diff -r c64241fa4dff -r f5db9e08effc pres/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/pres/ROOT.ML Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,6 @@ +(* + no_document use_thys ["This_Theory1", "This_Theory2"]; + use_thys ["That_Theory1", "That_Theory2", "That_Theory3"]; +*) + +use_thy "ListP";; diff -r c64241fa4dff -r f5db9e08effc pres/document/root.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/pres/document/root.tex Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,63 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +%\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage[greek,english]{babel} + %option greek for \ + %option english (default language) for \, \ + +%\usepackage[latin1]{inputenc} + %for \, \, \, \, + %\, \, \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} + + +\begin{document} + +\title{ListP} +\author{By xingyuan} +\maketitle + +\tableofcontents + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +\input{session} + +% optional bibliography +%\bibliographystyle{abbrv} +%\bibliography{root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff -r c64241fa4dff -r f5db9e08effc tphols-2011/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/tphols-2011/ROOT.ML Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,5 @@ +(* + no_document use_thys ["This_Theory1", "This_Theory2"]; + use_thys ["That_Theory1", "That_Theory2", "That_Theory3"]; +*) +use_thys ["../Myhill"];; diff -r c64241fa4dff -r f5db9e08effc tphols-2011/document/root.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/tphols-2011/document/root.tex Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,65 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} +\usepackage{amsmath} +%\usepackage[pdftex]{hyperref} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +%\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage[greek,english]{babel} + %option greek for \ + %option english (default language) for \, \ + +%\usepackage[latin1]{inputenc} + %for \, \, \, \, + %\, \, \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} + + +\begin{document} + +\title{tphols-2011} +\author{By xingyuan} +\maketitle + +\tableofcontents + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +\input{session} + +% optional bibliography +%\bibliographystyle{abbrv} +%\bibliography{root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff -r c64241fa4dff -r f5db9e08effc tphols-2011/generated/List_Prefix.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/tphols-2011/generated/List_Prefix.tex Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,1578 @@ +% +\begin{isabellebody}% +\def\isabellecontext{List{\isacharunderscore}Prefix}% +% +\isamarkupheader{List prefixes and postfixes% +} +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ List{\isacharunderscore}Prefix\isanewline +\isakeyword{imports}\ List\ Main\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupsubsection{Prefix order on lists% +} +\isamarkuptrue% +\isacommand{instantiation}\isamarkupfalse% +\ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}type{\isacharparenright}\ order\isanewline +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ prefix{\isacharunderscore}def\ {\isacharbrackleft}code\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isasymle}\ ys\ {\isacharequal}\ {\isacharparenleft}{\isasymexists}zs{\isachardot}\ ys\ {\isacharequal}\ xs\ {\isacharat}\ zs{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ strict{\isacharunderscore}prefix{\isacharunderscore}def\ {\isacharbrackleft}code\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isacharless}\ ys\ {\isacharequal}\ {\isacharparenleft}xs\ {\isasymle}\ ys\ {\isasymand}\ xs\ {\isasymnoteq}\ {\isacharparenleft}ys{\isacharcolon}{\isacharcolon}{\isacharprime}a\ list{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ intro{\isacharunderscore}classes\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ prefix{\isacharunderscore}def\ strict{\isacharunderscore}prefix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{end}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ prefixI\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}ys\ {\isacharequal}\ xs\ {\isacharat}\ zs\ {\isacharequal}{\isacharequal}{\isachargreater}\ xs\ {\isasymle}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ prefix{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ prefixE\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}xs\ {\isasymle}\ ys{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{obtains}\ zs\ \isakeyword{where}\ {\isachardoublequoteopen}ys\ {\isacharequal}\ xs\ {\isacharat}\ zs{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{using}\isamarkupfalse% +\ assms\ \isacommand{unfolding}\isamarkupfalse% +\ prefix{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ strict{\isacharunderscore}prefixI{\isacharprime}\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}ys\ {\isacharequal}\ xs\ {\isacharat}\ z\ {\isacharhash}\ zs\ {\isacharequal}{\isacharequal}{\isachargreater}\ xs\ {\isacharless}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ strict{\isacharunderscore}prefix{\isacharunderscore}def\ prefix{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ strict{\isacharunderscore}prefixE{\isacharprime}\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}xs\ {\isacharless}\ ys{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{obtains}\ z\ zs\ \isakeyword{where}\ {\isachardoublequoteopen}ys\ {\isacharequal}\ xs\ {\isacharat}\ z\ {\isacharhash}\ zs{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{from}\isamarkupfalse% +\ {\isacharbackquoteopen}xs\ {\isacharless}\ ys{\isacharbackquoteclose}\ \isacommand{obtain}\isamarkupfalse% +\ us\ \isakeyword{where}\ {\isachardoublequoteopen}ys\ {\isacharequal}\ xs\ {\isacharat}\ us{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}xs\ {\isasymnoteq}\ ys{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{unfolding}\isamarkupfalse% +\ strict{\isacharunderscore}prefix{\isacharunderscore}def\ prefix{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ blast\isanewline +\ \ \isacommand{with}\isamarkupfalse% +\ that\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ neq{\isacharunderscore}Nil{\isacharunderscore}conv{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ strict{\isacharunderscore}prefixI\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isasymle}\ ys\ {\isacharequal}{\isacharequal}{\isachargreater}\ xs\ {\isasymnoteq}\ ys\ {\isacharequal}{\isacharequal}{\isachargreater}\ xs\ {\isacharless}\ {\isacharparenleft}ys{\isacharcolon}{\isacharcolon}{\isacharprime}a\ list{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ strict{\isacharunderscore}prefix{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ strict{\isacharunderscore}prefixE\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \isakeyword{fixes}\ xs\ ys\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}xs\ {\isacharless}\ ys{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{obtains}\ {\isachardoublequoteopen}xs\ {\isasymle}\ ys{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}xs\ {\isasymnoteq}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{using}\isamarkupfalse% +\ assms\ \isacommand{unfolding}\isamarkupfalse% +\ strict{\isacharunderscore}prefix{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Basic properties of prefixes% +} +\isamarkuptrue% +\isacommand{theorem}\isamarkupfalse% +\ Nil{\isacharunderscore}prefix\ {\isacharbrackleft}iff{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}\ {\isasymle}\ xs{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ prefix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ prefix{\isacharunderscore}Nil\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}xs\ {\isasymle}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ xs{\isacharparenright}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ prefix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ prefix{\isacharunderscore}snoc\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}xs\ {\isasymle}\ ys\ {\isacharat}\ {\isacharbrackleft}y{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}xs\ {\isacharequal}\ ys\ {\isacharat}\ {\isacharbrackleft}y{\isacharbrackright}\ {\isasymor}\ xs\ {\isasymle}\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}xs\ {\isasymle}\ ys\ {\isacharat}\ {\isacharbrackleft}y{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{obtain}\isamarkupfalse% +\ zs\ \isakeyword{where}\ zs{\isacharcolon}\ {\isachardoublequoteopen}ys\ {\isacharat}\ {\isacharbrackleft}y{\isacharbrackright}\ {\isacharequal}\ xs\ {\isacharat}\ zs{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}xs\ {\isacharequal}\ ys\ {\isacharat}\ {\isacharbrackleft}y{\isacharbrackright}\ {\isasymor}\ xs\ {\isasymle}\ ys{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ append{\isacharunderscore}Nil{\isadigit{2}}\ butlast{\isacharunderscore}append\ butlast{\isacharunderscore}snoc\ prefixI\ zs{\isacharparenright}\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}xs\ {\isacharequal}\ ys\ {\isacharat}\ {\isacharbrackleft}y{\isacharbrackright}\ {\isasymor}\ xs\ {\isasymle}\ ys{\isachardoublequoteclose}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}xs\ {\isasymle}\ ys\ {\isacharat}\ {\isacharbrackleft}y{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ order{\isacharunderscore}eq{\isacharunderscore}iff\ strict{\isacharunderscore}prefixE\ strict{\isacharunderscore}prefixI{\isacharprime}\ xt{\isadigit{1}}{\isacharparenleft}{\isadigit{7}}{\isacharparenright}{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ Cons{\isacharunderscore}prefix{\isacharunderscore}Cons\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isacharhash}\ xs\ {\isasymle}\ y\ {\isacharhash}\ ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharequal}\ y\ {\isasymand}\ xs\ {\isasymle}\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ prefix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ same{\isacharunderscore}prefix{\isacharunderscore}prefix\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}xs\ {\isacharat}\ ys\ {\isasymle}\ xs\ {\isacharat}\ zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}ys\ {\isasymle}\ zs{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ same{\isacharunderscore}prefix{\isacharunderscore}nil\ {\isacharbrackleft}iff{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}xs\ {\isacharat}\ ys\ {\isasymle}\ xs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}ys\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ append{\isacharunderscore}Nil{\isadigit{2}}\ append{\isacharunderscore}self{\isacharunderscore}conv\ order{\isacharunderscore}eq{\isacharunderscore}iff\ prefixI{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ prefix{\isacharunderscore}prefix\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isasymle}\ ys\ {\isacharequal}{\isacharequal}{\isachargreater}\ xs\ {\isasymle}\ ys\ {\isacharat}\ zs{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ order{\isacharunderscore}le{\isacharunderscore}less{\isacharunderscore}trans\ prefixI\ strict{\isacharunderscore}prefixE\ strict{\isacharunderscore}prefixI{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ append{\isacharunderscore}prefixD{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isacharat}\ ys\ {\isasymle}\ zs\ {\isasymLongrightarrow}\ xs\ {\isasymle}\ zs{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ prefix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ prefix{\isacharunderscore}Cons{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}xs\ {\isasymle}\ y\ {\isacharhash}\ ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymor}\ {\isacharparenleft}{\isasymexists}zs{\isachardot}\ xs\ {\isacharequal}\ y\ {\isacharhash}\ zs\ {\isasymand}\ zs\ {\isasymle}\ ys{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}cases\ xs{\isacharparenright}\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ prefix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ prefix{\isacharunderscore}append{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}xs\ {\isasymle}\ ys\ {\isacharat}\ zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}xs\ {\isasymle}\ ys\ {\isasymor}\ {\isacharparenleft}{\isasymexists}us{\isachardot}\ xs\ {\isacharequal}\ ys\ {\isacharat}\ us\ {\isasymand}\ us\ {\isasymle}\ zs{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}induct\ zs\ rule{\isacharcolon}\ rev{\isacharunderscore}induct{\isacharparenright}\isanewline +\ \ \ \isacommand{apply}\isamarkupfalse% +\ force\isanewline +\ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}simp\ del{\isacharcolon}\ append{\isacharunderscore}assoc\ add{\isacharcolon}\ append{\isacharunderscore}assoc\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}metis\ append{\isacharunderscore}eq{\isacharunderscore}appendI{\isacharparenright}\isanewline +\ \ \isacommand{done}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ append{\isacharunderscore}one{\isacharunderscore}prefix{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}xs\ {\isasymle}\ ys\ {\isacharequal}{\isacharequal}{\isachargreater}\ length\ xs\ {\isacharless}\ length\ ys\ {\isacharequal}{\isacharequal}{\isachargreater}\ xs\ {\isacharat}\ {\isacharbrackleft}ys\ {\isacharbang}\ length\ xs{\isacharbrackright}\ {\isasymle}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ prefix{\isacharunderscore}def\isanewline +\ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ Cons{\isacharunderscore}eq{\isacharunderscore}appendI\ append{\isacharunderscore}eq{\isacharunderscore}appendI\ append{\isacharunderscore}eq{\isacharunderscore}conv{\isacharunderscore}conj\isanewline +\ \ \ \ eq{\isacharunderscore}Nil{\isacharunderscore}appendI\ nth{\isacharunderscore}drop{\isacharprime}{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ prefix{\isacharunderscore}length{\isacharunderscore}le{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isasymle}\ ys\ {\isacharequal}{\isacharequal}{\isachargreater}\ length\ xs\ {\isasymle}\ length\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ prefix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ prefix{\isacharunderscore}same{\isacharunderscore}cases{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}xs\isactrlisub {\isadigit{1}}{\isacharcolon}{\isacharcolon}{\isacharprime}a\ list{\isacharparenright}\ {\isasymle}\ ys\ {\isasymLongrightarrow}\ xs\isactrlisub {\isadigit{2}}\ {\isasymle}\ ys\ {\isasymLongrightarrow}\ xs\isactrlisub {\isadigit{1}}\ {\isasymle}\ xs\isactrlisub {\isadigit{2}}\ {\isasymor}\ xs\isactrlisub {\isadigit{2}}\ {\isasymle}\ xs\isactrlisub {\isadigit{1}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ prefix{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ append{\isacharunderscore}eq{\isacharunderscore}append{\isacharunderscore}conv{\isadigit{2}}{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ set{\isacharunderscore}mono{\isacharunderscore}prefix{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isasymle}\ ys\ {\isasymLongrightarrow}\ set\ xs\ {\isasymsubseteq}\ set\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ prefix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ take{\isacharunderscore}is{\isacharunderscore}prefix{\isacharcolon}\ {\isachardoublequoteopen}take\ n\ xs\ {\isasymle}\ xs{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ prefix{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ append{\isacharunderscore}take{\isacharunderscore}drop{\isacharunderscore}id{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ map{\isacharunderscore}prefixI{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isasymle}\ ys\ {\isasymLongrightarrow}\ map\ f\ xs\ {\isasymle}\ map\ f\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp{\isacharcolon}\ prefix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ prefix{\isacharunderscore}length{\isacharunderscore}less{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isacharless}\ ys\ {\isasymLongrightarrow}\ length\ xs\ {\isacharless}\ length\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp{\isacharcolon}\ strict{\isacharunderscore}prefix{\isacharunderscore}def\ prefix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ strict{\isacharunderscore}prefix{\isacharunderscore}simps\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}xs\ {\isacharless}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ False{\isachardoublequoteclose}\isanewline +\ \ \ \ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}\ {\isacharless}\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequoteclose}\isanewline +\ \ \ \ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharequal}\ y\ {\isasymand}\ xs\ {\isacharless}\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ strict{\isacharunderscore}prefix{\isacharunderscore}def\ cong{\isacharcolon}\ conj{\isacharunderscore}cong{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ take{\isacharunderscore}strict{\isacharunderscore}prefix{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isacharless}\ ys\ {\isasymLongrightarrow}\ take\ n\ xs\ {\isacharless}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}induct\ n\ arbitrary{\isacharcolon}\ xs\ ys{\isacharparenright}\isanewline +\ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}case{\isacharunderscore}tac\ ys{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}{\isacharbrackleft}{\isadigit{1}}{\isacharbrackright}\isanewline +\ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}metis\ order{\isacharunderscore}less{\isacharunderscore}trans\ strict{\isacharunderscore}prefixI\ take{\isacharunderscore}is{\isacharunderscore}prefix{\isacharparenright}\isanewline +\ \ \isacommand{done}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ not{\isacharunderscore}prefix{\isacharunderscore}cases{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ pfx{\isacharcolon}\ {\isachardoublequoteopen}{\isasymnot}\ ps\ {\isasymle}\ ls{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{obtains}\isanewline +\ \ \ \ {\isacharparenleft}c{\isadigit{1}}{\isacharparenright}\ {\isachardoublequoteopen}ps\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}ls\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isacharparenleft}c{\isadigit{2}}{\isacharparenright}\ a\ as\ x\ xs\ \isakeyword{where}\ {\isachardoublequoteopen}ps\ {\isacharequal}\ a{\isacharhash}as{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}ls\ {\isacharequal}\ x{\isacharhash}xs{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}x\ {\isacharequal}\ a{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}{\isasymnot}\ as\ {\isasymle}\ xs{\isachardoublequoteclose}\isanewline +\ \ {\isacharbar}\ {\isacharparenleft}c{\isadigit{3}}{\isacharparenright}\ a\ as\ x\ xs\ \isakeyword{where}\ {\isachardoublequoteopen}ps\ {\isacharequal}\ a{\isacharhash}as{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}ls\ {\isacharequal}\ x{\isacharhash}xs{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}x\ {\isasymnoteq}\ a{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}cases\ ps{\isacharparenright}\isanewline +\ \ \isacommand{case}\isamarkupfalse% +\ Nil\ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{using}\isamarkupfalse% +\ pfx\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{case}\isamarkupfalse% +\ {\isacharparenleft}Cons\ a\ as{\isacharparenright}\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ c\ {\isacharequal}\ {\isacharbackquoteopen}ps\ {\isacharequal}\ a{\isacharhash}as{\isacharbackquoteclose}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}cases\ ls{\isacharparenright}\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ Nil\ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ append{\isacharunderscore}Nil{\isadigit{2}}\ pfx\ c{\isadigit{1}}\ same{\isacharunderscore}prefix{\isacharunderscore}nil{\isacharparenright}\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ {\isacharparenleft}Cons\ x\ xs{\isacharparenright}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}cases\ {\isachardoublequoteopen}x\ {\isacharequal}\ a{\isachardoublequoteclose}{\isacharparenright}\isanewline +\ \ \ \ \ \ \isacommand{case}\isamarkupfalse% +\ True\isanewline +\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymnot}\ as\ {\isasymle}\ xs{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% +\ pfx\ c\ Cons\ True\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \ \ \isacommand{with}\isamarkupfalse% +\ c\ Cons\ True\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ c{\isadigit{2}}{\isacharparenright}\isanewline +\ \ \ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{case}\isamarkupfalse% +\ False\isanewline +\ \ \ \ \ \ \isacommand{with}\isamarkupfalse% +\ c\ Cons\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ c{\isadigit{3}}{\isacharparenright}\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ not{\isacharunderscore}prefix{\isacharunderscore}induct\ {\isacharbrackleft}consumes\ {\isadigit{1}}{\isacharcomma}\ case{\isacharunderscore}names\ Nil\ Neq\ Eq{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ np{\isacharcolon}\ {\isachardoublequoteopen}{\isasymnot}\ ps\ {\isasymle}\ ls{\isachardoublequoteclose}\isanewline +\ \ \ \ \isakeyword{and}\ base{\isacharcolon}\ {\isachardoublequoteopen}{\isasymAnd}x\ xs{\isachardot}\ P\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isakeyword{and}\ r{\isadigit{1}}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymAnd}x\ xs\ y\ ys{\isachardot}\ x\ {\isasymnoteq}\ y\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}y{\isacharhash}ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \isakeyword{and}\ r{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymAnd}x\ xs\ y\ ys{\isachardot}\ {\isasymlbrakk}\ x\ {\isacharequal}\ y{\isacharsemicolon}\ {\isasymnot}\ xs\ {\isasymle}\ ys{\isacharsemicolon}\ P\ xs\ ys\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharparenleft}y{\isacharhash}ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}P\ ps\ ls{\isachardoublequoteclose}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{using}\isamarkupfalse% +\ np\isanewline +\isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}induct\ ls\ arbitrary{\isacharcolon}\ ps{\isacharparenright}\isanewline +\ \ \isacommand{case}\isamarkupfalse% +\ Nil\ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}case\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp{\isacharcolon}\ neq{\isacharunderscore}Nil{\isacharunderscore}conv\ elim{\isacharbang}{\isacharcolon}\ not{\isacharunderscore}prefix{\isacharunderscore}cases\ intro{\isacharbang}{\isacharcolon}\ base{\isacharparenright}\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{case}\isamarkupfalse% +\ {\isacharparenleft}Cons\ y\ ys{\isacharparenright}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ npfx{\isacharcolon}\ {\isachardoublequoteopen}{\isasymnot}\ ps\ {\isasymle}\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{obtain}\isamarkupfalse% +\ x\ xs\ \isakeyword{where}\ pv{\isacharcolon}\ {\isachardoublequoteopen}ps\ {\isacharequal}\ x\ {\isacharhash}\ xs{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ not{\isacharunderscore}prefix{\isacharunderscore}cases{\isacharparenright}\ auto\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}case\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ Cons{\isachardot}hyps\ Cons{\isacharunderscore}prefix{\isacharunderscore}Cons\ npfx\ pv\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Parallel lists% +} +\isamarkuptrue% +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ parallel\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isacharequal}{\isachargreater}\ {\isacharprime}a\ list\ {\isacharequal}{\isachargreater}\ bool{\isachardoublequoteclose}\ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymparallel}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}xs\ {\isasymparallel}\ ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isasymnot}\ xs\ {\isasymle}\ ys\ {\isasymand}\ {\isasymnot}\ ys\ {\isasymle}\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ parallelI\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymnot}\ xs\ {\isasymle}\ ys\ {\isacharequal}{\isacharequal}{\isachargreater}\ {\isasymnot}\ ys\ {\isasymle}\ xs\ {\isacharequal}{\isacharequal}{\isachargreater}\ xs\ {\isasymparallel}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ parallel{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ parallelE\ {\isacharbrackleft}elim{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}xs\ {\isasymparallel}\ ys{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{obtains}\ {\isachardoublequoteopen}{\isasymnot}\ xs\ {\isasymle}\ ys\ {\isasymand}\ {\isasymnot}\ ys\ {\isasymle}\ xs{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{using}\isamarkupfalse% +\ assms\ \isacommand{unfolding}\isamarkupfalse% +\ parallel{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ prefix{\isacharunderscore}cases{\isacharcolon}\isanewline +\ \ \isakeyword{obtains}\ {\isachardoublequoteopen}xs\ {\isasymle}\ ys{\isachardoublequoteclose}\ {\isacharbar}\ {\isachardoublequoteopen}ys\ {\isacharless}\ xs{\isachardoublequoteclose}\ {\isacharbar}\ {\isachardoublequoteopen}xs\ {\isasymparallel}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ parallel{\isacharunderscore}def\ strict{\isacharunderscore}prefix{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{theorem}\isamarkupfalse% +\ parallel{\isacharunderscore}decomp{\isacharcolon}\isanewline +\ \ {\isachardoublequoteopen}xs\ {\isasymparallel}\ ys\ {\isacharequal}{\isacharequal}{\isachargreater}\ {\isasymexists}as\ b\ bs\ c\ cs{\isachardot}\ b\ {\isasymnoteq}\ c\ {\isasymand}\ xs\ {\isacharequal}\ as\ {\isacharat}\ b\ {\isacharhash}\ bs\ {\isasymand}\ ys\ {\isacharequal}\ as\ {\isacharat}\ c\ {\isacharhash}\ cs{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}induct\ xs\ rule{\isacharcolon}\ rev{\isacharunderscore}induct{\isacharparenright}\isanewline +\ \ \isacommand{case}\isamarkupfalse% +\ Nil\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ False\ \isacommand{by}\isamarkupfalse% +\ auto\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}case\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{case}\isamarkupfalse% +\ {\isacharparenleft}snoc\ x\ xs{\isacharparenright}\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}case\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}rule\ prefix{\isacharunderscore}cases{\isacharparenright}\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ le{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isasymle}\ ys{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{obtain}\isamarkupfalse% +\ ys{\isacharprime}\ \isakeyword{where}\ ys{\isacharcolon}\ {\isachardoublequoteopen}ys\ {\isacharequal}\ xs\ {\isacharat}\ ys{\isacharprime}{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}cases\ ys{\isacharprime}{\isacharparenright}\isanewline +\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}ys{\isacharprime}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ append{\isacharunderscore}Nil{\isadigit{2}}\ parallelE\ prefixI\ snoc{\isachardot}prems\ ys{\isacharparenright}\isanewline +\ \ \ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ c\ cs\ \isacommand{assume}\isamarkupfalse% +\ ys{\isacharprime}{\isacharcolon}\ {\isachardoublequoteopen}ys{\isacharprime}\ {\isacharequal}\ c\ {\isacharhash}\ cs{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\isanewline +\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ Cons{\isacharunderscore}eq{\isacharunderscore}appendI\ eq{\isacharunderscore}Nil{\isacharunderscore}appendI\ parallelE\ prefixI\isanewline +\ \ \ \ \ \ \ \ \ \ same{\isacharunderscore}prefix{\isacharunderscore}prefix\ snoc{\isachardot}prems\ ys{\isacharparenright}\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}ys\ {\isacharless}\ xs{\isachardoublequoteclose}\ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}ys\ {\isasymle}\ xs\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ strict{\isacharunderscore}prefix{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \ \ \isacommand{with}\isamarkupfalse% +\ snoc\ \isacommand{have}\isamarkupfalse% +\ False\ \isacommand{by}\isamarkupfalse% +\ blast\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}xs\ {\isasymparallel}\ ys{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{with}\isamarkupfalse% +\ snoc\ \isacommand{obtain}\isamarkupfalse% +\ as\ b\ bs\ c\ cs\ \isakeyword{where}\ neq{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}b{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isacharparenright}\ {\isasymnoteq}\ c{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isakeyword{and}\ xs{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isacharequal}\ as\ {\isacharat}\ b\ {\isacharhash}\ bs{\isachardoublequoteclose}\ \isakeyword{and}\ ys{\isacharcolon}\ {\isachardoublequoteopen}ys\ {\isacharequal}\ as\ {\isacharat}\ c\ {\isacharhash}\ cs{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{by}\isamarkupfalse% +\ blast\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ xs\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}xs\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}\ {\isacharequal}\ as\ {\isacharat}\ b\ {\isacharhash}\ {\isacharparenleft}bs\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \isacommand{with}\isamarkupfalse% +\ neq\ ys\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ blast\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ parallel{\isacharunderscore}append{\isacharcolon}\ {\isachardoublequoteopen}a\ {\isasymparallel}\ b\ {\isasymLongrightarrow}\ a\ {\isacharat}\ c\ {\isasymparallel}\ b\ {\isacharat}\ d{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}rule\ parallelI{\isacharparenright}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}erule\ parallelE{\isacharcomma}\ erule\ conjE{\isacharcomma}\isanewline +\ \ \ \ \ \ induct\ rule{\isacharcolon}\ not{\isacharunderscore}prefix{\isacharunderscore}induct{\isacharcomma}\ simp{\isacharplus}{\isacharparenright}{\isacharplus}\isanewline +\ \ \isacommand{done}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ parallel{\isacharunderscore}appendI{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isasymparallel}\ ys\ {\isasymLongrightarrow}\ x\ {\isacharequal}\ xs\ {\isacharat}\ xs{\isacharprime}\ {\isasymLongrightarrow}\ y\ {\isacharequal}\ ys\ {\isacharat}\ ys{\isacharprime}\ {\isasymLongrightarrow}\ x\ {\isasymparallel}\ y{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ parallel{\isacharunderscore}append{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ parallel{\isacharunderscore}commute{\isacharcolon}\ {\isachardoublequoteopen}a\ {\isasymparallel}\ b\ {\isasymlongleftrightarrow}\ b\ {\isasymparallel}\ a{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ parallel{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Postfix order on lists% +} +\isamarkuptrue% +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ postfix\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isacharequal}{\isachargreater}\ {\isacharprime}a\ list\ {\isacharequal}{\isachargreater}\ bool{\isachardoublequoteclose}\ \ {\isacharparenleft}{\isachardoublequoteopen}{\isacharparenleft}{\isacharunderscore}{\isacharslash}\ {\isachargreater}{\isachargreater}{\isacharequal}\ {\isacharunderscore}{\isacharparenright}{\isachardoublequoteclose}\ {\isacharbrackleft}{\isadigit{5}}{\isadigit{1}}{\isacharcomma}\ {\isadigit{5}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\isanewline +\ \ {\isachardoublequoteopen}{\isacharparenleft}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isasymexists}zs{\isachardot}\ xs\ {\isacharequal}\ zs\ {\isacharat}\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ postfixI\ {\isacharbrackleft}intro{\isacharquery}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isacharequal}\ zs\ {\isacharat}\ ys\ {\isacharequal}{\isacharequal}{\isachargreater}\ xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ postfix{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ postfixE\ {\isacharbrackleft}elim{\isacharquery}{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{obtains}\ zs\ \isakeyword{where}\ {\isachardoublequoteopen}xs\ {\isacharequal}\ zs\ {\isacharat}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{using}\isamarkupfalse% +\ assms\ \isacommand{unfolding}\isamarkupfalse% +\ postfix{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ postfix{\isacharunderscore}refl\ {\isacharbrackleft}iff{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ xs{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ postfix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{lemma}\isamarkupfalse% +\ postfix{\isacharunderscore}trans{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys{\isacharsemicolon}\ ys\ {\isachargreater}{\isachargreater}{\isacharequal}\ zs{\isasymrbrakk}\ {\isasymLongrightarrow}\ xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ zs{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ postfix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{lemma}\isamarkupfalse% +\ postfix{\isacharunderscore}antisym{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys{\isacharsemicolon}\ ys\ {\isachargreater}{\isachargreater}{\isacharequal}\ xs{\isasymrbrakk}\ {\isasymLongrightarrow}\ xs\ {\isacharequal}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ postfix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ Nil{\isacharunderscore}postfix\ {\isacharbrackleft}iff{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}simp\ add{\isacharcolon}\ postfix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{lemma}\isamarkupfalse% +\ postfix{\isacharunderscore}Nil\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharbrackleft}{\isacharbrackright}\ {\isachargreater}{\isachargreater}{\isacharequal}\ xs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ postfix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ postfix{\isacharunderscore}ConsI{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys\ {\isasymLongrightarrow}\ x{\isacharhash}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ postfix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{lemma}\isamarkupfalse% +\ postfix{\isacharunderscore}ConsD{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ y{\isacharhash}ys\ {\isasymLongrightarrow}\ xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ postfix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ postfix{\isacharunderscore}appendI{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys\ {\isasymLongrightarrow}\ zs\ {\isacharat}\ xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ postfix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{lemma}\isamarkupfalse% +\ postfix{\isacharunderscore}appendD{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ zs\ {\isacharat}\ ys\ {\isasymLongrightarrow}\ xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ postfix{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ postfix{\isacharunderscore}is{\isacharunderscore}subset{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys\ {\isacharequal}{\isacharequal}{\isachargreater}\ set\ ys\ {\isasymsubseteq}\ set\ xs{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys{\isachardoublequoteclose}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{obtain}\isamarkupfalse% +\ zs\ \isakeyword{where}\ {\isachardoublequoteopen}xs\ {\isacharequal}\ zs\ {\isacharat}\ ys{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ zs{\isacharparenright}\ auto\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ postfix{\isacharunderscore}ConsD{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}x{\isacharhash}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ y{\isacharhash}ys\ {\isacharequal}{\isacharequal}{\isachargreater}\ xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x{\isacharhash}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ y{\isacharhash}ys{\isachardoublequoteclose}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{obtain}\isamarkupfalse% +\ zs\ \isakeyword{where}\ {\isachardoublequoteopen}x{\isacharhash}xs\ {\isacharequal}\ zs\ {\isacharat}\ y{\isacharhash}ys{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}induct\ zs{\isacharparenright}\ {\isacharparenleft}auto\ intro{\isacharbang}{\isacharcolon}\ postfix{\isacharunderscore}appendI\ postfix{\isacharunderscore}ConsI{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ postfix{\isacharunderscore}to{\isacharunderscore}prefix{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys\ {\isasymlongleftrightarrow}\ rev\ ys\ {\isasymle}\ rev\ xs{\isachardoublequoteclose}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys{\isachardoublequoteclose}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{obtain}\isamarkupfalse% +\ zs\ \isakeyword{where}\ {\isachardoublequoteopen}xs\ {\isacharequal}\ zs\ {\isacharat}\ ys{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}rev\ xs\ {\isacharequal}\ rev\ ys\ {\isacharat}\ rev\ zs{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}rev\ ys\ {\isacharless}{\isacharequal}\ rev\ xs{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}rev\ ys\ {\isacharless}{\isacharequal}\ rev\ xs{\isachardoublequoteclose}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{obtain}\isamarkupfalse% +\ zs\ \isakeyword{where}\ {\isachardoublequoteopen}rev\ xs\ {\isacharequal}\ rev\ ys\ {\isacharat}\ zs{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}rev\ {\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ rev\ zs\ {\isacharat}\ rev\ {\isacharparenleft}rev\ ys{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}xs\ {\isacharequal}\ rev\ zs\ {\isacharat}\ ys{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ distinct{\isacharunderscore}postfix{\isacharcolon}\ {\isachardoublequoteopen}distinct\ xs\ {\isasymLongrightarrow}\ xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys\ {\isasymLongrightarrow}\ distinct\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}clarsimp\ elim{\isacharbang}{\isacharcolon}\ postfixE{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ postfix{\isacharunderscore}map{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys\ {\isasymLongrightarrow}\ map\ f\ xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ map\ f\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto\ elim{\isacharbang}{\isacharcolon}\ postfixE\ intro{\isacharcolon}\ postfixI{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ postfix{\isacharunderscore}drop{\isacharcolon}\ {\isachardoublequoteopen}as\ {\isachargreater}{\isachargreater}{\isacharequal}\ drop\ n\ as{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ postfix{\isacharunderscore}def\isanewline +\ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}rule\ exI\ {\isacharbrackleft}\isakeyword{where}\ x\ {\isacharequal}\ {\isachardoublequoteopen}take\ n\ as{\isachardoublequoteclose}{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \isacommand{apply}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{done}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ postfix{\isacharunderscore}take{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isachargreater}{\isachargreater}{\isacharequal}\ ys\ {\isasymLongrightarrow}\ xs\ {\isacharequal}\ take\ {\isacharparenleft}length\ xs\ {\isacharminus}\ length\ ys{\isacharparenright}\ xs\ {\isacharat}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}clarsimp\ elim{\isacharbang}{\isacharcolon}\ postfixE{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ parallelD{\isadigit{1}}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymparallel}\ y\ {\isasymLongrightarrow}\ {\isasymnot}\ x\ {\isasymle}\ y{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ parallelD{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymparallel}\ y\ {\isasymLongrightarrow}\ {\isasymnot}\ y\ {\isasymle}\ x{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ parallel{\isacharunderscore}Nil{\isadigit{1}}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymnot}\ x\ {\isasymparallel}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ parallel{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ parallel{\isacharunderscore}Nil{\isadigit{2}}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymnot}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymparallel}\ x{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ parallel{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ Cons{\isacharunderscore}parallelI{\isadigit{1}}{\isacharcolon}\ {\isachardoublequoteopen}a\ {\isasymnoteq}\ b\ {\isasymLongrightarrow}\ a\ {\isacharhash}\ as\ {\isasymparallel}\ b\ {\isacharhash}\ bs{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ Cons{\isacharunderscore}parallelI{\isadigit{2}}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}\ a\ {\isacharequal}\ b{\isacharsemicolon}\ as\ {\isasymparallel}\ bs\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ a\ {\isacharhash}\ as\ {\isasymparallel}\ b\ {\isacharhash}\ bs{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}metis\ Cons{\isacharunderscore}prefix{\isacharunderscore}Cons\ parallelE\ parallelI{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ not{\isacharunderscore}equal{\isacharunderscore}is{\isacharunderscore}parallel{\isacharcolon}\isanewline +\ \ \isakeyword{assumes}\ neq{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isasymnoteq}\ ys{\isachardoublequoteclose}\isanewline +\ \ \ \ \isakeyword{and}\ len{\isacharcolon}\ {\isachardoublequoteopen}length\ xs\ {\isacharequal}\ length\ ys{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequoteopen}xs\ {\isasymparallel}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{using}\isamarkupfalse% +\ len\ neq\isanewline +\isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}induct\ rule{\isacharcolon}\ list{\isacharunderscore}induct{\isadigit{2}}{\isacharparenright}\isanewline +\ \ \isacommand{case}\isamarkupfalse% +\ Nil\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}case\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{case}\isamarkupfalse% +\ {\isacharparenleft}Cons\ a\ as\ b\ bs{\isacharparenright}\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ ih{\isacharcolon}\ {\isachardoublequoteopen}as\ {\isasymnoteq}\ bs\ {\isasymLongrightarrow}\ as\ {\isasymparallel}\ bs{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\ \ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}case\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharparenleft}cases\ {\isachardoublequoteopen}a\ {\isacharequal}\ b{\isachardoublequoteclose}{\isacharparenright}\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ True\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}as\ {\isasymnoteq}\ bs{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% +\ Cons\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ Cons{\isacharunderscore}parallelI{\isadigit{2}}\ {\isacharbrackleft}OF\ True\ ih{\isacharbrackright}{\isacharparenright}\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ False\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ Cons{\isacharunderscore}parallelI{\isadigit{1}}{\isacharparenright}\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Executable code% +} +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ less{\isacharunderscore}eq{\isacharunderscore}code\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}{\isacharparenleft}{\isacharbrackleft}{\isacharbrackright}{\isasymColon}{\isacharprime}a{\isasymColon}{\isacharbraceleft}eq{\isacharcomma}\ ord{\isacharbraceright}\ list{\isacharparenright}\ {\isasymle}\ xs\ {\isasymlongleftrightarrow}\ True{\isachardoublequoteclose}\isanewline +\ \ \ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isasymColon}{\isacharbraceleft}eq{\isacharcomma}\ ord{\isacharbraceright}{\isacharparenright}\ {\isacharhash}\ xs\ {\isasymle}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongleftrightarrow}\ False{\isachardoublequoteclose}\isanewline +\ \ \ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isasymColon}{\isacharbraceleft}eq{\isacharcomma}\ ord{\isacharbraceright}{\isacharparenright}\ {\isacharhash}\ xs\ {\isasymle}\ y\ {\isacharhash}\ ys\ {\isasymlongleftrightarrow}\ x\ {\isacharequal}\ y\ {\isasymand}\ xs\ {\isasymle}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ simp{\isacharunderscore}all% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ less{\isacharunderscore}code\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline +\ \ \ \ {\isachardoublequoteopen}xs\ {\isacharless}\ {\isacharparenleft}{\isacharbrackleft}{\isacharbrackright}{\isasymColon}{\isacharprime}a{\isasymColon}{\isacharbraceleft}eq{\isacharcomma}\ ord{\isacharbraceright}\ list{\isacharparenright}\ {\isasymlongleftrightarrow}\ False{\isachardoublequoteclose}\isanewline +\ \ \ \ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}\ {\isacharless}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isasymColon}{\isacharbraceleft}eq{\isacharcomma}\ ord{\isacharbraceright}{\isacharparenright}{\isacharhash}\ xs\ {\isasymlongleftrightarrow}\ True{\isachardoublequoteclose}\isanewline +\ \ \ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isasymColon}{\isacharbraceleft}eq{\isacharcomma}\ ord{\isacharbraceright}{\isacharparenright}\ {\isacharhash}\ xs\ {\isacharless}\ y\ {\isacharhash}\ ys\ {\isasymlongleftrightarrow}\ x\ {\isacharequal}\ y\ {\isasymand}\ xs\ {\isacharless}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{unfolding}\isamarkupfalse% +\ strict{\isacharunderscore}prefix{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemmas}\isamarkupfalse% +\ {\isacharbrackleft}code{\isacharbrackright}\ {\isacharequal}\ postfix{\isacharunderscore}to{\isacharunderscore}prefix\isanewline +% +\isadelimtheory +\isanewline +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +\isanewline +% +\endisadelimtheory +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r c64241fa4dff -r f5db9e08effc tphols-2011/generated/Myhill.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/tphols-2011/generated/Myhill.tex Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,3856 @@ +% +\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: diff -r c64241fa4dff -r f5db9e08effc tphols-2011/generated/isabelle.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/tphols-2011/generated/isabelle.sty Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,215 @@ +%% +%% macros for Isabelle generated LaTeX output +%% + +%%% Simple document preparation (based on theory token language and symbols) + +% isabelle environments + +\newcommand{\isabellecontext}{UNKNOWN} + +\newcommand{\isastyle}{\UNDEF} +\newcommand{\isastyleminor}{\UNDEF} +\newcommand{\isastylescript}{\UNDEF} +\newcommand{\isastyletext}{\normalsize\rm} +\newcommand{\isastyletxt}{\rm} +\newcommand{\isastylecmt}{\rm} + +%symbol markup -- \emph achieves decent spacing via italic corrections +\newcommand{\isamath}[1]{\emph{$#1$}} +\newcommand{\isatext}[1]{\emph{#1}} +\DeclareRobustCommand{\isascriptstyle}{\def\isamath##1{##1}\def\isatext##1{\mbox{\isastylescript##1}}} +\newcommand{\isactrlsub}[1]{\emph{\isascriptstyle${}\sb{#1}$}} +\newcommand{\isactrlsup}[1]{\emph{\isascriptstyle${}\sp{#1}$}} +\newcommand{\isactrlisub}[1]{\emph{\isascriptstyle${}\sb{#1}$}} +\newcommand{\isactrlisup}[1]{\emph{\isascriptstyle${}\sp{#1}$}} +\DeclareRobustCommand{\isactrlbsub}{\emph\bgroup\math{}\sb\bgroup\mbox\bgroup\isastylescript} +\DeclareRobustCommand{\isactrlesub}{\egroup\egroup\endmath\egroup} +\DeclareRobustCommand{\isactrlbsup}{\emph\bgroup\math{}\sp\bgroup\mbox\bgroup\isastylescript} +\DeclareRobustCommand{\isactrlesup}{\egroup\egroup\endmath\egroup} +\newcommand{\isactrlbold}[1]{{\bfseries\upshape\boldmath#1}} +\newcommand{\isactrlloc}[1]{{\bfseries\upshape\boldmath#1}} + +\newenvironment{isaantiq}{{\isacharat\isacharbraceleft}}{{\isacharbraceright}} +\newcommand{\isaantiqopen}{\isakeyword{\isacharbraceleft}} +\newcommand{\isaantiqclose}{\isakeyword{\isacharbraceright}} + +\newdimen\isa@parindent\newdimen\isa@parskip + +\newenvironment{isabellebody}{% +\isamarkuptrue\par% +\isa@parindent\parindent\parindent0pt% +\isa@parskip\parskip\parskip0pt% +\isastyle}{\par} + +\newenvironment{isabelle} +{\begin{trivlist}\begin{isabellebody}\item\relax} +{\end{isabellebody}\end{trivlist}} + +\newcommand{\isa}[1]{\emph{\isastyleminor #1}} + +\newcommand{\isaindent}[1]{\hphantom{#1}} +\newcommand{\isanewline}{\mbox{}\par\mbox{}} +\newcommand{\isasep}{} +\newcommand{\isadigit}[1]{#1} + +\newcommand{\isachardefaults}{% +\chardef\isacharbang=`\!% +\chardef\isachardoublequote=`\"% +\chardef\isachardoublequoteopen=`\"% +\chardef\isachardoublequoteclose=`\"% +\chardef\isacharhash=`\#% +\chardef\isachardollar=`\$% +\chardef\isacharpercent=`\%% +\chardef\isacharampersand=`\&% +\chardef\isacharprime=`\'% +\chardef\isacharparenleft=`\(% +\chardef\isacharparenright=`\)% +\chardef\isacharasterisk=`\*% +\chardef\isacharplus=`\+% +\chardef\isacharcomma=`\,% +\chardef\isacharminus=`\-% +\chardef\isachardot=`\.% +\chardef\isacharslash=`\/% +\chardef\isacharcolon=`\:% +\chardef\isacharsemicolon=`\;% +\chardef\isacharless=`\<% +\chardef\isacharequal=`\=% +\chardef\isachargreater=`\>% +\chardef\isacharquery=`\?% +\chardef\isacharat=`\@% +\chardef\isacharbrackleft=`\[% +\chardef\isacharbackslash=`\\% +\chardef\isacharbrackright=`\]% +\chardef\isacharcircum=`\^% +\chardef\isacharunderscore=`\_% +\def\isacharunderscorekeyword{\_}% +\chardef\isacharbackquote=`\`% +\chardef\isacharbackquoteopen=`\`% +\chardef\isacharbackquoteclose=`\`% +\chardef\isacharbraceleft=`\{% +\chardef\isacharbar=`\|% +\chardef\isacharbraceright=`\}% +\chardef\isachartilde=`\~% +\def\isacharverbatimopen{\isacharbraceleft\isacharasterisk}% +\def\isacharverbatimclose{\isacharasterisk\isacharbraceright}% +} + + +% keyword and section markup + +\newcommand{\isakeyword}[1] +{\emph{\bf\def\isachardot{.}\def\isacharunderscore{\isacharunderscorekeyword}% +\def\isacharbraceleft{\{}\def\isacharbraceright{\}}#1}} +\newcommand{\isacommand}[1]{\isakeyword{#1}} + +\newcommand{\isamarkupheader}[1]{\section{#1}} +\newcommand{\isamarkupchapter}[1]{\chapter{#1}} +\newcommand{\isamarkupsection}[1]{\section{#1}} +\newcommand{\isamarkupsubsection}[1]{\subsection{#1}} +\newcommand{\isamarkupsubsubsection}[1]{\subsubsection{#1}} +\newcommand{\isamarkupsect}[1]{\section{#1}} +\newcommand{\isamarkupsubsect}[1]{\subsection{#1}} +\newcommand{\isamarkupsubsubsect}[1]{\subsubsection{#1}} + +\newif\ifisamarkup +\newcommand{\isabeginpar}{\par\ifisamarkup\relax\else\medskip\fi} +\newcommand{\isaendpar}{\par\medskip} +\newenvironment{isapar}{\parindent\isa@parindent\parskip\isa@parskip\isabeginpar}{\isaendpar} +\newenvironment{isamarkuptext}{\par\isastyletext\begin{isapar}}{\end{isapar}} +\newenvironment{isamarkuptxt}{\par\isastyletxt\begin{isapar}}{\end{isapar}} +\newcommand{\isamarkupcmt}[1]{{\isastylecmt--- #1}} + + +% styles + +\def\isabellestyle#1{\csname isabellestyle#1\endcsname} + +\newcommand{\isabellestyledefault}{% +\renewcommand{\isastyle}{\small\tt\slshape}% +\renewcommand{\isastyleminor}{\small\tt\slshape}% +\renewcommand{\isastylescript}{\footnotesize\tt\slshape}% +\isachardefaults% +} +\isabellestyledefault + +\newcommand{\isabellestylett}{% +\renewcommand{\isastyle}{\small\tt}% +\renewcommand{\isastyleminor}{\small\tt}% +\renewcommand{\isastylescript}{\footnotesize\tt}% +\isachardefaults% +} + +\newcommand{\isabellestyleit}{% +\renewcommand{\isastyle}{\small\it}% +\renewcommand{\isastyleminor}{\it}% +\renewcommand{\isastylescript}{\footnotesize\it}% +\renewcommand{\isacharunderscorekeyword}{\mbox{-}}% +\renewcommand{\isacharbang}{\isamath{!}}% +\renewcommand{\isachardoublequote}{}% +\renewcommand{\isachardoublequoteopen}{}% +\renewcommand{\isachardoublequoteclose}{}% +\renewcommand{\isacharhash}{\isamath{\#}}% +\renewcommand{\isachardollar}{\isamath{\$}}% +\renewcommand{\isacharpercent}{\isamath{\%}}% +\renewcommand{\isacharampersand}{\isamath{\&}}% +\renewcommand{\isacharprime}{\isamath{\mskip2mu{'}\mskip-2mu}}% +\renewcommand{\isacharparenleft}{\isamath{(}}% +\renewcommand{\isacharparenright}{\isamath{)}}% +\renewcommand{\isacharasterisk}{\isamath{*}}% +\renewcommand{\isacharplus}{\isamath{+}}% +\renewcommand{\isacharcomma}{\isamath{\mathord,}}% +\renewcommand{\isacharminus}{\isamath{-}}% +\renewcommand{\isachardot}{\isamath{\mathord.}}% +\renewcommand{\isacharslash}{\isamath{/}}% +\renewcommand{\isacharcolon}{\isamath{\mathord:}}% +\renewcommand{\isacharsemicolon}{\isamath{\mathord;}}% +\renewcommand{\isacharless}{\isamath{<}}% +\renewcommand{\isacharequal}{\isamath{=}}% +\renewcommand{\isachargreater}{\isamath{>}}% +\renewcommand{\isacharat}{\isamath{@}}% +\renewcommand{\isacharbrackleft}{\isamath{[}}% +\renewcommand{\isacharbackslash}{\isamath{\backslash}}% +\renewcommand{\isacharbrackright}{\isamath{]}}% +\renewcommand{\isacharunderscore}{\mbox{-}}% +\renewcommand{\isacharbraceleft}{\isamath{\{}}% +\renewcommand{\isacharbar}{\isamath{\mid}}% +\renewcommand{\isacharbraceright}{\isamath{\}}}% +\renewcommand{\isachartilde}{\isamath{{}\sp{\sim}}}% +\renewcommand{\isacharbackquoteopen}{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\langle$}}}% +\renewcommand{\isacharbackquoteclose}{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\rangle$}}}% +\renewcommand{\isacharverbatimopen}{\isamath{\langle\!\langle}}% +\renewcommand{\isacharverbatimclose}{\isamath{\rangle\!\rangle}}% +} + +\newcommand{\isabellestylesl}{% +\isabellestyleit% +\renewcommand{\isastyle}{\small\sl}% +\renewcommand{\isastyleminor}{\sl}% +\renewcommand{\isastylescript}{\footnotesize\sl}% +} + + +% tagged regions + +%plain TeX version of comment package -- much faster! +\let\isafmtname\fmtname\def\fmtname{plain} +\usepackage{comment} +\let\fmtname\isafmtname + +\newcommand{\isafold}[1]{\emph{$\langle\mathord{\mathit{#1}}\rangle$}} + +\newcommand{\isakeeptag}[1]% +{\includecomment{isadelim#1}\includecomment{isatag#1}\csarg\def{isafold#1}{}} +\newcommand{\isadroptag}[1]% +{\excludecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{}} +\newcommand{\isafoldtag}[1]% +{\includecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{\isafold{#1}}} + +\isakeeptag{theory} +\isakeeptag{proof} +\isakeeptag{ML} +\isakeeptag{visible} +\isadroptag{invisible} + +\IfFileExists{isabelletags.sty}{\usepackage{isabelletags}}{} diff -r c64241fa4dff -r f5db9e08effc tphols-2011/generated/isabellesym.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/tphols-2011/generated/isabellesym.sty Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,358 @@ +%% +%% definitions of standard Isabelle symbols +%% + +\newcommand{\isasymzero}{\isamath{\mathbf{0}}} %requires amssymb +\newcommand{\isasymone}{\isamath{\mathbf{1}}} %requires amssymb +\newcommand{\isasymtwo}{\isamath{\mathbf{2}}} %requires amssymb +\newcommand{\isasymthree}{\isamath{\mathbf{3}}} %requires amssymb +\newcommand{\isasymfour}{\isamath{\mathbf{4}}} %requires amssymb +\newcommand{\isasymfive}{\isamath{\mathbf{5}}} %requires amssymb +\newcommand{\isasymsix}{\isamath{\mathbf{6}}} %requires amssymb +\newcommand{\isasymseven}{\isamath{\mathbf{7}}} %requires amssymb +\newcommand{\isasymeight}{\isamath{\mathbf{8}}} %requires amssymb +\newcommand{\isasymnine}{\isamath{\mathbf{9}}} %requires amssymb +\newcommand{\isasymA}{\isamath{\mathcal{A}}} +\newcommand{\isasymB}{\isamath{\mathcal{B}}} +\newcommand{\isasymC}{\isamath{\mathcal{C}}} +\newcommand{\isasymD}{\isamath{\mathcal{D}}} +\newcommand{\isasymE}{\isamath{\mathcal{E}}} +\newcommand{\isasymF}{\isamath{\mathcal{F}}} +\newcommand{\isasymG}{\isamath{\mathcal{G}}} +\newcommand{\isasymH}{\isamath{\mathcal{H}}} +\newcommand{\isasymI}{\isamath{\mathcal{I}}} +\newcommand{\isasymJ}{\isamath{\mathcal{J}}} +\newcommand{\isasymK}{\isamath{\mathcal{K}}} +\newcommand{\isasymL}{\isamath{\mathcal{L}}} +\newcommand{\isasymM}{\isamath{\mathcal{M}}} +\newcommand{\isasymN}{\isamath{\mathcal{N}}} +\newcommand{\isasymO}{\isamath{\mathcal{O}}} +\newcommand{\isasymP}{\isamath{\mathcal{P}}} +\newcommand{\isasymQ}{\isamath{\mathcal{Q}}} +\newcommand{\isasymR}{\isamath{\mathcal{R}}} +\newcommand{\isasymS}{\isamath{\mathcal{S}}} +\newcommand{\isasymT}{\isamath{\mathcal{T}}} +\newcommand{\isasymU}{\isamath{\mathcal{U}}} +\newcommand{\isasymV}{\isamath{\mathcal{V}}} +\newcommand{\isasymW}{\isamath{\mathcal{W}}} +\newcommand{\isasymX}{\isamath{\mathcal{X}}} +\newcommand{\isasymY}{\isamath{\mathcal{Y}}} +\newcommand{\isasymZ}{\isamath{\mathcal{Z}}} +\newcommand{\isasyma}{\isamath{\mathrm{a}}} +\newcommand{\isasymb}{\isamath{\mathrm{b}}} +\newcommand{\isasymc}{\isamath{\mathrm{c}}} +\newcommand{\isasymd}{\isamath{\mathrm{d}}} +\newcommand{\isasyme}{\isamath{\mathrm{e}}} +\newcommand{\isasymf}{\isamath{\mathrm{f}}} +\newcommand{\isasymg}{\isamath{\mathrm{g}}} +\newcommand{\isasymh}{\isamath{\mathrm{h}}} +\newcommand{\isasymi}{\isamath{\mathrm{i}}} +\newcommand{\isasymj}{\isamath{\mathrm{j}}} +\newcommand{\isasymk}{\isamath{\mathrm{k}}} +\newcommand{\isasyml}{\isamath{\mathrm{l}}} +\newcommand{\isasymm}{\isamath{\mathrm{m}}} +\newcommand{\isasymn}{\isamath{\mathrm{n}}} +\newcommand{\isasymo}{\isamath{\mathrm{o}}} +\newcommand{\isasymp}{\isamath{\mathrm{p}}} +\newcommand{\isasymq}{\isamath{\mathrm{q}}} +\newcommand{\isasymr}{\isamath{\mathrm{r}}} +\newcommand{\isasyms}{\isamath{\mathrm{s}}} +\newcommand{\isasymt}{\isamath{\mathrm{t}}} +\newcommand{\isasymu}{\isamath{\mathrm{u}}} +\newcommand{\isasymv}{\isamath{\mathrm{v}}} +\newcommand{\isasymw}{\isamath{\mathrm{w}}} +\newcommand{\isasymx}{\isamath{\mathrm{x}}} +\newcommand{\isasymy}{\isamath{\mathrm{y}}} +\newcommand{\isasymz}{\isamath{\mathrm{z}}} +\newcommand{\isasymAA}{\isamath{\mathfrak{A}}} %requires eufrak +\newcommand{\isasymBB}{\isamath{\mathfrak{B}}} %requires eufrak +\newcommand{\isasymCC}{\isamath{\mathfrak{C}}} %requires eufrak +\newcommand{\isasymDD}{\isamath{\mathfrak{D}}} %requires eufrak +\newcommand{\isasymEE}{\isamath{\mathfrak{E}}} %requires eufrak +\newcommand{\isasymFF}{\isamath{\mathfrak{F}}} %requires eufrak +\newcommand{\isasymGG}{\isamath{\mathfrak{G}}} %requires eufrak +\newcommand{\isasymHH}{\isamath{\mathfrak{H}}} %requires eufrak +\newcommand{\isasymII}{\isamath{\mathfrak{I}}} %requires eufrak +\newcommand{\isasymJJ}{\isamath{\mathfrak{J}}} %requires eufrak +\newcommand{\isasymKK}{\isamath{\mathfrak{K}}} %requires eufrak +\newcommand{\isasymLL}{\isamath{\mathfrak{L}}} %requires eufrak +\newcommand{\isasymMM}{\isamath{\mathfrak{M}}} %requires eufrak +\newcommand{\isasymNN}{\isamath{\mathfrak{N}}} %requires eufrak +\newcommand{\isasymOO}{\isamath{\mathfrak{O}}} %requires eufrak +\newcommand{\isasymPP}{\isamath{\mathfrak{P}}} %requires eufrak +\newcommand{\isasymQQ}{\isamath{\mathfrak{Q}}} %requires eufrak +\newcommand{\isasymRR}{\isamath{\mathfrak{R}}} %requires eufrak +\newcommand{\isasymSS}{\isamath{\mathfrak{S}}} %requires eufrak +\newcommand{\isasymTT}{\isamath{\mathfrak{T}}} %requires eufrak +\newcommand{\isasymUU}{\isamath{\mathfrak{U}}} %requires eufrak +\newcommand{\isasymVV}{\isamath{\mathfrak{V}}} %requires eufrak +\newcommand{\isasymWW}{\isamath{\mathfrak{W}}} %requires eufrak +\newcommand{\isasymXX}{\isamath{\mathfrak{X}}} %requires eufrak +\newcommand{\isasymYY}{\isamath{\mathfrak{Y}}} %requires eufrak +\newcommand{\isasymZZ}{\isamath{\mathfrak{Z}}} %requires eufrak +\newcommand{\isasymaa}{\isamath{\mathfrak{a}}} %requires eufrak +\newcommand{\isasymbb}{\isamath{\mathfrak{b}}} %requires eufrak +\newcommand{\isasymcc}{\isamath{\mathfrak{c}}} %requires eufrak +\newcommand{\isasymdd}{\isamath{\mathfrak{d}}} %requires eufrak +\newcommand{\isasymee}{\isamath{\mathfrak{e}}} %requires eufrak +\newcommand{\isasymff}{\isamath{\mathfrak{f}}} %requires eufrak +\newcommand{\isasymgg}{\isamath{\mathfrak{g}}} %requires eufrak +\newcommand{\isasymhh}{\isamath{\mathfrak{h}}} %requires eufrak +\newcommand{\isasymii}{\isamath{\mathfrak{i}}} %requires eufrak +\newcommand{\isasymjj}{\isamath{\mathfrak{j}}} %requires eufrak +\newcommand{\isasymkk}{\isamath{\mathfrak{k}}} %requires eufrak +\newcommand{\isasymll}{\isamath{\mathfrak{l}}} %requires eufrak +\newcommand{\isasymmm}{\isamath{\mathfrak{m}}} %requires eufrak +\newcommand{\isasymnn}{\isamath{\mathfrak{n}}} %requires eufrak +\newcommand{\isasymoo}{\isamath{\mathfrak{o}}} %requires eufrak +\newcommand{\isasympp}{\isamath{\mathfrak{p}}} %requires eufrak +\newcommand{\isasymqq}{\isamath{\mathfrak{q}}} %requires eufrak +\newcommand{\isasymrr}{\isamath{\mathfrak{r}}} %requires eufrak +\newcommand{\isasymss}{\isamath{\mathfrak{s}}} %requires eufrak +\newcommand{\isasymtt}{\isamath{\mathfrak{t}}} %requires eufrak +\newcommand{\isasymuu}{\isamath{\mathfrak{u}}} %requires eufrak +\newcommand{\isasymvv}{\isamath{\mathfrak{v}}} %requires eufrak +\newcommand{\isasymww}{\isamath{\mathfrak{w}}} %requires eufrak +\newcommand{\isasymxx}{\isamath{\mathfrak{x}}} %requires eufrak +\newcommand{\isasymyy}{\isamath{\mathfrak{y}}} %requires eufrak +\newcommand{\isasymzz}{\isamath{\mathfrak{z}}} %requires eufrak +\newcommand{\isasymalpha}{\isamath{\alpha}} +\newcommand{\isasymbeta}{\isamath{\beta}} +\newcommand{\isasymgamma}{\isamath{\gamma}} +\newcommand{\isasymdelta}{\isamath{\delta}} +\newcommand{\isasymepsilon}{\isamath{\varepsilon}} +\newcommand{\isasymzeta}{\isamath{\zeta}} +\newcommand{\isasymeta}{\isamath{\eta}} +\newcommand{\isasymtheta}{\isamath{\vartheta}} +\newcommand{\isasymiota}{\isamath{\iota}} +\newcommand{\isasymkappa}{\isamath{\kappa}} +\newcommand{\isasymlambda}{\isamath{\lambda}} +\newcommand{\isasymmu}{\isamath{\mu}} +\newcommand{\isasymnu}{\isamath{\nu}} +\newcommand{\isasymxi}{\isamath{\xi}} +\newcommand{\isasympi}{\isamath{\pi}} +\newcommand{\isasymrho}{\isamath{\varrho}} +\newcommand{\isasymsigma}{\isamath{\sigma}} +\newcommand{\isasymtau}{\isamath{\tau}} +\newcommand{\isasymupsilon}{\isamath{\upsilon}} +\newcommand{\isasymphi}{\isamath{\varphi}} +\newcommand{\isasymchi}{\isamath{\chi}} +\newcommand{\isasympsi}{\isamath{\psi}} +\newcommand{\isasymomega}{\isamath{\omega}} +\newcommand{\isasymGamma}{\isamath{\Gamma}} +\newcommand{\isasymDelta}{\isamath{\Delta}} +\newcommand{\isasymTheta}{\isamath{\Theta}} +\newcommand{\isasymLambda}{\isamath{\Lambda}} +\newcommand{\isasymXi}{\isamath{\Xi}} +\newcommand{\isasymPi}{\isamath{\Pi}} +\newcommand{\isasymSigma}{\isamath{\Sigma}} +\newcommand{\isasymUpsilon}{\isamath{\Upsilon}} +\newcommand{\isasymPhi}{\isamath{\Phi}} +\newcommand{\isasymPsi}{\isamath{\Psi}} +\newcommand{\isasymOmega}{\isamath{\Omega}} +\newcommand{\isasymbool}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{B}}} +\newcommand{\isasymcomplex}{\isamath{\mathrm{C}\mkern-15mu{\phantom{\mathrm{t}}\vrule}\mkern9mu}} +\newcommand{\isasymnat}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{N}}} +\newcommand{\isasymrat}{\isamath{\mathrm{Q}\mkern-16mu{\phantom{\mathrm{t}}\vrule}\mkern10mu}} +\newcommand{\isasymreal}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{R}}} +\newcommand{\isasymint}{\isamath{\mathsf{Z}\mkern-7.5mu\mathsf{Z}}} +\newcommand{\isasymleftarrow}{\isamath{\leftarrow}} +\newcommand{\isasymlongleftarrow}{\isamath{\longleftarrow}} +\newcommand{\isasymrightarrow}{\isamath{\rightarrow}} +\newcommand{\isasymlongrightarrow}{\isamath{\longrightarrow}} +\newcommand{\isasymLeftarrow}{\isamath{\Leftarrow}} +\newcommand{\isasymLongleftarrow}{\isamath{\Longleftarrow}} +\newcommand{\isasymRightarrow}{\isamath{\Rightarrow}} +\newcommand{\isasymLongrightarrow}{\isamath{\Longrightarrow}} +\newcommand{\isasymleftrightarrow}{\isamath{\leftrightarrow}} +\newcommand{\isasymlongleftrightarrow}{\isamath{\longleftrightarrow}} +\newcommand{\isasymLeftrightarrow}{\isamath{\Leftrightarrow}} +\newcommand{\isasymLongleftrightarrow}{\isamath{\Longleftrightarrow}} +\newcommand{\isasymmapsto}{\isamath{\mapsto}} +\newcommand{\isasymlongmapsto}{\isamath{\longmapsto}} +\newcommand{\isasymmidarrow}{\isamath{\relbar}} +\newcommand{\isasymMidarrow}{\isamath{\Relbar}} +\newcommand{\isasymhookleftarrow}{\isamath{\hookleftarrow}} +\newcommand{\isasymhookrightarrow}{\isamath{\hookrightarrow}} +\newcommand{\isasymleftharpoondown}{\isamath{\leftharpoondown}} +\newcommand{\isasymrightharpoondown}{\isamath{\rightharpoondown}} +\newcommand{\isasymleftharpoonup}{\isamath{\leftharpoonup}} +\newcommand{\isasymrightharpoonup}{\isamath{\rightharpoonup}} +\newcommand{\isasymrightleftharpoons}{\isamath{\rightleftharpoons}} +\newcommand{\isasymleadsto}{\isamath{\leadsto}} %requires amssymb +\newcommand{\isasymdownharpoonleft}{\isamath{\downharpoonleft}} %requires amssymb +\newcommand{\isasymdownharpoonright}{\isamath{\downharpoonright}} %requires amssymb +\newcommand{\isasymupharpoonleft}{\isamath{\upharpoonleft}} %requires amssymb +\newcommand{\isasymupharpoonright}{\isamath{\upharpoonright}} %requires amssymb +\newcommand{\isasymrestriction}{\isamath{\restriction}} %requires amssymb +\newcommand{\isasymColon}{\isamath{\mathrel{::}}} +\newcommand{\isasymup}{\isamath{\uparrow}} +\newcommand{\isasymUp}{\isamath{\Uparrow}} +\newcommand{\isasymdown}{\isamath{\downarrow}} +\newcommand{\isasymDown}{\isamath{\Downarrow}} +\newcommand{\isasymupdown}{\isamath{\updownarrow}} +\newcommand{\isasymUpdown}{\isamath{\Updownarrow}} +\newcommand{\isasymlangle}{\isamath{\langle}} +\newcommand{\isasymrangle}{\isamath{\rangle}} +\newcommand{\isasymlceil}{\isamath{\lceil}} +\newcommand{\isasymrceil}{\isamath{\rceil}} +\newcommand{\isasymlfloor}{\isamath{\lfloor}} +\newcommand{\isasymrfloor}{\isamath{\rfloor}} +\newcommand{\isasymlparr}{\isamath{\mathopen{(\mkern-3mu\mid}}} +\newcommand{\isasymrparr}{\isamath{\mathclose{\mid\mkern-3mu)}}} +\newcommand{\isasymlbrakk}{\isamath{\mathopen{\lbrack\mkern-3mu\lbrack}}} +\newcommand{\isasymrbrakk}{\isamath{\mathclose{\rbrack\mkern-3mu\rbrack}}} +\newcommand{\isasymlbrace}{\isamath{\mathopen{\lbrace\mkern-4.5mu\mid}}} +\newcommand{\isasymrbrace}{\isamath{\mathclose{\mid\mkern-4.5mu\rbrace}}} +\newcommand{\isasymguillemotleft}{\isatext{\flqq}} %requires babel +\newcommand{\isasymguillemotright}{\isatext{\frqq}} %requires babel +\newcommand{\isasymbottom}{\isamath{\bot}} +\newcommand{\isasymtop}{\isamath{\top}} +\newcommand{\isasymand}{\isamath{\wedge}} +\newcommand{\isasymAnd}{\isamath{\bigwedge}} +\newcommand{\isasymor}{\isamath{\vee}} +\newcommand{\isasymOr}{\isamath{\bigvee}} +\newcommand{\isasymforall}{\isamath{\forall\,}} +\newcommand{\isasymexists}{\isamath{\exists\,}} +\newcommand{\isasymnexists}{\isamath{\nexists\,}} %requires amssymb +\newcommand{\isasymnot}{\isamath{\neg}} +\newcommand{\isasymbox}{\isamath{\Box}} %requires amssymb +\newcommand{\isasymdiamond}{\isamath{\Diamond}} %requires amssymb +\newcommand{\isasymturnstile}{\isamath{\vdash}} +\newcommand{\isasymTurnstile}{\isamath{\models}} +\newcommand{\isasymtturnstile}{\isamath{\vdash\!\!\!\vdash}} +\newcommand{\isasymTTurnstile}{\isamath{\mid\!\models}} +\newcommand{\isasymstileturn}{\isamath{\dashv}} +\newcommand{\isasymsurd}{\isamath{\surd}} +\newcommand{\isasymle}{\isamath{\le}} +\newcommand{\isasymge}{\isamath{\ge}} +\newcommand{\isasymlless}{\isamath{\ll}} +\newcommand{\isasymggreater}{\isamath{\gg}} +\newcommand{\isasymlesssim}{\isamath{\lesssim}} %requires amssymb +\newcommand{\isasymgreatersim}{\isamath{\gtrsim}} %requires amssymb +\newcommand{\isasymlessapprox}{\isamath{\lessapprox}} %requires amssymb +\newcommand{\isasymgreaterapprox}{\isamath{\gtrapprox}} %requires amssymb +\newcommand{\isasymin}{\isamath{\in}} +\newcommand{\isasymnotin}{\isamath{\notin}} +\newcommand{\isasymsubset}{\isamath{\subset}} +\newcommand{\isasymsupset}{\isamath{\supset}} +\newcommand{\isasymsubseteq}{\isamath{\subseteq}} +\newcommand{\isasymsupseteq}{\isamath{\supseteq}} +\newcommand{\isasymsqsubset}{\isamath{\sqsubset}} %requires amssymb +\newcommand{\isasymsqsupset}{\isamath{\sqsupset}} %requires amssymb +\newcommand{\isasymsqsubseteq}{\isamath{\sqsubseteq}} +\newcommand{\isasymsqsupseteq}{\isamath{\sqsupseteq}} +\newcommand{\isasyminter}{\isamath{\cap}} +\newcommand{\isasymInter}{\isamath{\bigcap\,}} +\newcommand{\isasymunion}{\isamath{\cup}} +\newcommand{\isasymUnion}{\isamath{\bigcup\,}} +\newcommand{\isasymsqunion}{\isamath{\sqcup}} +\newcommand{\isasymSqunion}{\isamath{\bigsqcup\,}} +\newcommand{\isasymsqinter}{\isamath{\sqcap}} +\newcommand{\isasymSqinter}{\isamath{\bigsqcap\,}} %requires stmaryrd +\newcommand{\isasymsetminus}{\isamath{\setminus}} +\newcommand{\isasympropto}{\isamath{\propto}} +\newcommand{\isasymuplus}{\isamath{\uplus}} +\newcommand{\isasymUplus}{\isamath{\biguplus\,}} +\newcommand{\isasymnoteq}{\isamath{\not=}} +\newcommand{\isasymsim}{\isamath{\sim}} +\newcommand{\isasymdoteq}{\isamath{\doteq}} +\newcommand{\isasymsimeq}{\isamath{\simeq}} +\newcommand{\isasymapprox}{\isamath{\approx}} +\newcommand{\isasymasymp}{\isamath{\asymp}} +\newcommand{\isasymcong}{\isamath{\cong}} +\newcommand{\isasymsmile}{\isamath{\smile}} +\newcommand{\isasymequiv}{\isamath{\equiv}} +\newcommand{\isasymfrown}{\isamath{\frown}} +\newcommand{\isasymJoin}{\isamath{\Join}} %requires amssymb +\newcommand{\isasymbowtie}{\isamath{\bowtie}} +\newcommand{\isasymprec}{\isamath{\prec}} +\newcommand{\isasymsucc}{\isamath{\succ}} +\newcommand{\isasympreceq}{\isamath{\preceq}} +\newcommand{\isasymsucceq}{\isamath{\succeq}} +\newcommand{\isasymparallel}{\isamath{\parallel}} +\newcommand{\isasymbar}{\isamath{\mid}} +\newcommand{\isasymplusminus}{\isamath{\pm}} +\newcommand{\isasymminusplus}{\isamath{\mp}} +\newcommand{\isasymtimes}{\isamath{\times}} +\newcommand{\isasymdiv}{\isamath{\div}} +\newcommand{\isasymcdot}{\isamath{\cdot}} +\newcommand{\isasymstar}{\isamath{\star}} +\newcommand{\isasymbullet}{\boldmath\isamath{\mathchoice{\displaystyle{\cdot}}{\textstyle{\cdot}}{\scriptstyle{\bullet}}{\scriptscriptstyle{\bullet}}}} +\newcommand{\isasymcirc}{\isamath{\circ}} +\newcommand{\isasymdagger}{\isamath{\dagger}} +\newcommand{\isasymddagger}{\isamath{\ddagger}} +\newcommand{\isasymlhd}{\isamath{\lhd}} %requires amssymb +\newcommand{\isasymrhd}{\isamath{\rhd}} %requires amssymb +\newcommand{\isasymunlhd}{\isamath{\unlhd}} %requires amssymb +\newcommand{\isasymunrhd}{\isamath{\unrhd}} %requires amssymb +\newcommand{\isasymtriangleleft}{\isamath{\triangleleft}} +\newcommand{\isasymtriangleright}{\isamath{\triangleright}} +\newcommand{\isasymtriangle}{\isamath{\triangle}} +\newcommand{\isasymtriangleq}{\isamath{\triangleq}} %requires amssymb +\newcommand{\isasymoplus}{\isamath{\oplus}} +\newcommand{\isasymOplus}{\isamath{\bigoplus\,}} +\newcommand{\isasymotimes}{\isamath{\otimes}} +\newcommand{\isasymOtimes}{\isamath{\bigotimes\,}} +\newcommand{\isasymodot}{\isamath{\odot}} +\newcommand{\isasymOdot}{\isamath{\bigodot\,}} +\newcommand{\isasymominus}{\isamath{\ominus}} +\newcommand{\isasymoslash}{\isamath{\oslash}} +\newcommand{\isasymdots}{\isamath{\dots}} +\newcommand{\isasymcdots}{\isamath{\cdots}} +\newcommand{\isasymSum}{\isamath{\sum\,}} +\newcommand{\isasymProd}{\isamath{\prod\,}} +\newcommand{\isasymCoprod}{\isamath{\coprod\,}} +\newcommand{\isasyminfinity}{\isamath{\infty}} +\newcommand{\isasymintegral}{\isamath{\int\,}} +\newcommand{\isasymointegral}{\isamath{\oint\,}} +\newcommand{\isasymclubsuit}{\isamath{\clubsuit}} +\newcommand{\isasymdiamondsuit}{\isamath{\diamondsuit}} +\newcommand{\isasymheartsuit}{\isamath{\heartsuit}} +\newcommand{\isasymspadesuit}{\isamath{\spadesuit}} +\newcommand{\isasymaleph}{\isamath{\aleph}} +\newcommand{\isasymemptyset}{\isamath{\emptyset}} +\newcommand{\isasymnabla}{\isamath{\nabla}} +\newcommand{\isasympartial}{\isamath{\partial}} +\newcommand{\isasymRe}{\isamath{\Re}} +\newcommand{\isasymIm}{\isamath{\Im}} +\newcommand{\isasymflat}{\isamath{\flat}} +\newcommand{\isasymnatural}{\isamath{\natural}} +\newcommand{\isasymsharp}{\isamath{\sharp}} +\newcommand{\isasymangle}{\isamath{\angle}} +\newcommand{\isasymcopyright}{\isatext{\rm\copyright}} +\newcommand{\isasymregistered}{\isatext{\rm\textregistered}} +\newcommand{\isasymhyphen}{\isatext{\rm-}} +\newcommand{\isasyminverse}{\isamath{{}^{-1}}} +\newcommand{\isasymonesuperior}{\isamath{\mathonesuperior}} %requires latin1 +\newcommand{\isasymonequarter}{\isatext{\rm\textonequarter}} %requires latin1 +\newcommand{\isasymtwosuperior}{\isamath{\mathtwosuperior}} %requires latin1 +\newcommand{\isasymonehalf}{\isatext{\rm\textonehalf}} %requires latin1 +\newcommand{\isasymthreesuperior}{\isamath{\maththreesuperior}} %requires latin1 +\newcommand{\isasymthreequarters}{\isatext{\rm\textthreequarters}} %requires latin1 +\newcommand{\isasymordfeminine}{\isatext{\rm\textordfeminine}} +\newcommand{\isasymordmasculine}{\isatext{\rm\textordmasculine}} +\newcommand{\isasymsection}{\isatext{\rm\S}} +\newcommand{\isasymparagraph}{\isatext{\rm\P}} +\newcommand{\isasymexclamdown}{\isatext{\rm\textexclamdown}} +\newcommand{\isasymquestiondown}{\isatext{\rm\textquestiondown}} +\newcommand{\isasymeuro}{\isatext{\textgreek{\euro}}} %requires greek babel +\newcommand{\isasympounds}{\isamath{\pounds}} +\newcommand{\isasymyen}{\isatext{\yen}} %requires amssymb +\newcommand{\isasymcent}{\isatext{\textcent}} %requires textcomp +\newcommand{\isasymcurrency}{\isatext{\textcurrency}} %requires textcomp +\newcommand{\isasymdegree}{\isatext{\rm\textdegree}} %requires latin1 +\newcommand{\isasymamalg}{\isamath{\amalg}} +\newcommand{\isasymmho}{\isamath{\mho}} %requires amssymb +\newcommand{\isasymlozenge}{\isamath{\lozenge}} %requires amssymb +\newcommand{\isasymwp}{\isamath{\wp}} +\newcommand{\isasymwrong}{\isamath{\wr}} +\newcommand{\isasymstruct}{\isamath{\diamond}} +\newcommand{\isasymacute}{\isatext{\'\relax}} +\newcommand{\isasymindex}{\isatext{\i}} +\newcommand{\isasymdieresis}{\isatext{\"\relax}} +\newcommand{\isasymcedilla}{\isatext{\c\relax}} +\newcommand{\isasymhungarumlaut}{\isatext{\H\relax}} +\newcommand{\isasymspacespace}{\isamath{~~}} +\newcommand{\isasymmodule}{\isamath{\langle}\isakeyword{module}\isamath{\rangle}} +\newcommand{\isasymsome}{\isamath{\epsilon\,}} diff -r c64241fa4dff -r f5db9e08effc tphols-2011/generated/pdfsetup.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/tphols-2011/generated/pdfsetup.sty Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,7 @@ +%% +%% default hyperref setup (both for pdf and dvi output) +%% + +\usepackage{color} +\definecolor{linkcolor}{rgb}{0,0,0.5} +\usepackage[colorlinks=true,linkcolor=linkcolor,citecolor=linkcolor,filecolor=linkcolor,pagecolor=linkcolor,urlcolor=linkcolor]{hyperref} diff -r c64241fa4dff -r f5db9e08effc tphols-2011/generated/root.aux --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/tphols-2011/generated/root.aux Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,25 @@ +\relax +\ifx\hyper@anchor\@undefined +\global \let \oldcontentsline\contentsline +\gdef \contentsline#1#2#3#4{\oldcontentsline{#1}{#2}{#3}} +\global \let \oldnewlabel\newlabel +\gdef \newlabel#1#2{\newlabelxx{#1}#2} +\gdef \newlabelxx#1#2#3#4#5#6{\oldnewlabel{#1}{{#2}{#3}}} +\AtEndDocument{\let \contentsline\oldcontentsline +\let \newlabel\oldnewlabel} +\else +\global \let \hyper@last\relax +\fi + +\@writefile{toc}{\contentsline {section}{\numberline {1}Preliminary definitions}{1}{section.1}} +\@writefile{toc}{\contentsline {section}{\numberline {2}Direction \emph {\it finite\ partition\ {\emph {$\Rightarrow $}}\ regular\ language}}{5}{section.2}} +\newlabel{example_eqns}{{1}{5}{Direction \isa {finite\ partition\ {\isasymRightarrow }\ regular\ language}\relax }{equation.1}{}} +\@writefile{toc}{\contentsline {subsection}{\numberline {2.1}Proof for this direction}{9}{subsection.2.1}} +\@writefile{toc}{\contentsline {section}{\numberline {3}Direction: \emph {\it regular\ language\ {\emph {$\Rightarrow $}}finite\ partition}}{21}{section.3}} +\@writefile{toc}{\contentsline {subsection}{\numberline {3.1}The scheme for this direction}{21}{subsection.3.1}} +\@writefile{toc}{\contentsline {subsection}{\numberline {3.2}A small theory for list difference}{22}{subsection.3.2}} +\@writefile{toc}{\contentsline {subsection}{\numberline {3.3}Lemmas for basic cases}{23}{subsection.3.3}} +\@writefile{toc}{\contentsline {subsection}{\numberline {3.4}The case for \emph {\it SEQ}}{24}{subsection.3.4}} +\@writefile{toc}{\contentsline {subsection}{\numberline {3.5}The case for \emph {\it ALT}}{26}{subsection.3.5}} +\@writefile{toc}{\contentsline {subsection}{\numberline {3.6}The case for \emph {\it STAR}}{27}{subsection.3.6}} +\@writefile{toc}{\contentsline {subsection}{\numberline {3.7}The main lemma}{29}{subsection.3.7}} diff -r c64241fa4dff -r f5db9e08effc tphols-2011/generated/root.dvi Binary file tphols-2011/generated/root.dvi has changed diff -r c64241fa4dff -r f5db9e08effc tphols-2011/generated/root.log --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/tphols-2011/generated/root.log Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,569 @@ +This is pdfTeXk, Version 3.141592-1.40.3 (Web2C 7.5.6) (format=pdflatex 2009.3.10) 24 JAN 2011 13:10 +entering extended mode + %&-line parsing enabled. +**\nonstopmode\input{root.tex} +(./root.tex (/usr/share/texmf/tex/latex/base/article.cls +Document Class: article 2005/09/16 v1.4f Standard LaTeX document class +(/usr/share/texmf/tex/latex/base/size11.clo +File: size11.clo 2005/09/16 v1.4f Standard LaTeX file (size option) +) +\c@part=\count79 +\c@section=\count80 +\c@subsection=\count81 +\c@subsubsection=\count82 +\c@paragraph=\count83 +\c@subparagraph=\count84 +\c@figure=\count85 +\c@table=\count86 +\abovecaptionskip=\skip41 +\belowcaptionskip=\skip42 +\bibindent=\dimen102 +) (./isabelle.sty +\isa@parindent=\dimen103 +\isa@parskip=\dimen104 + +(/usr/share/texmf/tex/latex/comment/comment.sty +\CommentStream=\write3 + Excluding comment 'comment') +Including comment 'isadelimtheory' Including comment 'isatagtheory' +Including comment 'isadelimproof' Including comment 'isatagproof' +Including comment 'isadelimML' Including comment 'isatagML' +Including comment 'isadelimvisible' Including comment 'isatagvisible' +Excluding comment 'isadeliminvisible' Excluding comment 'isataginvisible') +(./isabellesym.sty) (/usr/share/texmf/tex/latex/amsmath/amsmath.sty +Package: amsmath 2000/07/18 v2.13 AMS math features +\@mathmargin=\skip43 + +For additional information on amsmath, use the `?' option. +(/usr/share/texmf/tex/latex/amsmath/amstext.sty +Package: amstext 2000/06/29 v2.01 + +(/usr/share/texmf/tex/latex/amsmath/amsgen.sty +File: amsgen.sty 1999/11/30 v2.0 +\@emptytoks=\toks14 +\ex@=\dimen105 +)) +(/usr/share/texmf/tex/latex/amsmath/amsbsy.sty +Package: amsbsy 1999/11/29 v1.2d +\pmbraise@=\dimen106 +) +(/usr/share/texmf/tex/latex/amsmath/amsopn.sty +Package: amsopn 1999/12/14 v2.01 operator names +) +\inf@bad=\count87 +LaTeX Info: Redefining \frac on input line 211. +\uproot@=\count88 +\leftroot@=\count89 +LaTeX Info: Redefining \overline on input line 307. +\classnum@=\count90 +\DOTSCASE@=\count91 +LaTeX Info: Redefining \ldots on input line 379. +LaTeX Info: Redefining \dots on input line 382. +LaTeX Info: Redefining \cdots on input line 467. +\Mathstrutbox@=\box26 +\strutbox@=\box27 +\big@size=\dimen107 +LaTeX Font Info: Redeclaring font encoding OML on input line 567. +LaTeX Font Info: Redeclaring font encoding OMS on input line 568. +\macc@depth=\count92 +\c@MaxMatrixCols=\count93 +\dotsspace@=\muskip10 +\c@parentequation=\count94 +\dspbrk@lvl=\count95 +\tag@help=\toks15 +\row@=\count96 +\column@=\count97 +\maxfields@=\count98 +\andhelp@=\toks16 +\eqnshift@=\dimen108 +\alignsep@=\dimen109 +\tagshift@=\dimen110 +\tagwidth@=\dimen111 +\totwidth@=\dimen112 +\lineht@=\dimen113 +\@envbody=\toks17 +\multlinegap=\skip44 +\multlinetaggap=\skip45 +\mathdisplay@stack=\toks18 +LaTeX Info: Redefining \[ on input line 2666. +LaTeX Info: Redefining \] on input line 2667. +) (./pdfsetup.sty +(/usr/share/texmf/tex/latex/graphics/color.sty +Package: color 2005/11/14 v1.0j Standard LaTeX Color (DPC) + +(/usr/share/texmf/tex/latex/config/color.cfg +File: color.cfg 2007/01/18 v1.5 color configuration of teTeX/TeXLive +) +Package color Info: Driver file: pdftex.def on input line 130. + +(/usr/share/texmf/tex/latex/pdftex-def/pdftex.def +File: pdftex.def 2007/01/08 v0.04d Graphics/color for pdfTeX +\Gread@gobject=\count99 +))) +(/usr/share/texmf/tex/latex/hyperref/hyperref.sty +Package: hyperref 2007/02/07 v6.75r Hypertext links for LaTeX + +(/usr/share/texmf/tex/latex/graphics/keyval.sty +Package: keyval 1999/03/16 v1.13 key=value parser (DPC) +\KV@toks@=\toks19 +) +\@linkdim=\dimen114 +\Hy@linkcounter=\count100 +\Hy@pagecounter=\count101 + +(/usr/share/texmf/tex/latex/hyperref/pd1enc.def +File: pd1enc.def 2007/02/07 v6.75r Hyperref: PDFDocEncoding definition (HO) +) +(/usr/share/texmf/tex/latex/config/hyperref.cfg +File: hyperref.cfg 2002/06/06 v1.2 hyperref configuration of TeXLive +) +(/usr/share/texmf/tex/latex/oberdiek/kvoptions.sty +Package: kvoptions 2006/08/22 v2.4 Connects package keyval with LaTeX options ( +HO) +) +Package hyperref Info: Option `colorlinks' set `true' on input line 2238. +Package hyperref Info: Hyper figures OFF on input line 2288. +Package hyperref Info: Link nesting OFF on input line 2293. +Package hyperref Info: Hyper index ON on input line 2296. +Package hyperref Info: Plain pages OFF on input line 2303. +Package hyperref Info: Backreferencing OFF on input line 2308. + +Implicit mode ON; LaTeX internals redefined +Package hyperref Info: Bookmarks ON on input line 2444. +(/usr/share/texmf/tex/latex/ltxmisc/url.sty +\Urlmuskip=\muskip11 +Package: url 2005/06/27 ver 3.2 Verb mode for urls, etc. +) +LaTeX Info: Redefining \url on input line 2599. +\Fld@menulength=\count102 +\Field@Width=\dimen115 +\Fld@charsize=\dimen116 +\Choice@toks=\toks20 +\Field@toks=\toks21 +Package hyperref Info: Hyper figures OFF on input line 3102. +Package hyperref Info: Link nesting OFF on input line 3107. +Package hyperref Info: Hyper index ON on input line 3110. +Package hyperref Info: backreferencing OFF on input line 3117. +Package hyperref Info: Link coloring ON on input line 3120. +\Hy@abspage=\count103 +\c@Item=\count104 +\c@Hfootnote=\count105 +) +*hyperref using default driver hpdftex* +(/usr/share/texmf/tex/latex/hyperref/hpdftex.def +File: hpdftex.def 2007/02/07 v6.75r Hyperref driver for pdfTeX +\Fld@listcount=\count106 +) +No file root.aux. +\openout1 = `root.aux'. + +LaTeX Font Info: Checking defaults for OML/cmm/m/it on input line 42. +LaTeX Font Info: ... okay on input line 42. +LaTeX Font Info: Checking defaults for T1/cmr/m/n on input line 42. +LaTeX Font Info: ... okay on input line 42. +LaTeX Font Info: Checking defaults for OT1/cmr/m/n on input line 42. +LaTeX Font Info: ... okay on input line 42. +LaTeX Font Info: Checking defaults for OMS/cmsy/m/n on input line 42. +LaTeX Font Info: ... okay on input line 42. +LaTeX Font Info: Checking defaults for OMX/cmex/m/n on input line 42. +LaTeX Font Info: ... okay on input line 42. +LaTeX Font Info: Checking defaults for U/cmr/m/n on input line 42. +LaTeX Font Info: ... okay on input line 42. +LaTeX Font Info: Checking defaults for PD1/pdf/m/n on input line 42. +LaTeX Font Info: ... okay on input line 42. +Package hyperref Info: Link coloring ON on input line 42. +(/usr/share/texmf/tex/latex/hyperref/nameref.sty +Package: nameref 2006/12/27 v2.28 Cross-referencing by name of section + +(/usr/share/texmf/tex/latex/oberdiek/refcount.sty +Package: refcount 2006/02/20 v3.0 Data extraction from references (HO) +) +\c@section@level=\count107 +) +LaTeX Info: Redefining \ref on input line 42. +LaTeX Info: Redefining \pageref on input line 42. + (./root.out) (./root.out) +\@outlinefile=\write4 +\openout4 = `root.out'. + + + +Package hyperref Warning: old toc file detected, not used; run LaTeX again. + +\tf@toc=\write5 +\openout5 = `root.toc'. + +(./session.tex (./Myhill.tex [1 + +{/var/lib/texmf/fonts/map/pdftex/updmap/pdftex.map}] +LaTeX Font Info: Try loading font information for OMS+cmr on input line 272. + + (/usr/share/texmf/tex/latex/base/omscmr.fd +File: omscmr.fd 1999/05/25 v2.5h Standard LaTeX font definitions +) +LaTeX Font Info: Font shape `OMS/cmr/bx/n' in size <10> not available +(Font) Font shape `OMS/cmsy/b/n' tried instead on input line 272. + [2] [3] [4] + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@encoding' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@family' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@series' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@shape' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\font@name' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\OMX/cmex/m/n/7' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\ignorespaces' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `math shift' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\Rightarrow' on input line 533. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `math shift' on input line 533. + + +LaTeX Warning: Reference `example_eqns' on page 5 undefined on input line 539. + + +LaTeX Warning: Reference `example_eqns' on page 5 undefined on input line 553. + + +Overfull \hbox (32.24106pt too wide) in paragraph at lines 562--564 +[]\OT1/cmr/m/n/10.95 In this for-mal-iza-tion, pure reg-u-lar ex-pres-sions lik +e $\OML/cmm/m/it/10.95 ^^U$ \OT1/cmr/m/n/10.95 is repsented by \OT1/cmr/m/it/10 +.95 Lam$\OT1/cmr/m/n/10.95 ($\OT1/cmr/m/it/10.95 EMPTY$\OT1/cmr/m/n/10.95 )$, + [] + +[5] + +LaTeX Warning: Reference `example_eqns' on page 6 undefined on input line 628. + +[6] [7] [8] [9] [10] [11] [12] +Overfull \hbox (25.68793pt too wide) in paragraph at lines 1396--1398 +[][] \OT1/cmr/bx/n/10 by $\OT1/cmr/m/n/10 ($\OT1/cmr/m/it/10 auto simp add$\ +OT1/cmr/m/n/10 :$\OT1/cmr/m/it/10 ardenable[]def eqs[]def init[]rhs[]def rhs[]n +onempty[]def del$\OT1/cmr/m/n/10 :$\OT1/cmr/m/it/10 L[]rhs$\OML/cmm/m/it/10 :$\ +OT1/cmr/m/it/10 simps$\OT1/cmr/m/n/10 )$[] + [] + +[13] +Overfull \hbox (37.35094pt too wide) in paragraph at lines 1629--1642 +[][]\OT1/cmr/bx/n/10 by $\OT1/cmr/m/n/10 ($\OT1/cmr/m/it/10 simp only$\OT1/cmr/ +m/n/10 :$\OT1/cmr/m/it/10 rhs[]subst[]def append[]keeps[]nonempty nonempty[]se +t[]union nonempty[]set[]sub$\OT1/cmr/m/n/10 )$[] + [] + + +Overfull \hbox (6.14192pt too wide) in paragraph at lines 1669--1672 +[][] \OT1/cmr/bx/n/10 have \OT1/cmr/m/it/10 rhs $\OT1/cmr/m/n/10 =$ $($\OT1/ +cmr/m/it/10 rhs $\OMS/cmsy/m/n/10 ^^@$ \OT1/cmr/m/it/10 items[]of rhs X$\OT1/cm +r/m/n/10 )$ $\OMS/cmsy/m/n/10 [$ $\OT1/cmr/m/n/10 ($\OT1/cmr/m/it/10 items[]of +rhs X$\OT1/cmr/m/n/10 )$ \OT1/cmr/bx/n/10 by $\OT1/cmr/m/n/10 ($\OT1/cmr/m/it/1 +0 auto simp$\OT1/cmr/m/n/10 :$\OT1/cmr/m/it/10 items[]of[]def$\OT1/cmr/m/n/10 ) +$[] + [] + +[14] [15] [16] [17] [18] [19] [20] + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@encoding' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@family' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@series' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@shape' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\font@name' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\OMX/cmex/m/n/7' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\ignorespaces' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `math shift' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\Rightarrow' on input line 2581. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `math shift' on input line 2581. + +[21] [22] [23] + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3056. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@encoding' on input line 3056. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3056. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3056. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@family' on input line 3056. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3056. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@series' on input line 3056. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3056. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@shape' on input line 3056. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3056. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\font@name' on input line 3056. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\OT1/cmr/m/it/12' on input line 3056. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\ignorespaces' on input line 3056. + +[24] [25] + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3301. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@encoding' on input line 3301. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3301. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3301. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@family' on input line 3301. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3301. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@series' on input line 3301. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3301. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@shape' on input line 3301. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3301. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\font@name' on input line 3301. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\OT1/cmr/m/it/12' on input line 3301. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\ignorespaces' on input line 3301. + +[26] + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3375. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@encoding' on input line 3375. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3375. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3375. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@family' on input line 3375. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3375. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@series' on input line 3375. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3375. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\f@shape' on input line 3375. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\-command' on input line 3375. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\font@name' on input line 3375. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\OT1/cmr/m/it/12' on input line 3375. + + +Package hyperref Warning: Token not allowed in a PDFDocEncoded string, +(hyperref) removing `\ignorespaces' on input line 3375. + +[27] [28] [29])) [30] (./root.aux) + +LaTeX Warning: There were undefined references. + + +LaTeX Warning: Label(s) may have changed. Rerun to get cross-references right. + + ) +Here is how much of TeX's memory you used: + 3589 strings out of 256216 + 46991 string characters out of 1917072 + 109661 words of memory out of 1500000 + 6815 multiletter control sequences out of 10000+200000 + 13669 words of font info for 48 fonts, out of 1200000 for 2000 + 645 hyphenation exceptions out of 8191 + 31i,8n,42p,736b,457s stack positions out of 5000i,500n,6000p,200000b,15000s + +Output written on root.pdf (30 pages, 175317 bytes). +PDF statistics: + 257 PDF objects out of 1000 (max. 8388607) + 44 named destinations out of 1000 (max. 131072) + 105 words of extra memory for PDF output out of 10000 (max. 10000000) + diff -r c64241fa4dff -r f5db9e08effc tphols-2011/generated/root.out --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/tphols-2011/generated/root.out Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,11 @@ +\BOOKMARK [1][-]{section.1}{Preliminary definitions}{} +\BOOKMARK [1][-]{section.2}{Direction finite partition \040regular language}{} +\BOOKMARK [2][-]{subsection.2.1}{Proof for this direction}{section.2} +\BOOKMARK [1][-]{section.3}{Direction: regular language finite partition}{} +\BOOKMARK [2][-]{subsection.3.1}{The scheme for this direction}{section.3} +\BOOKMARK [2][-]{subsection.3.2}{A small theory for list difference}{section.3} +\BOOKMARK [2][-]{subsection.3.3}{Lemmas for basic cases}{section.3} +\BOOKMARK [2][-]{subsection.3.4}{The case for SEQ}{section.3} +\BOOKMARK [2][-]{subsection.3.5}{The case for ALT}{section.3} +\BOOKMARK [2][-]{subsection.3.6}{The case for STAR}{section.3} +\BOOKMARK [2][-]{subsection.3.7}{The main lemma}{section.3} diff -r c64241fa4dff -r f5db9e08effc tphols-2011/generated/root.pdf Binary file tphols-2011/generated/root.pdf has changed diff -r c64241fa4dff -r f5db9e08effc tphols-2011/generated/root.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/tphols-2011/generated/root.tex Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,65 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} +\usepackage{amsmath} +%\usepackage[pdftex]{hyperref} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +%\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage[greek,english]{babel} + %option greek for \ + %option english (default language) for \, \ + +%\usepackage[latin1]{inputenc} + %for \, \, \, \, + %\, \, \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} + + +\begin{document} + +\title{tphols-2011} +\author{By xingyuan} +\maketitle + +\tableofcontents + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +\input{session} + +% optional bibliography +%\bibliographystyle{abbrv} +%\bibliography{root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff -r c64241fa4dff -r f5db9e08effc tphols-2011/generated/root.toc --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/tphols-2011/generated/root.toc Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,11 @@ +\contentsline {section}{\numberline {1}Preliminary definitions}{1}{section.1} +\contentsline {section}{\numberline {2}Direction \emph {\it finite\ partition\ {\emph {$\Rightarrow $}}\ regular\ language}}{5}{section.2} +\contentsline {subsection}{\numberline {2.1}Proof for this direction}{9}{subsection.2.1} +\contentsline {section}{\numberline {3}Direction: \emph {\it regular\ language\ {\emph {$\Rightarrow $}}finite\ partition}}{21}{section.3} +\contentsline {subsection}{\numberline {3.1}The scheme for this direction}{21}{subsection.3.1} +\contentsline {subsection}{\numberline {3.2}A small theory for list difference}{22}{subsection.3.2} +\contentsline {subsection}{\numberline {3.3}Lemmas for basic cases}{23}{subsection.3.3} +\contentsline {subsection}{\numberline {3.4}The case for \emph {\it SEQ}}{24}{subsection.3.4} +\contentsline {subsection}{\numberline {3.5}The case for \emph {\it ALT}}{26}{subsection.3.5} +\contentsline {subsection}{\numberline {3.6}The case for \emph {\it STAR}}{27}{subsection.3.6} +\contentsline {subsection}{\numberline {3.7}The main lemma}{29}{subsection.3.7} diff -r c64241fa4dff -r f5db9e08effc tphols-2011/generated/session.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/tphols-2011/generated/session.tex Mon Jan 24 11:29:55 2011 +0000 @@ -0,0 +1,6 @@ +\input{Myhill.tex} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r c64241fa4dff -r f5db9e08effc tphols-2011/myhill.pdf Binary file tphols-2011/myhill.pdf has changed