diff --git a/thys/Combinatorics_Words/CoWBasic.thy b/thys/Combinatorics_Words/CoWBasic.thy --- a/thys/Combinatorics_Words/CoWBasic.thy +++ b/thys/Combinatorics_Words/CoWBasic.thy @@ -1,3509 +1,3509 @@ (* Title: CoW/CoWBasic.thy Author: Štěpán Holub, Charles University Author: Martin Raška, Charles University Author: Štěpán Starosta, CTU in Prague *) theory CoWBasic imports "HOL-Library.Sublist" Arithmetical_Hints Reverse_Symmetry -begin +begin chapter "Basics of Combinatorics on Words" text\Combinatorics on Words, as the name suggests, studies words, finite or infinite sequences of elements from a, usually finite, alphabet. -The essential operation on finite words is the concatenation of two words, which is associative and noncommutative. +The essential operation on finite words is the concatenation of two words, which is associative and noncommutative. This operation yields many simply formulated problems, often in terms of \emph{equations on words}, that are mathematically challenging. See for instance @{cite ChoKa97} for an introduction to Combinatorics on Words, and \cite{Lo,Lo2,Lo3} as another reference for Combinatorics on Words. This theory deals exclusively with finite words and provides basic facts of the field which can be considered as folklore. The most natural way to represent finite words is by the type @{typ "'a list"}. From an algebraic viewpoint, lists are free monoids. On the other hand, any free monoid is isomoporphic to a monoid of lists of its generators. The algebraic point of view and the combinatorial point of view therefore overlap significantly in Combinatorics on Words. \ section "Definitions and notations" text\First, we introduce elementary definitions and notations.\ text\The concatenation @{term append} of two finite lists/words is the very basic operation in Combinatorics on Words, its notation is usually omitted. In this field, a common notation for this operation is $\cdot$, which we use and add here.\ notation append (infixr "\" 65) lemmas rassoc = append_assoc lemmas lassoc = append_assoc[symmetric] text\We add a common notation for the length of a given word $|w|$.\ notation length ("\<^bold>|_\<^bold>|") \ \note that it's bold |\ notation (latex output) length ("\<^latex>\\\ensuremath{\\left| \_\<^latex>\\\right|}\") subsection \Empty and nonempty word\ text\As the word of length zero @{term Nil} or @{term "[]"} will be used often, we adopt its frequent notation $\varepsilon $ in this formalization.\ notation Nil ("\") abbreviation drop_emp :: "'a list set \ 'a list set" ( "_\<^sub>+" [1000] ) where "drop_emp S \ S - {\}" subsection \Prefix\ text\The property of being a prefix shall be frequently used, and we give it yet another frequent shorthand notation. Analogously, we introduce shorthand notations for non-empty prefix and strict prefix, and continue with suffixes and factors. \ notation prefix (infixl "\p" 50) notation (latex output) prefix ("\\<^sub>p") lemmas [simp] = prefix_def lemmas prefI'[intro] = prefixI lemma prefI[intro]: "p \ s = w \ p \p w" by auto lemma prefD: "u \p v \ \ z. v = u \ z" unfolding prefix_def. -definition prefix_comparable :: "'a list \ 'a list \ bool" (infixl "\" 50) +definition prefix_comparable :: "'a list \ 'a list \ bool" (infixl "\" 50) where prefix_comparable_def[simp]: "(prefix_comparable u v) \ u \p v \ v \p u" lemma pref_compIff[iff]: "u \ v \ u \p v \ v \p u" by simp definition nonempty_prefix (infixl "\np" 50) where nonempty_prefix_def[simp]: "u \np v \ u \ \ \ u \p v" notation (latex output) nonempty_prefix ("\\<^bsub>np\<^esub>" 50) lemma npI[intro]: "u \ \ \ u \p v \ u \np v" by auto lemma npI'[intro]: "u \ \ \ (\ z. u \ z = v) \ u \np v" by auto lemma npD: "u \np v \ u \p v" by simp lemma npD': "u \np v \ u \ \" by simp notation strict_prefix (infixl "p") lemmas [simp] = strict_prefix_def lemma sprefI1[intro]: "v = u \ z \ z \ \ \ u

z = v \ z \ \ \ u

p v \ length u < length v \ u

u \p v \ u \ v" by auto -lemmas sprefD1[elim] = prefix_order.order.strict_implies_order and +lemmas sprefD1[elim] = prefix_order.strict_implies_order and sprefD2[elim] = prefix_order.less_imp_neq lemma sprefE [elim]: assumes "u

z = v" and "z \ \" using assms by auto subsection \Suffix\ notation suffix (infixl "\s" 60) notation (latex output) suffix ("\\<^sub>s") lemmas [simp] = suffix_def lemma sufI[intro]: "p \ s = w \ s \s w" by auto lemma sufD[elim]: "u \s v \ \ z. z \ u = v" by auto notation strict_suffix (infixl "s") lemmas [simp] = strict_suffix_def lemmas [intro] = suffix_order.le_neq_trans lemma ssufI1[intro]: "u \ v = w \ u \ \ \ v s v \ length u < length v \ u u \s v \ u \ v" by auto -lemmas ssufD1[elim] = suffix_order.order.strict_implies_order and +lemmas ssufD1[elim] = suffix_order.strict_implies_order and ssufD2[elim] = suffix_order.less_imp_neq -definition suffix_comparable :: "'a list \ 'a list \ bool" (infixl "\\<^sub>s" 50) +definition suffix_comparable :: "'a list \ 'a list \ bool" (infixl "\\<^sub>s" 50) where suffix_comparable_def[simp]: "(suffix_comparable u v) \ (rev u) \ (rev v)" definition nonempty_suffix (infixl "\ns" 60) where nonempty_suffix_def[simp]: "u \ns v \ u \ \ \ u \s v" notation (latex output) nonempty_suffix ("\\<^bsub>ns\<^esub>" 50) lemma nsI[intro]: "u \ \ \ u \s v \ u \ns v" by auto lemma nsI'[intro]: "u \ \ \ (\ z. z \ u = v) \ u \ns v" by blast lemma nsD: "u \ns v \ u \s v" by simp lemma nsD': "u \ns v \ u \ \" by simp subsection \Factor\ text\A @{term sublist} of some word is in Combinatorics of Words called a factor. We adopt a common shorthand notation for the property of being a factor, strict factor and nonempty factor (the latter we also define).\ notation sublist (infixl "\f" 60) notation (latex output) sublist ("\\<^sub>f") lemmas factor_def[simp] = sublist_def notation strict_sublist (infixl "f\<^esub>") lemmas strict_factor_def[simp] = strict_sublist_def definition nonempty_factor (infixl "\nf" 60) where nonempty_factor_def[simp]: "u \nf v \ u \ \ \ (\ p s. p\u\s = v)" notation (latex output) nonempty_factor ("\\<^bsub>nf\<^esub>") lemma facI: "u \f p\u\s" using sublist_appendI. lemma facI': "a \ u \ b = w \ u \f w" by auto -lemma facE[elim]: assumes "u \f v" +lemma facE[elim]: assumes "u \f v" obtains p s where "v = p \ u \ s" using assms unfolding factor_def by blast -lemma facE'[elim]: assumes "u \f v" +lemma facE'[elim]: assumes "u \f v" obtains p s where "p \ u \ s = v" using assms unfolding factor_def by blast section "Various elementary lemmas" lemmas concat_morph = sym[OF concat_append] lemmas cancel = same_append_eq and cancel_right = append_same_eq lemma bij_lists: "bij_betw f X Y \ bij_betw (map f) (lists X) (lists Y)" proof- assume "bij_betw f X Y" hence "inj_on f X" by (simp add: bij_betw_def) have "\ x y. x \ lists X \ y \ lists X \ (set x \ set y) \ X" - by blast + by blast hence "\ x y. x \ lists X \ y \ lists X \ inj_on f (set x \ set y)" using subset_inj_on[OF \inj_on f X\] by meson - hence "\ x y. x \ lists X \ y \ lists X \ map f x = map f y \ x = y" + hence "\ x y. x \ lists X \ y \ lists X \ map f x = map f y \ x = y" by (simp add: inj_on_map_eq_map) hence "inj_on (map f) (lists X)" - by (simp add: inj_on_def) + by (simp add: inj_on_def) thus ?thesis using \bij_betw f X Y\ bij_betw_def lists_image by metis qed lemma concat_sing': "concat [r] = r" by simp lemma concat_sing: "s = [hd s] \ concat s = hd s" using concat_sing'[of "hd s"] by auto lemma rev_sing: "rev [x] = [x]" by simp lemma hd_word: "a#ws = [a] \ ws" by simp lemma hd_word': "w \ \ \ [hd w] \ tl w = w" by simp lemma hd_pref: "w \ \ \ [hd w] \p w" using hd_word' by blast lemma add_nth: assumes "n < \<^bold>|w\<^bold>|" shows "(take n w) \ [w!n] \p w" using take_is_prefix[of "Suc n" w, unfolded take_Suc_conv_app_nth[OF assms]]. lemma hd_pref': "w \ \ \ [w!0] \p w" using add_nth by fastforce lemma sub_lists_mono: "A \ B \ x \ lists A \ x \ lists B" by auto lemma lists_hd: "us \ \ \ us \ lists Q \ hd us \ Q" by fastforce lemma replicate_in_lists: "replicate k z \ lists {z}" by (induction k) auto lemma tl_lists: assumes "us \ lists A" shows "tl us \ lists A" using suffix_lists[OF suffix_tl assms]. lemma long_list_tl: assumes "1 < \<^bold>|us\<^bold>|" shows "tl us \ \" proof assume "tl us = \" from assms have "us \ \" and "\<^bold>|us\<^bold>| = Suc \<^bold>|tl us\<^bold>|" and "\<^bold>|us\<^bold>| \ Suc 0" by auto thus False using \tl us = \\ by simp qed lemma tl_set: "x \ set (tl a) \ x \ set a" using list.sel(2) list.set_sel(2) by metis lemma drop_take_inv: "n \ \<^bold>|u\<^bold>| \ drop n (take n u \ w) = w" by simp -lemma split_list_long: assumes "1 < \<^bold>|us\<^bold>|" and "u \ set us" +lemma split_list_long: assumes "1 < \<^bold>|us\<^bold>|" and "u \ set us" obtains xs ys where "us = xs \ [u] \ ys" and "xs\ys\\" proof- obtain xs ys where "us = xs \ [u] \ ys" using split_list_first[OF \u \ set us\] by auto - hence "xs\ys\\" + hence "xs\ys\\" using \1 < \<^bold>|us\<^bold>|\ by auto from that[OF \us = xs \ [u] \ ys\ this] show thesis. qed lemma flatten_lists: "G \ lists B \ xs \ lists G \ concat xs \ lists B" proof (induct xs, simp) case (Cons a xs) hence "concat xs \ lists B" and "a \ lists B" - by auto + by auto thus ?case by simp qed lemma concat_map_sing_ident: "concat (map (\ x. [x]) xs) = xs" - by auto + by auto lemma hd_concat_tl: assumes "ws \ \" shows "hd ws \ concat (tl ws) = concat ws" using concat.simps(2)[of "hd ws" "tl ws", unfolded list.collapse[OF \ws \ \\], symmetric]. lemma concat_butlast_last: assumes "ws \ \" shows "concat (butlast ws) \ last ws = concat ws" using concat_morph[of "butlast ws" "[last ws]", unfolded concat_sing' append_butlast_last_id[OF \ws \ \\]]. lemma concat_last_suf: "ws \ \ \ last ws \s concat ws" using concat_butlast_last by blast lemma concat_hd_pref: "ws \ \ \ hd ws \p concat ws" using hd_concat_tl by blast lemma set_nemp_concat_nemp: assumes "ws \ \" and "\ \ set ws" shows "concat ws \ \" using \\ \ set ws\ last_in_set[OF \ws \ \\] concat_butlast_last[OF \ws \ \\] by fastforce lemmas takedrop = append_take_drop_id -lemma comm_rev_iff: "rev u \ rev v = rev v \ rev u \ u \ v = v \ u" +lemma comm_rev_iff: "rev u \ rev v = rev v \ rev u \ u \ v = v \ u" unfolding rev_append[symmetric] rev_is_rev_conv eq_ac(1)[of "u \ v"] by blast lemma rev_induct2: "\ P [] []; \x xs. P (xs\[x]) []; \y ys. P [] (ys\[y]); \x xs y ys. P xs ys \ P (xs\[x]) (ys\[y]) \ \ P xs ys" proof (induct xs arbitrary: ys rule: rev_induct) case Nil then show ?case using rev_induct[of "P \"] by presburger next case (snoc x xs) hence "P xs ys'" for ys' by simp then show ?case by (simp add: rev_induct snoc.prems(2) snoc.prems(4)) qed subsection \Orderings on lists: prefix, suffix, factor\ -lemmas self_pref = prefix_order.order.refl -lemmas pref_antisym = prefix_order.order.antisym -lemmas pref_trans = prefix_order.order.trans -lemmas suf_trans = suffix_order.order.trans +lemmas self_pref = prefix_order.refl +lemmas pref_antisym = prefix_order.antisym +lemmas pref_trans = prefix_order.trans +lemmas suf_trans = suffix_order.trans subsection "On the empty word" lemma nemp_elem_setI[intro]: "u \ S \ u \ \ \ u \ S\<^sub>+" by simp lemma nel_drop_emp: "u \ \ \ u \ S \ u \ S\<^sub>+" by simp lemma drop_emp_nel: assumes "u \ S\<^sub>+" shows "u \ \" and "u \ S" using assms by simp+ lemma emp_concat_emp: "us \ lists S\<^sub>+ \ concat us = \ \ us = \" using DiffD2 by auto lemma take_nemp: "w \ \ \ n \ 0 \ take n w \ \" by simp lemma pref_nemp [intro]: "u \ \ \ u \ v \ \" unfolding append_is_Nil_conv by simp lemma suf_nemp [intro]: "v \ \ \ u \ v \ \" unfolding append_is_Nil_conv by simp section "Length and its properties" lemma lenarg: "u = v \ \<^bold>|u\<^bold>| = \<^bold>|v\<^bold>|" by simp lemma npos_len: "\<^bold>|u\<^bold>| \ 0 \ u = \" by simp lemma nemp_pos_len: "r \ \ \ 1 \ \<^bold>|r\<^bold>|" by (simp add: leI) lemma swap_len: "\<^bold>|u \ v\<^bold>| = \<^bold>|v \ u\<^bold>|" by simp lemma len_after_drop: "p + q \ \<^bold>|w\<^bold>| \ q \ \<^bold>|drop p w\<^bold>|" - by simp + by simp lemma short_take_append: "n \ \<^bold>|u\<^bold>|\ take n (u \ v) = take n u" by simp lemma sing_word: "\<^bold>|us\<^bold>| = 1 \ [hd us] = us" by (cases us) simp+ lemma sing_word_concat: assumes "\<^bold>|us\<^bold>| = 1" shows "[concat us] = us" by (simp add: assms concat_sing sing_word) lemma nonsing_concat_len: "\<^bold>|us\<^bold>| \ 1 \ concat us \ \ \ 1 < \<^bold>|us\<^bold>|" using nat_neq_iff by fastforce lemma sing_len: "\<^bold>|[a]\<^bold>| = 1" by simp - + lemma pref_len': "\<^bold>|u\<^bold>| \ \<^bold>|u \ z\<^bold>|" by auto lemma suf_len': "\<^bold>|u\<^bold>| \ \<^bold>|z \ u\<^bold>|" by auto lemma fac_len: "u \f v \ \<^bold>|u\<^bold>| \ \<^bold>|v\<^bold>|" by auto lemma fac_len': "\<^bold>|w\<^bold>| \ \<^bold>|u \ w \ v\<^bold>|" by simp lemma fac_len_eq: "u \f v \ \<^bold>|u\<^bold>| = \<^bold>|v\<^bold>| \ u = v" unfolding factor_def using length_append npos_len by fastforce lemma drop_len: "\<^bold>|u \ w\<^bold>| \ \<^bold>|u \ v \ w\<^bold>|" by simp lemma drop_pref: "drop \<^bold>|u\<^bold>| (u \ w) = w" by simp lemma take_len: "p \ \<^bold>|w\<^bold>| \ \<^bold>|take p w\<^bold>| = p" using trans[OF length_take min_absorb2]. lemma conj_len: "p \ x = x \ s \ \<^bold>|p\<^bold>| = \<^bold>|s\<^bold>|" using length_append[of p x] length_append[of x s] add.commute add_left_imp_eq by auto lemma take_nemp_len: "u \ \ \ r \ \ \ take \<^bold>|r\<^bold>| u \ \" by simp -lemma nemp_len: "u \ \ \ \<^bold>|u\<^bold>| \ 0" +lemma nemp_len: "u \ \ \ \<^bold>|u\<^bold>| \ 0" by simp lemma take_self: "take \<^bold>|w\<^bold>| w = w" using take_all[of w "\<^bold>|w\<^bold>|", OF order.refl]. lemma len_le_concat: "\ \ set ws \ \<^bold>|ws\<^bold>| \ \<^bold>|concat ws\<^bold>|" proof (induct ws, simp) case (Cons a ws) hence "1 \ \<^bold>|a\<^bold>|" - using list.set_intros(1)[of a ws] nemp_pos_len[of a] by blast - then show ?case - unfolding concat.simps(2) unfolding length_append hd_word[of a ws] sing_len + using list.set_intros(1)[of a ws] nemp_pos_len[of a] by blast + then show ?case + unfolding concat.simps(2) unfolding length_append hd_word[of a ws] sing_len using Cons.hyps Cons.prems by simp qed lemma eq_len_iff: assumes eq: "x \ y = u \ v" shows "\<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>| \ \<^bold>|v\<^bold>| \ \<^bold>|y\<^bold>|" using lenarg[OF eq] unfolding length_append by auto lemma eq_len_iff_less: assumes eq: "x \ y = u \ v" shows "\<^bold>|x\<^bold>| < \<^bold>|u\<^bold>| \ \<^bold>|v\<^bold>| < \<^bold>|y\<^bold>|" using lenarg[OF eq] unfolding length_append by auto section "Prefix and prefix comparability properties" -lemmas pref_emp = prefix_bot.bot.extremum_uniqueI +lemmas pref_emp = prefix_bot.extremum_uniqueI lemma triv_pref: "r \p r \ s" using prefI[OF refl]. lemma triv_spref: "s \ \ \ r

s" by simp lemma pref_cancel: "z \ u \p z \ v \ u \p v" by simp lemma pref_cancel': "u \p v \ z \ u \p z \ v" by simp lemmas pref_cancel_conv = same_prefix_prefix lemmas pref_ext = prefix_prefix \ \provided by Sublist.thy\ lemma spref_ext: "r

r

v" by force lemma pref_ext_nemp: "r \p u \ v \ \ \ r

v" by auto lemma pref_take: "p \p w \ take \<^bold>|p\<^bold>| w = p" by auto lemma pref_take_conv: "take (\<^bold>|r\<^bold>|) w = r \ r \p w" using pref_take[of r w] take_is_prefix[of "\<^bold>|r\<^bold>|" w] by argo lemma le_suf_drop: assumes "i \ j" shows "drop j w \s drop i w" using suffix_drop[of "j - i" "drop i w", unfolded drop_drop le_add_diff_inverse2[OF \i \ j\]]. lemma spref_take: "p

take \<^bold>|p\<^bold>| w = p" by auto lemma pref_same_len: "u \p v \ \<^bold>|u\<^bold>| = \<^bold>|v\<^bold>| \ u = v" - by auto + by auto lemma add_nth_pref: assumes "u

[w!\<^bold>|u\<^bold>|] \p w" using add_nth[OF prefix_length_less[OF \u

], unfolded spref_take[OF \u

]]. lemma index_pref: "\<^bold>|u\<^bold>| \ \<^bold>|w\<^bold>| \ (\ i < \<^bold>|u\<^bold>|. u!i = w!i) \ u \p w" - using trans[OF sym[OF take_all[OF order_refl]] nth_take_lemma[OF order_refl], of u w] + using trans[OF sym[OF take_all[OF order_refl]] nth_take_lemma[OF order_refl], of u w] take_is_prefix[of "\<^bold>|u\<^bold>|" w] by auto lemma pref_index: assumes "u \p w" "i < \<^bold>|u\<^bold>|" shows "u!i = w!i" using nth_take[OF \i < \<^bold>|u\<^bold>|\, of w, unfolded pref_take[OF \u \p w\]]. lemma pref_drop: "u \p v \ drop p u \p drop p v" using prefI[OF sym[OF drop_append]] by auto subsection "Prefix comparability" lemma pref_comp_sym[sym]: "u \ v \ v \ u" by blast lemmas ruler_le = prefix_length_prefix and ruler = prefix_same_cases and - ruler' = prefix_same_cases[folded prefix_comparable_def] + ruler' = prefix_same_cases[folded prefix_comparable_def] lemma ruler_equal: "u \p w \ v \p w \ \<^bold>|u\<^bold>| = \<^bold>|v\<^bold>| \ u = v" by auto lemma ruler_comp: "u \p v \ u' \p v' \ v \ v' \ u \ u'" unfolding prefix_comparable_def using disjE[OF _ ruler[OF pref_trans] ruler[OF _ pref_trans]]. lemma ruler_pref: "w \p v\z \ w \ v" unfolding prefix_comparable_def using ruler by blast lemma pref_prod_pref_short: "u \p z \ w \ v \p w \ \<^bold>|u\<^bold>| \ \<^bold>|z \ v\<^bold>| \ u \p z \ v" using ruler_le[OF _ pref_cancel']. lemma pref_prod_pref: "u \p z \ w \ u \p w \ u \p z \ u" using pref_prod_pref_short[OF _ _ suf_len']. lemma pref_prod_pref': assumes "u \p z\u\w" shows "u \p z\u" using pref_prod_pref[of u z "u \ w", OF \u \p z\u\w\ triv_pref]. lemma pref_prod_long: "u \p v \ w \ \<^bold>|v\<^bold>| \ \<^bold>|u\<^bold>| \ v \p u" using ruler_le[OF triv_pref]. lemma pref_keeps_root: "u \p r \ u \ v \p u \ v \p r \ v" using pref_prod_pref[of v r u] pref_trans[of v u "r\u"] by blast lemma pref_prolong: "w \p z \ r \ r \p s \ w \p z \ s" using pref_trans[OF _ pref_cancel']. lemma pref_prolong': assumes "u \p w \ z" "v \ u \p z" shows "u \p w \ v \ u" using prefix_length_prefix[OF \u \p w \ z\ pref_cancel'[OF \v \ u \p z\, of w] suf_len'[of u "w\v", unfolded rassoc]]. lemma pref_prolong_comp: "u \p w \ z \ v \ u \ z \ u \p w \ v \ u" using pref_prolong[of u w z "v \ u"] pref_prolong'[of u w z v] by blast lemma pref_prod_short: "u \p v \ w \ \<^bold>|u\<^bold>| \ \<^bold>|v\<^bold>| \ u \p v" using prefI prefix_length_prefix[of u "v\w" v] by blast lemma pref_prod_short': assumes "u \p v \ w" and "\<^bold>|u\<^bold>| < \<^bold>|v\<^bold>|" shows "u

u \p v \ w\ less_imp_le[OF \\<^bold>|u\<^bold>| < \<^bold>|v\<^bold>|\]] \\<^bold>|u\<^bold>| < \<^bold>|v\<^bold>|\ by blast lemma pref_prod_cancel: assumes "u \p p\w\q" and "\<^bold>|p\<^bold>| \ \<^bold>|u\<^bold>|" and "\<^bold>|u\<^bold>| \ \<^bold>|p\w\<^bold>|" obtains r where "u = p \ r" and "r \p w" proof- have "p \p u" using pref_prod_long[OF \u \p p\w\q\ \\<^bold>|p\<^bold>| \ \<^bold>|u\<^bold>|\]. - then obtain r where "u = p \ r" using prefD by blast - hence "r \p w" using \\<^bold>|u\<^bold>| \ \<^bold>|p\w\<^bold>|\ \u \p p\w\q\ + then obtain r where "u = p \ r" using prefD by blast + hence "r \p w" using \\<^bold>|u\<^bold>| \ \<^bold>|p\w\<^bold>|\ \u \p p\w\q\ unfolding \u = p \ r\ pref_cancel_conv length_append using pref_prod_short[of r w q] by simp from that[OF \u = p \ r\ this] show thesis. qed lemma pref_prod_cancel': assumes "u \p p\w\q" and "\<^bold>|p\<^bold>| < \<^bold>|u\<^bold>|" and "\<^bold>|u\<^bold>| \ \<^bold>|p\w\<^bold>|" - obtains r where "u = p \ r" and "r \p w" and "r \ \" + obtains r where "u = p \ r" and "r \p w" and "r \ \" proof- - obtain r where "u = p \ r" and "r \p w" + obtain r where "u = p \ r" and "r \p w" using pref_prod_cancel[OF \u \p p\w\q\ less_imp_le[OF \\<^bold>|p\<^bold>| < \<^bold>|u\<^bold>|\] \\<^bold>|u\<^bold>| \ \<^bold>|p\w\<^bold>|\]. moreover have "r \ \" using \u = p \ r\ less_not_refl3[OF \\<^bold>|p\<^bold>| < \<^bold>|u\<^bold>|\, folded self_append_conv] by simp ultimately show thesis using that by simp qed lemma pref_comp_eq: "u \ v \\<^bold>|u\<^bold>| = \<^bold>|v\<^bold>| \ u = v" by auto lemma non_comp_parallel: "\ u \ v \ u \ v" unfolding prefix_comparable_def parallel_def de_Morgan_disj.. lemma comp_refl: "u \ u" by simp lemma incomp_cancel: "\ p\u \ p\v \ \ u \ v" by simp lemma comp_cancel: "z \ u \ z \ v \ u \ v" by simp lemma comm_ruler: "r \ s \p w1 \ s \ r \p w2 \ w1 \ w2 \ r \ s = s \ r" using pref_comp_eq[OF ruler_comp swap_len]. lemma pref_share_take: "u \p v \ q \ \<^bold>|u\<^bold>| \ take q u = take q v" by auto -lemma pref_prod_longer: "u \p z \ w \ v \p w \ \<^bold>|z \ v\<^bold>| \ \<^bold>|u\<^bold>| \ z \ v \p u" +lemma pref_prod_longer: "u \p z \ w \ v \p w \ \<^bold>|z \ v\<^bold>| \ \<^bold>|u\<^bold>| \ z \ v \p u" using ruler_le[OF pref_cancel']. lemma pref_comp_not_pref: "u \ v \ \ v \p u \ u

v \ \ u

v \p u" using contrapos_np[OF _ pref_comp_not_pref]. lemma hd_prod: "u \ \ \ (u \ v)!0 = u!0" by (cases u) (blast, simp) lemma distinct_first: assumes "w \ \" "z \ \" "w!0 \ z!0" shows "w \ w' \ z \ z'" - using hd_prod[of w w', OF \w \ \\] hd_prod[of z z', OF \z \ \\] \w!0 \ z!0\ by auto + using hd_prod[of w w', OF \w \ \\] hd_prod[of z z', OF \z \ \\] \w!0 \ z!0\ by auto lemmas last_no_split = prefix_snoc lemma last_no_split': assumes "u

p u \ [a]" shows "w = u \ [a]" using assms(1)[unfolded prefix_order.less_le_not_le] assms(2)[unfolded last_no_split] by presburger lemma pcomp_shorter: "v \ w \ \<^bold>|v\<^bold>| \ \<^bold>|w\<^bold>| \ v \p w" by auto lemma pref_comp_len_trans: "w \p v \ u \ v \ \<^bold>|w\<^bold>| \ \<^bold>|u\<^bold>| \ w \p u" unfolding prefix_comparable_def - using prefix_length_prefix[of w v u] prefix_order.order.trans[of w v u] + using prefix_length_prefix[of w v u] prefix_order.trans[of w v u] by argo lemma comp_ext: "z \ w1 \ z \ w2 \ w1 \ w2" using pref_cancel by auto lemma emp_pref: "\ \p u" by simp lemma emp_spref: "u \ \ \ \

p v \ \<^bold>|v\<^bold>| \ \<^bold>|u\<^bold>| \ u = v" by auto lemma incomp_ext: "\ w1 \ w2 \ \ w1 \ z \ w2 \ z'" using contrapos_nn[OF _ ruler_comp[OF triv_pref triv_pref]]. lemma mismatch_incopm: "\<^bold>|u\<^bold>| = \<^bold>|v\<^bold>| \ x \ y \ \ u \ [x] \ v \ [y]" by simp lemma comp_prefs_comp: "u \ z \ v \ w \ u \ v" using ruler_comp[OF prefI[of _ z] prefI[of _ w], OF refl refl]. lemma comp_hd_eq: "u \ v \ u \ \ \ v \ \ \ hd u = hd v" - by auto + by auto lemma pref_hd_eq': "p \p u \ p \p v \ p \ \ \ hd u = hd v" by auto -lemma pref_hd_eq: "u \p v \ u \ \ \ hd u = hd v" +lemma pref_hd_eq: "u \p v \ u \ \ \ hd u = hd v" by auto lemma suf_last_eq: "p \s u \ p \s v \ p \ \ \ last u = last v" by auto lemma comp_hd_eq': assumes "u \ r \ v \ s" "u \ \" "v \ \" shows "hd u = hd v" - using comp_prefs_comp[OF \u \ r \ v \ s\] \u \ \\ \v \ \\ by auto + using comp_prefs_comp[OF \u \ r \ v \ s\] \u \ \\ \v \ \\ by auto section "Suffix and suffix comparability properties" -lemmas suf_emp = suffix_bot.bot.extremum_uniqueI +lemmas suf_emp = suffix_bot.extremum_uniqueI lemma triv_suf: "u \s v \ u" by simp lemma emp_ssuf: "u \ \ \ \ v \s w\v \ u \s w" - by simp + by simp lemma suf_cancel': "u \s w \ u\v \s w\v" - by simp + by simp lemmas suf_cancel_eq = same_suffix_suffix \ \provided by Sublist.thy\ text\Straightforward relations of suffix and prefix follow.\ lemmas suf_rev_pref_iff = suffix_to_prefix \ \provided by Sublist.thy\ lemmas ssuf_rev_pref_iff = strict_suffix_to_prefix \ \provided by Sublist.thy\ lemma pref_rev_suf_iff: "u \p v \ rev u \s rev v" using suffix_to_prefix[of "rev u" "rev v"] unfolding rev_rev_ident by blast lemma spref_rev_suf_iff: "s

rev s ns w \ rev s \np rev w" unfolding nonempty_prefix_def nonempty_suffix_def suffix_to_prefix by fast lemma npref_rev_suf_iff: "s \np w \ rev s \ns rev w" unfolding nonempty_prefix_def nonempty_suffix_def pref_rev_suf_iff by fast -lemmas [reversal_rule] = +lemmas [reversal_rule] = suf_rev_pref_iff[symmetric] pref_rev_suf_iff[symmetric] nsuf_rev_pref_iff[symmetric] npref_rev_suf_iff[symmetric] ssuf_rev_pref_iff[symmetric] spref_rev_suf_iff[symmetric] lemmas suf_ext = suffix_appendI \ \provided by Sublist.thy\ lemmas ssuf_ext = spref_ext[reversed] and suf_ext_nem = pref_ext_nemp[reversed] and suf_same_len = pref_same_len[reversed] and suf_take = pref_drop[reversed] and suf_share_take = pref_share_take[reversed] and long_suf = long_pref[reversed] subsection "Suffix comparability" lemma pref_comp_rev_suf_comp[reversal_rule]: "(rev w) \\<^sub>s (rev v) \ w \ v" by simp lemma suf_comp_rev_pref_comp[reversal_rule]: "(rev w) \ (rev v) \ w \\<^sub>s v" by simp lemmas suf_ruler_le = suffix_length_suffix \ \provided by Sublist.thy, same as ruler\_le[reversed]\ lemmas suf_ruler = suffix_same_cases \ \provided by Sublist.thy, same as ruler[reversed]\ lemmas suf_ruler_equal = ruler_equal[reversed] and suf_ruler_comp = ruler_comp[reversed] and ruler_suf = ruler_pref[reversed] and suf_prod_short = pref_prod_short[reversed] and suf_prod_short' = pref_prod_short'[reversed] and suf_prod_cancel = pref_prod_cancel[reversed] and suf_prod_cancel' = pref_prod_cancel'[reversed] and suf_prod_suf_short = pref_prod_pref_short[reversed] and suf_prod_suf = pref_prod_pref[reversed] and suf_prod_suf' = pref_prod_pref'[reversed, unfolded rassoc] and suf_prolong = pref_prolong[reversed] and suf_prolong' = pref_prolong'[reversed, unfolded rassoc] and suf_prolong_comp = pref_prolong_comp[reversed, unfolded rassoc] and suf_prod_long = pref_prod_long[reversed] and suf_prod_longer = pref_prod_longer[reversed] and suf_keeps_root = pref_keeps_root[reversed] and comm_suf_ruler = comm_ruler[reversed] lemmas comp_sufs_comp = comp_prefs_comp[reversed] and suf_comp_not_suf = pref_comp_not_pref[reversed] and suf_comp_not_ssuf = pref_comp_not_spref[reversed] and (* hd_no_split = last_no_split[reversed] *) (* this is suffix_Cons *) suf_comp_ext = comp_ext[reversed] and suf_incomp_ext = incomp_ext[reversed] and mismatch_suf_incopm = mismatch_incopm[reversed] and suf_comp_sym[sym] = pref_comp_sym[reversed] lemma suf_comp_last_eq: assumes "u \\<^sub>s v" "u \ \" "v \ \" - shows "last u = last v" - using comp_hd_eq[reversed, OF assms] unfolding hd_rev[OF \u \ \\] hd_rev[OF \v \ \\]. + shows "last u = last v" + using comp_hd_eq[reversed, OF assms] unfolding hd_rev hd_rev . lemma suf_comp_last_eq': "r \ u \\<^sub>s s \ v \ u \ \ \ v \ \ \ last u = last v" using comp_sufs_comp suf_comp_last_eq by blast section "Left and Right Quotient" text\A useful function of left quotient is given. Note that the function is sometimes undefined.\ definition left_quotient:: "'a list \ 'a list \ 'a list" ("(_\\<^sup>>)(_)" [75,74] 74) - where left_quotient_def[simp]: "left_quotient u v = (if u \p v then (THE z. u \ z = v) else undefined)" + where left_quotient_def[simp]: "left_quotient u v = (if u \p v then (THE z. u \ z = v) else undefined)" notation (latex output) left_quotient ("\<^latex>\\\ensuremath{ {\_ \<^latex>\}^{-1} \\cdot {\ _ \<^latex>\}}\") text\Analogously, we define the right quotient.\ definition right_quotient :: "'a list \ 'a list \ 'a list" ("(_)(\<^sup><\_) " [76,77] 76) where right_quotient_def[simp]: "right_quotient u v = rev ((rev v)\\<^sup>>(rev u))" notation (latex output) right_quotient ("\<^latex>\\\ensuremath{ {\_ \<^latex>\} \\cdot {\ _ \<^latex>\}^{-1}}\") text\Priorities of these operations are as follows:\ lemma "u\<^sup><\v\<^sup><\w = (u\<^sup><\v)\<^sup><\w" by simp lemma "u\\<^sup>>v\\<^sup>>w = u\\<^sup>>(v\\<^sup>>w)" by simp lemma "u\\<^sup>>v\<^sup><\w = u\\<^sup>>(v\<^sup><\w)" by simp lemma "r \ u\\<^sup>>w\<^sup><\v \ s = r \ (u\\<^sup>>w\<^sup><\v) \ s" by simp lemma rq_rev_lq[reversal_rule]: "(rev v)\<^sup><\(rev u) = rev (u\\<^sup>>v)" by simp lemma lq_rev_rq[reversal_rule]: "(rev v)\\<^sup>>rev u = rev (u\<^sup><\v)" by simp subsection \Left Quotient\ lemma lqI: "u \ z = v \ u\\<^sup>>v = z" by auto lemma lq_triv[simp]: "u\\<^sup>>(u \ z) = z" using lqI[OF refl]. lemma lq_triv'[simp]: "u \ u\\<^sup>>(u \ z) = u \z" by simp -lemma lq_self[simp]: "u\\<^sup>>u = \" +lemma lq_self[simp]: "u\\<^sup>>u = \" by auto -lemma lq_emp[simp]: "\\\<^sup>>u = u" - by auto +lemma lq_emp[simp]: "\\\<^sup>>u = u" + by auto lemma lq_pref[simp]: "u \p v \ u \ (u\\<^sup>>v) = v" by auto lemma lq_the[simp]: "u \p v \ (u\\<^sup>>v) = (THE z. u \ z = v)" by simp lemma lq_reassoc: "u \p v \ (u\\<^sup>>v)\w = u\\<^sup>>(v\w)" by auto lemma lq_trans: "u \p v \ v \p w \ (u\\<^sup>>v) \ (v\\<^sup>>w) = u\\<^sup>>w" by auto lemma lq_rq_reassoc_suf: "u \p z \ u \s w \ w\u\\<^sup>>z = w\<^sup><\u \ z" using lq_pref[reversed] by fastforce lemma lq_ne: "p \p u\p \ u \ \ \ p\\<^sup>>(u\p) \ \" - using lq_pref[of p "u \ p"] by fastforce + using lq_pref[of p "u \ p"] by fastforce lemma lq_spref: "u

u\\<^sup>>v \ \" using lq_pref by auto lemma lq_suf_suf: "r \p s \ (r\\<^sup>>s) \s s" by auto lemma lq_len: "r \p s \ \<^bold>|r\<^bold>| + \<^bold>|r\\<^sup>>s\<^bold>| = \<^bold>|s\<^bold>|" by auto lemma pref_lq: "u \p v \ v \p w \ u\\<^sup>>v \p u\\<^sup>>w" by auto lemma spref_lq: "u \p v \ v

u\\<^sup>>v

\<^sup>>w" by force lemma conjug_lq: "x \ z = z \ y \ y = z\\<^sup>>(x \ z)" by simp lemma conjug_emp_emp: "p \p u \ p \ p\\<^sup>>(u \ p) = \ \ u = \" - using lq_ne by blast + using lq_ne by blast lemma lq_drop: "u \p v \ u\\<^sup>>v = drop \<^bold>|u\<^bold>| v" by fastforce lemma lq_code [code]: - "left_quotient \ v = v" + "left_quotient \ v = v" "left_quotient (a#u) \ = undefined" "left_quotient (a#u) (b#v) = (if a=b then left_quotient u v else undefined)" by simp_all subsection "Right quotient" lemmas rqI = lqI[reversed] and rq_triv = lq_triv[reversed] and rq_triv' = lq_triv'[reversed] and rq_sefl = lq_self[reversed] and rq_emp = lq_emp[reversed] and rq_suf = lq_pref[reversed] and rq_ssuf = lq_spref[reversed] and rq_reassoc = lq_reassoc[reversed] and rq_len = lq_len[reversed] and rq_trans = lq_trans[reversed] and rq_lq_reassoc_suf = lq_rq_reassoc_suf[reversed] and rq_ne = lq_ne[reversed] and rq_suf_suf = lq_suf_suf[reversed] and suf_rq = pref_lq[reversed] and ssuf_rq = spref_lq[reversed] and conjug_rq = conjug_lq[reversed] and conjug_emp_emp' = conjug_emp_emp[reversed] and rq_take = lq_drop[reversed] subsection \Left and right quotients combined\ lemma rev_lq': "r \p s \ rev (r\\<^sup>>s) = (rev s)\<^sup><\(rev r)" by auto lemma pref_rq_suf_lq: "s \s u \ r \p (u\<^sup><\s) \ s \s (r\\<^sup>>u)" using lq_reassoc[of r "u\<^sup><\s" s] rq_suf[of s u] triv_suf[of s "r\\<^sup>>u\<^sup><\s"] by presburger lemmas suf_lq_pref_rq = pref_rq_suf_lq[reversed] lemma "w\s = v \ v\<^sup><\s = w" using rqI. lemma lq_rq_assoc: "s \s u \ r \p (u\<^sup><\s) \ (r\\<^sup>>u)\<^sup><\s = r\\<^sup>>(u\<^sup><\s)" using lq_reassoc[of r "u\<^sup><\s" s] rq_suf[of s u] rqI[of "r\\<^sup>>u\<^sup><\s" s "r\\<^sup>>u"] by argo lemmas rq_lq_assoc = lq_rq_assoc[reversed] lemma lq_prod: "u \p v\u \ u \p w \ u\\<^sup>>(v\u)\u\\<^sup>>w = u\\<^sup>>(v\w)" using lq_reassoc[of u "v \ u" "u\\<^sup>>w"] lq_rq_reassoc_suf[of u w "v \ u", unfolded rq_triv[of v u]] by auto lemmas rq_prod = lq_prod[reversed] section \Equidivisibility\ text\Equidivisibility is the following property: if \[ xy = uv, \] then there exists a word $t$ such that $xt = u$ and $ty = v$, or $ut = x$ and $y = tv$. For monoids over words, this property is equivalent to the freeness of the monoid. As the monoid of all words is free, we can prove that it is equidivisible. Related lemmas based on this property follow. \ lemma eqd: "x \ y = u \ v \ \<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>| \ \ t. x \ t = u \ t \ v = y" by (simp add: append_eq_conv_conj) -lemma eqdE: assumes "x \ y = u \ v" and "\<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>|" +lemma eqdE: assumes "x \ y = u \ v" and "\<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>|" obtains t where "x \ t = u" and "t \ v = y" using eqd[OF assms] by blast -lemma eqdE': assumes "x \ y = u \ v" and "\<^bold>|v\<^bold>| \ \<^bold>|y\<^bold>|" +lemma eqdE': assumes "x \ y = u \ v" and "\<^bold>|v\<^bold>| \ \<^bold>|y\<^bold>|" obtains t where "x \ t = u" and "t \ v = y" using eqdE[OF assms(1)] lenarg[OF assms(1), unfolded length_append] assms(2) - by auto + by auto lemma eqd_pref: "x \ y = u \ v \ \<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>| \ x \ (x\\<^sup>>u) = u \ (x\\<^sup>>u) \ v = y" using eqd lq_triv by blast -lemma eqd_prefE: assumes "x \ y = u \ v" and "\<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>|" +lemma eqd_prefE: assumes "x \ y = u \ v" and "\<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>|" obtains t where "x \ t = u" and "t \ v = y" using eqd_pref assms by blast lemma eqd_pref1: "x \ y = u \ v \ \<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>| \ x \ (x\\<^sup>>u) = u" using eqd_pref by blast lemma eqd_pref2: "x \ y = u \ v \ \<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>| \ (x\\<^sup>>u) \ v = y" using eqd_pref by blast lemma eqd_equal: "x \ y = u \ v \ \<^bold>|x\<^bold>| = \<^bold>|u\<^bold>| \ x = u \ y = v" - by simp + by simp lemma pref_equal: "u \p v \ w \ \<^bold>|u\<^bold>| = \<^bold>|v\<^bold>| \ u = v" by simp lemma eqd_equal_suf: "x \ y = u \ v \ \<^bold>|y\<^bold>| = \<^bold>|v\<^bold>| \ x = u \ y = v" - by simp + by simp lemma eqd_comp: assumes "x \ y = u \ v" shows "x \ u" - using le_cases[of "\<^bold>|x\<^bold>|" "\<^bold>|u\<^bold>|" "x \ u"] - eqd_pref1[of x y u v, THEN prefI[of x "x\\<^sup>>u" u], OF assms] + using le_cases[of "\<^bold>|x\<^bold>|" "\<^bold>|u\<^bold>|" "x \ u"] + eqd_pref1[of x y u v, THEN prefI[of x "x\\<^sup>>u" u], OF assms] eqd_pref1[of u v x y, THEN prefI[of u "u\\<^sup>>x" x], OF assms[symmetric]] by auto \ \not equal to eqd\_pref1[reversed]\ lemma eqd_suf1: "x \ y = u \ v \ \<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>| \ (y\<^sup><\v)\v = y" using eqd_pref2 rq_triv by blast -\ \not equal to eqd\_pref2[reversed]\ +\ \not equal to eqd\_pref2[reversed]\ lemma eqd_suf2: assumes "x \ y = u \ v" "\<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>|" shows "x \ (y\<^sup><\v) = u" - using rq_reassoc[OF sufI[OF eqd_suf1[OF \x \ y = u \ v\ \\<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>|\]], of x, unfolded \x \ y = u \ v\ rq_triv[of u v]]. + using rq_reassoc[OF sufI[OF eqd_suf1[OF \x \ y = u \ v\ \\<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>|\]], of x, unfolded \x \ y = u \ v\ rq_triv[of u v]]. \ \ not equal to eqd\_pref[reversed] \ -lemma eqd_suf: assumes "x \ y = u \ v" and "\<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>|" +lemma eqd_suf: assumes "x \ y = u \ v" and "\<^bold>|x\<^bold>| \ \<^bold>|u\<^bold>|" shows "(y\<^sup><\v)\v = y \ x \ (y\<^sup><\v) = u" - using eqd_suf1[OF assms] eqd_suf2[OF assms] by blast + using eqd_suf1[OF assms] eqd_suf2[OF assms] by blast section \Longest common prefix\ notation longest_common_prefix ("_ \\<^sub>p _" [61,62] 64) \ \provided by Sublist.thy\ lemmas lcp_simps = longest_common_prefix.simps \ \provided by Sublist.thy\ -lemma lcp_sym: "u \\<^sub>p v = v \\<^sub>p u" +lemma lcp_sym: "u \\<^sub>p v = v \\<^sub>p u" by (induct u v rule: list_induct2') auto \ \provided by Sublist.thy\ lemmas lcp_pref = longest_common_prefix_prefix1 lemmas lcp_pref' = longest_common_prefix_prefix2 lemmas pref_pref_lcp = longest_common_prefix_max_prefix lemma lcp_take_eq: "take (\<^bold>|u \\<^sub>p v\<^bold>|) u = take (\<^bold>|u \\<^sub>p v\<^bold>|) v" using pref_take[OF lcp_pref[of u v]] pref_take[OF lcp_pref'[of u v]] by simp lemma lcp_pref_conv: "u \\<^sub>p v = u \ u \p v" unfolding prefix_order.eq_iff[of "u \\<^sub>p v" u] - using lcp_pref'[of u v] + using lcp_pref'[of u v] lcp_pref[of u v] longest_common_prefix_max_prefix[OF self_pref[of u], of v] by auto lemma pref_lcp_pref: "w \p u \\<^sub>p v \ w \p u" - using lcp_pref pref_trans by blast + using lcp_pref pref_trans by blast lemma pref_lcp_pref': "w \p u \\<^sub>p v \ w \p v" using pref_lcp_pref[of w v u, unfolded lcp_sym[of v u]]. lemma lcp_self[simp]: "w \\<^sub>p w = w" using lcp_pref_conv by blast lemma lcp_eq: "\<^bold>|u\<^bold>| = \<^bold>|u \\<^sub>p v\<^bold>| \ u = u \\<^sub>p v" - using long_pref[OF lcp_pref, of u v] by auto + using long_pref[OF lcp_pref, of u v] by auto lemma lcp_len: "\<^bold>|u\<^bold>| \ \<^bold>|u \\<^sub>p v\<^bold>| \ u \p v" - using long_pref[OF lcp_pref, of u v] unfolding lcp_pref_conv[symmetric]. + using long_pref[OF lcp_pref, of u v] unfolding lcp_pref_conv[symmetric]. lemma lcp_len': "\ u \p v \ \<^bold>|u \\<^sub>p v\<^bold>| < \<^bold>|u\<^bold>|" using not_le_imp_less[OF contrapos_nn[OF _ lcp_len]]. lemma incomp_lcp_len: assumes "\ u \ v" shows "\<^bold>|u \\<^sub>p v\<^bold>| < min \<^bold>|u\<^bold>| \<^bold>|v\<^bold>|" unfolding min_less_iff_conj[of "\<^bold>|u \\<^sub>p v\<^bold>|" "\<^bold>|u\<^bold>|" "\<^bold>|v\<^bold>|"] - using assms lcp_len'[of u v] lcp_len'[of v u, folded lcp_sym[of u v]] + using assms lcp_len'[of u v] lcp_len'[of v u, folded lcp_sym[of u v]] min_less_iff_conj[of "\<^bold>|u \\<^sub>p v\<^bold>|" "\<^bold>|u\<^bold>|" "\<^bold>|v\<^bold>|"] by blast lemma lcp_ext_right [case_names comp non_comp]: obtains "r \ r'" | "(r \ u) \\<^sub>p (r' \ v) = r \\<^sub>p r'" proof- have "\ r \ r' \ r \ u \\<^sub>p r' \ v = r \\<^sub>p r'" by (induct r r' rule: list_induct2', simp+) thus ?thesis using that(1) that(2) by linarith qed lemma lcp_same_len: "\<^bold>|u\<^bold>| = \<^bold>|v\<^bold>| \ u \ v \ u \\<^sub>p v = u \ w \\<^sub>p v \ w'" using lcp_ext_right[of u v _ w w'] pref_comp_eq[of u v] by argo lemma lcp_mismatch: "\<^bold>|u \\<^sub>p v\<^bold>| < \<^bold>|u\<^bold>| \ \<^bold>|u \\<^sub>p v\<^bold>| < \<^bold>|v\<^bold>| \ u! \<^bold>|u \\<^sub>p v\<^bold>| \ v! \<^bold>|u \\<^sub>p v\<^bold>|" by (induct u v rule: list_induct2') auto lemma lcp_mismatch': assumes "\ u \ v" shows "u! \<^bold>|u \\<^sub>p v\<^bold>| \ v! \<^bold>|u \\<^sub>p v\<^bold>|" using incomp_lcp_len[OF assms, unfolded min_less_iff_conj] lcp_mismatch by blast lemma lcp_ext_left: "(z \ u) \\<^sub>p (z \ v) = z \ (u \\<^sub>p v)" by (induct z) auto lemma lcp_first_letters: "u!0 \ v!0 \ u \\<^sub>p v = \" by (induct u v rule: list_induct2') auto lemma lcp_first_mismatch: "a \ b \ w \ [a] \ u \\<^sub>p w \ [b] \ v = w" by (simp add: lcp_ext_left) lemma lcp_first_mismatch': "a \ b \ u \ [a] \\<^sub>p u \ [b] = u" using lcp_first_mismatch[of a b u \ \] by simp lemma lcp_mismatch_shorter: assumes "\<^bold>|u\<^bold>| = \<^bold>|v\<^bold>|" "x \ y" shows "u \ [x] \\<^sub>p v \ [y] = u \\<^sub>p v" - by (cases "u = v") + by (cases "u = v") (simp add: lcp_self[of v] lcp_first_mismatch'[OF \x \ y\, of v], use lcp_same_len[OF \\<^bold>|u\<^bold>| = \<^bold>|v\<^bold>|\, of "[x]" "[y]"] in auto) lemma lcp_rulers: "r \p s \ r' \p s' \ (r \ r' \ r \\<^sub>p r' = s \\<^sub>p s')" using lcp_ext_right prefD[of r s] prefD[of r' s'] by metis lemma pref_pref_lcp': "w \p r \ w' \p s \ w \\<^sub>p w' \p (r \\<^sub>p s)" using pref_pref_lcp lcp_pref lcp_sym pref_trans by metis lemma lcp_distinct_hd: "hd u \ hd v \ u \\<^sub>p v = \" proof- assume "hd u \ hd v" hence "(u \ \ \ v \ \) \ hd u \ hd v \ u \\<^sub>p v = \" by (simp add: lcp_first_letters hd_conv_nth) moreover have "u = \ \ v = \ \ u \\<^sub>p v = \" - using lcp_pref' by auto + using lcp_pref' by auto ultimately show ?thesis using \hd u \ hd v\ by blast -qed - -lemma lcp_lenI: assumes "i < min \<^bold>|u\<^bold>| \<^bold>|v\<^bold>|" and "take i u = take i v" and "u!i \ v!i" - shows "i = \<^bold>|u \\<^sub>p v\<^bold>|" +qed + +lemma lcp_lenI: assumes "i < min \<^bold>|u\<^bold>| \<^bold>|v\<^bold>|" and "take i u = take i v" and "u!i \ v!i" + shows "i = \<^bold>|u \\<^sub>p v\<^bold>|" proof- - have u: "take i u \ [u ! i] \ (drop (Suc i) u) = u" - using \i < min \<^bold>|u\<^bold>| \<^bold>|v\<^bold>|\ id_take_nth_drop[of i u] by simp + have u: "take i u \ [u ! i] \ (drop (Suc i) u) = u" + using \i < min \<^bold>|u\<^bold>| \<^bold>|v\<^bold>|\ id_take_nth_drop[of i u] by simp have v: "take i u \ [v ! i] \ drop (Suc i) v = v" using \i < min \<^bold>|u\<^bold>| \<^bold>|v\<^bold>|\ - unfolding \take i u = take i v\ using id_take_nth_drop by fastforce + unfolding \take i u = take i v\ using id_take_nth_drop by fastforce from lcp_first_mismatch[OF \u!i \ v!i\, of "take i u" "drop (Suc i) u" "drop (Suc i) v", unfolded u v] have "u \\<^sub>p v = take i u". thus ?thesis - using \i < min \<^bold>|u\<^bold>| \<^bold>|v\<^bold>|\ by auto + using \i < min \<^bold>|u\<^bold>| \<^bold>|v\<^bold>|\ by auto qed lemma lcp_prefs: "\<^bold>|u \ w \\<^sub>p v \ w'\<^bold>| < \<^bold>|u\<^bold>| \ \<^bold>|u \ w \\<^sub>p v \ w'\<^bold>| < \<^bold>|v\<^bold>| \ u \\<^sub>p v = u \ w \\<^sub>p v \ w'" by (induct u v rule: list_induct2') auto subsection "Longest common prefix and prefix comparability" lemma lexord_cancel_right: "(u \ z, v \ w) \ lexord r \ \ u \ v \ (u,v) \ lexord r" by (induction rule: list_induct2', simp+, auto) lemma lcp_ruler: "r \ w1 \ r \ w2 \ \ w1 \ w2 \ r \p w1 \\<^sub>p w2" unfolding prefix_comparable_def by (meson pref_pref_lcp pref_trans ruler) lemma comp_monotone: "w \ r \ u \p w \ u \ r" using pref_trans[of u w r] ruler[of u w r] by blast lemma comp_monotone': "w \ r \ w \\<^sub>p w' \ r" using comp_monotone[of w r "w \\<^sub>p w'", OF _ longest_common_prefix_prefix1]. -lemma double_ruler: assumes "w \ r" and "w' \ r'" and "\ r \ r'" +lemma double_ruler: assumes "w \ r" and "w' \ r'" and "\ r \ r'" shows "w \\<^sub>p w' \p r \\<^sub>p r'" using comp_monotone'[OF \w' \ r'\, of w] unfolding lcp_sym[of w' w] using lcp_ruler[OF comp_monotone'[OF \w \ r\, of w'] _ \\ r \ r'\] by blast -lemma pref_comp_ruler: assumes "w \ u \ [x]" and "w \ v \ [y]" and "x \ y" and "\<^bold>|u\<^bold>| = \<^bold>|v\<^bold>|" +lemma pref_comp_ruler: assumes "w \ u \ [x]" and "w \ v \ [y]" and "x \ y" and "\<^bold>|u\<^bold>| = \<^bold>|v\<^bold>|" shows "w \p u \ w \p v" using double_ruler[OF \w \ u \ [x]\ \w \ v \ [y]\ mismatch_incopm[OF \\<^bold>|u\<^bold>| = \<^bold>|v\<^bold>|\ \x \ y\]] unfolding lcp_self[of w] lcp_mismatch_shorter[OF \\<^bold>|u\<^bold>| = \<^bold>|v\<^bold>|\ \x \ y\] using pref_lcp_pref pref_lcp_pref' by blast lemmas suf_comp_ruler = pref_comp_ruler[reversed] section "Mismatch" text \The first pair of letters on which two words/lists disagree\ fun mismatch_pair :: "'a list \ 'a list \ ('a \ 'a)" where "mismatch_pair \ v = (\!0, v!0)" | "mismatch_pair v \ = (v!0, \!0)" | "mismatch_pair (a#u) (b#v) = (if a=b then mismatch_pair u v else (a,b))" text \Alternatively, mismatch pair may be defined using the longest common prefix as follows.\ lemma mismatch_pair_lcp: "mismatch_pair u v = (u!\<^bold>|u\\<^sub>pv\<^bold>|,v!\<^bold>|u\\<^sub>pv\<^bold>|)" proof(induction u v rule: mismatch_pair.induct, simp+) qed text \For incomparable words the pair is out of diagonal.\ lemma incomp_neq: "\ u \ v \ (mismatch_pair u v) \ Id" unfolding mismatch_pair_lcp by (simp add: lcp_mismatch') -lemma mismatch_ext_left: "\ u \ v \ mismatch_pair u v = mismatch_pair (p\u) (p\v)" +lemma mismatch_ext_left: "\ u \ v \ mismatch_pair u v = mismatch_pair (p\u) (p\v)" unfolding mismatch_pair_lcp by (simp add: lcp_ext_left) -lemma mismatch_ext_right: assumes "\ u \ v" +lemma mismatch_ext_right: assumes "\ u \ v" shows "mismatch_pair u v = mismatch_pair (u\z) (v\w)" -proof- +proof- have less1: "\<^bold>|u \\<^sub>p v\<^bold>| < \<^bold>|u\<^bold>|" and less2: "\<^bold>|v \\<^sub>p u\<^bold>| < \<^bold>|v\<^bold>|" - using lcp_len'[of u v] lcp_len'[of v u] assms by auto - show ?thesis + using lcp_len'[of u v] lcp_len'[of v u] assms by auto + show ?thesis unfolding mismatch_pair_lcp unfolding pref_index[OF triv_pref less1, of z] pref_index[OF triv_pref less2, of w, unfolded lcp_sym[of v]] using assms lcp_ext_right[of u v _ z w] by metis -qed +qed lemma mismatchI: "\ u \ v \ i < min \<^bold>|u\<^bold>| \<^bold>|v\<^bold>| \ take i u = take i v \ u!i \ v!i \ mismatch_pair u v = (u!i,v!i)" - unfolding mismatch_pair_lcp using lcp_lenI by blast + unfolding mismatch_pair_lcp using lcp_lenI by blast text \For incomparable words, the mismatch letters work in a similar way as the lexicographic order\ -lemma mismatch_lexord: assumes "\ u \ v" and "mismatch_pair u v \ r" +lemma mismatch_lexord: assumes "\ u \ v" and "mismatch_pair u v \ r" shows "(u,v) \ lexord r" unfolding lexord_take_index_conv mismatch_pair_lcp using \mismatch_pair u v \ r\[unfolded mismatch_pair_lcp] incomp_lcp_len[OF assms(1)] lcp_take_eq by blast -text \However, the equivalence requires r to be irreflexive. +text \However, the equivalence requires r to be irreflexive. (Due to the definition of lexord which is designed for irreflexive relations.)\ -lemma lexord_mismatch: assumes "\ u \ v" and "irrefl r" +lemma lexord_mismatch: assumes "\ u \ v" and "irrefl r" shows "mismatch_pair u v \ r \ (u,v) \ lexord r" proof assume "(u,v) \ lexord r" obtain i where "i < min \<^bold>|u\<^bold>| \<^bold>|v\<^bold>|" and "take i u = take i v" and "(u ! i, v ! i) \ r" using \(u,v) \ lexord r\[unfolded lexord_take_index_conv] \\ u \ v\ pref_take_conv by blast have "u!i \ v!i" using \irrefl r\[unfolded irrefl_def] \(u ! i, v ! i) \ r\ by fastforce from \(u ! i, v ! i) \ r\[folded mismatchI[OF \\ u \ v\ \i < min \<^bold>|u\<^bold>| \<^bold>|v\<^bold>|\ \take i u = take i v\ \u!i \ v!i\]] show "mismatch_pair u v \ r". next from mismatch_lexord[OF \\ u \ v\] show "mismatch_pair u v \ r \ (u, v) \ lexord r". qed section "Factor properties" lemma rev_fac[reversal_rule]: "rev u \f rev v \ u \f v" using Sublist.sublist_rev. lemma fac_pref: "u \f v \ \ p. p \ u \p v" by simp lemma fac_pref_suf: "u \f v \ \ p. p \p v \ u \s p" using sublist_altdef by blast lemma pref_suf_fac: "r \p v \ u \s r \ u \f v" using sublist_altdef by blast -lemmas +lemmas fac_suf = fac_pref[reversed] and fac_suf_pref = fac_pref_suf[reversed] and suf_pref_fac = pref_suf_fac[reversed] lemma suf_pref_eq: "s \s p \ p \p s \ p = s" - using sublist_order.eq_iff by blast + using long_pref suffix_length_le by blast lemma fac_triv': assumes "p\x\q = x" shows "q = \" - using prefI[of "p\x" q "p\x\q"] sufI[of \ "p\x\q" "x", THEN suf_ext[of "p\x\q" x p]] - suf_pref_eq[of x "p\x"] self_append_conv[of "p\x" q] - unfolding assms append_Nil rassoc + using prefI[of "p\x" q "p\x\q"] sufI[of \ "p\x\q" "x", THEN suf_ext[of "p\x\q" x p]] + suf_pref_eq[of x "p\x"] self_append_conv[of "p\x" q] + unfolding assms append_Nil rassoc by blast lemma fac_triv: "p\x\q = x \ p = \" - using fac_triv' by force - -lemmas + using fac_triv' by force + +lemmas suf_fac = suffix_imp_sublist and pref_fac = prefix_imp_sublist lemma fac_Cons_E: assumes "u \f (a#v)" obtains "u \p (a#v)" | "u \f v" using assms[unfolded sublist_Cons_right] by fast lemmas fac_snoc_E = fac_Cons_E[reversed] section "Power and its properties" text\Word powers are often investigated in Combinatorics on Words. We thus interpret words as @{term monoid_mult} and adopt a notation for the word power. \ -declare power.power.simps [code] +declare power.power.simps [code] interpretation monoid_mult "\" "append" by standard simp+ notation power (infixr "\<^sup>@" 80) \ \inherited power properties\ lemma pow_zero [simp]: "u\<^sup>@0 = \" using power.power_0. lemma emp_pow: "\\<^sup>@n = \" using power_one. lemma pow_Suc_list: "u\<^sup>@(Suc n) = u \ u\<^sup>@n" using power.power_Suc. -lemma pow_commutes_list: "u\<^sup>@n \ u = u \ u\<^sup>@n" +lemma pow_commutes_list: "u\<^sup>@n \ u = u \ u\<^sup>@n" using power_commutes. lemma pow_add_list: "x\<^sup>@(a+b) = x\<^sup>@a\x\<^sup>@b" using power_add. lemma pow_Suc2_list: "u\<^sup>@Suc n = u\<^sup>@n \ u" using power_Suc2. lemma pow_eq_if_list: "p\<^sup>@m = (if m = 0 then \ else p \ p\<^sup>@(m-1))" using power_eq_if. -lemma pow_one_id: "u\<^sup>@1 = u" +lemma pow_one_id: "u\<^sup>@1 = u" using power_one_right. lemma pow2_list: "u\<^sup>@2 = u \ u" using power2_eq_square. lemma comm_add_exp: "u \ v = v \ u \ u \<^sup>@ n \ v = v \ u \<^sup>@ n" using power_commuting_commutes. lemma pow_mult_list: "u\<^sup>@(m*n) = (u\<^sup>@m)\<^sup>@n" using power_mult. lemma pow_rev_emp_conv[reversal_rule]: "power.power (rev \) (\) = (\<^sup>@)" by simp \ \more power properties\ lemma zero_exp: "n = 0 \ r\<^sup>@n = \" by simp lemma nemp_pow[elim]: "t\<^sup>@m \ \ \ m \ 0" using zero_exp by blast lemma nemp_pow'[elim]: "t\<^sup>@m \ \ \ t \ \" using emp_pow by auto lemma sing_pow:"i < m \ ([a]\<^sup>@m) ! i = a" by (induct i m rule: diff_induct) auto lemma pow_is_concat_replicate: "u\<^sup>@n = concat (replicate n u)" by (induct n) auto lemma pow_slide: "u \ (v \ u)\<^sup>@n \ v = (u \ v)\<^sup>@(Suc n)" by (induct n) simp+ lemma pop_pow_one: "m \ 0 \ r\<^sup>@m = r \ r\<^sup>@(m-1)" by (simp add: pow_eq_if_list) lemma hd_pow: assumes "n \ 0" shows "hd(u\<^sup>@n) = hd u" - unfolding pop_pow_one[OF \n \ 0\] using hd_append2 by (cases "u = \", simp) + unfolding pop_pow_one[OF \n \ 0\] using hd_append2 by (cases "u = \", simp) lemma pop_pow: "m \ k \u\<^sup>@m \ u\<^sup>@(k-m) = u\<^sup>@k" - using le_add_diff_inverse pow_add_list by metis + using le_add_diff_inverse pow_add_list by metis lemma pop_pow_cancel: "u\<^sup>@k \ v = u\<^sup>@m \ w \ m \ k \ u\<^sup>@(k-m) \ v = w" - using lassoc pop_pow[of m k u] same_append_eq[of "u\<^sup>@m" "u\<^sup>@(k-m)\v" w, unfolded lassoc] by argo + using lassoc pop_pow[of m k u] same_append_eq[of "u\<^sup>@m" "u\<^sup>@(k-m)\v" w, unfolded lassoc] by argo lemma pow_comm: "t\<^sup>@k \ t\<^sup>@m = t\<^sup>@m \ t\<^sup>@k" unfolding pow_add_list[symmetric] add.commute[of k].. lemma comm_add_exps: assumes "r \ u = u \ r" shows "r\<^sup>@m \ u\<^sup>@k = u\<^sup>@k \ r\<^sup>@m" using comm_add_exp[OF comm_add_exp[OF assms, symmetric], symmetric]. lemma rev_pow: "rev (x\<^sup>@m) = (rev x)\<^sup>@m" by (induct m, simp, simp add: pow_commutes_list) lemmas [reversal_rule] = rev_pow[symmetric] lemmas pow_eq_if_list' = pow_eq_if_list[reversed] and pop_pow_one' = pop_pow_one[reversed] and pop_pow' = pop_pow[reversed] and pop_pow_cancel' = pop_pow_cancel[reversed] -lemma pow_len: "\<^bold>|u\<^sup>@k\<^bold>| = k * \<^bold>|u\<^bold>|" +lemma pow_len: "\<^bold>|u\<^sup>@k\<^bold>| = k * \<^bold>|u\<^bold>|" by (induct k) simp+ lemma eq_pow_exp: assumes "u \ \" shows "u\<^sup>@k = u\<^sup>@m \ k = m" proof assume "k = m" thus "u\<^sup>@k = u\<^sup>@m" by simp next - assume "u\<^sup>@k = u\<^sup>@m" + assume "u\<^sup>@k = u\<^sup>@m" from lenarg[OF this, unfolded pow_len mult_cancel2] show "k = m" using \u \ \\[folded length_0_conv] by blast qed lemma nemp_emp_power: assumes "u \ \" shows "u\<^sup>@m = \ \ m = 0" using eq_pow_exp[OF assms] by fastforce lemma nonzero_pow_emp: assumes "m \ 0" shows "u\<^sup>@m = \ \ u = \" by (meson assms nemp_emp_power nemp_pow') lemma pow_eq_eq: assumes "u\<^sup>@k = v\<^sup>@k" and "k \ 0" shows "u = v" proof- have "\<^bold>|u\<^bold>| = \<^bold>|v\<^bold>|" using lenarg[OF \u\<^sup>@k = v\<^sup>@k\, unfolded pow_len] \k \ 0\ by simp - from eqd_equal[of u "u\<^sup>@(k-1)" v "v\<^sup>@(k-1)", OF _ this] + from eqd_equal[of u "u\<^sup>@(k-1)" v "v\<^sup>@(k-1)", OF _ this] show ?thesis using \u\<^sup>@k = v\<^sup>@k\ unfolding pop_pow_one[OF \k \ 0\] by blast qed lemma sing_pow_empty: "[a]\<^sup>@n = \ \ n = 0" by (simp add: nemp_emp_power) lemma sing_pow_lists: "a \ A \ [a]\<^sup>@n \ lists A" by (induct n, auto) lemma long_power: "r \ \ \ \<^bold>|x\<^bold>| \ \<^bold>|r\<^sup>@\<^bold>|x\<^bold>|\<^bold>|" unfolding pow_len[of r "\<^bold>|x\<^bold>|"] using nemp_pos_len by auto lemma long_power': "r \ \ \ \<^bold>|x\<^bold>| < \<^bold>|r\<^sup>@(Suc \<^bold>|x\<^bold>|)\<^bold>|" - unfolding pow_Suc_list length_append - by (simp add: long_power add_strict_increasing) + unfolding pow_Suc_list length_append + by (simp add: long_power add_strict_increasing) lemma long_pow_exp: "r \ \ \ m \ \<^bold>|r\<^sup>@m\<^bold>|" - unfolding pow_len[of r m] using nemp_pos_len[of r] by simp + unfolding pow_len[of r m] using nemp_pos_len[of r] by simp lemma long_pow_ex: assumes "r \ \" obtains n where "m \ \<^bold>|r\<^sup>@n\<^bold>|" and "n \ 0" proof- obtain x :: "'a list" where "\<^bold>|x\<^bold>| = m" - using Ex_list_of_length by auto + using Ex_list_of_length by auto show thesis using that[of m, OF long_power[OF \r \ \\, of x, unfolded \\<^bold>|x\<^bold>| = m\]] that[of "Suc 1"] by auto qed lemma pref_pow_ext: "x \p r\<^sup>@k \ x \p r\<^sup>@Suc k" using pref_trans[OF _ prefI[OF pow_Suc2_list[symmetric]]]. lemma pref_pow_ext': "u \p r\<^sup>@k \ u \p r \ r\<^sup>@k" using pref_pow_ext[unfolded pow_Suc_list]. lemma pref_pow_root_ext: "x \p r\<^sup>@k \ r \ x \p r\<^sup>@Suc k" by simp lemma pref_prod_root: "u \p r\<^sup>@k \ u \p r \ u" using pref_pow_ext'[THEN pref_prod_pref]. lemma pref_exps_pow: "k \ l \ r\<^sup>@k \p r\<^sup>@l" using leI pop_pow[of k l r] by blast lemma pref_exp_le: assumes "u \ \" "u\<^sup>@m \p u\<^sup>@n" shows "m \ n" - using mult_cancel_le[OF nemp_len[OF \u \ \\], of m n] + using mult_cancel_le[OF nemp_len[OF \u \ \\], of m n] prefix_length_le[OF \u\<^sup>@m \p u\<^sup>@n\, unfolded pow_len[of u m] pow_len[of u n]] - by blast + by blast lemmas suf_pow_ext = pref_pow_ext[reversed] and suf_pow_ext'= pref_pow_ext'[reversed] and suf_pow_root_ext = pref_pow_root_ext[reversed] and suf_prod_root = pref_prod_root[reversed] and suf_exps_pow = pref_exps_pow[reversed] and suf_exp_le = pref_exp_le[reversed] lemma comm_common_power: assumes "r \ u = u \ r" shows "r\<^sup>@\<^bold>|u\<^bold>| = u\<^sup>@\<^bold>|r\<^bold>|" using eqd_equal[OF comm_add_exps[OF \r \ u = u \ r\], of "\<^bold>|u\<^bold>|" "\<^bold>|r\<^bold>|"] - unfolding pow_len by fastforce + unfolding pow_len by fastforce lemma one_generated_list_power: "u \ lists {x} \ \k. concat u = x\<^sup>@k" proof(induction u) case Nil then show ?case by (simp add: pow_is_concat_replicate) next case (Cons a u) then show ?case - unfolding Cons_in_lists_iff concat.simps(2) + unfolding Cons_in_lists_iff concat.simps(2) using singletonD[of a x] pow_Suc_list[of a] by metis qed lemma pow_lists: "0 < k \ u\<^sup>@k \ lists B \ u \ lists B" by (simp add: pow_eq_if_list) lemma concat_morph_power: "xs \ lists B \ xs = ts\<^sup>@k \ concat ts\<^sup>@k = concat xs" by (induct k arbitrary: xs ts) simp+ lemma pref_not_idem: "z \ \ \ z \ x \ z \ x\<^sup>@k \ x" using fac_triv by (cases k, simp, auto) lemma per_exp_pref: "u \p r \ u \ u \p r\<^sup>@k \ u" proof(induct k, simp) case (Suc k) show ?case unfolding pow_Suc_list rassoc using Suc.hyps Suc.prems pref_prolong by blast qed lemmas suf_not_idem = pref_not_idem[reversed] and per_exp_suf = per_exp_pref[reversed] lemma hd_sing_power: "k \ 0 \ hd ([a]\<^sup>@k) = a" by (induction k) simp+ lemma root_pref_cancel: assumes "t\<^sup>@m \ y = t\<^sup>@k" shows "y = t\<^sup>@(k - m)" using lqI[of "t\<^sup>@m" "t\<^sup>@(k-m)" "t\<^sup>@k"] unfolding lqI[OF \t\<^sup>@m \ y = t\<^sup>@k\] - using nat_le_linear[of m k] pop_pow[of m k t] diff_is_0_eq[of k m] append.right_neutral[of "t\<^sup>@k"] pow_zero[of t] - pref_antisym[of "t\<^sup>@m" "t\<^sup>@k", OF prefI[OF \t\<^sup>@m \ y = t\<^sup>@k\] pref_exps_pow[of k m t]] + using nat_le_linear[of m k] pop_pow[of m k t] diff_is_0_eq[of k m] append.right_neutral[of "t\<^sup>@k"] pow_zero[of t] + pref_antisym[of "t\<^sup>@m" "t\<^sup>@k", OF prefI[OF \t\<^sup>@m \ y = t\<^sup>@k\] pref_exps_pow[of k m t]] by presburger lemmas root_suf_cancel = root_pref_cancel[reversed] lemma index_pow_mod: "i < \<^bold>|r\<^sup>@k\<^bold>| \ (r\<^sup>@k)!i = r!(i mod \<^bold>|r\<^bold>|)" proof(induction k) have aux: "\<^bold>|r\<^sup>@(Suc l)\<^bold>| = \<^bold>|r\<^sup>@l\<^bold>| + \<^bold>|r\<^bold>|" for l by simp have aux1: "\<^bold>|(r\<^sup>@l)\<^bold>| \ i \ i < \<^bold>|r\<^sup>@l\<^bold>| + \<^bold>|r\<^bold>| \ i mod \<^bold>|r\<^bold>| = i - \<^bold>|r\<^sup>@l\<^bold>|" for l unfolding pow_len[of r l] using less_diff_conv2[of "l * \<^bold>|r\<^bold>|" i "\<^bold>|r\<^bold>|", unfolded add.commute[of "\<^bold>|r\<^bold>|" "l * \<^bold>|r\<^bold>|"]] - get_mod[of "i - l * \<^bold>|r\<^bold>|" "\<^bold>|r\<^bold>|" l] le_add_diff_inverse[of "l*\<^bold>|r\<^bold>|" i] by argo + get_mod[of "i - l * \<^bold>|r\<^bold>|" "\<^bold>|r\<^bold>|" l] le_add_diff_inverse[of "l*\<^bold>|r\<^bold>|" i] by argo case (Suc k) show ?case - unfolding aux sym[OF pow_Suc2_list[symmetric]] nth_append le_mod_geq + unfolding aux sym[OF pow_Suc2_list[symmetric]] nth_append le_mod_geq using aux1[ OF _ Suc.prems[unfolded aux]] Suc.IH pow_Suc2_list[symmetric] Suc.prems[unfolded aux] leI[of i "\<^bold>|r \<^sup>@ k\<^bold>|"] by presburger qed auto lemma sing_pow_len: "\<^bold>|[r]\<^sup>@l\<^bold>| = l" by (induct l) auto lemma concat_take_sing: "k \ l \ concat (take k ([r]\<^sup>@l)) = r\<^sup>@k" proof(induct k, simp) case (Suc k) - then show ?case + then show ?case using concat_morph[of "take k ((r # \) \<^sup>@ l)""(r # \)", unfolded - sym[OF take_Suc_conv_app_nth[of k "[r]\<^sup>@l", unfolded sing_pow_len[of r l] less_eq_Suc_le - sing_pow[OF iffD2[OF less_eq_Suc_le Suc.prems], of r], OF \Suc k \ l\]] - concat_sing'[of r] + sym[OF take_Suc_conv_app_nth[of k "[r]\<^sup>@l", unfolded sing_pow_len[of r l] less_eq_Suc_le + sing_pow[OF iffD2[OF less_eq_Suc_le Suc.prems], of r], OF \Suc k \ l\]] + concat_sing'[of r] Suc.hyps[OF Suc_leD[OF Suc.prems]] pow_Suc2_list[symmetric]] by argo qed lemma concat_sing_pow: "concat ([a]\<^sup>@k) = a\<^sup>@k" proof(induct k) show "concat ((a # \) \<^sup>@ 0) = a \<^sup>@ 0" by simp next fix k assume "concat ((a # \) \<^sup>@ k) = a \<^sup>@ k" thus "concat ((a # \) \<^sup>@ Suc k) = a \<^sup>@ Suc k" by simp qed lemma unique_letter_word: "(\ c. c \ set w \ c = a) \ \ k. w = [a]\<^sup>@k" proof (induct w) case Nil then show ?case - by (metis pow_zero) + by (metis pow_zero) next - case (Cons b w) - then show ?case + case (Cons b w) + then show ?case proof- - obtain k where "w = [a]\<^sup>@k" using Cons.hyps Cons.prems by auto + obtain k where "w = [a]\<^sup>@k" using Cons.hyps Cons.prems by auto hence "b#w = [a]\<^sup>@Suc k" - by (simp add: \w = (a # \)\<^sup>@k\ Cons.prems) - thus ?thesis by blast + by (simp add: \w = (a # \)\<^sup>@k\ Cons.prems) + thus ?thesis by blast qed qed lemma unique_letter_wordE[elim]: assumes "(\ c. c \ set w \ c = a)" obtains k where "w = [a]\<^sup>@k" using unique_letter_word assms by metis - + lemma unique_letter_wordE'[elim]: assumes "set w \ {a}" obtains k where "w = [a] \<^sup>@ k" using assms unique_letter_word[of w a] by blast lemma conjug_pow: "x \ z = z \ y \ x\<^sup>@k \ z = z \ y\<^sup>@k" by (induct k) fastforce+ lemma shift_pow: "(u\v)\<^sup>@k\u = u\(v\u)\<^sup>@k" by (simp add: conjug_pow) lemma lq_conjug_pow: assumes "p \p x \ p" shows "p\\<^sup>>(x\<^sup>@k \ p) = (p\\<^sup>>(x \ p))\<^sup>@k" - using lqI[OF sym[OF conjug_pow[of x p "p\\<^sup>>(x \ p)", OF sym[OF lq_pref[OF \p \p x \ p\]], of k]]]. + using lqI[OF sym[OF conjug_pow[of x p "p\\<^sup>>(x \ p)", OF sym[OF lq_pref[OF \p \p x \ p\]], of k]]]. lemmas rq_conjug_pow = lq_conjug_pow[reversed] section "Total morphisms" locale morphism = fixes f :: "'a list \ 'b list" assumes morph: "f (u \ v) = f u \ f v" begin lemma emp_to_emp[simp]: "f \ = \" using morph[of \ \] self_append_conv2[of "f \" "f \"] by simp lemma pow_morph: "f (x\<^sup>@k) = (f x)\<^sup>@k" by (induction k) (simp add: morph)+ lemma pop_hd: "f (a#u) = f [a] \ f u" - unfolding hd_word[of a u] using morph. + unfolding hd_word[of a u] using morph. lemma pop_hd_nemp: "u \ \ \ f (u) = f [hd u] \ f (tl u)" using list.exhaust_sel pop_hd[of "hd u" "tl u"] by force lemma pop_last_nemp: "u \ \ \ f (u) = f (butlast u) \ f [last u]" unfolding morph[symmetric] append_butlast_last_id .. lemma pref_mono: "u \p v \ f u \p f v" using morph by auto lemma morph_concat_map: "f x = concat (map (\ x. f [x]) x)" proof (induction x, simp) case (Cons a x) - then show ?case - unfolding pop_hd[of a x] by auto -qed - -end + then show ?case + unfolding pop_hd[of a x] by auto +qed + +end locale two_morphisms = morphg: morphism g + morphh: morphism h for g h :: "'a list \ 'b list" begin lemma def_on_sings: assumes "\a. g [a] = h [a]" shows "g u = h u" proof (induct u, simp) next case (Cons a u) then show ?case unfolding morphg.pop_hd[of a u] morphh.pop_hd[of a u] using \\a. g [a] = h [a]\ by presburger qed end section \Reversed morphism and composition\ -definition rev_morph :: "('a list \ 'b list) \ ('a list \ 'b list)" where - "rev_morph f = rev \ f \ rev" +definition rev_morph :: "('a list \ 'b list) \ ('a list \ 'b list)" where + "rev_morph f = rev \ f \ rev" lemma rev_morph_idemp[simp]: "rev_morph (rev_morph f) = f" unfolding rev_morph_def by auto lemma morph_compose: "morphism f \ morphism g \ morphism (f \ g)" by (simp add: morphism_def) lemma rev_morph_arg: "rev_morph f u = rev (f (rev u))" by (simp add: rev_morph_def) lemmas rev_morph_arg_rev[reversal_rule] = rev_morph_arg[reversed add: rev_rev_ident] lemma rev_morph_sing: "rev_morph f [a] = rev (f [a])" unfolding rev_morph_def by simp context morphism begin lemma rev_morph_morph: "morphism (rev_morph f)" by (standard, auto simp add: rev_morph_def morph) lemma morph_rev_len: "\<^bold>|f (rev u)\<^bold>| = \<^bold>|f u\<^bold>|" proof (induction u, simp) case (Cons a u) - then show ?case + then show ?case unfolding rev.simps(2) pop_hd[of a u] morph length_append by force qed lemma rev_morph_len: "\<^bold>|rev_morph f u\<^bold>| = \<^bold>|f u\<^bold>|" unfolding rev_morph_def - by (simp add: morph_rev_len) + by (simp add: morph_rev_len) end section \Rotation\ lemma rotate_comp_eq:"w \ rotate n w \ rotate n w = w" using pref_same_len[OF _ length_rotate[of n w]] pref_same_len[OF _ length_rotate[of n w, symmetric], symmetric] by blast corollary mismatch_iff_lexord: assumes "rotate n w \ w" and "irrefl r" shows "mismatch_pair w (rotate n w) \ r \ (w,rotate n w) \ lexord r" proof- have "\ w \ rotate n w" - using rotate_comp_eq \rotate n w \ w\ + using rotate_comp_eq \rotate n w \ w\ unfolding prefix_comparable_def by blast from lexord_mismatch[OF this \irrefl r\] show ?thesis. qed lemma rotate_back: obtains m where "rotate m (rotate n u) = u" proof(cases "u = \", simp) assume "u \ \" show ?thesis - using that[of "\<^bold>|u\<^bold>| - n mod \<^bold>|u\<^bold>|"] + using that[of "\<^bold>|u\<^bold>| - n mod \<^bold>|u\<^bold>|"] unfolding rotate_rotate[of "\<^bold>|u\<^bold>| - n mod \<^bold>|u\<^bold>|" "n mod \<^bold>|u\<^bold>|" u] - le_add_diff_inverse2[OF + le_add_diff_inverse2[OF less_imp_le_nat[OF mod_less_divisor[OF nemp_len[OF \u \ \\, unfolded neq0_conv], of n]]] arg_cong[OF rotate_conv_mod[of n u], of "rotate (\<^bold>|u\<^bold>| - n mod \<^bold>|u\<^bold>|)"] by simp qed lemma rotate_class_rotate': "(\n. rotate n w = u) \ (\n. rotate n (rotate l w) = u)" -proof +proof obtain m where rot_m: "rotate m (rotate l w) = w" using rotate_back. assume "\n. rotate n w = u" - then obtain n where rot_n: "rotate n w = u" by blast + then obtain n where rot_n: "rotate n w = u" by blast show "\n. rotate n (rotate l w) = u" - using exI[of "\ x. rotate x (rotate l w) = u" "n+m", OF + using exI[of "\ x. rotate x (rotate l w) = u" "n+m", OF rotate_rotate[symmetric, of n m "rotate l w", unfolded rot_m rot_n]]. next show "\n. rotate n (rotate l w) = u \ \n. rotate n w = u" - using rotate_rotate[symmetric] by blast + using rotate_rotate[symmetric] by blast qed lemma rotate_class_rotate: "{u . \n. rotate n w = u} = {u . \n. rotate n (rotate l w) = u}" using rotate_class_rotate' by blast lemma rotate_pow_self: "rotate (l*\<^bold>|u\<^bold>|) (u\<^sup>@k) = u\<^sup>@k" proof(induct l, simp) case (Suc l) then show ?case proof(cases "k = 0", simp) - assume "k \ 0" + assume "k \ 0" show ?thesis unfolding rotate_rotate[of "\<^bold>|u\<^bold>|" "l * \<^bold>|u\<^bold>|" "u\<^sup>@k", unfolded mult_Suc[symmetric] Suc.hyps, symmetric] using rotate_append[of u "u\<^sup>@(k-1)", folded pop_pow_one[OF \k \ 0\, of u] pop_pow_one'[OF \k \ 0\, of u]]. qed qed lemma rotate_root_self: "rotate \<^bold>|r\<^bold>| (r\<^sup>@k) = r\<^sup>@k" using rotate_pow_self[of 1 r k] by auto lemma rotate_pow_mod: "rotate n (u\<^sup>@k) = rotate (n mod \<^bold>|u\<^bold>|) (u\<^sup>@k)" using rotate_rotate[of "n mod \<^bold>|u\<^bold>|" "n div \<^bold>|u\<^bold>| * \<^bold>|u\<^bold>|" "u\<^sup>@k", symmetric] unfolding rotate_pow_self[of "n div \<^bold>|u\<^bold>|" u k] div_mult_mod_eq[of n "\<^bold>|u\<^bold>|", unfolded add.commute[of "n div \<^bold>|u\<^bold>| * \<^bold>|u\<^bold>|" "n mod \<^bold>|u\<^bold>|"]]. lemma rotate_conj_pow: "rotate \<^bold>|u\<^bold>| ((u\v)\<^sup>@k) = (v\u)\<^sup>@k" by (induct k, simp, simp add: rotate_append shift_pow) lemma rotate_pow_comm: "rotate n (u\<^sup>@k) = (rotate n u)\<^sup>@k" proof (cases "u = \", simp) assume "u \ \" show ?thesis unfolding rotate_drop_take[of n u] rotate_pow_mod[of n u k] using rotate_conj_pow[of "take (n mod \<^bold>|u\<^bold>|) u" "drop (n mod \<^bold>|u\<^bold>|) u" k, unfolded append_take_drop_id[of "n mod \<^bold>|u\<^bold>|" u]] unfolding mod_le_divisor[of "\<^bold>|u\<^bold>|" n, THEN take_len, OF \u\\\[unfolded length_greater_0_conv[symmetric]]]. qed subsection \Lists of words and their concatenation\ text\The helpful lemmas of this section deal with concatenation of a list of words @{term concat}. The main objective is to cover elementary facts needed to study factorizations of words. \ lemma append_in_lists: "u \ lists A \ v \ lists A \ u \ v \ lists A" by simp lemma concat_take_is_prefix: "concat(take n ws) \p concat ws" - using concat_morph[of "take n ws" "drop n ws", unfolded append_take_drop_id[of n ws], THEN prefI]. + using concat_morph[of "take n ws" "drop n ws", unfolded append_take_drop_id[of n ws], THEN prefI]. lemma concat_take_suc: assumes "j < \<^bold>|ws\<^bold>|" shows "concat(take j ws) \ ws!j = concat(take (Suc j) ws)" - unfolding take_Suc_conv_app_nth[OF \j < \<^bold>|ws\<^bold>|\] + unfolding take_Suc_conv_app_nth[OF \j < \<^bold>|ws\<^bold>|\] using sym[OF concat_append[of "(take j ws)" "[ws ! j]", unfolded concat.simps(2)[of "ws!j" \, unfolded concat.simps(1) append_Nil2]]]. -lemma pref_mod_list': assumes "u \np concat ws" - obtains j r where "j < \<^bold>|ws\<^bold>|" and "r \np ws!j" and "concat (take j ws) \ r = u" +lemma pref_mod_list': assumes "u \np concat ws" + obtains j r where "j < \<^bold>|ws\<^bold>|" and "r \np ws!j" and "concat (take j ws) \ r = u" proof- have "\<^bold>|ws\<^bold>| \ 0" using assms by fastforce then obtain l where "Suc l = \<^bold>|ws\<^bold>|" using Suc_pred by blast let ?P = "\ j. u \p concat(take (Suc j) ws)" have "?P l" using assms \Suc l = \<^bold>|ws\<^bold>|\ by auto define j where "j = (LEAST j. ?P j)" \ \largest j such that concat (take j ws)

have "u \p concat(take (Suc j) ws)" and "j < \<^bold>|ws\<^bold>|" using Least_le[of ?P, OF \?P l\] LeastI[of ?P, OF \?P l\] \Suc l = \<^bold>|ws\<^bold>|\ unfolding sym[OF j_def] by auto have "concat(take j ws)

0" hence "j - 1 < j" and "Suc (j-1) = j" by auto hence "\ ?P (j-1)" - using j_def not_less_Least by blast + using j_def not_less_Least by blast from this[unfolded \Suc (j-1) = j\] show "concat(take j ws)

?r = u" and "?r \ \" - using \concat (take j ws)

by auto + using \concat (take j ws)

by auto have "?r \p ws!j" using pref_cancel[of "concat (take j ws)" "concat (take j ws)\\<^sup>>u" "ws ! j"] \u \p concat (take (Suc j) ws)\[unfolded sym[OF concat_take_suc[OF \j < \<^bold>|ws\<^bold>|\]]] - \concat (take j ws) \ concat (take j ws)\\<^sup>>u = u\ by argo + \concat (take j ws) \ concat (take j ws)\\<^sup>>u = u\ by argo show thesis - using that[OF \j < \<^bold>|ws\<^bold>|\ npI[OF \?r \ \\ \?r \p ws!j\] \concat(take j ws) \ ?r = u\]. + using that[OF \j < \<^bold>|ws\<^bold>|\ npI[OF \?r \ \\ \?r \p ws!j\] \concat(take j ws) \ ?r = u\]. qed -lemma pref_mod_list_suf: assumes "u \np concat ws" - obtains j s where "j < \<^bold>|ws\<^bold>|" and "s s" +lemma pref_mod_list_suf: assumes "u \np concat ws" + obtains j s where "j < \<^bold>|ws\<^bold>|" and "s s" proof- obtain j r where "j < \<^bold>|ws\<^bold>|" and "r \np ws!j" and "concat (take j ws) \ r = u" using pref_mod_list'[OF assms] by blast have "r\\<^sup>>(ws!j) r \np ws ! j\]] lq_pref[OF npD[OF \r \np ws ! j\]]. + using ssufI1[OF _ npD'[OF \r \np ws ! j\]] lq_pref[OF npD[OF \r \np ws ! j\]]. have "concat (take (Suc j) ws) = u \ r\\<^sup>>(ws!j)" using lq_pref[OF npD[OF \r \np ws ! j\], symmetric, unfolded same_append_eq[of "concat (take j ws)" "ws ! j" "r \ r\\<^sup>>ws ! j", symmetric] lassoc \concat (take j ws) \ r = u\ concat_take_suc[OF \j < \<^bold>|ws\<^bold>|\] \r \np ws!j\]. from that[OF \j < \<^bold>|ws\<^bold>|\ \r\\<^sup>>(ws!j) this] show thesis. qed lemma suf_mod_list': "s \ns concat ws \ ws \ \ \ \j r. j < \<^bold>|ws\<^bold>| \ r \ (concat (drop (Suc j) ws)) = s \ r \ns ws!j" proof- assume "s \ns concat ws" "ws \ \" have "rev s \np concat (rev (map (\u. rev u) ws))" using \s \ns concat ws\[unfolded nsuf_rev_pref_iff] - by (simp add: rev_concat rev_map) - have "rev ws \ \" + by (simp add: rev_concat rev_map) + have "rev ws \ \" by (simp add: \ws \ \\) then obtain j r where "j < \<^bold>|rev (map rev ws)\<^bold>|" "concat (take j (rev (map rev ws))) \ r = rev s" "r \np rev (map rev ws) ! j" using pref_mod_list'[of "rev s" "rev (map (\u. rev u) ws)" thesis] \rev s \np concat (rev (map rev ws))\ by blast have "rev ws ! j = rev (rev (map rev ws) ! j)" using \j < \<^bold>|rev (map rev ws)\<^bold>|\ length_map nth_map[of j "rev ws" rev, unfolded rev_map[of rev ws, symmetric]] rev_rev_ident by simp from \j < \<^bold>|rev (map rev ws)\<^bold>|\ rev_nth[of j ws, unfolded this] have "rev (rev (map rev ws) ! j) = ws!(\<^bold>|ws\<^bold>|- Suc j)" by fastforce from \r \np rev (map rev ws) ! j\[unfolded npref_rev_suf_iff, unfolded this] have p2: "rev r \ns ws!(\<^bold>|ws\<^bold>|-Suc j)". have "concat (take j (rev (map rev ws))) = rev (concat (drop (\<^bold>|ws\<^bold>|-j) ws))" by (simp add: rev_concat rev_map take_map take_rev) from arg_cong[OF this, of "\x. x\r", unfolded \concat (take j (rev (map rev ws))) \ r = rev s\] have p1: "s = rev r \ (concat (drop (\<^bold>|ws\<^bold>|-j) ws))" using rev_append[of "rev (concat (drop (\<^bold>|ws\<^bold>| - j) ws))" r] rev_rev_ident[of s] rev_rev_ident[of "(concat (drop (\<^bold>|ws\<^bold>| - j) ws))"] by argo have p0: "\<^bold>|ws\<^bold>| - Suc j < \<^bold>|ws\<^bold>|" by (simp add: \ws \ \\) have "\<^bold>|ws\<^bold>| - j = Suc (\<^bold>|ws\<^bold>| - Suc j)" using Suc_diff_Suc[OF \j < \<^bold>|rev (map rev ws)\<^bold>|\] by auto from p0 p1[unfolded this] p2 show ?thesis by blast qed -lemma pref_mod_list: assumes "u

|ws\<^bold>|" and "r

r = u" +lemma pref_mod_list: assumes "u

|ws\<^bold>|" and "r

r = u" proof- have "\<^bold>|ws\<^bold>| \ 0" using assms by auto then obtain l where "Suc l = \<^bold>|ws\<^bold>|" using Suc_pred by blast let ?P = "\ j. u

Suc l = \<^bold>|ws\<^bold>|\ by auto define j where "j = (LEAST j. ?P j)" \ \smallest j such that concat (take (Suc j) ws)

have "u

?P l\] unfolding sym[OF j_def]. + using LeastI[of ?P, OF \?P l\] unfolding sym[OF j_def]. have "j < \<^bold>|ws\<^bold>|" using Least_le[of ?P, OF \?P l\] \Suc l = \<^bold>|ws\<^bold>|\ unfolding sym[OF j_def] by auto have "0 < j \ concat(take j ws) \p u" - using Least_le[of ?P "(j- Suc 0)", unfolded sym[OF j_def]] + using Least_le[of ?P "(j- Suc 0)", unfolded sym[OF j_def]] ruler[OF concat_take_is_prefix sprefD1[OF assms], of j] - by force + by force hence "concat(take j ws) \p u" by fastforce let ?r = "concat(take j ws)\\<^sup>>u" have "concat(take j ws) \ ?r = u" - using \concat (take j ws) \p u\ by auto + using \concat (take j ws) \p u\ by auto have "?r

concat (take j ws) \p u\ \u

, unfolded + using spref_lq[OF \concat (take j ws) \p u\ \u

, unfolded lq_triv[of "concat (take j ws)" "ws!j", unfolded concat_take_suc[OF \j < \<^bold>|ws\<^bold>|\]]]. show thesis - using that[OF \j < \<^bold>|ws\<^bold>|\ \?r

\concat(take j ws) \ ?r = u\]. + using that[OF \j < \<^bold>|ws\<^bold>|\ \?r

\concat(take j ws) \ ?r = u\]. qed lemma pref_mod_power: assumes "u

@l" obtains k z where "k < l" and "z

@k\z = u" - using pref_mod_list[of u "[w]\<^sup>@l", unfolded sing_pow_len concat_sing_pow, OF \u

@l\, of thesis] + using pref_mod_list[of u "[w]\<^sup>@l", unfolded sing_pow_len concat_sing_pow, OF \u

@l\, of thesis] sing_pow[of _ l w] concat_take_sing less_imp_le_nat by metis lemma get_pow_exp: assumes "z

|t\<^sup>@m\z\<^bold>| div \<^bold>|t\<^bold>|" unfolding length_append[of "t\<^sup>@m" z, unfolded pow_len] using get_div[OF prefix_length_less[OF assms]]. lemma get_pow_remainder: assumes "z

|t\<^sup>@m\z\<^bold>| div \<^bold>|t\<^bold>|)*\<^bold>|t\<^bold>|) (t\<^sup>@m\z)" using drop_pref[of "t\<^sup>@m" z] pow_len[of t m] get_pow_exp[OF assms, of m] by simp lemma pref_mod_power': assumes "u \np w\<^sup>@l" obtains k z where "k < l" and "z \np w" and "w\<^sup>@k\z = u" - using pref_mod_list'[of u "[w]\<^sup>@l", unfolded sing_pow_len concat_sing_pow, OF \u \np w\<^sup>@l\] + using pref_mod_list'[of u "[w]\<^sup>@l", unfolded sing_pow_len concat_sing_pow, OF \u \np w\<^sup>@l\] sing_pow[of _ l w] concat_take_sing[of _ l w] less_imp_le_nat[of _ l] by metis -lemma pref_power: assumes "t \ \" and "u \p t\<^sup>@k" +lemma pref_power: assumes "t \ \" and "u \p t\<^sup>@k" shows "\ m. t\<^sup>@m \p u \ u

@m \ t" proof (cases "u = t\<^sup>@k") show "u = t \<^sup>@ k \ \m. t \<^sup>@ m \p u \ u

@ m \ t" using \t \ \\ by blast next - assume "u \ t \<^sup>@ k" + assume "u \ t \<^sup>@ k" obtain m z where "m < k" "z

@ m \ z = u" using pref_mod_power[of u t k] \u \p t\<^sup>@k\[unfolded prefix_order.dual_order.order_iff_strict] \u \ t\<^sup>@k\ by blast - hence "t \<^sup>@ m \p u" and "u

@ m \ t" by auto - thus ?thesis by blast -qed - -lemma pref_powerE: assumes "t \ \" and "u \p t\<^sup>@k" + hence "t \<^sup>@ m \p u" and "u

@ m \ t" by auto + thus ?thesis by blast +qed + +lemma pref_powerE: assumes "t \ \" and "u \p t\<^sup>@k" obtains m where "t\<^sup>@m \p u" "u

@m \ t" - using assms pref_power by blast - -lemma pref_power': assumes "u \ \" and "u \p t\<^sup>@k" + using assms pref_power by blast + +lemma pref_power': assumes "u \ \" and "u \p t\<^sup>@k" shows "\ m. t\<^sup>@m

u \p t\<^sup>@m \ t" proof- - obtain m z where "m < k" "z \np t" "t \<^sup>@ m \ z = u" + obtain m z where "m < k" "z \np t" "t \<^sup>@ m \ z = u" using pref_mod_power'[OF npI[OF \u \ \\ \u \p t\<^sup>@k\]]. - thus ?thesis - by auto -qed + thus ?thesis + by auto +qed lemma suf_power: assumes "t \ \" and "u \s t\<^sup>@k" shows "\m. t\<^sup>@m \s u \ u t\<^sup>@m" proof- - have "rev u \p (rev t)\<^sup>@k" - using \u \s t\<^sup>@k\[unfolded suffix_to_prefix rev_pow]. + have "rev u \p (rev t)\<^sup>@k" + using \u \s t\<^sup>@k\[unfolded suffix_to_prefix rev_pow]. from pref_power[OF _ this] obtain m where "(rev t)\<^sup>@m \p rev u" and "rev u

@m \ rev t" - using \t \ \\ by blast + using \t \ \\ by blast have "t\<^sup>@m \s u" - using \(rev t)\<^sup>@m \p rev u\[folded suffix_to_prefix rev_pow]. - moreover have "u t\<^sup>@m" + using \(rev t)\<^sup>@m \p rev u\[folded suffix_to_prefix rev_pow]. + moreover have "u t\<^sup>@m" using sprefD1[OF \rev u

@m \ rev t\, folded rev_pow rev_append suffix_to_prefix] sprefD2[OF \rev u

@m \ rev t\, folded rev_pow rev_append] by blast ultimately show ?thesis by blast qed -lemma suf_powerE: assumes "t \ \" and "u \s t\<^sup>@k" +lemma suf_powerE: assumes "t \ \" and "u \s t\<^sup>@k" obtains m where "t\<^sup>@m \s u" "u t\<^sup>@m" - using assms suf_power by blast + using assms suf_power by blast lemma del_emp_concat: "concat us = concat (filter (\x. x \ \) us)" by (induct us) simp+ lemma lists_drop_emp: "us \ lists C\<^sub>+ \ us \ lists C" by blast lemma lists_drop_emp': "us \ lists C \ (filter (\x. x \ \) us) \ lists C\<^sub>+" by (simp add: in_lists_conv_set) -lemma pref_concat_pref: "us \p ws \ concat us \p concat ws" +lemma pref_concat_pref: "us \p ws \ concat us \p concat ws" by auto lemma le_take_is_prefix: assumes "k \ n" shows "take k ws \p take n ws" - using take_add[of k "(n-k)" ws, unfolded le_add_diff_inverse[OF \k \ n\]] + using take_add[of k "(n-k)" ws, unfolded le_add_diff_inverse[OF \k \ n\]] by force lemma take_pp_less: assumes "take k ws

Root\ definition root :: "'a list \ 'a list \ bool" ("_ \ _*" [51,51] 60 ) where "u \ r* = (\ k. r\<^sup>@k = u)" notation (latex output) root ("_ \ _\<^sup>*") text\Empty word has all roots, including the empty root.\ lemma "\ \ r*" unfolding root_def using power_0 by blast lemma rootI: "r\<^sup>@k \ r*" using root_def by auto lemma self_root: "u \ u*" using rootI[of u "Suc 0"] by simp lemma rootE: assumes "u \ r*" obtains k where "r\<^sup>@k = u" using assms root_def by auto lemma empty_all_roots[simp]: "\ \ r*" using rootI[of r 0, unfolded pow_zero]. lemma take_root: "k \ 0 \ r = take (length r) (r\<^sup>@k)" by (simp add: pow_eq_if_list) lemma root_nemp: "u \ \ \ u \ r* \ r \ \" - unfolding root_def using emp_pow by auto + unfolding root_def using emp_pow by auto lemma root_shorter: "u \ \ \ u \ r* \ u \ r \ \<^bold>|r\<^bold>| < \<^bold>|u\<^bold>|" by (metis root_def leI take_all take_root pow_zero) lemma root_shorter_eq: "u \ \ \ u \ r* \ \<^bold>|r\<^bold>| \ \<^bold>|u\<^bold>|" using root_shorter by fastforce -lemma root_trans[trans]: "\v \ u*; u \ t*\ \ v \ t*" +lemma root_trans[trans]: "\v \ u*; u \ t*\ \ v \ t*" by (metis root_def pow_mult_list) lemma root_pow_root[trans]: "v \ u* \ v\<^sup>@n \ u*" using rootI root_trans by blast lemma root_len: "u \ q* \ \k. \<^bold>|u\<^bold>| = k*\<^bold>|q\<^bold>|" unfolding root_def using pow_len by auto lemma root_len_dvd: "u \ t* \ \<^bold>|t\<^bold>| dvd \<^bold>|u\<^bold>|" - using root_len root_def by fastforce + using root_len root_def by fastforce lemma common_root_len_gcd: "u \ t* \ v \ t* \ \<^bold>|t\<^bold>| dvd ( gcd \<^bold>|u\<^bold>| \<^bold>|v\<^bold>| )" by (simp add: root_len_dvd) lemma add_root[simp]: "z \ w \ z* \ w \ z*" proof - assume "w \ z*" thus "z \ w \ z*" - unfolding root_def using pow_Suc_list by blast + assume "w \ z*" thus "z \ w \ z*" + unfolding root_def using pow_Suc_list by blast next assume "z \ w \ z*" thus "w \ z*" unfolding root_def using root_pref_cancel[of z 1 w, unfolded power_one_right] by metis -qed +qed lemma add_roots: "w \ z* \ w' \ z* \ w \ w' \ z*" - unfolding root_def using pow_add_list by blast + unfolding root_def using pow_add_list by blast lemma concat_sing_list_pow: "ws \ lists {u} \ \<^bold>|ws\<^bold>| = k \ concat ws = u\<^sup>@k" proof(induct k arbitrary: ws) case (Suc k) have "ws \ \" - using list.size(3) nat.distinct(2)[of k, folded \\<^bold>|ws\<^bold>| = Suc k\] by blast - from hd_Cons_tl[OF this] + using list.size(3) nat.distinct(2)[of k, folded \\<^bold>|ws\<^bold>| = Suc k\] by blast + from hd_Cons_tl[OF this] have "ws = hd ws # tl ws" and "\<^bold>|tl ws\<^bold>| = k" using \ \<^bold>|ws\<^bold>| = Suc k\ by simp+ then show ?case - unfolding pow_Suc_list hd_concat_tl[OF \ws \ \\, symmetric] + unfolding pow_Suc_list hd_concat_tl[OF \ws \ \\, symmetric] using Suc.hyps[OF tl_lists[OF \ ws \ lists {u}\] \\<^bold>|tl ws\<^bold>| = k\] Nitpick.size_list_simp(2) lists_hd[of "ws" "{u}"] \ws \ lists{u}\ by blast qed simp lemma concat_sing_list_pow': "ws \ lists{u} \ concat ws = u\<^sup>@\<^bold>|ws\<^bold>|" - by (simp add: concat_sing_list_pow) - -lemma root_pref_cancel': assumes "x\y \ t*" and "x \ t*" shows "y \ t*" + by (simp add: concat_sing_list_pow) + +lemma root_pref_cancel': assumes "x\y \ t*" and "x \ t*" shows "y \ t*" proof- obtain n m where "t\<^sup>@m = x \ y" and "t\<^sup>@n = x" - using \x\y \ t*\[unfolded root_def] \x \ t*\[unfolded root_def] by blast + using \x\y \ t*\[unfolded root_def] \x \ t*\[unfolded root_def] by blast from root_pref_cancel[of t n y m, unfolded this] show "y \ t*" using rootI by auto qed lemma root_rev_iff[reversal_rule]: "rev u \ rev t* \ u \ t*" unfolding root_def[reversed] using root_def.. section Commutation text\The solution of the easiest nontrivial word equation, @{term "x \ y = y \ x"}, is in fact already contained in List.thy as the fact @{thm comm_append_are_replicate[no_vars]}.\ theorem comm: "x \ y = y \ x \ (\ t m k. x = t\<^sup>@k \ y = t\<^sup>@m)" - using comm_append_are_replicate[of x y, folded pow_is_concat_replicate] pow_comm by auto - + using comm_append_are_replicate[of x y, folded pow_is_concat_replicate] pow_comm by auto + corollary comm_root: "x \ y = y \ x \ (\ t. x \ t* \ y \ t*)" unfolding root_def comm by fast lemma commD[elim]: "x \ t* \ y \ t* \ x \ y = y \ x" using comm_root by auto lemma commE[elim]: assumes "x \ y = y \ x" obtains t m k where "x = t\<^sup>@k" and "y = t\<^sup>@m" - using assms[unfolded comm] by blast + using assms[unfolded comm] by blast lemma comm_rootE: assumes "x \ y = y \ x" obtains t where "x \ t*" and "y \ t*" - using assms[unfolded comm_root] by blast + using assms[unfolded comm_root] by blast section \Periods\ text\Periodicity is probably the most studied property of words. It captures the fact that a word overlaps with itself. -Another possible point of view is that the periodic word is a prefix of an (infinite) power of some nonempty +Another possible point of view is that the periodic word is a prefix of an (infinite) power of some nonempty word, which can be called its period word. Both these points of view are expressed by the following definition. \ subsection "Periodic root" definition period_root :: "'a list \ 'a list \ bool" ("_ \p _\<^sup>\" [51,51] 60 ) where [simp]: "period_root u r = (u \p r \ u \ r \ \)" lemma per_rootI[simp,intro]: "u \p r \ u \ r \ \ \ u \p r\<^sup>\" - by simp + by simp lemma per_rootI': assumes "u \p r\<^sup>@k" and "r \ \" shows "u \p r\<^sup>\" using per_rootI[OF pref_prod_pref[OF pref_pow_ext'[OF \u \p r\<^sup>@k\] \u \p r\<^sup>@k\] \r\\\]. lemma per_rootD[elim]: "u \p r\<^sup>\ \ u \p r \ u" by simp lemma per_rootD': "u \p r\<^sup>\ \ r \ \" by simp text \Empty word is not a periodic root but it has all nonempty periodic roots.\ lemma emp_any_per: "r \ \ \ (\ \p r\<^sup>\ )" by simp lemma emp_not_per: "\ (x \p \\<^sup>\)" - by simp + by simp text \Any nonempty word is its own periodic root.\ lemma root_self: "w \ \ \ w \p w\<^sup>\" by simp text\"Short roots are prefixes"\ lemma root_is_take: "w \p r\<^sup>\ \ \<^bold>|r\<^bold>| \ \<^bold>|w\<^bold>| \ r \p w" unfolding period_root_def using pref_prod_long[of w r w] by fastforce text \Periodic words are prefixes of the power of the root, which motivates the notation\ lemma pref_pow_ext_take: assumes "u \p r\<^sup>@k" shows "u \p take \<^bold>|r\<^bold>| u \ r\<^sup>@k" proof (rule le_cases[of "\<^bold>|u\<^bold>|" "\<^bold>|r\<^bold>|"]) assume "\<^bold>|r\<^bold>| \ \<^bold>|u\<^bold>|" show "u \p take \<^bold>|r\<^bold>| u \ r \<^sup>@ k" unfolding pref_take[OF pref_prod_long[OF pref_pow_ext'[OF \u \p r\<^sup>@k\] \\<^bold>|r\<^bold>| \ \<^bold>|u\<^bold>|\]] using pref_pow_ext'[OF \u \p r\<^sup>@k\]. qed simp lemma pref_pow_take: assumes "u \p r\<^sup>@k" shows "u \p take \<^bold>|r\<^bold>| u \ u" using pref_prod_pref[of u "take \<^bold>|r\<^bold>| u" "r\<^sup>@k", OF pref_pow_ext_take \u \p r\<^sup>@k\, OF \u \p r\<^sup>@k\]. - + lemma per_exp: assumes "u \p r\<^sup>\" shows "u \p r\<^sup>@k \ u" using per_exp_pref[OF per_rootD[OF assms]]. -lemma per_root_spref_powE: assumes "u \p r\<^sup>\" +lemma per_root_spref_powE: assumes "u \p r\<^sup>\" obtains k where "u

@k" using pref_prod_short'[OF per_exp[OF assms, of "Suc \<^bold>|u\<^bold>|"] long_power'[of r u, OF per_rootD'[OF assms]]] by blast lemma period_rootE [elim]: assumes "u \p t\<^sup>\" obtains n r where "r

@n \ r = u" proof- obtain m where "u

@m" using \u \p t\<^sup>\\ using per_root_spref_powE by blast from pref_mod_power[OF this that, of "\ x y. y" "\ x y. x"] show ?thesis by blast qed lemma per_add_exp: assumes "u \p r\<^sup>\" and "m \ 0" shows "u \p (r\<^sup>@m)\<^sup>\" using per_exp_pref[OF per_rootD, OF \u \p r\<^sup>\\, of m] per_rootD'[OF \u \p r\<^sup>\\, folded nonzero_pow_emp[OF \m \ 0\, of r]] unfolding period_root_def.. -lemma per_pref_ex: assumes "u \p r\<^sup>\" +lemma per_pref_ex: assumes "u \p r\<^sup>\" obtains n where "u \p r\<^sup>@n" and "n \ 0" using pcomp_shorter[OF ruler_pref[OF per_exp[OF \u \p r\<^sup>\\]]] long_pow_ex[of r "\<^bold>|u\<^bold>|", OF per_rootD'[OF \u \p r\<^sup>\\], of thesis] by blast theorem per_pref: "x \p r\<^sup>\ \ (\ k. x \p r\<^sup>@k) \ r \ \" using per_pref_ex period_root_def pref_pow_ext' pref_prod_pref by metis -lemma pref_pow_conv: "(\ k. x \p r\<^sup>@k) \ (\ k z. r\<^sup>@k\z = x \ z \p r)" +lemma pref_pow_conv: "(\ k. x \p r\<^sup>@k) \ (\ k z. r\<^sup>@k\z = x \ z \p r)" proof assume "\k z. r \<^sup>@ k \ z = x \ z \p r" - then obtain k z where "r\<^sup>@k \ z = x" and "z \p r" by blast + then obtain k z where "r\<^sup>@k \ z = x" and "z \p r" by blast thus "\ k. x \p r\<^sup>@k" - using pref_cancel'[OF \z \p r\, of "r\<^sup>@k", unfolded \r\<^sup>@k \ z = x\, folded pow_Suc2_list] by fast -next - assume "\ k. x \p r\<^sup>@k" then obtain k where "x \p r\<^sup>@k" by blast - {assume "r = \" + using pref_cancel'[OF \z \p r\, of "r\<^sup>@k", unfolded \r\<^sup>@k \ z = x\, folded pow_Suc2_list] by fast +next + assume "\ k. x \p r\<^sup>@k" then obtain k where "x \p r\<^sup>@k" by blast + {assume "r = \" have "x = \" using pref_emp[OF \x \p r \<^sup>@ k\[unfolded \r = \\ emp_pow]]. hence "\ k z. r\<^sup>@k\z = x \ z \p r" - using \r = \\ emp_pow by auto} + using \r = \\ emp_pow by auto} moreover {assume "r \ \" have "x

@(Suc k)" - using pref_ext_nemp[OF \x \p r\<^sup>@k\ \r \ \\, folded pow_Suc2_list]. + using pref_ext_nemp[OF \x \p r\<^sup>@k\ \r \ \\, folded pow_Suc2_list]. then have "\ k z. r\<^sup>@k\z = x \ z \p r" - using pref_mod_power[OF pref_ext_nemp[OF \x \p r\<^sup>@k\ \r \ \\, folded pow_Suc2_list], of "\ k z. r\<^sup>@k\z = x \ z \p r"] + using pref_mod_power[OF pref_ext_nemp[OF \x \p r\<^sup>@k\ \r \ \\, folded pow_Suc2_list], of "\ k z. r\<^sup>@k\z = x \ z \p r"] by auto} ultimately show "\ k z. r\<^sup>@k\z = x \ z \p r" by blast qed lemma per_eq: "x \p r\<^sup>\ \ (\ k z. r\<^sup>@k\z = x \ z \p r) \ r \ \" - using per_pref[unfolded pref_pow_conv]. + using per_pref[unfolded pref_pow_conv]. text\The previous theorem allows to prove some basic relations between powers, periods and commutation\ -lemma per_drop_exp: "u \p (r\<^sup>@m)\<^sup>\ \ u \p r\<^sup>\" - using per_pref[of u "r\<^sup>@m"] pow_mult_list[of r m, symmetric] unfolding per_pref[of u r] +lemma per_drop_exp: "u \p (r\<^sup>@m)\<^sup>\ \ u \p r\<^sup>\" + using per_pref[of u "r\<^sup>@m"] pow_mult_list[of r m, symmetric] unfolding per_pref[of u r] using nemp_pow'[of r m] by auto lemma per_drop_exp': "i \ 0 \ p \p e\<^sup>@i \ p \ p \p e \ p" - by (metis period_root_def eq_pow_exp per_drop_exp pow_zero) + by (metis period_root_def eq_pow_exp per_drop_exp pow_zero) lemma per_drop_exp_rev: assumes "i \ 0" "p \s p \ e\<^sup>@i" shows "p \s p \ e" using per_drop_exp'[OF \i \ 0\ \p \s p \ e\<^sup>@i\[unfolded suffix_to_prefix rev_append rev_pow], folded rev_append suffix_to_prefix]. corollary comm_drop_exp: assumes "m \ 0" and "u \ r\<^sup>@m = r\<^sup>@m \ u" shows "r \ u = u \ r" proof- {assume "r \ \" "u \ \" hence "u\r \p u\r\<^sup>@m" using pop_pow_one \m \ 0\ by (simp add: pop_pow_one) - have "u\r \p r\<^sup>@m\(u\r)" + have "u\r \p r\<^sup>@m\(u\r)" using pref_ext[of "u \ r" "r\<^sup>@m \ u" r, unfolded append_assoc[of "r\<^sup>@m" u r], OF \u\r \p u\r\<^sup>@m\[unfolded \u \ r\<^sup>@m = r\<^sup>@m \ u\]]. - hence "u\r \p r\(u\r)" + hence "u\r \p r\(u\r)" using per_drop_exp[of "u\r" r m] \m \ 0\ per_drop_exp' by blast from comm_ruler[OF self_pref[of "r \ u"], of "r \ u \ r", OF this] have "r \ u = u \ r" by auto } thus "r \ u = u \ r" - by fastforce + by fastforce qed corollary pow_comm_comm: assumes "x\<^sup>@j = y\<^sup>@k" and "j \ 0" shows "x\y = y\x" proof(cases "k = 0") assume "k = 0" from \x\<^sup>@j = y\<^sup>@k\[unfolded zero_exp[OF this, of y], unfolded nonzero_pow_emp[OF \j \ 0\]] show "x \ y = y \ x" by simp next assume "k \ 0" show "x \ y = y \ x" using comm_drop_exp[of j y x, OF \j \ 0\, OF comm_drop_exp[of k "x\<^sup>@j" y, OF \k \ 0\ , unfolded \x\<^sup>@j = y\<^sup>@k\, OF refl, folded \x\<^sup>@j = y\<^sup>@k\]]. -qed - -corollary comm_pow_roots: assumes "m \ 0" "k \ 0" +qed + +corollary comm_pow_roots: assumes "m \ 0" "k \ 0" shows "u\<^sup>@m \ v\<^sup>@k = v\<^sup>@k \ u\<^sup>@m \ u \ v = v \ u" using comm_drop_exp[OF \k \ 0\, of "u\<^sup>@m" v, THEN comm_drop_exp[OF \m \ 0\, of v u]] - comm_add_exps[of u v m k] by blast + comm_add_exps[of u v m k] by blast lemma pow_comm_comm': assumes comm: "u\<^sup>@(Suc k) = v\<^sup>@(Suc l)" shows "u \ v = v \ u" using comm pow_comm_comm by blast lemma comm_trans: assumes uv: "u\v = v\u" and vw: "w\v = v\w" and nemp: "v \ \" shows "u \ w = w \ u" proof - consider (u_emp) "u = \" | (w_emp) "w = \" | (nemp') "u \ \" and "w \ \" by blast then show ?thesis proof (cases) case nemp' have eq: "u\<^sup>@(\<^bold>|v\<^bold>| * \<^bold>|w\<^bold>|) = w\<^sup>@(\<^bold>|v\<^bold>| * \<^bold>|u\<^bold>|)" unfolding power_mult comm_common_power[OF uv] comm_common_power[OF vw] unfolding pow_mult_list[symmetric] mult.commute[of "\<^bold>|u\<^bold>|"].. obtain k l where k: "\<^bold>|v\<^bold>| * \<^bold>|w\<^bold>| = Suc k" and l: "\<^bold>|v\<^bold>| * \<^bold>|u\<^bold>| = Suc l" using nemp nemp' unfolding length_0_conv[symmetric] using not0_implies_Suc[OF no_zero_divisors] by presburger show ?thesis using pow_comm_comm'[OF eq[unfolded k l]]. qed simp+ qed theorem per_all_exps: "\ m \ 0; k \ 0 \ \ (u \p (r\<^sup>@m)\<^sup>\) \ (u \p (r\<^sup>@k)\<^sup>\)" using per_drop_exp[of u r m, THEN per_add_exp[of u r k]] per_drop_exp[of u r k, THEN per_add_exp[of u r m]] by blast lemma drop_per_pref: assumes "w \p u\<^sup>\" shows "drop \<^bold>|u\<^bold>| w \p w" using pref_drop[OF per_rootD[OF \w \p u\<^sup>\\], of "\<^bold>|u\<^bold>|", unfolded drop_pref[of u w]]. lemma per_root_trans[trans]: "w \p u\<^sup>\ \ u \ t* \ w \p t\<^sup>\" using root_def[of u t] per_drop_exp[of w t] by blast -text\Note that +text\Note that @{term "w \p u\<^sup>\ \ u \p t\<^sup>\ \ w \p t\<^sup>\"} does not hold. \ lemma per_root_same_prefix:"w \p r\<^sup>\ \ w' \p r\<^sup>\ \ w \ w" by blast lemma take_after_drop: "\<^bold>|u\<^bold>| + q \ \<^bold>|w\<^bold>| \ w \p u\<^sup>\ \ take q (drop \<^bold>|u\<^bold>| w) = take q w" using pref_share_take[OF drop_per_pref[of w u] len_after_drop[of "\<^bold>|u\<^bold>|" q w]]. text\The following lemmas are a weak version of the Periodicity lemma\ lemma two_pers': assumes pu: "w \p u \ w" and pv: "w \p v \ w" and len: "\<^bold>|u\<^bold>| + \<^bold>|v\<^bold>| \ \<^bold>|w\<^bold>|" - shows "u \ v = v \ u" + shows "u \ v = v \ u" proof- have uv: "w \p (u \ v) \ w" using pref_prolong[OF pu pv] unfolding lassoc. have vu: "w \p (v \ u) \ w" using pref_prolong[OF pv pu] unfolding lassoc. - have "u \ v \p w" using len pref_prod_long[OF uv] by simp + have "u \ v \p w" using len pref_prod_long[OF uv] by simp moreover have "v \ u \p w" using len pref_prod_long[OF vu] by simp ultimately show "u \ v = v \ u" by (rule pref_comp_eq[unfolded prefix_comparable_def, OF ruler swap_len]) qed lemma two_pers: assumes "w \p u\<^sup>\" and "w \p v\<^sup>\" and "\<^bold>|u\<^bold>|+\<^bold>|v\<^bold>| \ \<^bold>|w\<^bold>|" shows "u\v = v\u" - using two_pers'[OF per_rootD[OF assms(1)] per_rootD[OF assms(2)] assms(3)]. + using two_pers'[OF per_rootD[OF assms(1)] per_rootD[OF assms(2)] assms(3)]. subsection "Period - numeric" -text\Definition of a period as the length of the periodic root is often offered as the basic one. From our point of view, +text\Definition of a period as the length of the periodic root is often offered as the basic one. From our point of view, it is secondary, and less convenient for reasoning.\ -definition periodN :: "'a list \ nat \ bool" +definition periodN :: "'a list \ nat \ bool" where [simp]: "periodN w n = w \p (take n w)\<^sup>\" lemma periodN_I: "w \ \ \ n \ 0 \ w \p (take n w) \ w \ periodN w n" unfolding periodN_def period_root_def by fastforce text\The numeric definition respects the following convention about empty words and empty periods.\ lemma emp_no_periodN: "\ periodN \ n" by simp lemma zero_not_per: "\ periodN w 0" by simp (* lemma periodN_I [intro]: assumes "u \p r\<^sup>@k" and "u \ \" shows "periodN u \<^bold>|r\<^bold>|" *) (* unfolding periodN_def period_root_def *) (* using pref_pow_take[OF \u \p r\<^sup>@k\] take_nemp_len[OF \u \ \\] \u \p r\<^sup>@k\ by force *) (* lemma periodNI' [intro]: "u \np r\<^sup>@k \ periodN u \<^bold>|r\<^bold>|" *) (* unfolding nonempty_prefix_def by blast *) lemma periodN_D1: "periodN w n \ w \ \" by simp lemma periodN_D2: "periodN w n \ n \ 0" by simp lemma periodN_D3: "periodN w n \ w \p take n w \ w" by simp text\A nonempty word has all "long" periods\ lemma all_long_pers: "\ w \ \; \<^bold>|w\<^bold>| \ n \ \ periodN w n" by simp lemmas per_nemp = periodN_D1 lemmas per_positive = periodN_D2 text\The standard numeric definition of a period uses indeces.\ lemma periodN_indeces: assumes "periodN w n" and "i + n < \<^bold>|w\<^bold>|" shows "w!i = w!(i+n)" proof- have "w ! i = (take n w \ w) ! (n + i)" using nth_append_length_plus[of "take n w" w i, symmetric] - unfolding take_len[OF less_imp_le[OF add_lessD2[OF \i + n < \<^bold>|w\<^bold>|\]]]. + unfolding take_len[OF less_imp_le[OF add_lessD2[OF \i + n < \<^bold>|w\<^bold>|\]]]. also have "... = w ! (i + n)" - using pref_index[OF periodN_D3[OF \periodN w n\] \i + n < \<^bold>|w\<^bold>|\, symmetric] unfolding add.commute[of n i]. + using pref_index[OF periodN_D3[OF \periodN w n\] \i + n < \<^bold>|w\<^bold>|\, symmetric] unfolding add.commute[of n i]. finally show ?thesis. qed -lemma indeces_periodN: - assumes "w \ \" and "n \ 0" and forall: "\ i. i + n < \<^bold>|w\<^bold>| \ w!i = w!(i+n)" +lemma indeces_periodN: + assumes "w \ \" and "n \ 0" and forall: "\ i. i + n < \<^bold>|w\<^bold>| \ w!i = w!(i+n)" shows "periodN w n" proof- - have "\<^bold>|w\<^bold>| \ \<^bold>|take n w \ w\<^bold>|" + have "\<^bold>|w\<^bold>| \ \<^bold>|take n w \ w\<^bold>|" by auto {fix j assume "j < \<^bold>|w\<^bold>|" have "w ! j = (take n w \ w) ! j" proof (cases "j < \<^bold>|take n w\<^bold>|") assume "j < \<^bold>|take n w\<^bold>|" show "w ! j = (take n w \ w) ! j" - using pref_index[OF take_is_prefix \j < \<^bold>|take n w\<^bold>|\, symmetric] + using pref_index[OF take_is_prefix \j < \<^bold>|take n w\<^bold>|\, symmetric] unfolding pref_index[OF triv_pref \j < \<^bold>|take n w\<^bold>|\, of w]. next - assume "\ j < \<^bold>|take n w\<^bold>|" + assume "\ j < \<^bold>|take n w\<^bold>|" from leI[OF this] \j < \<^bold>|w\<^bold>|\ - have "\<^bold>|take n w\<^bold>| = n" + have "\<^bold>|take n w\<^bold>| = n" by force hence "j = (j - n) + n" and "(j - n) + n < \<^bold>|w\<^bold>|" using leI[OF \\ j < \<^bold>|take n w\<^bold>|\] \j < \<^bold>|w\<^bold>|\ by simp+ hence "w!j = w!(j - n)" using forall by simp from this[folded nth_append_length_plus[of "take n w" w "j-n", unfolded \\<^bold>|take n w\<^bold>| = n\]] - show "w ! j = (take n w \ w) ! j" - using \j = (j - n) + n\ by simp + show "w ! j = (take n w \ w) ! j" + using \j = (j - n) + n\ by simp qed} with index_pref[OF \\<^bold>|w\<^bold>| \ \<^bold>|take n w \ w\<^bold>|\] have "w \p take n w \ w" by blast - thus ?thesis + thus ?thesis using assms by force qed text\In some cases, the numeric definition is more useful than the definition using the period root.\ -lemma periodN_rev: assumes "periodN w p" shows "periodN (rev w) p" +lemma periodN_rev: assumes "periodN w p" shows "periodN (rev w) p" proof (rule indeces_periodN[of "rev w" p, OF _ periodN_D2[OF assms]]) show "rev w \ \" - using assms[unfolded periodN_def period_root_def] by force + using assms[unfolded periodN_def period_root_def] by force next - fix i assume "i + p < \<^bold>|rev w\<^bold>|" + fix i assume "i + p < \<^bold>|rev w\<^bold>|" from this[unfolded length_rev] add_lessD1 have "i < \<^bold>|w\<^bold>|" and "i + p < \<^bold>|w\<^bold>|" by blast+ have e: "\<^bold>|w\<^bold>| - Suc (i + p) + p = \<^bold>|w\<^bold>| - Suc i" using \i + p < \<^bold>|rev w\<^bold>|\ by simp - have "\<^bold>|w\<^bold>| - Suc (i + p) + p < \<^bold>|w\<^bold>|" using \i + p < \<^bold>|w\<^bold>|\ by linarith - from periodN_indeces[OF assms this] rev_nth[OF \i < \<^bold>|w\<^bold>|\, folded e] rev_nth[OF \i + p < \<^bold>|w\<^bold>|\] + have "\<^bold>|w\<^bold>| - Suc (i + p) + p < \<^bold>|w\<^bold>|" using \i + p < \<^bold>|w\<^bold>|\ by linarith + from periodN_indeces[OF assms this] rev_nth[OF \i < \<^bold>|w\<^bold>|\, folded e] rev_nth[OF \i + p < \<^bold>|w\<^bold>|\] show "rev w ! i = rev w !(i+p)" by presburger qed lemma periodN_rev_conv [reversal_rule]: "periodN (rev w) n \ periodN w n" using periodN_rev periodN_rev[of "rev w"] unfolding rev_rev_ident by (intro iffI) lemma periodN_fac: assumes "periodN (u\w\v) p" and "w \ \" shows "periodN w p" proof (rule indeces_periodN, simp add: \w \ \\) - show "p \ 0" using periodN_D2[OF \periodN (u\w\v) p\]. + show "p \ 0" using periodN_D2[OF \periodN (u\w\v) p\]. fix i assume "i + p < \<^bold>|w\<^bold>|" - hence "\<^bold>|u\<^bold>| + i + p < \<^bold>|u\w\v\<^bold>|" + hence "\<^bold>|u\<^bold>| + i + p < \<^bold>|u\w\v\<^bold>|" by simp from periodN_indeces[OF \periodN (u\w\v) p\ this] have "(u\w\v)!(\<^bold>|u\<^bold>| + i) = (u\w\v)! (\<^bold>|u\<^bold>| + (i + p))" - by (simp add: add.assoc) - thus "w!i = w!(i+p)" + by (simp add: add.assoc) + thus "w!i = w!(i+p)" using nth_append_length_plus[of u "w\v" , unfolded lassoc] \i + p < \<^bold>|w\<^bold>|\ add_lessD1[OF \i + p < \<^bold>|w\<^bold>|\] - nth_append[of w v] by auto + nth_append[of w v] by auto qed lemma periodN_fac': "periodN v p \ u \f v \ u \ \ \ periodN u p" by (elim facE, hypsubst, rule periodN_fac) lemma assumes "y \ \" and "k \ 0" shows "y\<^sup>@k \ \" by (simp add: assms(1) assms(2) nemp_emp_power) lemma pow_per: assumes "y \ \" and "k \ 0" shows "periodN (y\<^sup>@k) \<^bold>|y\<^bold>|" - using periodN_I[OF \k \ 0\[folded nemp_emp_power[OF \y \ \\]] nemp_len[OF \y \ \\]] pref_pow_ext_take by blast + using periodN_I[OF \k \ 0\[folded nemp_emp_power[OF \y \ \\]] nemp_len[OF \y \ \\]] pref_pow_ext_take by blast lemma per_fac: assumes "y \ \" and "v \ \" and "y\<^sup>@k = u\v\w" shows "periodN v \<^bold>|y\<^bold>|" proof- have "k \ 0" using nemp_pow suf_nemp[OF pref_nemp[OF \v \ \\, of w], of u, folded \y\<^sup>@k = u\v\w\] by blast from pow_per[OF \y \ \\ this, unfolded \y\<^sup>@k = u\v\w\] have "periodN (u \ v \ w) \<^bold>|y\<^bold>|". from periodN_fac[OF this \v \ \\] show "periodN v \<^bold>|y\<^bold>|". qed text\The numeric definition is equivalent to being prefix of a power.\ theorem periodN_pref: "periodN w n \ (\k r. w \np r\<^sup>@k \ \<^bold>|r\<^bold>| = n)" (is "_ \ ?R") proof(cases "w = \", simp) assume "w \ \" show "periodN w n \ ?R" proof - assume "periodN w n" + assume "periodN w n" consider (short) "\<^bold>|w\<^bold>| \ n" | (long) "n < \<^bold>|w\<^bold>|" by linarith then show ?R proof(cases) assume "\<^bold>|w\<^bold>| \ n" from le_add_diff_inverse[OF this] - obtain z where "\<^bold>|w \ z\<^bold>| = n" - unfolding length_append using exE[OF Ex_list_of_length[of "n - \<^bold>|w\<^bold>|"]] by metis + obtain z where "\<^bold>|w \ z\<^bold>| = n" + unfolding length_append using exE[OF Ex_list_of_length[of "n - \<^bold>|w\<^bold>|"]] by metis thus ?R using pow_one_id npI'[OF \w \ \\] by metis next assume "n < \<^bold>|w\<^bold>|" - then show ?R - using \periodN w n\[unfolded periodN_def per_pref[of w "take n w"]] + then show ?R + using \periodN w n\[unfolded periodN_def per_pref[of w "take n w"]] \w \ \\ take_len[OF less_imp_le[OF \n < \<^bold>|w\<^bold>|\]] by blast - qed - next - assume ?R + qed + next + assume ?R then obtain k r where "w \np r\<^sup>@k" and "n = \<^bold>|r\<^bold>|" by blast have "w \p take n w \ w" using pref_pow_take[OF npD[OF \w \np r \<^sup>@ k\], folded \n = \<^bold>|r\<^bold>|\]. - have "n \ 0" + have "n \ 0" unfolding length_0_conv[of r, folded \n = \<^bold>|r\<^bold>|\] using \w \np r \<^sup>@ k\ by force - hence "take n w \ \" + hence "take n w \ \" unfolding \n = \<^bold>|r\<^bold>|\ using \w \ \\ by simp - thus "periodN w n" + thus "periodN w n" unfolding periodN_def period_root_def using \w \p take n w \ w\ by blast qed qed text \Two more characterizations of a period\ -theorem per_shift: assumes "w \ \" "n \ 0" +theorem per_shift: assumes "w \ \" "n \ 0" shows "periodN w n \ drop n w \p w" proof assume "periodN w n" show "drop n w \p w" using drop_per_pref[OF \periodN w n\[unfolded periodN_def]] append_take_drop_id[of n w, unfolded append_eq_conv_conj] by argo next - assume "drop n w \p w" + assume "drop n w \p w" show "periodN w n" using conjI[OF pref_cancel'[OF \drop n w \p w\, of "take n w"] take_nemp[OF \w \ \\ \n \ 0\]] unfolding append_take_drop_id period_root_def by force qed -lemma rotate_per_root: assumes "w \ \" and "n \ 0" and "w = rotate n w" - shows "periodN w n" +lemma rotate_per_root: assumes "w \ \" and "n \ 0" and "w = rotate n w" + shows "periodN w n" proof (cases "\<^bold>|w\<^bold>| \ n") assume "\<^bold>|w\<^bold>| \ n" from all_long_pers[OF \w \ \\, OF this] show "periodN w n". next assume not: "\ \<^bold>|w\<^bold>| \ n" have "drop (n mod \<^bold>|w\<^bold>|) w \p w" using prefI[OF rotate_drop_take[symmetric, of n w]] unfolding \w = rotate n w\[symmetric]. from per_shift[OF \w \ \\ \n \ 0\] this[unfolded mod_less[OF not[unfolded not_le]]] - show "periodN w n".. + show "periodN w n".. qed subsubsection \Various lemmas on periods\ lemma periodN_drop: assumes "periodN w p" and "p < \<^bold>|w\<^bold>|" shows "periodN (drop p w) p" - using periodN_fac[of "take p w" "drop p w" \ p] \p < \<^bold>|w\<^bold>|\ \periodN w p\ - unfolding append_take_drop_id drop_eq_Nil not_le append_Nil2 by blast - -lemma ext_per_left: assumes "periodN w p" and "p \ \<^bold>|w\<^bold>|" + using periodN_fac[of "take p w" "drop p w" \ p] \p < \<^bold>|w\<^bold>|\ \periodN w p\ + unfolding append_take_drop_id drop_eq_Nil not_le append_Nil2 by blast + +lemma ext_per_left: assumes "periodN w p" and "p \ \<^bold>|w\<^bold>|" shows "periodN (take p w \ w) p" proof- - have f: "take p (take p w \ w) = take p w" + have f: "take p (take p w \ w) = take p w" using \p \ \<^bold>|w\<^bold>|\ by simp show ?thesis using \periodN w p\ pref_cancel'[of w "take p w \ w" "take p w" ] unfolding f periodN_def period_root_def by blast qed lemma ext_per_left_power: "periodN w p \ p \ \<^bold>|w\<^bold>| \ periodN ((take p w)\<^sup>@k \ w) p" proof (induction k) case (Suc k) - show ?case + show ?case using ext_per_left[OF Suc.IH[OF \periodN w p\ \p \ \<^bold>|w\<^bold>|\]] \p \ \<^bold>|w\<^bold>|\ - unfolding pref_share_take[OF per_exp_pref[OF periodN_D3[OF \periodN w p\]] \p \ \<^bold>|w\<^bold>|\,symmetric] + unfolding pref_share_take[OF per_exp_pref[OF periodN_D3[OF \periodN w p\]] \p \ \<^bold>|w\<^bold>|\,symmetric] lassoc pow_Suc_list[symmetric] by fastforce qed auto lemma take_several_pers: assumes "periodN w n" and "m*n \ \<^bold>|w\<^bold>|" shows "(take n w)\<^sup>@m = take (m*n) w" proof (cases "m = 0", simp) - assume "m \ 0" + assume "m \ 0" have "\<^bold>|(take n w)\<^sup>@m\<^bold>| = m*n" - unfolding pow_len nat_prod_le[OF \m \ 0\ \m*n \ \<^bold>|w\<^bold>|\, THEN take_len] by blast + unfolding pow_len nat_prod_le[OF \m \ 0\ \m*n \ \<^bold>|w\<^bold>|\, THEN take_len] by blast have "(take n w)\<^sup>@m \p w" - using \periodN w n\[unfolded periodN_def, THEN per_exp[of w "take n w" m], THEN + using \periodN w n\[unfolded periodN_def, THEN per_exp[of w "take n w" m], THEN ruler_le[of "take n w\<^sup>@m" "take n w\<^sup>@m \ w" w, OF triv_pref], OF \m * n \ \<^bold>|w\<^bold>|\[folded \\<^bold>|take n w\<^sup>@m\<^bold>| = m * n\]]. - show ?thesis + show ?thesis using pref_take[OF \take n w\<^sup>@m \p w\, unfolded \\<^bold>|take n w\<^sup>@m\<^bold>| = m * n\, symmetric]. qed lemma per_mult: assumes "periodN w n" and "m \ 0" shows "periodN w (m*n)" proof (cases "m*n \ \<^bold>|w\<^bold>|") - have "w \ \" using periodN_D1[OF \periodN w n\]. + have "w \ \" using periodN_D1[OF \periodN w n\]. assume "\ m * n \ \<^bold>|w\<^bold>|" thus "periodN w (m*n)" using all_long_pers[of w "m * n", OF \w \ \\] by linarith -next +next assume "m * n \ \<^bold>|w\<^bold>|" show "periodN w (m*n)" - using per_add_exp[of w "take n w", OF _ \m \ 0\] \periodN w n\ + using per_add_exp[of w "take n w", OF _ \m \ 0\] \periodN w n\ unfolding periodN_def period_root_def take_several_pers[OF \periodN w n\ \m*n \ \<^bold>|w\<^bold>|\, symmetric] by blast qed lemma root_periodN: assumes "w \ \" and "w \p r\<^sup>\" shows "periodN w \<^bold>|r\<^bold>|" - unfolding periodN_def period_root_def using per_pref_ex[OF \w \p r\<^sup>\\ - pref_pow_take[of w r], of "\ x. x"] take_nemp_len[OF \w \ \\ per_rootD'[OF \w \p r\<^sup>\\]] by blast + unfolding periodN_def period_root_def using per_pref_ex[OF \w \p r\<^sup>\\ + pref_pow_take[of w r], of "\ x. x"] take_nemp_len[OF \w \ \\ per_rootD'[OF \w \p r\<^sup>\\]] by blast theorem two_periodsN: - assumes "periodN w p" "periodN w q" "p + q \ \<^bold>|w\<^bold>|" + assumes "periodN w p" "periodN w q" "p + q \ \<^bold>|w\<^bold>|" shows "periodN w (gcd p q)" proof- - obtain t where "take p w \ t*" "take q w \ t*" + obtain t where "take p w \ t*" "take q w \ t*" using two_pers[OF \periodN w p\[unfolded periodN_def] \periodN w q\[unfolded periodN_def], unfolded take_len[OF add_leD1[OF \p + q \ \<^bold>|w\<^bold>|\]] take_len[OF add_leD2[OF \p + q \ \<^bold>|w\<^bold>|\]], OF \p + q \ \<^bold>|w\<^bold>|\, unfolded comm_root[of "take p w" "take q w"]] by blast hence "w \p t\<^sup>\" using \periodN w p\ periodN_def per_root_trans by blast have "periodN w \<^bold>|t\<^bold>|" - using root_periodN[OF per_nemp[OF \periodN w p\] \w \p t\<^sup>\\]. + using root_periodN[OF per_nemp[OF \periodN w p\] \w \p t\<^sup>\\]. have "\<^bold>|t\<^bold>| dvd (gcd p q)" - using gcd_nat.boundedI[OF root_len_dvd[OF \take p w \ t*\] root_len_dvd[OF \take q w \ t*\]] - unfolding take_len[OF add_leD1[OF \p + q \ \<^bold>|w\<^bold>|\]] take_len[OF add_leD2[OF \p + q \ \<^bold>|w\<^bold>|\]]. - thus ?thesis + using gcd_nat.boundedI[OF root_len_dvd[OF \take p w \ t*\] root_len_dvd[OF \take q w \ t*\]] + unfolding take_len[OF add_leD1[OF \p + q \ \<^bold>|w\<^bold>|\]] take_len[OF add_leD2[OF \p + q \ \<^bold>|w\<^bold>|\]]. + thus ?thesis using per_mult[OF \periodN w \<^bold>|t\<^bold>|\, of "gcd p q div \<^bold>|t\<^bold>|", unfolded dvd_div_mult_self[OF \\<^bold>|t\<^bold>| dvd (gcd p q)\]] dvd_div_mult_self[OF \\<^bold>|t\<^bold>| dvd (gcd p q)\] gcd_eq_0_iff[of p q] mult_zero_left[of "\<^bold>|t\<^bold>|"] periodN_D2[OF \periodN w p\] by argo qed lemma index_mod_per_root: assumes "r \ \" and i: "\ i < \<^bold>|w\<^bold>|. w!i = r!(i mod \<^bold>|r\<^bold>|)" shows "w \p r\<^sup>\" proof- have "i < \<^bold>|w\<^bold>| \ (r \ w) ! i = r ! (i mod \<^bold>|r\<^bold>|)" for i by (simp add: i mod_if nth_append) hence "w \p r \ w" using index_pref[of w "r \ w"] i - by simp - thus ?thesis unfolding period_root_def using \r \ \\ by auto + by simp + thus ?thesis unfolding period_root_def using \r \ \\ by auto qed lemma index_pref_pow_mod: "w \p r\<^sup>@k \ i < \<^bold>|w\<^bold>| \ w!i = r!(i mod \<^bold>|r\<^bold>| )" using index_pow_mod[of i r k] less_le_trans[of i "\<^bold>|w\<^bold>|" "\<^bold>|r\<^sup>@k\<^bold>|"] prefix_length_le[of w "r\<^sup>@k"] pref_index[of w "r\<^sup>@k" i] by argo lemma index_per_root_mod: "w \p r\<^sup>\ \ i < \<^bold>|w\<^bold>| \ w!i = r!(i mod \<^bold>|r\<^bold>|)" using index_pref_pow_mod[of w r _ i] per_pref[of w r ] by blast lemma root_divisor: assumes "periodN w k" and "k dvd \<^bold>|w\<^bold>|" shows "w \ (take k w)*" - using rootI[of "take k w" "(\<^bold>|w\<^bold>| div k)"] unfolding - take_several_pers[OF \periodN w k\, of "\<^bold>|w\<^bold>| div k", unfolded dvd_div_mult_self[OF \k dvd \<^bold>|w\<^bold>|\] take_self, OF , OF order_refl]. - -lemma per_pref': assumes "u \ \" and "periodN v k" and "u \p v" shows "periodN u k" + using rootI[of "take k w" "(\<^bold>|w\<^bold>| div k)"] unfolding + take_several_pers[OF \periodN w k\, of "\<^bold>|w\<^bold>| div k", unfolded dvd_div_mult_self[OF \k dvd \<^bold>|w\<^bold>|\] take_self, OF , OF order_refl]. + +lemma per_pref': assumes "u \ \" and "periodN v k" and "u \p v" shows "periodN u k" proof- { assume "k \ \<^bold>|u\<^bold>|" have "take k v = take k u" using pref_share_take[OF \u \p v\ \k \ \<^bold>|u\<^bold>|\] by auto hence "take k v \ \" using \periodN v k\ by auto hence "take k u \ \" by (simp add: \take k v = take k u\) - have "u \p take k u \ v" - using \periodN v k\ - unfolding periodN_def period_root_def \take k v = take k u\ + have "u \p take k u \ v" + using \periodN v k\ + unfolding periodN_def period_root_def \take k v = take k u\ using pref_trans[OF \u \p v\, of "take k u \ v"] - by blast + by blast hence "u \p take k u \ u" - using \u \p v\ pref_prod_pref by blast + using \u \p v\ pref_prod_pref by blast hence ?thesis using \take k u \ \\ periodN_def by fast } thus ?thesis using \u \ \\ all_long_pers nat_le_linear by blast qed subsection "Period: overview" notepad begin - fix w r::"'a list" + fix w r::"'a list" fix n::nat assume "w \ \" "r \ \" "n > 0" have "\ w \p \\<^sup>\" by simp have "\ \ \p \\<^sup>\" by simp have "\ \p r\<^sup>\" by (simp add: \r \ \\) have "\ periodN w 0" by simp have "\ periodN \ 0" by simp have "\ periodN \ n" by simp end subsection \Singleton and its power\ lemma concat_len_one: assumes "\<^bold>|us\<^bold>| = 1" shows "concat us = hd us" - using concat_sing[OF sing_word[OF \\<^bold>|us\<^bold>| = 1\, symmetric]]. + using concat_sing[OF sing_word[OF \\<^bold>|us\<^bold>| = 1\, symmetric]]. lemma sing_pow_hd_tl: "c # w \ [a]* \ c = a \ w \ [a]*" proof assume "c = a \ w \ [a]*" thus "c # w \ [a]*" unfolding hd_word[of _ w] using add_root[of "[c]" w] by simp next assume "c # w \ [a]*" - then obtain k where "[a]\<^sup>@k = c # w" unfolding root_def by blast + then obtain k where "[a]\<^sup>@k = c # w" unfolding root_def by blast thus "c = a \ w \ [a]*" proof (cases "k = 0", simp) assume "[a] \<^sup>@ k = c # w" and "k \ 0" from eqd_equal[of "[a]", OF this(1)[unfolded hd_word[of _ w] pop_pow_one[OF \k \ 0\]]] - show ?thesis + show ?thesis unfolding root_def by auto qed qed lemma pref_sing_power: assumes "w \p [a]\<^sup>@m" shows "w = [a]\<^sup>@(\<^bold>|w\<^bold>|)" proof- have "[a]\<^sup>@m = [a]\<^sup>@(\<^bold>|w\<^bold>|)\[a]\<^sup>@(m-\<^bold>|w\<^bold>|)" using pop_pow[OF prefix_length_le[OF assms, unfolded sing_pow_len], of "[a]", symmetric]. show ?thesis - using conjunct1[OF eqd_equal[of w "w\\<^sup>>[a]\<^sup>@m" "[a]\<^sup>@(\<^bold>|w\<^bold>|)""[a]\<^sup>@(m-\<^bold>|w\<^bold>|)", - unfolded lq_pref[OF assms] sing_pow_len, - OF \[a]\<^sup>@m = [a]\<^sup>@(\<^bold>|w\<^bold>|)\[a]\<^sup>@(m-\<^bold>|w\<^bold>|)\ refl]]. + using conjunct1[OF eqd_equal[of w "w\\<^sup>>[a]\<^sup>@m" "[a]\<^sup>@(\<^bold>|w\<^bold>|)""[a]\<^sup>@(m-\<^bold>|w\<^bold>|)", + unfolded lq_pref[OF assms] sing_pow_len, + OF \[a]\<^sup>@m = [a]\<^sup>@(\<^bold>|w\<^bold>|)\[a]\<^sup>@(m-\<^bold>|w\<^bold>|)\ refl]]. qed lemma sing_pow_palindrom: assumes "w = [a]\<^sup>@k" shows "rev w = w" - using rev_pow[of "[a]" "\<^bold>|w\<^bold>|", unfolded rev_sing] + using rev_pow[of "[a]" "\<^bold>|w\<^bold>|", unfolded rev_sing] unfolding pref_sing_power[of w a k, unfolded assms[unfolded root_def, symmetric], OF self_pref, symmetric]. lemma suf_sing_power: assumes "w \s [a]\<^sup>@m" shows "w \ [a]*" using sing_pow_palindrom[of "rev w" a "\<^bold>|rev w\<^bold>|", unfolded rev_rev_ident] - pref_sing_power[of "rev w" a m, OF \w \s [a]\<^sup>@m\[unfolded suffix_to_prefix rev_pow rev_rev_ident rev_sing]] + pref_sing_power[of "rev w" a m, OF \w \s [a]\<^sup>@m\[unfolded suffix_to_prefix rev_pow rev_rev_ident rev_sing]] rootI[of "[a]" "\<^bold>|rev w\<^bold>|"] by auto lemma sing_fac_pow: assumes "w \ [a]*" and "v \f w" shows "v \ [a]*" proof- obtain k where "w = [a]\<^sup>@k" using \w \ [a]*\[unfolded root_def] by blast - obtain p where "p \p w" and "v \s p" + obtain p where "p \p w" and "v \s p" using fac_pref_suf[OF \ v \f w\] by blast hence "v \s [a]\<^sup>@ \<^bold>|p\<^bold>|" using pref_sing_power[OF \p \p w\[unfolded \w = [a]\<^sup>@k\]] by argo from suf_sing_power[OF this] show ?thesis. qed lemma sing_pow_fac': assumes "a \ b" and "w \ [a]*" shows "\ ([b] \f w)" - using sing_fac_pow[OF \ w \ [a]*\, of "[b]"] unfolding sing_pow_hd_tl[of b \] - using \a \ b\ by auto + using sing_fac_pow[OF \ w \ [a]*\, of "[b]"] unfolding sing_pow_hd_tl[of b \] + using \a \ b\ by auto lemma all_set_sing_pow: "(\ b. b \ set w \ b = a) \ w \ [a]*" (is "?All \ _") proof assume ?All then show "w \ [a]*" proof (induct w, simp) case (Cons c w) then show ?case by (simp add: sing_pow_hd_tl) qed next assume "w \ [a]*" then show ?All proof (induct w, simp) case (Cons c w) then show ?case - unfolding sing_pow_hd_tl by simp + unfolding sing_pow_hd_tl by simp qed qed lemma sing_pow_set: "u \ [a]* \ u \ \ \ set u = {a}" unfolding all_set_sing_pow[symmetric] using hd_in_set[of u] is_singletonI'[unfolded is_singleton_the_elem, of "set u"] singleton_iff[of a "the_elem (set u)"] by auto lemma takeWhile_sing_root: "takeWhile (\ x. x = a) w \ [a]*" unfolding all_set_sing_pow[symmetric] using set_takeWhileD[of _ "\ x. x = a" w] by blast lemma takeWhile_sing_pow: "takeWhile (\ x. x = a) w = w \ w = [a]\<^sup>@\<^bold>|w\<^bold>|" by(induct w, auto) lemma dropWhile_sing_pow: "dropWhile (\ x. x = a) w = \ \ w = [a]\<^sup>@\<^bold>|w\<^bold>|" by(induct w, auto) -lemma distinct_letter_in: assumes "\ w \ [a]*" +lemma distinct_letter_in: assumes "\ w \ [a]*" obtains m b q where "[a]\<^sup>@m \ [b] \ q = w" and "b \ a" proof- have "dropWhile (\ x. x = a) w \ \" - unfolding dropWhile_sing_pow using assms rootI[of "[a]" "\<^bold>|w\<^bold>|"] by auto + unfolding dropWhile_sing_pow using assms rootI[of "[a]" "\<^bold>|w\<^bold>|"] by auto hence eq: "takeWhile (\ x. x = a) w \ [hd (dropWhile (\ x. x = a) w)] \ tl (dropWhile (\ x. x = a) w) = w" by simp have root:"takeWhile (\ x. x = a) w \ [a]*" - by (simp add: takeWhile_sing_root) - have "hd (dropWhile (\ x. x = a) w) \ a" + by (simp add: takeWhile_sing_root) + have "hd (dropWhile (\ x. x = a) w) \ a" using hd_dropWhile[OF \dropWhile (\x. x = a) w \ \\]. - from that[OF _ this] + from that[OF _ this] show thesis using eq root unfolding root_def by metis qed lemma distinct_letter_in_hd: assumes "\ w \ [hd w]*" obtains m b q where "[hd w]\<^sup>@m \ [b] \ q = w" and "b \ hd w" and "m \ 0" proof- obtain m b q where a1: "[hd w]\<^sup>@m \ [b] \ q = w" and a2: "b \ hd w" - using distinct_letter_in[OF assms]. + using distinct_letter_in[OF assms]. hence "m \ 0" - using power_eq_if[of "[hd w]" m] list.sel(1) by fastforce - from that a1 a2 this + using power_eq_if[of "[hd w]" m] list.sel(1) by fastforce + from that a1 a2 this show thesis by blast qed -lemma distinct_letter_in_suf: assumes "\ w \ [a]*" shows "\ m b. [b] \ [a]\<^sup>@m \s w \ b \ a" +lemma distinct_letter_in_suf: assumes "\ w \ [a]*" shows "\ m b. [b] \ [a]\<^sup>@m \s w \ b \ a" proof- have "\ rev w \ [a]*" using rev_pow[of "[a]"] unfolding rev_sing using assms root_def rev_rev_ident[of w] by metis obtain m b q where "[a]\<^sup>@m \ [b] \ q = rev w" and "b \ a" using distinct_letter_in[OF \\ rev w \ [a]*\] by blast have eq: "rev ([b] \ [a]\<^sup>@m) = [a]\<^sup>@m \ [b]" - by (simp add: rev_pow) + by (simp add: rev_pow) have "[b] \ [a]\<^sup>@m \s w" - using \[a]\<^sup>@m \ [b] \ q = rev w\ unfolding suf_rev_pref_iff eq lassoc + using \[a]\<^sup>@m \ [b] \ q = rev w\ unfolding suf_rev_pref_iff eq lassoc using prefI by blast thus ?thesis using \b \ a\ by blast qed lemma sing_pow_exp: assumes "w \ [a]*" shows "w = [a]\<^sup>@\<^bold>|w\<^bold>|" proof- obtain k where "[a] \<^sup>@ k = w" using rootE[OF assms]. from this[folded sing_pow_len[of a k, folded this, unfolded this], symmetric] show ?thesis. qed lemma sing_power': assumes "w \ [a]*" and "i < \<^bold>|w\<^bold>|" shows "w ! i = a" - using sing_pow[OF \i < \<^bold>|w\<^bold>|\, of a, folded sing_pow_exp[OF \w \ [a]*\]]. + using sing_pow[OF \i < \<^bold>|w\<^bold>|\, of a, folded sing_pow_exp[OF \w \ [a]*\]]. lemma rev_sing_power: "x \ [a]* \ rev x = x" unfolding root_def using rev_pow rev_singleton_conv by metis -lemma lcp_letter_power: - assumes "w \ \" and "w \ [a]*" and "[a]\<^sup>@m \ [b] \p z" and "a \ b" - shows "w \ z \\<^sub>p z \ w = [a]\<^sup>@m" +lemma lcp_letter_power: + assumes "w \ \" and "w \ [a]*" and "[a]\<^sup>@m \ [b] \p z" and "a \ b" + shows "w \ z \\<^sub>p z \ w = [a]\<^sup>@m" proof- obtain k z' where "w = [a]\<^sup>@k" "z = [a]\<^sup>@m \ [b] \ z'" "k \ 0" - using \w \ [a]*\ \[a]\<^sup>@m \ [b] \p z\ \w \ \\ nemp_pow[of "[a]"] + using \w \ [a]*\ \[a]\<^sup>@m \ [b] \p z\ \w \ \\ nemp_pow[of "[a]"] unfolding root_def - by auto + by auto hence eq1: "w \ z = [a]\<^sup>@m \ ([a]\<^sup>@k \ [b] \ z')" and eq2: "z \ w = [a]\<^sup>@m \ ([b] \ z'\ [a]\<^sup>@k)" by (simp add: \w = [a]\<^sup>@k\ \z = [a]\<^sup>@m \ [b] \ z'\ pow_comm)+ - have "hd ([a]\<^sup>@k \ [b] \ z') = a" - using hd_append2[OF \w \ \\, of "[b]\z'", + have "hd ([a]\<^sup>@k \ [b] \ z') = a" + using hd_append2[OF \w \ \\, of "[b]\z'", unfolded \w = (a # \)\<^sup>@k\ hd_sing_power[OF \k \ 0\, of a]]. moreover have "hd([b] \ z'\ [a]\<^sup>@k) = b" - by simp + by simp ultimately have "[a]\<^sup>@k \ [b] \ z' \\<^sub>p [b] \ z'\ [a]\<^sup>@k = \" - by (simp add: \a \ b\ lcp_distinct_hd) + by (simp add: \a \ b\ lcp_distinct_hd) thus ?thesis using eq1 eq2 lcp_ext_left[of "[a]\<^sup>@m" "[a]\<^sup>@k \ [b] \ z'" "[b] \ z'\ [a]\<^sup>@k"] by simp qed lemma per_one: assumes "w \p [a]\<^sup>\" shows "w \ [a]*" proof- have "w \p (a # \) \<^sup>@ n \ n \ 0 \ w \ [a]*" for n using pref_sing_power[of w a] rootI[of "[a]" "\<^bold>|w\<^bold>|"] by auto with per_pref_ex[OF \w \p [a]\<^sup>\\] show ?thesis by auto -qed +qed lemma per_one': "w \ [a]* \ w \p [a]\<^sup>\" by (metis append_Nil2 not_Cons_self2 per_pref prefI root_def) lemma per_sing_one: assumes "w \ \" "w \p [a]\<^sup>\" shows "periodN w 1" using root_periodN[OF \w \ \\ \w \p [a]\<^sup>\\] unfolding sing_len[of a]. section "Border" text\A non-empty word $x \neq w$ is a \emph{border} of a word $w$ if it is both its prefix and suffix. This elementary property captures how much the word $w$ overlaps with itself, and it is in the obvious way related to a period of $w$. However, in many cases it is much easier to reason about borders than about periods.\ definition border :: "'a list \ 'a list \ bool" ("_ \b _" [51,51] 60 ) where [simp]: "border x w = (x \p w \ x \s w \ x \ w \ x \ \)" -definition bordered :: "'a list \ bool" +definition bordered :: "'a list \ bool" where [simp]: "bordered w = (\b. b \b w)" lemma borderI[intro]: "x \p w \ x \s w \ x \ w \ x \ \ \ x \b w" unfolding border_def by blast lemma borderD_pref: "x \b w \ x \p w" unfolding border_def by blast lemma borderD_spref: "x \b w \ x

b w \ x \s w" unfolding border_def by blast lemma borderD_ssuf: "x \b w \ x b w \ x \ \" using border_def by blast lemma borderD_neq: "x \b w \ x \ w" unfolding border_def by blast lemma border_lq_nemp: assumes "x \b w" shows "x\\<^sup>>w \ \" - using assms borderD_spref lq_spref by blast + using assms borderD_spref lq_spref by blast lemma border_rq_nemp: assumes "x \b w" shows "w\<^sup><\x \ \" using assms borderD_ssuf rq_ssuf by blast lemma border_trans[trans]: assumes "t \b x" "x \b w" shows "t \b w" - using assms unfolding border_def - using suffix_order.order.antisym pref_trans[of t x w] suf_trans[of t x w] by blast - + using assms unfolding border_def + using suffix_order.antisym pref_trans[of t x w] suf_trans[of t x w] by blast + lemma border_rev_conv[reversal_rule]: "rev x \b rev w \ x \b w" unfolding border_def using rev_is_Nil_conv[of x] rev_swap[of w] rev_swap[of x] suf_rev_pref_iff[of x w] pref_rev_suf_iff[of x w] by fastforce lemma border_lq_comp: "x \b w \ (w\<^sup><\x) \ x" unfolding border_def using rq_suf ruler by blast lemmas border_lq_suf_comp = border_lq_comp[reversed] subsection "The shortest border" lemma border_len: assumes "x \b w" - shows "1 < \<^bold>|w\<^bold>|" and "0 < \<^bold>|x\<^bold>|" and "\<^bold>|x\<^bold>| < \<^bold>|w\<^bold>|" + shows "1 < \<^bold>|w\<^bold>|" and "0 < \<^bold>|x\<^bold>|" and "\<^bold>|x\<^bold>| < \<^bold>|w\<^bold>|" proof- show "\<^bold>|x\<^bold>| < \<^bold>|w\<^bold>|" using assms prefix_length_less unfolding border_def strict_prefix_def by blast show "0 < \<^bold>|x\<^bold>|" using assms unfolding border_def by blast thus "1 < \<^bold>|w\<^bold>|" using assms \\<^bold>|x\<^bold>| < \<^bold>|w\<^bold>|\ unfolding border_def - by linarith + by linarith qed lemma borders_compare: assumes "x \b w" and "x' \b w" and "\<^bold>|x'\<^bold>| < \<^bold>|x\<^bold>|" shows "x' \b x" using ruler_le[OF borderD_pref[OF \x' \b w\] borderD_pref[OF \x \b w\] less_imp_le_nat[OF \\<^bold>|x'\<^bold>| < \<^bold>|x\<^bold>|\]] suf_ruler_le[OF borderD_suf[OF \x' \b w\] borderD_suf[OF \x \b w\] less_imp_le_nat[OF \\<^bold>|x'\<^bold>| < \<^bold>|x\<^bold>|\]] borderD_nemp[OF \x' \b w\] \\<^bold>|x'\<^bold>| < \<^bold>|x\<^bold>|\ borderI by blast -lemma unbordered_border: +lemma unbordered_border: "bordered w \ \ x. x \b w \ \ bordered x" proof (induction "\<^bold>|w\<^bold>|" arbitrary: w rule: less_induct) case less obtain x' where "x' \b w" using bordered_def less.prems by blast - show ?case + show ?case proof (cases "bordered x'") - assume "\ bordered x'" - thus ?case + assume "\ bordered x'" + thus ?case using \x' \b w\ by blast - next + next assume "bordered x'" from less.hyps[OF border_len(3)[OF \x' \b w\] this] - show ?case - using border_trans[of _ x' w] \x' \b w\ by blast + show ?case + using border_trans[of _ x' w] \x' \b w\ by blast qed qed lemma unbordered_border_shortest: "x \b w \ \ bordered x \ y \b w \ \<^bold>|x\<^bold>| \ \<^bold>|y\<^bold>|" using bordered_def[of x] borders_compare[of x w y] not_le_imp_less[of "\<^bold>|x\<^bold>|" "\<^bold>|y\<^bold>|"] by blast -lemma long_border_bordered: assumes long: "\<^bold>|w\<^bold>| < \<^bold>|x\<^bold>| + \<^bold>|x\<^bold>|" and border: "x \b w" +lemma long_border_bordered: assumes long: "\<^bold>|w\<^bold>| < \<^bold>|x\<^bold>| + \<^bold>|x\<^bold>|" and border: "x \b w" shows "(w\<^sup><\x)\\<^sup>>x \b x" proof- define p s where "p = w\<^sup><\x" and "s = x\\<^sup>>w" - hence eq: "p\x = x\s" + hence eq: "p\x = x\s" using assms unfolding border_def using lq_pref[of x w] rq_suf[of x w] by simp have "\<^bold>|p\<^bold>| < \<^bold>|x\<^bold>|" using p_def long[folded rq_len[OF borderD_suf[OF border]]] by simp - have px: "p \ p\\<^sup>>x = x" and sx: "p\\<^sup>>x \ s = x" + have px: "p \ p\\<^sup>>x = x" and sx: "p\\<^sup>>x \ s = x" using eqd_pref[OF eq less_imp_le, OF \\<^bold>|p\<^bold>| < \<^bold>|x\<^bold>|\] by blast+ have "p\\<^sup>>x \ \" using \\<^bold>|p\<^bold>| < \<^bold>|x\<^bold>|\ px by fastforce have "p \ \" using border_rq_nemp[OF border] p_def by presburger have "p\\<^sup>>x \ x" using \p \ \\ px by force - have "(p\\<^sup>>x) \b x" - unfolding border_def + have "(p\\<^sup>>x) \b x" + unfolding border_def using eqd_pref[OF eq less_imp_le, OF \\<^bold>|p\<^bold>| < \<^bold>|x\<^bold>|\] \p\\<^sup>>x \ \\ \p\\<^sup>>x \ x\ by blast - thus ?thesis + thus ?thesis unfolding p_def. qed thm long_border_bordered[reversed] lemma border_short_dec: assumes border: "x \b w" and short: "\<^bold>|x\<^bold>| + \<^bold>|x\<^bold>| \ \<^bold>|w\<^bold>|" shows "x \ x\\<^sup>>(w\<^sup><\x) \ x = w" proof- have eq: "x\x\\<^sup>>w = w\<^sup><\x\x" - using lq_pref[OF borderD_pref[OF border]] rq_suf[OF borderD_suf[OF border]] by simp + using lq_pref[OF borderD_pref[OF border]] rq_suf[OF borderD_suf[OF border]] by simp have "\<^bold>|x\<^bold>| \ \<^bold>|w\<^sup><\x\<^bold>|" using short[folded rq_len[OF borderD_suf[OF border]]] by simp from lq_pref[of x w, OF borderD_pref[OF border], folded conjunct2[OF eqd_pref[OF eq this]]] show ?thesis. qed -lemma bordered_dec: assumes "bordered w" +lemma bordered_dec: assumes "bordered w" obtains u v where "u\v\u = w" and "u \ \" proof- obtain u where "u \b w" and "\ bordered u" using unbordered_border[OF assms] by blast have "\<^bold>|u\<^bold>| + \<^bold>|u\<^bold>| \ \<^bold>|w\<^bold>|" - using long_border_bordered[OF _ \u \b w\] \\ bordered u\ bordered_def leI by blast + using long_border_bordered[OF _ \u \b w\] \\ bordered u\ bordered_def leI by blast from border_short_dec[OF \u \b w\ this, THEN that, OF borderD_nemp[OF \u \b w\]] show thesis. qed subsection "Relation to period and conjugation" lemma border_conjug_eq: "x \b w \ (w\<^sup><\x) \ w = w \ (x\\<^sup>>w)" using lq_rq_reassoc_suf[OF borderD_pref borderD_suf, symmetric] by blast lemma border_per_root: "x \b w \ w \p (w\<^sup><\x) \ w" using border_conjug_eq by blast -lemma per_root_border: assumes "\<^bold>|r\<^bold>| < \<^bold>|w\<^bold>|" and "r \ \" and "w \p r \ w" - shows "r\\<^sup>>w \b w" -proof +lemma per_root_border: assumes "\<^bold>|r\<^bold>| < \<^bold>|w\<^bold>|" and "r \ \" and "w \p r \ w" + shows "r\\<^sup>>w \b w" +proof have "\<^bold>|r\<^bold>| \ \<^bold>|w\<^bold>|" and "r \p w" using less_imp_le[OF \\<^bold>|r\<^bold>| < \<^bold>|w\<^bold>|\] pref_prod_long[OF \w \p r \ w\] by blast+ show "r\\<^sup>>w \p w" using pref_lq[OF \r \p w\ \w \p r \ w\] unfolding lq_triv. show "r\\<^sup>>w \s w" using \r \p w\ by auto show "r\\<^sup>>w \ w" - using \r \p w\ \r \ \\ by force + using \r \p w\ \r \ \\ by force show "r\\<^sup>>w \ \" using lq_pref[OF \r \p w\] \\<^bold>|r\<^bold>| < \<^bold>|w\<^bold>|\ by force qed -lemma border_per: assumes "x \b w" shows "periodN w (\<^bold>|w\<^bold>|-\<^bold>|x\<^bold>|)" +lemma border_per: assumes "x \b w" shows "periodN w (\<^bold>|w\<^bold>|-\<^bold>|x\<^bold>|)" proof- - have "w = (w\<^sup><\x)\x" + have "w = (w\<^sup><\x)\x" using rq_suf[OF borderD_suf[OF assms]] by simp - have "w = x\(x\\<^sup>>w)" + have "w = x\(x\\<^sup>>w)" using lq_pref[OF borderD_pref[OF assms]] by simp have take: "take (\<^bold>|w\<^bold>|-\<^bold>|x\<^bold>|) w = w\<^sup><\x" using borderD_suf[OF assms] by auto have nemp: "take (\<^bold>|w\<^bold>|-\<^bold>|x\<^bold>|) w \ \" - using assms by auto + using assms by auto have "w \p take (\<^bold>|w\<^bold>|-\<^bold>|x\<^bold>|) w \ w" - unfolding take lassoc[of "w\<^sup><\x" x "x\\<^sup>>w", folded \w = x \ x\\<^sup>>w\ \w = w\<^sup><\x \ x\] + unfolding take lassoc[of "w\<^sup><\x" x "x\\<^sup>>w", folded \w = x \ x\\<^sup>>w\ \w = w\<^sup><\x \ x\] using triv_pref[of w "x\\<^sup>>w"]. - thus "periodN w (\<^bold>|w\<^bold>|-\<^bold>|x\<^bold>|)" - unfolding periodN_def period_root_def using nemp by blast + thus "periodN w (\<^bold>|w\<^bold>|-\<^bold>|x\<^bold>|)" + unfolding periodN_def period_root_def using nemp by blast qed -lemma per_border: assumes "n < \<^bold>|w\<^bold>|" and "periodN w n" +lemma per_border: assumes "n < \<^bold>|w\<^bold>|" and "periodN w n" shows "take (\<^bold>|w\<^bold>| - n) w \b w" proof- - have eq: "take (\<^bold>|w\<^bold>| - n) w = drop n w" + have eq: "take (\<^bold>|w\<^bold>| - n) w = drop n w" using pref_take[OF \periodN w n\[unfolded per_shift[OF periodN_D1[OF \periodN w n\] per_positive[OF \periodN w n\]]], unfolded length_drop]. have "take (\<^bold>|w\<^bold>| - n) w \ \" using \n < \<^bold>|w\<^bold>|\ take_eq_Nil by fastforce - moreover have "take (\<^bold>|w\<^bold>| - n) w \ w" + moreover have "take (\<^bold>|w\<^bold>| - n) w \ w" using periodN_D2[OF \periodN w n\] \n < \<^bold>|w\<^bold>|\ unfolding take_all_iff[of "\<^bold>|w\<^bold>|-n" w] by fastforce - ultimately show ?thesis + ultimately show ?thesis unfolding border_def using take_is_prefix[of "\<^bold>|w\<^bold>|-n" w] suffix_drop[of n w, folded eq] by blast qed section \Primitive words\ text\If a word $w$ is not a non-trivial power of some other word, we say it is primitive.\ -definition primitive :: "'a list \ bool" +definition primitive :: "'a list \ bool" where "primitive u = (\ r k. r\<^sup>@k = u \ k = 1)" lemma primI: "(\ r k. r\<^sup>@k = u \ k = 1) \ primitive u" by (simp add: primitive_def) lemma prim_nemp: "primitive u \ u \ \" proof- have "u = \ \ \\<^sup>@0 = u" by simp thus "primitive u \ u \ \" using primitive_def zero_neq_one by blast qed lemma prim_exp_one: "primitive u \ r\<^sup>@k = u \ k = 1" using primitive_def by blast lemma prim_exp_eq: "primitive u \ r\<^sup>@k = u \ u = r" using prim_exp_one power_one_right by blast lemma prim_triv_root: "primitive u \ u \ t* \ t = u" - using prim_exp_eq unfolding root_def + using prim_exp_eq unfolding root_def unfolding primitive_def root_def by fastforce lemma prim_comm_exp[elim]: assumes "primitive v" and "v\u = u\v" obtains k where "u = v\<^sup>@k" using \v\u = u\v\[unfolded comm] prim_exp_eq[OF \primitive v\] by metis lemma pow_non_prim: "1 < k \ \ primitive (w\<^sup>@k)" using prim_exp_one by auto lemma comm_non_prim: assumes "u \ \" "v \ \" "u\v = v\u" shows "\ primitive (u\v)" proof- - obtain t k m where "u = t\<^sup>@k" "v = t\<^sup>@m" + obtain t k m where "u = t\<^sup>@k" "v = t\<^sup>@m" using \u\v = v\u\[unfolded comm] by blast show ?thesis using pow_non_prim[of "k+m" "t"] - unfolding \u = t\<^sup>@k\ \v = t\<^sup>@m\ pow_add_list[of t k m] - using nemp_pow[OF \u \ \\[unfolded \u = t\<^sup>@k\]] nemp_pow[OF \v \ \\[unfolded \v = t\<^sup>@m\]] + unfolding \u = t\<^sup>@k\ \v = t\<^sup>@m\ pow_add_list[of t k m] + using nemp_pow[OF \u \ \\[unfolded \u = t\<^sup>@k\]] nemp_pow[OF \v \ \\[unfolded \v = t\<^sup>@m\]] by linarith qed lemma prim_rotate_conv: "primitive w \ primitive (rotate n w)" proof assume "primitive w" show "primitive (rotate n w)" proof (rule primI) - fix r k assume "r\<^sup>@k = rotate n w" + fix r k assume "r\<^sup>@k = rotate n w" obtain l where "(rotate l r)\<^sup>@k = w" using rotate_back[of n w, folded \r\<^sup>@k = rotate n w\, unfolded rotate_pow_comm] by blast from prim_exp_one[OF \primitive w\ this] show "k = 1". qed -next +next assume "primitive (rotate n w)" show "primitive w" proof (rule primI) fix r k assume "r\<^sup>@k = w" from prim_exp_one[OF \primitive (rotate n w)\, OF rotate_pow_comm[of n r k, unfolded this, symmetric]] show "k = 1". qed qed lemma non_prim: assumes "\ primitive w" and "w \ \" obtains r k where "r \ \" and "1 < k" and "r\<^sup>@k = w" and "w \ r" proof- from \\ primitive w\[unfolded primitive_def] obtain r k where "k \ 1" and "r\<^sup>@k = w" by blast have "r \ \" - using \w \ \\ \r\<^sup>@k = w\ emp_pow by blast + using \w \ \\ \r\<^sup>@k = w\ emp_pow by blast have "k \ 0" using \w \ \\ \r\<^sup>@k = w\ pow_zero[of r] by meson have "w \ r" using \k \ 1\[folded eq_pow_exp[OF \r \ \\, of k 1, unfolded \r \<^sup>@ k = w\]] by simp show thesis using that[OF \r \ \\ _ \r\<^sup>@k = w\ \w \ r\] \k \ 0\ \k \ 1\ less_linear by blast qed -lemma prim_no_rotate: assumes "primitive w" and "0 < n" and "n < \<^bold>|w\<^bold>|" +lemma prim_no_rotate: assumes "primitive w" and "0 < n" and "n < \<^bold>|w\<^bold>|" shows "rotate n w \ w" proof assume "rotate n w = w" have "take n w \ drop n w = drop n w \ take n w" - using rotate_append[of "take n w" "drop n w"] + using rotate_append[of "take n w" "drop n w"] unfolding take_len[OF less_imp_le_nat[OF \n < \<^bold>|w\<^bold>|\]] append_take_drop_id \rotate n w = w\. have "take n w \ \" "drop n w \ \" using \0 < n\ \n < \<^bold>|w\<^bold>|\ by auto+ from \primitive w\ show False using comm_non_prim[OF \take n w \ \\ \drop n w \ \\ \take n w \ drop n w = drop n w \ take n w\, unfolded append_take_drop_id] by simp qed lemma no_rotate_prim: assumes "w \ \" and "\ n. 0 < n \ n < \<^bold>|w\<^bold>| \ rotate n w \ w" shows "primitive w" proof (rule ccontr) assume "\ primitive w" from non_prim[OF this \w \ \\] obtain r l where "r \ \" and "1 < l" and "r\<^sup>@l = w" and "w \ r" by blast have "rotate \<^bold>|r\<^bold>| w = w" using rotate_root_self[of r l, unfolded \r\<^sup>@l = w\]. moreover have "0 < \<^bold>|r\<^bold>|" by (simp add: \r \ \\) - moreover have "\<^bold>|r\<^bold>| < \<^bold>|w\<^bold>|" + moreover have "\<^bold>|r\<^bold>| < \<^bold>|w\<^bold>|" unfolding pow_len[of r l, unfolded \r\<^sup>@l = w\] using \1 < l\ \0 < \<^bold>|r\<^bold>|\ by auto ultimately show False using assms(2) by blast qed -corollary prim_iff_rotate: assumes "w \ \" shows +corollary prim_iff_rotate: assumes "w \ \" shows "primitive w \ (\ n. 0 < n \ n < \<^bold>|w\<^bold>| \ rotate n w \ w)" using no_rotate_prim[OF \w \ \\] prim_no_rotate by blast -lemma prim_sing: "primitive [a]" +lemma prim_sing: "primitive [a]" using prim_iff_rotate[of "[a]"] by fastforce lemma prim_rev_iff[reversal_rule]: "primitive (rev u) \ primitive u" unfolding primitive_def[reversed] using primitive_def.. section \Primitive root\ text\Given a non-empty word $w$ which is not primitive, it is natural to look for the shortest $u$ such that $w = u^k$. Such a word is primitive, and it is the primitive root of $w$.\ definition primitive_rootP :: "'a list \ 'a list \ bool" ("_ \\<^sub>p _ *" [51,51] 60) where "primitive_rootP x r = (x \ \ \ x \ r* \ primitive r)" lemma prim_rootD [dest]: "x \\<^sub>p r* \ x \ r*" unfolding primitive_rootP_def by (elim conjE) lemma prim_rootI [intro]: "u \ \ \ u \ r* \ primitive r \ u \\<^sub>p r*" unfolding primitive_rootP_def by (intro conjI) lemma prim_root_rev_conv [reversal_rule]: "rev x \\<^sub>p rev r* \ x \\<^sub>p r*" unfolding primitive_rootP_def[reversed] using primitive_rootP_def.. -fun primitive_root :: "'a list \ 'a list" ("\") where "primitive_root x = (THE r. x \\<^sub>p r*)" +fun primitive_root :: "'a list \ 'a list" ("\") where "primitive_root x = (THE r. x \\<^sub>p r*)" lemma primrootE: assumes "x \\<^sub>p r*" obtains k where "k \ 0" and "r\<^sup>@k = x" using assms unfolding primitive_rootP_def root_def using nemp_pow[of r] by auto lemma primroot_of_root: "\ x \ \; x \ u*; u \\<^sub>p r*\ \ x \\<^sub>p r*" - unfolding primitive_rootP_def using root_trans by blast + unfolding primitive_rootP_def using root_trans by blast lemma comm_prim: assumes "primitive r" and "primitive s" and "r\s = s\r" shows "r = s" using \r\s = s\r\[unfolded comm] assms[unfolded primitive_def, rule_format] by metis lemma primroot_ex: assumes "x \ \" shows "\ r k. x \\<^sub>p r* \ k \ 0 \ x = r\<^sup>@k" using \x \ \\ proof(induction "\<^bold>|x\<^bold>|" arbitrary: x rule: less_induct) case less then show "\ r k. x \\<^sub>p r* \ k \ 0 \ x = r\<^sup>@k" proof (cases "primitive x") assume "\ primitive x" from non_prim[OF this \x \ \\] obtain r l where "r \ \" and "1 < l" and "r\<^sup>@l = x" and "x \ r" by blast then obtain pr k where "r \\<^sub>p pr*" "k \ 0" "r = pr\<^sup>@k" - using \x \ \\ less.hyps rootI root_shorter by blast + using \x \ \\ less.hyps rootI root_shorter by blast hence "x \\<^sub>p pr*" using \r \<^sup>@ l = x\ less.prems primroot_of_root rootI by blast show "\ r k. x \\<^sub>p r* \ k \ 0 \ x = r\<^sup>@k" using \x \\<^sub>p pr*\[unfolded primitive_rootP_def root_def] \x \\<^sub>p pr *\ nemp_pow by blast next assume "primitive x" have "x \\<^sub>p x*" - by (simp add: \primitive x\ less.prems prim_rootI self_root) + by (simp add: \primitive x\ less.prems prim_rootI self_root) thus "\ r k. x \\<^sub>p r* \ k \ 0 \ x = r\<^sup>@k" by force qed qed -lemma primroot_exE: assumes"x \ \" +lemma primroot_exE: assumes"x \ \" obtains r k where "primitive r" and "k \ 0" and "x = r\<^sup>@k" using assms primitive_rootP_def primroot_ex[OF \ x \ \\] by blast text\Uniqueness of the primitive root follows from the following lemma\ -lemma primroot_unique: assumes "u \\<^sub>p r*" shows "\ u = r" +lemma primroot_unique: assumes "u \\<^sub>p r*" shows "\ u = r" proof- obtain kr where "kr \ 0" and "r\<^sup>@kr = u" using primrootE[OF \u \\<^sub>p r*\]. have "u \\<^sub>p s* \ s = r" for s proof- fix s assume "u \\<^sub>p s*" obtain ks where "ks \ 0" and "s\<^sup>@ks = u" using primrootE[OF \u \\<^sub>p s*\]. obtain t where "s \ t*" and "r \ t*" - using comm_rootE[OF pow_comm_comm[of r kr s ks, OF _ \kr \ 0\, unfolded \r\<^sup>@kr = u\ \s\<^sup>@ks = u\, OF refl]]. + using comm_rootE[OF pow_comm_comm[of r kr s ks, OF _ \kr \ 0\, unfolded \r\<^sup>@kr = u\ \s\<^sup>@ks = u\, OF refl]]. have "primitive r" and "primitive s" using \u \\<^sub>p r *\ \u \\<^sub>p s *\ primitive_rootP_def by blast+ from prim_exp_eq[OF \primitive r\, of t] prim_exp_eq[OF \primitive s\, of t] show "s = r" using rootE[OF \s \ t*\, of "s=r"] rootE[OF \r \ t*\, of "r = t"] by fastforce qed from the_equality[of "\ r. u \\<^sub>p r*",OF \u \\<^sub>p r*\ this] show "\ u = r" - by auto + by auto qed lemma prim_self_root: "primitive x \ \ x = x" using prim_nemp prim_rootI primroot_unique self_root by blast text\Existence and uniqueness of the primitive root justifies the function @{term "\"}: it indeed yields the primitive root of a nonempty word.\ lemma primroot_is_primroot[intro]: assumes "x \ \" shows "x \\<^sub>p (\ x)*" using primroot_ex[OF \x \ \\] primroot_unique[of x] by force lemma primroot_is_root[intro]: "x \ \ \ x \ (\ x)*" using primroot_is_primroot by auto lemma primrootI [intro]: assumes "x \ \" shows primroot_prim: "primitive (\ x)" and primroot_nemp: "\ x \ \" using assms prim_nemp primitive_rootP_def by blast+ lemma primroot_root: assumes "u \ \" "u \ q*" shows "\ q = \ u" using primroot_unique[OF primroot_of_root[OF \u \ \\ \u \ q*\ primroot_is_primroot, OF root_nemp[OF \u \ \\ \u \ q*\]], symmetric]. lemma primroot_len_mult: assumes "u \ \" "u \ q*" obtains k where "\<^bold>|q\<^bold>| = k*\<^bold>|\ u\<^bold>|" using primroot_is_primroot[OF root_nemp[OF \u \ \\ \u \ q*\], unfolded primroot_root[OF \u \ \\ \u \ q*\] - primitive_rootP_def] root_len[of q "\ u"] by blast + primitive_rootP_def] root_len[of q "\ u"] by blast lemma primroot_shorter_root: assumes "u \ \" "u \ q*" shows "\<^bold>|\ u\<^bold>| \ \<^bold>|q\<^bold>|" - using quotient_smaller[OF root_nemp[OF \u \ \\ \u \ q*\, folded length_0_conv], of _ "\<^bold>|\ u\<^bold>|"] + using quotient_smaller[OF root_nemp[OF \u \ \\ \u \ q*\, folded length_0_conv], of _ "\<^bold>|\ u\<^bold>|"] primroot_len_mult[OF \u \ \\ \u \ q*\] by blast lemma primroot_shortest_root: assumes "u \ \" shows "\<^bold>|\ u\<^bold>| = (LEAST d. (\ r. (u \ r*) \ \<^bold>|r\<^bold>| = d))" using Least_equality[of "\ k. (\ r. (u \ r*) \ \<^bold>|r\<^bold>| = k)" "\<^bold>|\ u\<^bold>|"] proof show "\r. u \ r* \ \<^bold>|r\<^bold>| = \<^bold>|\ u\<^bold>|" - using assms primitive_rootP_def primroot_is_primroot by blast + using assms primitive_rootP_def primroot_is_primroot by blast show "\y. \r. u \ r* \ \<^bold>|r\<^bold>| = y \ \<^bold>|\ u\<^bold>| \ y" using assms primroot_shorter_root by fastforce qed lemma primroot_shorter_eq: "u \ \ \ \<^bold>|\ u\<^bold>| \ \<^bold>|u\<^bold>|" using primroot_shorter_root self_root by auto lemma primroot_take: assumes "u \ \" shows "\ u = (take ( \<^bold>|\ u\<^bold>| ) u)" proof- obtain k where "(\ u)\<^sup>@k = u" and "k \ 0" using primrootE[OF primroot_is_primroot[OF \u \ \\]]. show "\ u = (take ( \<^bold>|\ u\<^bold>| ) u)" using take_root[of _ "(\ u)", OF \k \ 0\, unfolded \(\ u)\<^sup>@k = u\]. qed lemma primroot_take_shortest: assumes "u \ \" shows "\ u = (take (LEAST d. (\ r. (u \ r*) \ \<^bold>|r\<^bold>| = d)) u)" using primroot_take[OF assms, unfolded primroot_shortest_root[OF assms]]. lemma primroot_rotate_comm: assumes "w \ \" shows "\ (rotate n w) = rotate n (\ w)" proof- obtain l where "w = (\ w)\<^sup>@l" using pow_zero primrootE primroot_is_primroot by metis hence "rotate n w \ (rotate n (\ w))*" using rotate_pow_comm root_def by metis moreover have "rotate n w \ \" using assms by auto moreover have "primitive (rotate n (\ w))" using assms prim_rotate_conv primitive_rootP_def primroot_is_primroot by blast ultimately have "rotate n w \\<^sub>p (rotate n (\ w))*" unfolding primitive_rootP_def by blast thus ?thesis using primroot_unique by blast qed lemma primrootI1 [intro]: assumes pow: "u = r\<^sup>@(Suc k)" and prim: "primitive r" shows "\ u = r" proof- have "u \ \" - using pow prim prim_nemp by auto + using pow prim prim_nemp by auto have "u \ r*" using pow rootI by blast show "\ u = r" using primroot_unique[OF prim_rootI[OF \u \ \\ \u \ r*\ \primitive r\]]. qed lemma prim_root_power [elim]: assumes "x \ \" obtains i where "(\ x)\<^sup>@(Suc i) = x" using prim_rootD[OF primroot_is_primroot[OF \x \ \\], unfolded root_def] assms pow_zero[of "\ x"] not0_implies_Suc by metis lemma prim_root_cases: obtains "u = \" | "primitive u" | "\<^bold>|\ u\<^bold>| < \<^bold>|u\<^bold>|" - using primroot_is_primroot[THEN prim_rootD[of u "\ u"]] + using primroot_is_primroot[THEN prim_rootD[of u "\ u"]] primroot_prim[of u] root_shorter[of u "\ u"] by fastforce text\We also have the standard characterization of commutation for nonempty words.\ theorem comm_primroots: assumes "u \ \" "v \ \" shows "u \ v = v \ u \ \ u = \ v" proof assume "u \ v = v \ u" then obtain t where "u \ t*" and "v \ t*" using comm_root by blast show "\ u = \ v" - using primroot_root[OF \v \ \\ \v \ t*\, unfolded primroot_root[OF \u \ \\ \u \ t*\]]. + using primroot_root[OF \v \ \\ \v \ t*\, unfolded primroot_root[OF \u \ \\ \u \ t*\]]. next assume "\ u = \ v" show "u \ v = v \ u" using primroot_is_primroot[OF \u \ \\, unfolded \\ u = \ v\] primroot_is_primroot[OF \v \ \\] unfolding primitive_rootP_def comm_root by blast qed lemma prim_comm_short_emp: assumes "primitive p" and "u\p=p\u" and "\<^bold>|u\<^bold>| < \<^bold>|p\<^bold>|" shows "u = \" proof (rule ccontr) assume "u \ \" from \u \ p = p \ u\ have "\ u = \ p" unfolding comm_primroots[OF \u \ \\ prim_nemp, OF \primitive p\]. have "\ u = p" using prim_self_root[OF \primitive p\, folded \\ u = \ p\]. from \\<^bold>|u\<^bold>| < \<^bold>|p\<^bold>|\[folded this] show False using primroot_shorter_eq[OF \u \ \\] by auto qed section \Conjugation\ text\Two words $x$ and $y$ are conjugated if one is a rotation of the other. Or, equivalently, there exists $z$ such that \[ xz = zy. \] \ definition conjugate ("_ \ _" [50,50] 51) where "u \ v \ \r s. r \ s = u \ s \ r = v" lemma conjug_rev_conv [reversal_rule]: "rev u \ rev v \ u \ v" unfolding conjugate_def[reversed] using conjugate_def by blast lemma conjug_rotate_iff: "u \ v \ (\ n. v = rotate n u)" unfolding conjugate_def using rotate_drop_take[of _ u] takedrop[of _ u] rotate_append by metis lemma conjugI [intro]: "r \ s = u \ s \ r = v \ u \ v" unfolding conjugate_def by (intro exI conjI) lemma conjugI' [intro!]: "r \ s \ s \ r" unfolding conjugate_def by (intro exI conjI, standard+) -lemma conjugE [elim]: +lemma conjugE [elim]: assumes "u \ v" obtains r s where "r \ s = u" and "s \ r = v" using assms unfolding conjugate_def by (elim exE conjE) lemma conjugE1 [elim]: assumes "u \ v" obtains r where "u \ r = r \ v" proof - obtain r s where u: "r \ s = u" and v: "s \ r = v" using assms.. have "u \ r = r \ v" unfolding u[symmetric] v[symmetric] using rassoc. then show thesis by fact qed lemma conjug_refl: "u \ u" by standard+ lemma conjug_sym [sym]: "u \ v \ v \ u" by (elim conjugE, intro conjugI) assumption+ lemma conjug_nemp_iff: "u \ v \ u = \ \ v = \" by (elim conjugE1, intro iffI) simp+ lemma conjug_len: "u \ v \ \<^bold>|u\<^bold>| = \<^bold>|v\<^bold>|" by (elim conjugE, hypsubst, rule swap_len) lemma pow_conjug: assumes eq: "t\<^sup>@i \ r \ u = t\<^sup>@k" and t: "r \ s = t" shows "u \ t\<^sup>@i \ r = (s \ r)\<^sup>@k" proof - - have "t\<^sup>@i \ r \ u \ t\<^sup>@i \ r = t\<^sup>@i \ t\<^sup>@k \ r" unfolding eq[unfolded lassoc] lassoc append_same_eq pow_comm.. + have "t\<^sup>@i \ r \ u \ t\<^sup>@i \ r = t\<^sup>@i \ t\<^sup>@k \ r" unfolding eq[unfolded lassoc] lassoc append_same_eq pow_comm.. also have "\ = t\<^sup>@i \ r \ (s \ r)\<^sup>@k" unfolding conjug_pow[OF rassoc, symmetric] t.. - finally show "u \ t\<^sup>@i \ r = (s \ r)\<^sup>@k" unfolding same_append_eq. + finally show "u \ t\<^sup>@i \ r = (s \ r)\<^sup>@k" unfolding same_append_eq. qed -text\The solution of the equation +text\The solution of the equation \[ xz = zy \] is given by the next lemma. \ lemma conjug_eqE [elim, consumes 2]: assumes eq: "x \ z = z \ y" and "x \ \" obtains u v k where "u \ v = x" and "v \ u = y" and "(u \ v)\<^sup>@k \ u = z" and "v \ \" proof - have "z \p x \ z" using eq[symmetric].. from this \x \ \\ have "z \p x\<^sup>\".. then obtain k u where "x\<^sup>@k \ u = z" and "u

u

obtain v where x: "u \ v = x" and "v \ \".. have z: "(u\v)\<^sup>@k \ u = z" unfolding x \x\<^sup>@k \ u = z\.. have "z \ y = (u\v) \ ((u\v)\<^sup>@k \ u)" unfolding z unfolding x eq.. also have "\ = (u\v)\<^sup>@k \ u \ (v \ u)" unfolding lassoc pow_commutes_list[symmetric].. finally have y: "v \ u = y" unfolding z[symmetric] rassoc same_append_eq.. from x y z \v \ \\ show thesis.. qed -theorem conjugation: assumes "x\z = z\y" and "x \ \" +theorem conjugation: assumes "x\z = z\y" and "x \ \" shows "\ u v k. u \ v = x \ v \ u = y \ (u \ v)\<^sup>@k \ u = z" using assms by blast lemma conjug_eq_prim_rootE [elim, consumes 2]: assumes eq: "x \ z = z \ y" and "x \ \" obtains r s i n where "(r \ s)\<^sup>@Suc i = x" and - "(s \ r)\<^sup>@Suc i = y" and + "(s \ r)\<^sup>@Suc i = y" and "(r \ s)\<^sup>@n \ r = z" and "s \ \" and "primitive (r \ s)" proof - from \x \ \\ obtain i where "(\ x)\<^sup>@(Suc i) = x".. also have "z \p x\<^sup>\" using prefI[OF \x \ z = z \ y\[symmetric]] \x \ \\.. finally have "z \p (\ x)\<^sup>\" by (elim per_drop_exp) then obtain n r where "(\ x)\<^sup>@n \ r = z" and "r

x".. from \r

x\ obtain s where "r \ s = \ x" and "s \ \".. define j where "j = Suc i" have x: "(r\s)\<^sup>@j = x" unfolding \r \ s = \ x\ \j = Suc i\ \(\ x)\<^sup>@(Suc i) = x\.. have z: "(r\s)\<^sup>@n \ r = z" unfolding \r \ s = \ x\ \(\ x)\<^sup>@n \ r = z\.. have "z \ y = ((r\s)\<^sup>@j \ (r\s)\<^sup>@n) \ r" using eq[symmetric] unfolding rassoc x z. also have "(r\s)\<^sup>@j \ (r\s)\<^sup>@n = (r\s)\<^sup>@n \ (r\s)\<^sup>@j" using pow_comm. also have "\ \ r = (r\s)\<^sup>@n \ r \ (s\r)\<^sup>@j" unfolding rassoc unfolding shift_pow.. - finally have y: "y = (s\r)\<^sup>@j" unfolding z[symmetric] rassoc cancel. + finally have y: "y = (s\r)\<^sup>@j" unfolding z[symmetric] rassoc cancel. from \x \ \\ have "primitive (r \ s)" unfolding \r \ s = \ x\.. - with that x y z \s \ \\ show thesis unfolding \j = Suc i\ by blast + with that x y z \s \ \\ show thesis unfolding \j = Suc i\ by blast qed lemma conjug_eq_prim_root: assumes "x \ z = z \ y" and "x \ \" shows "\ r s i n. (r \ s)\<^sup>@(Suc i) = x \ (s \ r)\<^sup>@(Suc i) = y \ (r \ s)\<^sup>@n \ r = z \ s \ \ \ primitive (r \ s)" using conjug_eq_prim_rootE[OF assms, of ?thesis] by blast lemma conjugI1 [intro]: assumes eq: "u \ r = r \ v" shows "u \ v" proof (cases) assume "u = \" have "v = \" using eq unfolding \u = \\ by simp show "u \ v" unfolding \u = \\ \v = \\ using conjug_refl. next assume "u \ \" show "u \ v" using eq \u \ \\ by (cases rule: conjug_eqE, intro conjugI) qed lemma conjug_trans [trans]: assumes uv: "u \ v" and vw: "v \ w" shows "u \ w" using assms unfolding conjug_rotate_iff using rotate_rotate by blast lemma conjug_trans': assumes uv': "u \ r = r \ v" and vw': "v \ s = s \ w" shows "u \ (r \ s) = (r \ s) \ w" proof - have "u \ (r \ s) = (r \ v) \ s" unfolding uv'[symmetric] rassoc.. also have "\ = r \ (s \ w)" unfolding vw'[symmetric] rassoc.. finally show "u \ (r \ s) = (r \ s) \ w" unfolding rassoc. qed lemma nconjug_neq: "\ u \ v \ u \ v" by blast lemma prim_conjug: assumes prim: "primitive u" and conjug: "u \ v" shows "primitive v" proof - have "v \ \" using prim_nemp[OF prim] unfolding conjug_nemp_iff[OF conjug]. from conjug[symmetric] obtain t where "v \ t = t \ u".. from this \v \ \\ obtain r s i where v: "(r \ s)\<^sup>@(Suc i) = v" and u: "(s \ r)\<^sup>@(Suc i) = u" and prim': "primitive (r \ s)".. have "r \ s = v" using v unfolding prim_exp_one[OF prim u] pow_one_id. show "primitive v" using prim' unfolding \r \ s = v\. qed lemma root_conjug: "u \p r \ u \ u\\<^sup>>(r\u) \ r" using conjugI1 conjug_sym lq_pref by metis lemma conjug_prim_root: assumes conjug: "u \ v" and "u \ \" shows "\ u \ \ v" proof - from conjug obtain t where "u \ t = t \ v".. from this \u \ \\ obtain r s i where u: "(r \ s)\<^sup>@(Suc i) = u" and v: "(s \ r)\<^sup>@(Suc i) = v" and prim: "primitive (r \ s)".. have rs: "\ u = r \ s" and sr: "\ v = s \ r" using prim prim_conjug u v by blast+ show "\ u \ \ v" using conjugI' unfolding rs sr. qed lemma conjug_add_exp: "u \ v \ u\<^sup>@k \ v\<^sup>@k" by (elim conjugE1, intro conjugI1, rule conjug_pow) lemma conjug_prim_root_iff: assumes nemp:"u \ \" and len: "\<^bold>|u\<^bold>| = \<^bold>|v\<^bold>|" shows "\ u \ \ v \ u \ v" proof show "u \ v \ \ u \ \ v" using conjug_prim_root[OF _ nemp]. assume conjug: "\ u \ \ v" have "v \ \" using nemp_len[OF nemp] unfolding len length_0_conv. with nemp obtain k l where roots: "(\ u)\<^sup>@k = u" "(\ v)\<^sup>@l = v" by (elim prim_root_power) have "\<^bold>|(\ u)\<^sup>@k\<^bold>| = \<^bold>|(\ v)\<^sup>@l\<^bold>|" using len unfolding roots. then have "k = l" using primroot_nemp[OF \v \ \\] unfolding pow_len conjug_len[OF conjug] by simp show "u \ v" using conjug_add_exp[OF conjug, of l] unfolding roots[unfolded \k = l\]. qed lemma fac_pow_pref_conjug: assumes "u \f t\<^sup>@k" obtains t' where "t \ t'" and "u \p t'\<^sup>@k" proof (cases "u = \") assume "u \ \" obtain p q where eq: "p \ u \ q = t\<^sup>@k" using facE'[OF assms]. obtain i r where "i < k" and "r

@i \ r = p" using pref_mod_power[OF sprefI1'[OF eq pref_nemp[OF \u \ \\]]]. from \r

obtain s where t: "r \ s = t".. have eq': "t\<^sup>@i \ r \ (u \ q) = t\<^sup>@k" using eq unfolding lassoc p. have "u \p (s \ r)\<^sup>@k" using pow_conjug[OF eq' t] unfolding rassoc.. with conjugI'[of r s] show thesis unfolding t.. qed blast lemma fac_pow_len_conjug: assumes "\<^bold>|u\<^bold>| = \<^bold>|v\<^bold>|" and "u \f v\<^sup>@k" shows "v \ u" proof- - obtain t where "v \ t" and "u \p t\<^sup>@k" + obtain t where "v \ t" and "u \p t\<^sup>@k" using fac_pow_pref_conjug assms by blast have "u = t" using pref_equal[OF pref_prod_root[OF \u \p t\<^sup>@k\] conjug_len[OF \v \ t\,folded \\<^bold>|u\<^bold>| = \<^bold>|v\<^bold>|\]]. from \v \ t\[folded this] show "v \ u". qed lemma border_conjug: "x \b w \ w\<^sup><\x \ x\\<^sup>>w" using border_conjug_eq conjugI1 by blast lemmas fac_pow_suf_conjug = fac_pow_pref_conjug[reversed] - + end \ No newline at end of file diff --git a/thys/Combinatorics_Words/Lyndon_Schutzenberger.thy b/thys/Combinatorics_Words/Lyndon_Schutzenberger.thy --- a/thys/Combinatorics_Words/Lyndon_Schutzenberger.thy +++ b/thys/Combinatorics_Words/Lyndon_Schutzenberger.thy @@ -1,333 +1,333 @@ (* Title: CoW/Lyndon_Schutzenberger.thy Author: Štěpán Holub, Charles University *) theory Lyndon_Schutzenberger - imports CoWBasic + imports CoWBasic begin chapter \Lyndon-Schützenberger Equation\ text\The Lyndon-Schützenberger equation is the following equation on words: \[ x^ay^b = z^c, \] in this formalization denoted as @{term "x\<^sup>@a\y\<^sup>@b = z\<^sup>@c"}, with $2 \leq a,b,c$. -We formalize here a proof that the equation has periodic solutions only in free monoids, that is, that any three words +We formalize here a proof that the equation has periodic solutions only in free monoids, that is, that any three words $x$, $y$ and $z$ satisfying the equality pairwise commute. The result was first proved in @{cite LySch62} in a more general setting of free groups. There are several proofs to be found in the literature (for instance @{cite Lo83 and Dmsi2006}). The presented proof is the author's proof. \ text\We set up a locale representing the Lyndon-Schützenberger Equation.\ locale LS = fixes x a y b z c - assumes a: "2 \ a" and b: "2 \ b" and c: "2 \ c" and eq: "x\<^sup>@a \ y\<^sup>@b = z\<^sup>@c" + assumes a: "2 \ a" and b: "2 \ b" and c: "2 \ c" and eq: "x\<^sup>@a \ y\<^sup>@b = z\<^sup>@c" begin -lemma a0: "a \ 0" and b0: "b \ 0" +lemma a0: "a \ 0" and b0: "b \ 0" using a b by auto text\If $x^a$ or $y^b$ is sufficiently long, then the calim follows from the Periodicity Lemma.\ lemma per_lemma_case: assumes "\<^bold>|z\<^bold>| + \<^bold>|x\<^bold>| \ \<^bold>|x\<^sup>@a\<^bold>|" shows "x\y=y\x" proof (cases "x = \") assume "x = \" - thus "x\y=y\x" by simp -next + thus "x\y=y\x" by simp +next assume "x \ \" - hence "z\<^sup>@c \ \" + hence "z\<^sup>@c \ \" using eq assms emp_pow[of c] by auto hence "x\<^sup>@a \p (z\<^sup>@c)\<^sup>\" - unfolding period_root_def using - pref_ext[OF triv_pref[of "x\<^sup>@a" "y\<^sup>@b", unfolded eq], of "x\<^sup>@a"] by blast + unfolding period_root_def using + pref_ext[OF triv_pref[of "x\<^sup>@a" "y\<^sup>@b", unfolded eq], of "x\<^sup>@a"] by blast have "x \<^sup>@ a \p x\<^sup>\" using \x \ \\ a0 root_self[THEN per_drop_exp] by blast from two_pers[OF per_drop_exp[OF \x\<^sup>@a \p (z\<^sup>@c)\<^sup>\\] this assms] have "z \ x = x \ z". hence "z\<^sup>@c\x\<^sup>@a = x\<^sup>@a\z\<^sup>@c" by (simp add: comm_add_exps) from this[folded eq, unfolded rassoc cancel, symmetric] have "x\<^sup>@a \ y\<^sup>@b = y\<^sup>@b \ x\<^sup>@a". from this[unfolded comm_pow_roots[OF a0 b0]] show "x \ y = y \ x". qed text\The most challenging case is when $c = 3$.\ lemma core_case: - assumes - "c = 3" and - "b*\<^bold>|y\<^bold>| \ a*\<^bold>|x\<^bold>|" and "x \ \" and "y \ \" and + assumes + "c = 3" and + "b*\<^bold>|y\<^bold>| \ a*\<^bold>|x\<^bold>|" and "x \ \" and "y \ \" and lenx: "a*\<^bold>|x\<^bold>| < \<^bold>|z\<^bold>| + \<^bold>|x\<^bold>|" and leny: "b*\<^bold>|y\<^bold>| < \<^bold>|z\<^bold>| + \<^bold>|y\<^bold>|" - shows "x\y = y\x" -proof- - have "a \ 0" and "b \ 0" + shows "x\y = y\x" +proof- + have "a \ 0" and "b \ 0" using a b by auto \\We first show that a = 2\ have "a*\<^bold>|x\<^bold>|+b*\<^bold>|y\<^bold>| = 3*\<^bold>|z\<^bold>|" using \c = 3\ eq length_append[of "x\<^sup>@a" "y\<^sup>@b"] by (simp add: pow_len) hence "3*\<^bold>|z\<^bold>| \ a*\<^bold>|x\<^bold>| + a*\<^bold>|x\<^bold>|" using \b*\<^bold>|y\<^bold>| \ a*\<^bold>|x\<^bold>|\ by simp hence "3*\<^bold>|z\<^bold>| < 2*\<^bold>|z\<^bold>| + 2*\<^bold>|x\<^bold>|" using lenx by linarith hence "\<^bold>|z\<^bold>| + \<^bold>|x\<^bold>| < 3 * \<^bold>|x\<^bold>|" by simp from less_trans[OF lenx this, unfolded mult_less_cancel2] have "a = 2" using a by force - hence "\<^bold>|y\<^bold>| \ \<^bold>|x\<^bold>|" using \b*\<^bold>|y\<^bold>| \ a*\<^bold>|x\<^bold>|\ b + hence "\<^bold>|y\<^bold>| \ \<^bold>|x\<^bold>|" using \b*\<^bold>|y\<^bold>| \ a*\<^bold>|x\<^bold>|\ b pow_len[of x 2] pow_len[of y b] mult_le_less_imp_less[of a b "\<^bold>|x\<^bold>|" "\<^bold>|y\<^bold>|"] not_le by auto have "x\x\y\<^sup>@b = z\z\z" using a eq \c=3\ \a=2\ - by (simp add: numeral_2_eq_2 numeral_3_eq_3) + by (simp add: numeral_2_eq_2 numeral_3_eq_3) \ \Find words u, v, w\ have "\<^bold>|z\<^bold>| < \<^bold>|x\x\<^bold>|" - using \\<^bold>|z\<^bold>| + \<^bold>|x\<^bold>| < 3 * \<^bold>|x\<^bold>|\ add.commute by auto + using \\<^bold>|z\<^bold>| + \<^bold>|x\<^bold>| < 3 * \<^bold>|x\<^bold>|\ add.commute by auto from ruler_le[THEN prefD, OF triv_pref[of z "z\z"] _ less_imp_le[OF this]] - obtain w where "z\w = x\x" - using prefI[of "x\x" "y\<^sup>@b" "z\z\z", unfolded rassoc, OF \x\x\y\<^sup>@b = z\z\z\] by fastforce + obtain w where "z\w = x\x" + using prefI[of "x\x" "y\<^sup>@b" "z\z\z", unfolded rassoc, OF \x\x\y\<^sup>@b = z\z\z\] by fastforce have "\<^bold>|x\<^bold>| < \<^bold>|z\<^bold>|" using \a = 2\ lenx by auto from ruler_le[THEN prefD, OF _ _ less_imp_le[OF this], of "x\x\y\<^sup>@b", OF triv_pref, unfolded \x\x\y\<^sup>@b = z\z\z\, OF triv_pref] - obtain u :: "'a list" where "x\u=z" + obtain u :: "'a list" where "x\u=z" by blast have "u \ \" - using \\<^bold>|x\<^bold>| < \<^bold>|z\<^bold>|\ \x\u = z\ by auto + using \\<^bold>|x\<^bold>| < \<^bold>|z\<^bold>|\ \x\u = z\ by auto have "x = u\w" using \z\w = x\x\ \x\u = z\ by auto have "\<^bold>|x\x\<^bold>| < \<^bold>|z\z\<^bold>|" by (simp add: \\<^bold>|x\<^bold>| < \<^bold>|z\<^bold>|\ add_less_mono) from ruler_le[OF triv_pref[of "x\x" "y\<^sup>@b", unfolded rassoc \x\x\y\<^sup>@b = z\z\z\, unfolded lassoc] triv_pref, OF less_imp_le[OF this]] have "z\w \p z\z" unfolding \z\w = x\x\. obtain v :: "'a list" where "w \ v = x" - using lq_pref[of w x] - pref_prod_pref'[OF pref_cancel[OF \z\w \p z\z\, folded \x \ u = z\, unfolded \x = u \ w\ rassoc], folded \x = u \ w\] by blast + using lq_pref[of w x] + pref_prod_pref'[OF pref_cancel[OF \z\w \p z\z\, folded \x \ u = z\, unfolded \x = u \ w\ rassoc], folded \x = u \ w\] by blast have "u\w\v \ \" by (simp add: \u \ \\) \ \Express x, y and z in terms of u, v and w\ hence "z = w\v\u" using \w \ v = x\ \x \ u = z\ by auto from \x \ x \ y\<^sup>@b = z \ z \ z\[unfolded this lassoc, folded \z \ w = x \ x\, unfolded this rassoc] have "w\v \ u\w \ y\<^sup>@b = w\v\u\w\v\u\w\v\u". hence "y\<^sup>@b = v\u\w\v\u" using pref_cancel by auto \ \Double period of uwv\ from periodN_fac[OF _ \u\w\v \ \\, of v u "\<^bold>|y\<^bold>|", unfolded rassoc, folded this] have "periodN (u\w\v) \<^bold>|y\<^bold>|" using pow_per[OF \y \ \\ b0] by blast have "u\w\v = x \v" - by (simp add: \x = u \ w\) + by (simp add: \x = u \ w\) have "u\w\v = u\ x" by (simp add: \w \ v = x\) have "u\w\v \p u\<^sup>\" unfolding period_root_def using \u \ w \ v = u \ x\[unfolded \x = u \ w\] \u \ \\ triv_pref[of "u \ u \ w" v] by force have "periodN (u\w\v) \<^bold>|u\<^bold>|" using \u \ w \ v \p u \<^sup>\\ by auto \ \Common period d\ obtain d::nat where "d=gcd \<^bold>|y\<^bold>| \<^bold>|u\<^bold>|" by simp have "\<^bold>|y\<^bold>| + \<^bold>|u\<^bold>| \ \<^bold>|u\w\v\<^bold>|" using \\<^bold>|y\<^bold>| \ \<^bold>|x\<^bold>|\ length_append \u\w\v = u\ x\ by simp - hence "periodN (u\w\v) d" + hence "periodN (u\w\v) d" using \periodN (u \ w \ v) \<^bold>|u\<^bold>|\ \periodN (u \ w \ v) \<^bold>|y\<^bold>|\ \d = gcd \<^bold>|y\<^bold>| \<^bold>|u\<^bold>|\ two_periodsN by blast \ \Divisibility\ have "v\u\z=y\<^sup>@b" by (simp add: \y\<^sup>@b = v \ u \ w \ v \ u\ \z = w \ v \ u\) have "\<^bold>|u\<^bold>| = \<^bold>|v\<^bold>|" using \x = u \ w\ \w \ v = x\ length_append[of u w] length_append[of w v] add.commute[of "\<^bold>|u\<^bold>|" "\<^bold>|w\<^bold>|"] add_left_cancel by simp hence "d dvd \<^bold>|v\<^bold>|" using gcd_nat.cobounded1[of "\<^bold>|v\<^bold>|" "\<^bold>|y\<^bold>|"] gcd.commute[of "\<^bold>|y\<^bold>|" "\<^bold>|u\<^bold>|"] by (simp add: \d = gcd \<^bold>|y\<^bold>| \<^bold>|u\<^bold>|\) have "d dvd \<^bold>|u\<^bold>|" by (simp add: \d = gcd \<^bold>|y\<^bold>| \<^bold>|u\<^bold>|\) have "\<^bold>|z\<^bold>| + \<^bold>|u\<^bold>| + \<^bold>|v\<^bold>| = b*\<^bold>|y\<^bold>|" using lenarg[OF \v\u\z=y\<^sup>@b\, unfolded length_append pow_len] by auto from dvd_add_left_iff[OF \d dvd \<^bold>|v\<^bold>|\, of "\<^bold>|z\<^bold>|+\<^bold>|u\<^bold>|", unfolded this dvd_add_left_iff[OF \d dvd \<^bold>|u\<^bold>|\, of "\<^bold>|z\<^bold>|"]] - have "d dvd \<^bold>|z\<^bold>|" - using \d = gcd \<^bold>|y\<^bold>| \<^bold>|u\<^bold>|\ dvd_mult by blast + have "d dvd \<^bold>|z\<^bold>|" + using \d = gcd \<^bold>|y\<^bold>| \<^bold>|u\<^bold>|\ dvd_mult by blast from lenarg[OF \z = w \ v \ u\, unfolded length_append pow_len] have "d dvd \<^bold>|w\<^bold>|" using \d dvd \<^bold>|z\<^bold>|\ \d dvd \<^bold>|u\<^bold>|\ \d dvd \<^bold>|v\<^bold>|\ by (simp add: dvd_add_left_iff) hence "d dvd \<^bold>|x\<^bold>|" using \d dvd \<^bold>|v\<^bold>|\ \w \ v = x\ by force \ \x and y commute\ have "x \p u\w\v" - by (simp add: \x = u \ w\) + by (simp add: \x = u \ w\) have "periodN x d" using per_pref'[OF \x\\\ \periodN (u\w\v) d \ \x \p u \w\v\]. hence "x \ (take d x)*" using \d dvd \<^bold>|x\<^bold>|\ - using root_divisor by blast + using root_divisor by blast hence "periodN u d " using \x = u \ w\ per_pref' using \periodN x d\ \u \ \\ by blast have " take d x = take d u" using \u\\\ \x = u \ w\ pref_share_take by (simp add: \d = gcd \<^bold>|y\<^bold>| \<^bold>|u\<^bold>|\) from root_divisor[OF \periodN u d\ \d dvd \<^bold>|u\<^bold>|\, folded this] have "u \ (take d x)*". - hence "z \ (take d x)*" + hence "z \ (take d x)*" using \x\u=z\ \x \ (take d x)*\ add_roots by blast from root_pref_cancel'[OF _ root_pow_root[OF \x \ take d x*\, of a],of "y\<^sup>@b", unfolded eq, OF root_pow_root[OF this, of c]] - have "y\<^sup>@b \ (take d x)*". - from commD[OF root_pow_root[OF \x \ take d x*\, of a] this] - show "x \ y = y \ x" + have "y\<^sup>@b \ (take d x)*". + from commD[OF root_pow_root[OF \x \ take d x*\, of a] this] + show "x \ y = y \ x" unfolding comm_pow_roots[OF \a \ 0\ \b \ 0\, of x y]. qed end \ \locale LS\ -text\The main proof is by induction on the length of $z$. It also uses the reverse symmetry of the equation which is -exploited by two interpretations of the locale @{term LS}. Note also that the case $|x^a| < |y^b|$ is solved by +text\The main proof is by induction on the length of $z$. It also uses the reverse symmetry of the equation which is +exploited by two interpretations of the locale @{term LS}. Note also that the case $|x^a| < |y^b|$ is solved by using induction on $|z| + |y^b|$ instead of just on $|z|$. \ lemma Lyndon_Schutzenberger': "\ x\<^sup>@a\y\<^sup>@b = z\<^sup>@c; 2 \ a; 2 \ b; 2 \ c \ \ x\y = y\x" proof (induction "\<^bold>|z\<^bold>| + b* \<^bold>|y\<^bold>|" arbitrary: x y z a b c rule: less_induct) case less interpret LS x a y b z c using \ x\<^sup>@a\y\<^sup>@b = z\<^sup>@c\ \ 2 \a \ \ 2 \ b\ \ 2 \ c\ by (simp add: LS.intro) - interpret LSrev: LS "rev y" b "rev x" a "rev z" c + interpret LSrev: LS "rev y" b "rev x" a "rev z" c using LS.intro[OF b a c, of "rev y" "rev x" "rev z", folded rev_append rev_pow, unfolded rev_is_rev_conv, OF \x\<^sup>@a \ y\<^sup>@b = z\<^sup>@c\]. - have leneq: "a * \<^bold>|x\<^bold>| + b*\<^bold>|y\<^bold>| = c * \<^bold>|z\<^bold>|" + have leneq: "a * \<^bold>|x\<^bold>| + b*\<^bold>|y\<^bold>| = c * \<^bold>|z\<^bold>|" using eq unfolding pow_len[symmetric] length_append[symmetric] by simp - have "a \ 0" and "b \ 0" + have "a \ 0" and "b \ 0" using a b by auto show "x \ y = y \ x" - proof (cases "x = \ \ y = \") + proof (cases "x = \ \ y = \") show "x = \ \ y = \ \ x \ y = y \ x" by auto next assume "\ (x = \ \ y = \)" hence "x \ \" and "y \ \" by blast+ show "x \ y = y \ x" proof (cases "\<^bold>|x \<^sup>@ a\<^bold>| < \<^bold>|y \<^sup>@ b\<^bold>|") \ \WLOG assumption\ assume "\<^bold>|x\<^sup>@a\<^bold>| < \<^bold>|y\<^sup>@b\<^bold>|" have "\<^bold>|rev z\<^bold>| + a* \<^bold>|rev x\<^bold>| < \<^bold>|z\<^bold>| + b* \<^bold>|y\<^bold>|" using \\<^bold>|x\<^sup>@a\<^bold>| < \<^bold>|y\<^sup>@b\<^bold>|\ by (simp add: pow_len) from "less.hyps"[OF this LSrev.eq b a c, symmetric] show "x \ y = y \ x" - unfolding rev_append[symmetric] rev_is_rev_conv by simp + unfolding rev_append[symmetric] rev_is_rev_conv by simp next assume " \ \<^bold>|x\<^sup>@a\<^bold>| < \<^bold>|y\<^sup>@b\<^bold>|" hence "\<^bold>|y\<^sup>@b\<^bold>| \ \<^bold>|x\<^sup>@a\<^bold>|" by force \ \case solved by the Periodicity lemma\ consider (with_Periodicity_lemma) "\<^bold>|z\<^bold>| + \<^bold>|x\<^bold>| \ \<^bold>|x \<^sup>@ a\<^bold>| \ \<^bold>|z\<^bold>| + \<^bold>|y\<^bold>| \ \<^bold>|y \<^sup>@ b\<^bold>|" | - (without_Periodicity_lemma) + (without_Periodicity_lemma) "\<^bold>|x\<^sup>@a\<^bold>| < \<^bold>|z\<^bold>| + \<^bold>|x\<^bold>|" and "\<^bold>|y\<^sup>@b\<^bold>| < \<^bold>|z\<^bold>| + \<^bold>|y\<^bold>|" - using not_le_imp_less by blast + using not_le_imp_less by blast thus "x \ y = y \ x" proof (cases) case with_Periodicity_lemma assume short: "\<^bold>|z\<^bold>| + \<^bold>|x\<^bold>| \ \<^bold>|x\<^sup>@a\<^bold>| \ \<^bold>|z\<^bold>| + \<^bold>|y\<^bold>| \ \<^bold>|y\<^sup>@b\<^bold>|" have "x = \ \ rev y = \ \ x \ y = y \ x" by auto thus "x \ y = y \ x" using per_lemma_case LSrev.per_lemma_case short unfolding length_rev rev_append[symmetric] rev_is_rev_conv rev_pow[symmetric] by blast next case without_Periodicity_lemma assume lenx: "\<^bold>|x\<^sup>@a\<^bold>| < \<^bold>|z\<^bold>| + \<^bold>|x\<^bold>|" and leny: "\<^bold>|y\<^sup>@b\<^bold>| < \<^bold>|z\<^bold>| + \<^bold>|y\<^bold>|" have ex: "\ k. d = Suc (Suc k)" if "2 \ d" for d - using nat_le_iff_add that by auto - have "a * \<^bold>|x\<^bold>| + b*\<^bold>|y\<^bold>| < 4 * \<^bold>|z\<^bold>|" + using nat_le_iff_add that by auto + have "a * \<^bold>|x\<^bold>| + b*\<^bold>|y\<^bold>| < 4 * \<^bold>|z\<^bold>|" using ex[OF \2 \ a\] ex[OF \2 \ b\] lenx leny unfolding pow_len by auto hence "c < 4" using leneq by auto consider (c_is_3) "c = 3" | (c_is_2) "c = 2" using \c < 4\ c by linarith then show "x \ y = y \ x" proof(cases) - case c_is_3 - show "x \ y = y \ x" - using - core_case[OF \c = 3\ \\<^bold>|y\<^sup>@b\<^bold>| \ \<^bold>|x\<^sup>@a\<^bold>|\[unfolded pow_len] _ _ lenx[unfolded pow_len] leny[unfolded pow_len]] - \x \ \\ \y \ \\ + case c_is_3 + show "x \ y = y \ x" + using + core_case[OF \c = 3\ \\<^bold>|y\<^sup>@b\<^bold>| \ \<^bold>|x\<^sup>@a\<^bold>|\[unfolded pow_len] _ _ lenx[unfolded pow_len] leny[unfolded pow_len]] + \x \ \\ \y \ \\ by blast - next - assume "c = 2" + next + assume "c = 2" hence eq2: "x\<^sup>@a \ y\<^sup>@b = z \ z" by (simp add: eq pow2_list) - from dual_order.trans le_cases[of "\<^bold>|x\<^sup>@a\<^bold>|" "\<^bold>|z\<^bold>|" "\<^bold>|z\<^bold>| \ \<^bold>|x\<^sup>@a\<^bold>|", unfolded eq_len_iff[OF this]] - have "\<^bold>|z\<^bold>| \ \<^bold>|x\<^sup>@a\<^bold>|" - using \\<^bold>|y\<^sup>@b\<^bold>| \ \<^bold>|x\<^sup>@a\<^bold>|\ by blast + from dual_order.trans le_cases[of "\<^bold>|x\<^sup>@a\<^bold>|" "\<^bold>|z\<^bold>|" "\<^bold>|z\<^bold>| \ \<^bold>|x\<^sup>@a\<^bold>|", unfolded eq_len_iff[OF this]] + have "\<^bold>|z\<^bold>| \ \<^bold>|x\<^sup>@a\<^bold>|" + using \\<^bold>|y\<^sup>@b\<^bold>| \ \<^bold>|x\<^sup>@a\<^bold>|\ by blast obtain a' where "Suc a' = a" and "1 \ a'" using \2 \ a\ ex by auto from eq2[folded \Suc a' = a\, unfolded pow_Suc2_list rassoc] pow_Suc2_list[of x a', unfolded this, symmetric] have eq3: "x \<^sup>@ a' \ x \ y \<^sup>@ b = z \ z" and aa':"x \<^sup>@ a' \ x = x \<^sup>@ a ". - hence "\<^bold>|x\<^sup>@a'\<^bold>| < \<^bold>|z\<^bold>|" + hence "\<^bold>|x\<^sup>@a'\<^bold>| < \<^bold>|z\<^bold>|" using \Suc a' = a\ lenx pow_len by auto - hence "\<^bold>|x\<^bold>| < \<^bold>|z\<^bold>|" - using mult_le_mono[of 1 a' "\<^bold>|z\<^bold>|" "\<^bold>|x\<^bold>|", OF \1 \ a'\, THEN leD] unfolding pow_len + hence "\<^bold>|x\<^bold>| < \<^bold>|z\<^bold>|" + using mult_le_mono[of 1 a' "\<^bold>|z\<^bold>|" "\<^bold>|x\<^bold>|", OF \1 \ a'\, THEN leD] unfolding pow_len by linarith obtain u w where "x\<^sup>@a'\u = z" and "w \ y\<^sup>@b = z" using eqd_prefE[OF eq3[unfolded rassoc] less_imp_le[OF \\<^bold>|x\<^sup>@a'\<^bold>| < \<^bold>|z\<^bold>|\], of thesis] eqd_prefE[OF eq2[symmetric] \\<^bold>|z\<^bold>| \ \<^bold>|x\<^sup>@a\<^bold>|\, of thesis] by fast have "x\<^sup>@a'\x\y\<^sup>@b = x\<^sup>@a'\u\w\y\<^sup>@b" - unfolding lassoc \x \<^sup>@ a' \ u = z\ \w \ y\<^sup>@b = z\ aa' cancel eq2 by simp + unfolding lassoc \x \<^sup>@ a' \ u = z\ \w \ y\<^sup>@b = z\ aa' cancel eq2 by simp hence "u\w=x" by auto hence "\<^bold>|w\u\<^bold>| = \<^bold>|x\<^bold>|" using swap_len by blast \ \Induction step: new equation with shorter z\ have "w\<^sup>@2\y\<^sup>@b = (w\u)\<^sup>@a" unfolding pow2_list using \w \ y \<^sup>@ b = z\ \x \<^sup>@ a' \ u = z\ \u\w=x\ pow_slide[of w u a', unfolded \Suc a' = a\] by simp from "less.hyps"[OF _ this _ b a, unfolded \\<^bold>|w\u\<^bold>| = \<^bold>|x\<^bold>|\] - have "y\w = w\y" + have "y\w = w\y" using \\<^bold>|x\<^bold>| < \<^bold>|z\<^bold>|\ by force have "y \ z = z \ y" unfolding \w \ y\<^sup>@b = z\[symmetric] lassoc \y\w = w\y\ - by (simp add: pow_commutes_list) + by (simp add: pow_commutes_list) hence "z\<^sup>@c\y\<^sup>@b = y\<^sup>@b\z\<^sup>@c" - by (simp add: comm_add_exps) + by (simp add: comm_add_exps) from this[folded eq, unfolded lassoc] have "x\<^sup>@a\y\<^sup>@b = y\<^sup>@b\x\<^sup>@a" using cancel_right by blast from this[unfolded comm_pow_roots[OF \a\0\ \b \ 0\]] show "x \ y = y \ x". qed qed qed qed qed theorem Lyndon_Schutzenberger: assumes "x\<^sup>@a\y\<^sup>@b = z\<^sup>@c" and "2 \ a" and "2 \ b" and "2 \ c" shows "x\y = y\x" and "x\z = z\x" and "y\z = z\y" proof- show "x \ y = y \ x" using Lyndon_Schutzenberger'[OF assms]. have "c \ 0" and "b \ 0" using \2 \ c\ \2 \ b\ by auto have "x \ x\<^sup>@a \ y\<^sup>@b = x\<^sup>@a \ y\<^sup>@b \ x" and "y \ x\<^sup>@a \ y\<^sup>@b = x\<^sup>@a \ y\<^sup>@b \ y" unfolding comm_add_exp[OF \x \ y = y \ x\[symmetric], of b] comm_add_exp[OF \x \ y = y \ x\, symmetric, of a] - lassoc power_commutes by blast+ + lassoc power_commutes by blast+ thus "x\z = z\x" and "y\z = z\y" using comm_drop_exp[OF \c \ 0\] unfolding lassoc \x\<^sup>@a\y\<^sup>@b = z\<^sup>@c\ by force+ qed end \ No newline at end of file diff --git a/thys/Combinatorics_Words/Periodicity_Lemma.thy b/thys/Combinatorics_Words/Periodicity_Lemma.thy --- a/thys/Combinatorics_Words/Periodicity_Lemma.thy +++ b/thys/Combinatorics_Words/Periodicity_Lemma.thy @@ -1,437 +1,437 @@ (* Title: CoW/Periodicity_Lemma.thy Author: Štěpán Holub, Charles University *) theory Periodicity_Lemma imports CoWBasic begin chapter "The Periodicity Lemma" -text\The Periodicity Lemma says that if a sufficiently long word has two periods p and q, -then the period can be refined to @{term "gcd p q"}. -The consequence is equivalent to the fact that the corresponding periodic roots commute. +text\The Periodicity Lemma says that if a sufficiently long word has two periods p and q, +then the period can be refined to @{term "gcd p q"}. +The consequence is equivalent to the fact that the corresponding periodic roots commute. ``Sufficiently long'' here means at least @{term "p + q - gcd p q"}. It is also known as the Fine and Wilf theorem due to its authors @{cite FineWilf}.\ text\ If we relax the requirement to @{term "p + q"}, then the claim becomes easy, and it is proved in theory @{theory Combinatorics_Words.CoWBasic} as @{term two_pers}: @{thm[names_long] two_pers[no_vars]}. \ theorem per_lemma_relaxed: assumes "periodN w p" and "periodN w q" and "p + q \ \<^bold>|w\<^bold>|" shows "(take p w)\(take q w) = (take q w)\(take p w)" using two_pers[OF \periodN w p\[unfolded periodN_def[of w p]] - \periodN w q\[unfolded periodN_def[of w q]], unfolded - take_len[OF add_leD1[OF \p + q \ \<^bold>|w\<^bold>|\]] - take_len[OF add_leD2[OF \p + q \ \<^bold>|w\<^bold>|\]], OF \p + q \ \<^bold>|w\<^bold>|\]. + \periodN w q\[unfolded periodN_def[of w q]], unfolded + take_len[OF add_leD1[OF \p + q \ \<^bold>|w\<^bold>|\]] + take_len[OF add_leD2[OF \p + q \ \<^bold>|w\<^bold>|\]], OF \p + q \ \<^bold>|w\<^bold>|\]. text\Also in terms of the numeric period:\ thm two_periodsN section \Main claim\ text\We first formulate the claim of the Periodicity lemma in terms of commutation of two periodic roots. For trivial reasons we can also drop the requirement that the roots are nonempty. \ lemma per_lemma_comm: - assumes "w \p r \ w" and "w \p s \ w" + assumes "w \p r \ w" and "w \p s \ w" and len: "\<^bold>|s\<^bold>| + \<^bold>|r\<^bold>| - (gcd \<^bold>|s\<^bold>| \<^bold>|r\<^bold>|) \ \<^bold>|w\<^bold>|" shows "s \ r = r \ s" using assms proof (induction "\<^bold>|s\<^bold>| + \<^bold>|s\<^bold>| + \<^bold>|r\<^bold>|" arbitrary: w r s rule: less_induct) case less consider (empty) "s = \" | (short) "\<^bold>|r\<^bold>| < \<^bold>|s\<^bold>|" | (step) "s \ \ \ \<^bold>|s\<^bold>| \ \<^bold>|r\<^bold>|" by force - then show ?case + then show ?case proof (cases) - case (empty) + case (empty) thus "s \ r = r \ s" by fastforce next case (short) thus "s \ r = r \ s" - using "less.hyps"[OF _ \ w \p s \ w\ \ w \p r \ w\ + using "less.hyps"[OF _ \ w \p s \ w\ \ w \p r \ w\ \\<^bold>|s\<^bold>| + \<^bold>|r\<^bold>| - (gcd \<^bold>|s\<^bold>| \<^bold>|r\<^bold>|) \ \<^bold>|w\<^bold>|\[unfolded gcd.commute[of "\<^bold>|s\<^bold>|"] add.commute[of "\<^bold>|s\<^bold>|"]]] by fastforce next case (step) hence "s \ \" and "\<^bold>|s\<^bold>| \ \<^bold>|r\<^bold>|" by blast+ from le_add_diff[OF gcd_le2_nat[OF \s \ \\[folded length_0_conv], of "\<^bold>|r\<^bold>|"], unfolded gcd.commute[of "\<^bold>|r\<^bold>|"]] - have "\<^bold>|r\<^bold>| \ \<^bold>|w\<^bold>|" + have "\<^bold>|r\<^bold>| \ \<^bold>|w\<^bold>|" using \\<^bold>|s\<^bold>| + \<^bold>|r\<^bold>| - (gcd \<^bold>|s\<^bold>| \<^bold>|r\<^bold>|) \ \<^bold>|w\<^bold>|\ order.trans by fast hence "\<^bold>|s\<^bold>| \ \<^bold>|w\<^bold>|" using \\<^bold>|s\<^bold>| \ \<^bold>|r\<^bold>|\ order.trans by blast - from pref_prod_long[OF \w \p s \ w\ this] + from pref_prod_long[OF \w \p s \ w\ this] have "s \p w". - + obtain w' where "s \ w' = w" and "\<^bold>|w'\<^bold>| < \<^bold>|w\<^bold>|" using \s \ \\ \s \p w\ by auto have "w' \p w" using \w \p s \ w\ unfolding \s \ w' = w\[symmetric] pref_cancel_conv. from this[folded \s \ w' = w\] have "w' \p s\w'". have "s \p r" - using pref_prod_short[OF prefix_order.order.trans[OF \s \p w\ \w \p r \ w\] \\<^bold>|s\<^bold>| \ \<^bold>|r\<^bold>|\]. + using pref_prod_short[OF prefix_order.trans[OF \s \p w\ \w \p r \ w\] \\<^bold>|s\<^bold>| \ \<^bold>|r\<^bold>|\]. hence "w' \p (s\\<^sup>>r) \ w'" using \w \p r \ w\ \s \ w' = w\ pref_prod_pref[OF _ \w' \p w\, of "s\\<^sup>>r"] by fastforce have ind_len: "\<^bold>|s\<^bold>| + \<^bold>|s\\<^sup>>r\<^bold>| - (gcd \<^bold>|s\<^bold>| \<^bold>|s\\<^sup>>r\<^bold>|) \ \<^bold>|w'\<^bold>|" using \\<^bold>|s\<^bold>| + \<^bold>|r\<^bold>| - (gcd \<^bold>|s\<^bold>| \<^bold>|r\<^bold>|) \ \<^bold>|w\<^bold>|\ \s \ w' = w\ \s \p r\ by auto - have "s \ s\\<^sup>>r = s\\<^sup>>r \ s" - using "less.hyps"[OF _ \w' \p (s\\<^sup>>r) \ w'\ \w' \p s \ w'\ ind_len] \s \p r\ \\<^bold>|w'\<^bold>| < \<^bold>|w\<^bold>|\ by force + have "s \ s\\<^sup>>r = s\\<^sup>>r \ s" + using "less.hyps"[OF _ \w' \p (s\\<^sup>>r) \ w'\ \w' \p s \ w'\ ind_len] \s \p r\ \\<^bold>|w'\<^bold>| < \<^bold>|w\<^bold>|\ by force thus "s \ r = r \ s" using \s \p r\ by auto - qed + qed qed text\We can now prove the numeric version.\ theorem per_lemma: assumes "periodN w p" and "periodN w q" and len: "p + q - gcd p q \ \<^bold>|w\<^bold>|" shows "periodN w (gcd p q)" proof- have takep: "w \p (take p w) \ w" and takeq: "w \p (take q w) \ w" using \periodN w p\ \periodN w q\ periodN_D3 by blast+ have lenp: "\<^bold>|take p w\<^bold>| = p" using gcd_le2_nat[OF per_positive[OF \periodN w q\], of p] len take_len by auto have lenq: "\<^bold>|take q w\<^bold>| = q" using gcd_le1_nat[OF per_positive[OF \periodN w p\], of q] len take_len by simp - obtain t where "take p w \ t*" and "take q w \ t*" + obtain t where "take p w \ t*" and "take q w \ t*" using comm_rootE[OF per_lemma_comm[OF takeq takep, unfolded lenp lenq, OF len], of thesis] by blast hence "w \p t\<^sup>\" using \periodN w p\ periodN_def per_root_trans by blast have "periodN w \<^bold>|t\<^bold>|" - using root_periodN[OF per_nemp[OF \periodN w p\] \w \p t\<^sup>\\]. + using root_periodN[OF per_nemp[OF \periodN w p\] \w \p t\<^sup>\\]. have "\<^bold>|t\<^bold>| dvd (gcd p q)" using common_root_len_gcd[OF \take p w \ t*\ \take q w \ t*\] unfolding lenp lenq. from dvd_div_mult_self[OF this] have "gcd p q div \<^bold>|t\<^bold>| * \<^bold>|t\<^bold>| = gcd p q". have "gcd p q \ 0" - using \periodN w p\ by auto + using \periodN w p\ by auto from this[folded dvd_div_eq_0_iff[OF \\<^bold>|t\<^bold>| dvd (gcd p q)\]] show "periodN w (gcd p q)" using per_mult[OF \periodN w \<^bold>|t\<^bold>|\, of "gcd p q div \<^bold>|t\<^bold>|", unfolded dvd_div_mult_self[OF \\<^bold>|t\<^bold>| dvd (gcd p q)\]] by blast qed section \Optimality\ -text\@{term "FW_word"} (where FW stands for Fine and Wilf) yields a word which show the optimality of the bound in the Periodicity lemma. +text\@{term "FW_word"} (where FW stands for Fine and Wilf) yields a word which show the optimality of the bound in the Periodicity lemma. Moreover, the obtained word has maximum possible letters (each equality of letters is forced by periods). The latter is not proved here.\ term "butlast ([0..<(gcd p q)]\<^sup>@(p div (gcd p q)))\[gcd p q]\(butlast ([0..<(gcd p q)]\<^sup>@(p div (gcd p q))))" \ \an auxiliary claim\ -lemma ext_per_sum: assumes "periodN w p" and "periodN w q" and "p \ \<^bold>|w\<^bold>|" +lemma ext_per_sum: assumes "periodN w p" and "periodN w q" and "p \ \<^bold>|w\<^bold>|" shows "periodN ((take p w) \ w) (p+q)" proof- have nemp: "take p w \ take q w \ \" using \periodN w p\ by auto have "take (p + q) (take p w \ w) = take p (take p w \ w) \ take q (drop p (take p w \ w))" using take_add by blast also have "... = take p w \ take q w" by (simp add: \p \ \<^bold>|w\<^bold>|\) ultimately have sum: "take (p + q) (take p w \ w) = take p w \ take q w" by simp show ?thesis using assms(1) assms(2) nemp unfolding periodN_def period_root_def sum rassoc same_prefix_prefix using pref_prolong by blast qed abbreviation "fw_p_per p q \ butlast ([0..<(gcd p q)]\<^sup>@(p div (gcd p q)))" abbreviation "fw_base p q \ fw_p_per p q \ [gcd p q]\ fw_p_per p q" fun FW_word :: "nat \ nat \ nat list" where - FW_word_def: "FW_word p q = -\\symmetry\ (if q < p then FW_word q p else -\\artificial value\ if p = 0 \ p = q then \ else + FW_word_def: "FW_word p q = +\\symmetry\ (if q < p then FW_word q p else +\\artificial value\ if p = 0 \ p = q then \ else \\base case\ if gcd p q = q - p then fw_base p q \\step\ else (take p (FW_word p (q-p))) \ FW_word p (q-p) )" lemma FW_sym: "FW_word p q = FW_word q p" by (cases rule: linorder_cases[of p q], simp+) theorem fw_word': "\ p dvd q \ \ q dvd p \ \<^bold>|FW_word p q\<^bold>| = p + q - gcd p q - 1 \ periodN (FW_word p q) p \ periodN (FW_word p q) q \ \ periodN (FW_word p q) (gcd p q)" proof (induction "p + p + q" arbitrary: p q rule: less_induct) case less have "p \ 0" - using \\ q dvd p\ dvd_0_right[of q] by meson + using \\ q dvd p\ dvd_0_right[of q] by meson have "p \ q" using \\ p dvd q\ by auto - then consider "q < p" | "p < q" + then consider "q < p" | "p < q" by linarith - then show ?case + then show ?case proof (cases) assume "q < p" have less: "q + q + p < p + p + q" by (simp add: \q < p\) thus ?case - using "less.hyps"[OF _ \\ q dvd p\ \\ p dvd q\] - unfolding FW_sym[of p q] gcd.commute[of p q] add.commute[of p q] by blast + using "less.hyps"[OF _ \\ q dvd p\ \\ p dvd q\] + unfolding FW_sym[of p q] gcd.commute[of p q] add.commute[of p q] by blast next let ?d = "gcd p q" let ?dw = "[0..<(gcd p q)]" let ?pd = "p div (gcd p q)" - assume "p < q" + assume "p < q" thus ?thesis proof (cases "?d = q - p") assume "?d = q - p" hence "p + ?d = q" using \p < q\ by auto have fw: "FW_word p q = fw_base p q" - using FW_word_def \p \ 0\ \gcd p q = q - p\ \p < q\ by auto + using FW_word_def \p \ 0\ \gcd p q = q - p\ \p < q\ by auto have ppref: "\<^bold>|butlast (?dw\<^sup>@?pd)\[?d]\<^bold>| = p" using length_append \p \ 0\ pow_len[of "?dw" "?pd"] by auto note ppref' = this[unfolded length_append] have qpref: "\<^bold>|butlast (?dw\<^sup>@?pd)\[?d]\?dw\<^bold>| = q" - unfolding lassoc length_append ppref' using \p + gcd p q = q\ by simp + unfolding lassoc length_append ppref' using \p + gcd p q = q\ by simp have "butlast (?dw\<^sup>@?pd)\[?d] \p FW_word p q" unfolding fw by force from pref_take[OF this] have takep: "take p (FW_word p q) = butlast (?dw\<^sup>@?pd)\[?d]" unfolding ppref. have "?dw \ \" and "\<^bold>|?dw\<^bold>| = ?d" using \p \ 0\ by auto have "?pd \ 0" - by (simp add: \p \ 0\ dvd_div_eq_0_iff) + by (simp add: \p \ 0\ dvd_div_eq_0_iff) from not0_implies_Suc[OF this] obtain e where "?pd = Suc e" by blast have "gcd p q \ p" - using \\ p dvd q\ gcd_dvd2[of p q] by force + using \\ p dvd q\ gcd_dvd2[of p q] by force hence "Suc e \ 1" using dvd_mult_div_cancel[OF gcd_dvd1[of p q], unfolded \?pd = Suc e\] by force hence "e \ 0" by simp have "[0..@ e \ \" using \[0.. \\ \e \ 0\ nonzero_pow_emp by blast hence but_dec: "butlast (?dw\<^sup>@?pd) = ?dw \ butlast (?dw\<^sup>@e)" unfolding \?pd = Suc e\ pow_Suc_list butlast_append if_not_P[OF \[0..@ e \ \\] by blast - have but_dec_b: "butlast (?dw\<^sup>@?pd) = ?dw\<^sup>@e \ butlast ?dw" + have but_dec_b: "butlast (?dw\<^sup>@?pd) = ?dw\<^sup>@e \ butlast ?dw" unfolding \?pd = Suc e\ pow_Suc2_list butlast_append if_not_P[OF \?dw \ \\] by blast have "butlast (?dw\<^sup>@?pd)\[?d]\?dw \p FW_word p q" using fw but_dec by auto note takeq = pref_take[OF this, unfolded qpref] have "\<^bold>|FW_word p q\<^bold>| = p + q - gcd p q - 1" proof- have "\<^bold>|?dw\<^bold>| = ?d" by auto have "gcd p q dvd p" by auto hence "\<^bold>|?dw\<^sup>@?pd\<^bold>| = p" using pow_len[of "?dw" "?pd"] by auto hence "\<^bold>|butlast (?dw\<^sup>@?pd)\<^bold>| = p - 1" by simp hence "\<^bold>|FW_word p q\<^bold>| = (p - 1) + 1 + (p - 1)" using fw unfolding length_append by auto thus "\<^bold>|FW_word p q\<^bold>| = p + q - gcd p q - 1" unfolding \gcd p q = q - p\ using \p + gcd p q = q\ \p \ 0\ add.assoc by auto - qed + qed have "periodN (FW_word p q) p" proof- have "take p (FW_word p q) \ \" using \p \ 0\ unfolding length_0_conv[symmetric] ppref[folded takep]. thus "periodN (FW_word p q) p" unfolding periodN_def period_root_def takep unfolding fw lassoc by auto qed have "periodN (FW_word p q) q" - unfolding periodN_def period_root_def + unfolding periodN_def period_root_def proof show "take q (FW_word p q) \ \" unfolding length_0_conv[symmetric] qpref[folded takeq] using \p \ 0\ \p < q\ by linarith next show "FW_word p q \p take q (FW_word p q) \ FW_word p q" - unfolding takeq + unfolding takeq unfolding fw rassoc pref_cancel_conv but_dec but_dec_b \?pd = Suc e\ pow_Suc2_list butlast_append pow_Suc_list if_not_P[OF \?dw \ \\] unfolding lassoc power_commutes[symmetric] unfolding rassoc pref_cancel_conv - using pref_ext[OF prefixeq_butlast, of "?dw"] by blast + using pref_ext[OF prefixeq_butlast, of "?dw"] by blast qed have "\ periodN (FW_word p q) ?d" proof- have last_a: "last (take p (FW_word p q)) = ?d" unfolding takep nth_append_length[of "butlast (?dw \<^sup>@ ?pd)" "?d" \] last_snoc by blast have "?dw \p FW_word p q" using fw but_dec by simp from pref_take[OF this, unfolded \\<^bold>|?dw\<^bold>| = ?d\] have takegcd: "take (gcd p q) (FW_word p q) = [0..@e \ butlast ([0.. [?d] \ (butlast ([0..@(p div gcd p q)))" using fw but_dec_b by simp hence gcdepref: "[0..@ Suc e \p take (gcd p q) (FW_word p q) \ FW_word p q" using takegcd by simp have "\<^bold>|[0..@ Suc e\<^bold>| = p" - unfolding pow_len \\<^bold>|?dw\<^bold>| = ?d\ \?pd = Suc e\[symmetric] using + unfolding pow_len \\<^bold>|?dw\<^bold>| = ?d\ \?pd = Suc e\[symmetric] using dvd_div_mult_self[OF gcd_dvd1]. - from pref_take[OF gcdepref, unfolded this] + from pref_take[OF gcdepref, unfolded this] have takegcdp: "take p (take (gcd p q) (FW_word p q) \ (FW_word p q)) = [0..@e \ [0..p \ 0\) from last_upt[OF this] have last_b: "last (take p (take (gcd p q) (FW_word p q) \ (FW_word p q))) = gcd p q - 1" unfolding takegcdp last_appendR[of "[0..@e", OF \[0.. \\]. have "p \ \<^bold>|FW_word p q\<^bold>|" - using \\<^bold>|FW_word p q\<^bold>| = p + q - gcd p q - 1\ \gcd p q = q - p\ \p < q\ by linarith + using \\<^bold>|FW_word p q\<^bold>| = p + q - gcd p q - 1\ \gcd p q = q - p\ \p < q\ by linarith have "gcd p q \ gcd p q - 1" using \gcd p q = q - p\ \p < q\ by linarith hence "take p (FW_word p q) \ take p (take (gcd p q) (FW_word p q) \ (FW_word p q))" unfolding last_b[symmetric] unfolding last_a[symmetric] using arg_cong[of _ _ last] by blast - hence "\ FW_word p q \p take (gcd p q) (FW_word p q) \ FW_word p q " + hence "\ FW_word p q \p take (gcd p q) (FW_word p q) \ FW_word p q " using pref_share_take[OF _ \p \ \<^bold>|FW_word p q\<^bold>|\, of "take (gcd p q) (FW_word p q) \ FW_word p q"] by blast - thus "\ periodN (FW_word p q) (gcd p q)" + thus "\ periodN (FW_word p q) (gcd p q)" unfolding periodN_def period_root_def by blast - qed + qed show ?thesis using \periodN (FW_word p q) p\ \periodN (FW_word p q) q\ \\<^bold>|FW_word p q\<^bold>| = p + q - gcd p q - 1\ \\ periodN (FW_word p q) (gcd p q)\ by blast next - let ?d' = "gcd p (q-p)" + let ?d' = "gcd p (q-p)" assume "gcd p q \ q - p" - hence fw: "FW_word p q = (take p (FW_word p (q-p))) \ FW_word p (q-p)" + hence fw: "FW_word p q = (take p (FW_word p (q-p))) \ FW_word p (q-p)" using FW_word_def \p \ 0\ \p \ q\ \p < q\ by (meson less_Suc_eq not_less_eq) have divhyp1: "\ p dvd q - p" using \\ p dvd q\ \p < q\ dvd_minus_self by auto have divhyp2: "\ q - p dvd p" proof (rule notI) assume "q - p dvd p" have "q = p + (q - p)" by (simp add: \p < q\ less_or_eq_imp_le) - from gcd_add2[of p "q - p", folded this, unfolded gcd_nat.absorb2[of "q - p" p, OF \q - p dvd p\]] + from gcd_add2[of p "q - p", folded this, unfolded gcd_nat.absorb2[of "q - p" p, OF \q - p dvd p\]] show "False" using \gcd p q \ q - p\ by blast qed have lenhyp: "p + p + (q - p) < p + p + q" - using \p < q\ \p \ 0\ by linarith + using \p < q\ \p \ 0\ by linarith -\ \induction assumption\ +\ \induction assumption\ have "\<^bold>|FW_word p (q - p)\<^bold>| = p + (q - p) - ?d' - 1" and "periodN (FW_word p (q-p)) p" and "periodN (FW_word p (q-p)) (q-p)" and - "\ periodN (FW_word p (q-p)) (gcd p (q-p))" + "\ periodN (FW_word p (q-p)) (gcd p (q-p))" using "less.hyps"[OF _ divhyp1 divhyp2] lenhyp - by blast+ + by blast+ \ \auxiliary facts\ have "p + (q - p) = q" using divhyp1 dvd_minus_self by auto have "?d = ?d'" using gcd_add2[of p "q-p", unfolded le_add_diff_inverse[OF less_imp_le[OF \p < q\]]]. have "?d \ q" using \\ q dvd p\ gcd_dvd2[of q p, unfolded gcd.commute[of q]] by force from this[unfolded nat_neq_iff] have "?d < q" using gr_implies_not0 \p < q\ nat_dvd_not_less by blast hence "1 \ q - ?d" by linarith have "?d' < q - p" using gcd_le2_nat[OF per_positive[OF \periodN (FW_word p (q - p)) (q - p)\], of p] divhyp2[unfolded gcd_nat.absorb_iff2] nat_less_le by blast hence "p \ \<^bold>|(FW_word p (q - p))\<^bold>|" unfolding \\<^bold>|FW_word p (q - p)\<^bold>| = p + (q - p) - ?d' - 1\ diff_diff_left discrete by linarith - have "FW_word p (q - p) \ \" - unfolding length_0_conv[symmetric] using \p \ \<^bold>|FW_word p (q - p)\<^bold>|\ \p \ 0\[folded le_zero_eq] + have "FW_word p (q - p) \ \" + unfolding length_0_conv[symmetric] using \p \ \<^bold>|FW_word p (q - p)\<^bold>|\ \p \ 0\[folded le_zero_eq] by linarith \ \claim 1\ - have "\<^bold>|FW_word p q\<^bold>| = p + q - ?d - 1" + have "\<^bold>|FW_word p q\<^bold>| = p + q - ?d - 1" proof- have "\<^bold>|FW_word p q\<^bold>| = \<^bold>|take p (FW_word p (q - p))\<^bold>| + \<^bold>|FW_word p (q - p)\<^bold>|" using fw length_append[of "take p (FW_word p (q - p))" "FW_word p (q - p)"] by presburger also have "... = p + (p + (q - p) - ?d' - 1)" - unfolding \\<^bold>|FW_word p (q - p)\<^bold>| = p + (q - p) - ?d' - 1\ + unfolding \\<^bold>|FW_word p (q - p)\<^bold>| = p + (q - p) - ?d' - 1\ take_len[OF \p \ \<^bold>|FW_word p (q - p)\<^bold>|\] by blast also have "... = p + (q - ?d - 1)" unfolding \?d = ?d'\ using \p < q\ by auto also have "... = p + (q - ?d) - 1" using Nat.add_diff_assoc[OF \1 \ q - ?d\]. also have "... = p + q - ?d - 1" by (simp add: \?d < q\ less_or_eq_imp_le) - finally show "\<^bold>|FW_word p q\<^bold>| = p + q - ?d - 1" + finally show "\<^bold>|FW_word p q\<^bold>| = p + q - ?d - 1" by presburger qed \ \claim 2\ have "periodN (FW_word p q) p" using fw ext_per_left[OF \periodN (FW_word p (q-p)) p\ \p \ \<^bold>|FW_word p (q - p)\<^bold>|\] by presburger \ \claim 3\ have "periodN (FW_word p q) q" using ext_per_sum[OF \periodN (FW_word p (q - p)) p\ \periodN (FW_word p (q - p)) (q - p)\ \p \ \<^bold>|FW_word p (q - p)\<^bold>|\, folded fw, unfolded \p + (q-p) = q\]. \ \claim 4\ - have "\ periodN (FW_word p q) ?d" - using \\ periodN (FW_word p (q -p)) (gcd p (q-p))\ - unfolding \?d = ?d'\[symmetric] - using periodN_fac[of "take p (FW_word p (q - p))" "FW_word p (q - p)" \ "?d", unfolded append_Nil2, + have "\ periodN (FW_word p q) ?d" + using \\ periodN (FW_word p (q -p)) (gcd p (q-p))\ + unfolding \?d = ?d'\[symmetric] + using periodN_fac[of "take p (FW_word p (q - p))" "FW_word p (q - p)" \ "?d", unfolded append_Nil2, OF _ \FW_word p (q - p) \ \\, folded fw] by blast thus ?thesis - using \periodN (FW_word p q) p\ \periodN (FW_word p q) q\ \\<^bold>|FW_word p q\<^bold>| = p + q - ?d - 1\ by blast + using \periodN (FW_word p q) p\ \periodN (FW_word p q) q\ \\<^bold>|FW_word p q\<^bold>| = p + q - ?d - 1\ by blast qed qed qed theorem fw_word: assumes "\ p dvd q" "\ q dvd p" shows "\<^bold>|FW_word p q\<^bold>| = p + q - gcd p q - 1" and "periodN (FW_word p q) p" and "periodN (FW_word p q) q" and "\ periodN (FW_word p q) (gcd p q)" using fw_word'[OF assms] by blast+ text\Calculation examples\ value "FW_word 3 7" value "FW_word 5 7" value "FW_word 5 13" value "FW_word 4 6" value "FW_word 12 18" section "Other variants of the periodicity lemma" text \Periodicity lemma is one of the most frequent tools in Combinatorics on words. Here are some useful variants.\ lemma fac_two_conjug_prim_root: assumes facs: "u \f r\<^sup>@k" "u \f s\<^sup>@l" and nemps: "r \ \" "s \ \" and len: "\<^bold>|r\<^bold>| + \<^bold>|s\<^bold>| - gcd (\<^bold>|r\<^bold>|) (\<^bold>|s\<^bold>|) \ \<^bold>|u\<^bold>|" shows "\ r \ \ s" proof - obtain r' s' where prefr': "u \p r'\<^sup>@k" and prefs': "u \p s'\<^sup>@l" and conjugs: "r \ r'" "s \ s'" using facs by (elim fac_pow_pref_conjug) have rootr': "u \p r' \ u" and roots': "u \p s' \ u" - using pref_prod_root[OF prefr'] pref_prod_root[OF prefs']. + using pref_prod_root[OF prefr'] pref_prod_root[OF prefs']. have nemps': "r' \ \" "s'\ \" using nemps conjugs conjug_nemp_iff by blast+ have "\<^bold>|r'\<^bold>| + \<^bold>|s'\<^bold>| - gcd (\<^bold>|r'\<^bold>|) (\<^bold>|s'\<^bold>|) \ \<^bold>|u\<^bold>|" using len unfolding conjug_len[OF \r \ r'\] conjug_len[OF \s \ s'\]. from per_lemma_comm[OF roots' rootr' this] have "r' \ s' = s' \ r'". - then have "\ r' = \ s'" using comm_primroots[OF nemps'] by force + then have "\ r' = \ s'" using comm_primroots[OF nemps'] by force also have "\ s \ \ s'" using conjug_prim_root[OF \s \ s'\ \s \ \\]. also have [symmetric]: "\ r \ \ r'" using conjug_prim_root[OF \r \ r'\ \r \ \\]. finally show "\ r \ \ s".. qed lemma fac_two_conjug_prim_root': assumes facs: "u \f r\<^sup>@k" "u \f s\<^sup>@l" and "u \ \" and len: "\<^bold>|r\<^bold>| + \<^bold>|s\<^bold>| - gcd (\<^bold>|r\<^bold>|) (\<^bold>|s\<^bold>|) \ \<^bold>|u\<^bold>|" shows "\ r \ \ s" proof - have nemps: "r \ \" "s \ \" using facs \u \ \\ by auto show "conjugate (\ r) (\ s)" using fac_two_conjug_prim_root[OF facs nemps len]. qed lemma fac_two_nconj_prim_pow: assumes prims: "primitive r" "primitive s" and "\ r \ s" and facs: "u \f r\<^sup>@k" "u \f s\<^sup>@l" shows "\<^bold>|u\<^bold>| < \<^bold>|r\<^bold>| + \<^bold>|s\<^bold>| - gcd (\<^bold>|r\<^bold>|) (\<^bold>|s\<^bold>|)" using \\ r \ s\ fac_two_conjug_prim_root[OF facs prim_nemp prim_nemp leI, OF prims] unfolding prim_self_root[OF \primitive r\] prim_self_root[OF \primitive s\] by (rule contrapos_np) end diff --git a/thys/Combinatorics_Words/Submonoids.thy b/thys/Combinatorics_Words/Submonoids.thy --- a/thys/Combinatorics_Words/Submonoids.thy +++ b/thys/Combinatorics_Words/Submonoids.thy @@ -1,1058 +1,1058 @@ (* Title: CoW/Submonoids.thy Author: Štěpán Holub, Charles University Author: Štěpán Starosta, CTU in Prague *) theory Submonoids imports CoWBasic begin chapter \Submonoids of a free monoid\ text\This chapter deals with properties of submonoids of a free monoid, that is, with monoids of words. See more in Chapter 1 of @{cite Lo83}. \ section \Hull\ text\First, we define the hull of a set of words, that is, the monoid generated by them.\ inductive_set hull :: "'a list set \ 'a list set" ("\_\") for G where emp_in: "\ \ \G\" | prod_cl: "w1 \ G \ w2 \ \G\ \ w1 \ w2 \ \G\" lemmas [intro] = hull.intros lemma hull_closed: "w1 \ \G\ \ w2 \ \G\ \ w1 \ w2 \ \G\" - by (rule hull.induct[of w1 G "\ x. (x\w2)\\G\"]) auto+ + by (rule hull.induct[of w1 G "\ x. (x\w2)\\G\"]) auto+ lemma gen_in [intro]: "w \ G \ w \ \G\" - using hull.prod_cl by fastforce + using hull.prod_cl by fastforce -lemma hull_induct: assumes "x \ \G\" "P \" "\w. w \ G \ P w" +lemma hull_induct: assumes "x \ \G\" "P \" "\w. w \ G \ P w" "\w1 w2. w1 \ \G\ \ P w1 \ w2 \ \G\ \ P w2 \ P (w1 \ w2)" shows "P x" - using hull.induct[of _ _ P, OF \x \ \G\\ \P \\] + using hull.induct[of _ _ P, OF \x \ \G\\ \P \\] assms by (simp add: gen_in) - + lemma genset_sub: "G \ \G\" using gen_in .. lemma in_lists_conv_set_subset: "set ws \ G \ ws \ lists G" by blast lemma concat_in_hull [intro]: assumes "set ws \ G" shows "concat ws \ \G\" using assms by (induction ws) auto lemma concat_in_hull' [intro]: assumes "ws \ lists G" shows "concat ws \ \G\" using assms by (induction ws) auto -lemma hull_concat_lists0: "w \ \G\ \ (\ ws \ lists G. concat ws = w)" +lemma hull_concat_lists0: "w \ \G\ \ (\ ws \ lists G. concat ws = w)" proof(rule hull.induct[of _ G], simp) show "\ws\lists G. concat ws = \" using concat.simps(1) lists.Nil[of G] exI[of "\ x. concat x = \", OF concat.simps(1)] by blast show " \w1 w2. w1 \ G \ w2 \ \G\ \ \ws\lists G. concat ws = w2 \ \ws\lists G. concat ws = w1 \ w2" - by (metis Cons_in_lists_iff concat.simps(2)) + by (metis Cons_in_lists_iff concat.simps(2)) qed lemma hull_concat_lists: "\G\ = concat ` lists G" unfolding image_iff using hull_concat_lists0 by blast lemma concat_tl: "x # xs \ lists G \ concat xs \ \G\" by (simp add: hull_concat_lists) lemma nemp_concat_hull: assumes "us \ \" and "us \ lists G\<^sub>+" shows "concat us \ \G\" and "concat us \ \" using assms by fastforce+ -lemma hull_mon: "A \ B \ \A\ \ \B\" -proof +lemma hull_mon: "A \ B \ \A\ \ \B\" +proof fix x assume "A \ B" "x \ \A\" - thus "x \ \B\" + thus "x \ \B\" unfolding image_def hull_concat_lists using sub_lists_mono[OF \A \ B\] by blast qed lemma emp_gen_set: "\{}\ = {\}" - unfolding hull_concat_lists by auto + unfolding hull_concat_lists by auto lemma hull_drop_one: "\G\ = \G\<^sub>+\" proof (intro equalityI subsetI) - fix x assume "x \ \G\" thus "x \ \G\<^sub>+\" + fix x assume "x \ \G\" thus "x \ \G\<^sub>+\" unfolding hull_concat_lists using del_emp_concat lists_drop_emp' by blast next fix x assume "x \ \G\<^sub>+\" thus "x \ \G\" unfolding hull_concat_lists image_iff by auto -qed +qed -lemma sing_gen_power: "u \ \{x}\ \ \k. u = x\<^sup>@k" - unfolding hull_concat_lists using one_generated_list_power by auto +lemma sing_gen_power: "u \ \{x}\ \ \k. u = x\<^sup>@k" + unfolding hull_concat_lists using one_generated_list_power by auto lemma sing_gen: "w \ \{z}\ \ w \ z*" using rootI sing_gen_power by blast - + lemma lists_gen_to_hull: "us \ lists G\<^sub>+ \ us \ lists \G\\<^sub>+" using lists_mono genset_sub by force -lemma rev_hull0: "x \ rev ` \G\ \ x \ \rev ` G\" +lemma rev_hull0: "x \ rev ` \G\ \ x \ \rev ` G\" proof- assume "x \ rev ` \G\" then obtain xs where "x = rev (concat xs)" and "xs \ lists G" unfolding hull_concat_lists by auto thus "x \ \rev ` G\" unfolding image_iff hull_concat_lists using rev_concat[of xs] - by fastforce + by fastforce qed lemma rev_hull1: "x \ \rev ` G\ \ x \ rev ` \G\" proof- assume "x \ \rev ` G\" then obtain xs where "x = concat xs" and "xs \ lists (rev ` G)" unfolding hull_concat_lists by blast hence "rev x \ \G\" unfolding hull_concat_lists using rev_concat by fastforce thus "x \ rev ` \G\" by (simp add: rev_in_conv) qed lemma rev_hull: "rev`\G\ = \rev`G\" by (simp add: rev_hull0 rev_hull1 set_eq_subset subsetI) lemma power_in: "x \ \G\ \ x\<^sup>@k \ \G\" by (induction k, auto, simp add: hull_closed) lemma hull_closed_lists: "us \ lists \G\ \ concat us \ \G\" proof (induct us, auto) show "\a us. concat us \ \G\ \ a \ \G\ \ \x\set us. x \ \G\ \ a \ concat us \ \G\ " by (simp add: hull_closed) qed lemma self_gen: "\\G\\ = \G\" using image_subsetI[of "lists \G\" concat "\G\", unfolded hull_concat_lists[of "\G\", symmetric], - THEN subset_antisym[OF _ genset_sub[of "\G\"]]] hull_closed_lists[of _ G] by blast + THEN subset_antisym[OF _ genset_sub[of "\G\"]]] hull_closed_lists[of _ G] by blast text\Intersection of hulls is a hull.\ lemma hulls_inter: "\\ {\G\ | G. G \ S}\ = \ {\G\ | G. G \ S}" proof {fix G assume "G \ S" - hence "\\ {\G\ |G. G \ S}\ \ \G\" - using Inter_lower[of "\G\" "{\G\ |G. G \ S}"] mem_Collect_eq[of "\G\" "\ A. \ G. G \ S \ A = \G\"] + hence "\\ {\G\ |G. G \ S}\ \ \G\" + using Inter_lower[of "\G\" "{\G\ |G. G \ S}"] mem_Collect_eq[of "\G\" "\ A. \ G. G \ S \ A = \G\"] hull_mon[of "\ {\G\ |G. G \ S}" "\G\"] unfolding self_gen by auto} - thus "\\ {\G\ |G. G \ S}\ \ \ {\G\ |G. G \ S}" by blast -next + thus "\\ {\G\ |G. G \ S}\ \ \ {\G\ |G. G \ S}" by blast +next show "\ {\G\ |G. G \ S} \ \\ {\G\ |G. G \ S}\" by (simp add: genset_sub) qed section "Factorization into generators" -text\We define a decomposition (or a factorization) of a into elements of a given generating set. Such a decomposition is well defined only +text\We define a decomposition (or a factorization) of a into elements of a given generating set. Such a decomposition is well defined only if the decomposed word is an element of the hull. Even int that case, however, the decomposition need not be unique.\ fun decompose :: "'a list set \ 'a list \ 'a list list" ("Dec _ _" [51,51] 64) where "decompose G u = (SOME us. us \ lists G\<^sub>+ \ concat us = u)" -lemma dec_ex: assumes "u \ \G\" shows "\ us. (us \ lists G\<^sub>+ \ concat us = u)" - using assms unfolding image_def hull_concat_lists[of G] mem_Collect_eq - using del_emp_concat lists_drop_emp' by metis +lemma dec_ex: assumes "u \ \G\" shows "\ us. (us \ lists G\<^sub>+ \ concat us = u)" + using assms unfolding image_def hull_concat_lists[of G] mem_Collect_eq + using del_emp_concat lists_drop_emp' by metis lemma decI': "u \ \G\ \ (Dec G u) \ lists G\<^sub>+" unfolding decompose.simps using someI_ex[OF dec_ex] by blast lemma decI: "u \ \G\ \ concat (Dec G u) = u" unfolding decompose.simps using someI_ex[OF dec_ex] by blast -lemma dec_emp: "Dec G \ = \" +lemma dec_emp: "Dec G \ = \" proof- have ex: "\ \ lists G\<^sub>+ \ concat \ = \" by simp have all: "(us \ lists G\<^sub>+ \ concat us = \) \ us = \" for us using emp_concat_emp by auto - show ?thesis + show ?thesis unfolding decompose.simps using all[OF someI[of "\ x. x \ lists G\<^sub>+ \ concat x = \", OF ex]]. qed lemma dec_nemp: "u \ \G\\<^sub>+ \ Dec G u \ \" - using decI[of u G] by force + using decI[of u G] by force lemma dec_nemp': "u \ \ \ u \ \G\ \ Dec G u \ \" using dec_nemp by blast lemma dec_dom': "u \ \G\ \ Dec G u \ lists G" - using decI' by auto + using decI' by auto lemma dec_hd: assumes "u \ \" "u \ \G\" shows "hd (Dec G u) \ G" using dec_nemp'[OF assms] dec_dom'[OF \u \ \G\\] lists_hd[of "Dec G u" G] by blast lemma non_gen_dec: "u \ \G\ \ u \ G \ Dec G u \ [u]" using decI' Cons_in_lists_iff by fastforce subsection \Refinement into a specific decomposition\ text\We extend the decomposition to lists of words. This can be seen as a refinement of a previous decomposition of some word.\ fun refine :: "'a list set \ 'a list list \ 'a list list" ("Ref _ _" [51,51] 65) where "refine G us = concat(map (decompose G) us)" lemma ref_morph: "us \ lists \G\ \ vs \ lists \G\ \ refine G (us \ vs) = refine G us \ refine G vs" using refine.simps by simp lemma ref_morph_plus: "us \ lists \G\\<^sub>+ \ vs \ lists \G\\<^sub>+ \ refine G (us \ vs) = refine G us \ refine G vs" - using refine.simps by simp + using refine.simps by simp lemma ref_pop_hd: "us \ \ \ us \ lists \G\ \ refine G us = decompose G (hd us) \ refine G (tl us)" - unfolding refine.simps using list.simps(9)[of "decompose G" "hd us" "tl us"] by simp + unfolding refine.simps using list.simps(9)[of "decompose G" "hd us" "tl us"] by simp lemma ref_in: "us \ lists \G\ \ (Ref G us) \ lists G\<^sub>+" proof (induction us, simp) case (Cons a us) then show ?case using Cons.IH Cons.prems decI' by auto -qed +qed lemma ref: "us \ lists \G\ \ concat (Ref G us) = concat us" proof (induction us, simp) case (Cons a us) then show ?case using Cons.IH Cons.prems decI by auto qed -lemma ref_gen: "us \ lists B \ B \ \G\ \ Ref G us \ \decompose G ` B\" +lemma ref_gen: "us \ lists B \ B \ \G\ \ Ref G us \ \decompose G ` B\" by (induct us, auto) lemma emp_ref: assumes "us \ lists \G\\<^sub>+" and "Ref G us = \" shows "us = \" using emp_concat_emp[OF \us \ lists \G\\<^sub>+\] - ref[OF lists_drop_emp[OF assms(1)], unfolded \Ref G us = \\ concat.simps(1),symmetric] by blast + ref[OF lists_drop_emp[OF assms(1)], unfolded \Ref G us = \\ concat.simps(1),symmetric] by blast -lemma sing_ref_sing: - assumes "us \ lists \G\\<^sub>+" and "refine G us = [b]" +lemma sing_ref_sing: + assumes "us \ lists \G\\<^sub>+" and "refine G us = [b]" shows "us = [b]" proof- have "us \ \" using \refine G us = [b]\ by auto have "tl us \ lists \G\\<^sub>+" and "hd us \ \G\\<^sub>+" using list.collapse[OF \us \ \\] \us \ lists \G\\<^sub>+\ Cons_in_lists_iff[of "hd us" "tl us" "\G\\<^sub>+"] by auto - have "Dec G (hd us) \ \" + have "Dec G (hd us) \ \" using dec_nemp[OF \hd us \ \G\\<^sub>+\]. have "us \ lists \G\" using \us \ lists \G\\<^sub>+\ lists_drop_emp by auto have "concat us = b" - using \us \ lists \G\\ assms(2) ref by force + using \us \ lists \G\\ assms(2) ref by force have "refine G (tl us) = \" - using ref_pop_hd[OF \us \ \\ \us \ lists \G\\] unfolding \refine G us = [b]\ + using ref_pop_hd[OF \us \ \\ \us \ lists \G\\] unfolding \refine G us = [b]\ using \Dec G (hd us) \ \\ Cons_eq_append_conv[of b \ "(Dec G (hd us))" "(Ref G (tl us))"] Cons_eq_append_conv[of b \ "(Dec G (hd us))" "(Ref G (tl us))"] append_is_Nil_conv[of _ "(Ref G (tl us))"] by blast from emp_ref[OF \tl us \ lists \G\\<^sub>+\ this, symmetric] have "\ = tl us". from this[unfolded Nil_tl] show ?thesis - using \us \ \\ \concat us = b\ by auto + using \us \ \\ \concat us = b\ by auto qed -lemma ref_ex: assumes "Q \ \G\" and "us \ lists Q" +lemma ref_ex: assumes "Q \ \G\" and "us \ lists Q" shows "Ref G us \ lists G\<^sub>+" and "concat (Ref G us) = concat us" using ref_in[OF sub_lists_mono[OF assms]] ref[OF sub_lists_mono[OF assms]]. section "Basis" text\An important property of monoids of words is that they have a unique minimal generating set. Which is the set consisting of indecomposable elements.\ text\The simple element is defined as a word which has only trivial decomposition into generators: a singleton.\ function simple_element :: "'a list \ 'a list set \ bool" (" _ \B _ " [51,51] 50) where "simple_element b G = (b \ G \ (\ us. us \ lists G\<^sub>+ \ concat us = b \ \<^bold>|us\<^bold>| = 1))" - using prod.exhaust by auto + using prod.exhaust by auto termination - using "termination" by blast + using "termination" by blast lemma simp_el_el: "b \B G \ b \ G" unfolding simple_element.simps by blast lemma simp_elD: "b \B G \ us \ lists G\<^sub>+ \ concat us = b \ \<^bold>|us\<^bold>| = 1" unfolding simple_element.simps by blast lemma simp_el_sing: assumes "b \B G" "us \ lists G\<^sub>+" "concat us = b" shows "us = [b]" - using simp_elD[OF assms] \concat us = b\ concat_len_one sing_word by fastforce + using simp_elD[OF assms] \concat us = b\ concat_len_one sing_word by fastforce lemma nonsimp: "us \ lists G\<^sub>+ \ concat us \B G \ us = [concat us]" - using simp_el_sing[of "concat us" G us] unfolding simple_element.simps + using simp_el_sing[of "concat us" G us] unfolding simple_element.simps by blast -lemma emp_nonsimp: "\ \ \B G" +lemma emp_nonsimp: "\ \ \B G" unfolding simple_element.simps using list.size(3) concat.simps(1) lists.Nil[of "G\<^sub>+"] by fastforce lemma basis_no_fact: assumes "u \ \G\" and "v \ \G\" and "u \ v \B G" shows "u = \ \ v = \" proof- have eq1: "concat ((Dec G u) \ (Dec G v)) = u \ v" - using concat_morph[of "Dec G u" "Dec G v",symmetric] + using concat_morph[of "Dec G u" "Dec G v",symmetric] unfolding decI[OF \u \ \G\\] decI[OF \v \ \G\\]. have eq2: "(Dec G u) \ (Dec G v) = [u \ v]" - using \u \ v \B G\ nonsimp[of "(Dec G u) \ (Dec G v)"] + using \u \ v \B G\ nonsimp[of "(Dec G u) \ (Dec G v)"] unfolding eq1 append_in_lists_conv[of "(Dec G u)" "(Dec G v)" "G\<^sub>+"] using decI'[OF \u \ \G\\] decI'[OF \v \ \G\\] - by (meson append_in_lists_conv) + by (meson append_in_lists_conv) have "Dec G u = \ \ Dec G v = \" - using butlast_append[of "Dec G u" "Dec G v"] unfolding eq2 butlast.simps(2)[of "u\v" \] - using Nil_is_append_conv[of "Dec G u" "butlast (Dec G v)"] by auto + using butlast_append[of "Dec G u" "Dec G v"] unfolding eq2 butlast.simps(2)[of "u\v" \] + using Nil_is_append_conv[of "Dec G u" "butlast (Dec G v)"] by auto thus ?thesis using decI[OF \u \ \G\\] decI[OF \v \ \G\\] concat.simps(1) - by auto + by auto qed lemma simp_elI: assumes "b \ G" and "b \ \" and all: "\ u v. u \ \ \ u \ \G\ \ v \ \ \ v \ \G\ \ u \ v \ b" shows "b \B G" unfolding simple_element.simps proof(simp add: \b \ G\, standard, standard, elim conjE) fix us assume "us \ lists G\<^sub>+" "concat us = b" hence "us \ \" using \b \ \\ concat.simps(1) by blast hence "hd us \ \G\" and "hd us \ \" - using \us \ lists G\<^sub>+\ lists_hd gen_in by auto + using \us \ lists G\<^sub>+\ lists_hd gen_in by auto have "tl us = \" proof(rule ccontr) assume "tl us \ \" from nemp_concat_hull[of "tl us", OF this tl_lists[OF \us \ lists G\<^sub>+\]] show False - using all \hd us \ \\ \hd us \ \G\\ concat.simps(2)[of "hd us" "tl us", symmetric] + using all \hd us \ \\ \hd us \ \G\\ concat.simps(2)[of "hd us" "tl us", symmetric] unfolding list.collapse[OF \us \ \\] \concat us = b\ by blast qed hence "\<^bold>|us\<^bold>| = 1" using \concat us = b\ assms(2) long_list_tl nonsing_concat_len by blast thus "\<^bold>|us\<^bold>| = Suc 0" by (simp add: \b \ G\) qed -lemma simp_el_indecomp: - assumes "b \B G" +lemma simp_el_indecomp: + assumes "b \B G" shows "b \ G" and "b \ \" and "\ u v. u \ \ \ u \ \G\ \ v \ \ \ v \ \G\ \ u \ v \ b" using assms basis_no_fact emp_nonsimp simple_element.simps by blast+ text\We are ready to define the \emph{basis} as the set of all simple elements.\ fun basis :: "'a list set \ 'a list set" ("\ _" [51] ) where basisdef: "basis G = {x. x \B G}" lemma basisI: "x \B G \ x \ \ G" by simp lemma basisD: "x \ \ G \ x \B G" by simp lemma emp_not_basis: "x \ \ G \ x \ \" - using basisD emp_nonsimp by blast + using basisD emp_nonsimp by blast lemma basis_sub: "\ G \ G" using basisdef by simp lemma basis_drop_emp: "(\ G)\<^sub>+ = \ G" using emp_not_basis by blast -lemma simp_el_hull': assumes "b \B \G\" shows "b \B G" +lemma simp_el_hull': assumes "b \B \G\" shows "b \B G" proof- have all: "\us. us \ lists G\<^sub>+ \ concat us = b \ \<^bold>|us\<^bold>| = 1" using assms lists_gen_to_hull unfolding simple_element.simps by metis have "b \ \G\" using assms simp_elD by auto obtain bs where "bs \ lists G\<^sub>+" and "concat bs = b" using dec_ex[OF \b \ \G\\] by blast have "b \ G" - using lists_drop_emp[OF \bs \ lists G\<^sub>+\] - lists_gen_to_hull[OF \bs \ lists G\<^sub>+\, THEN nonsimp[of bs "\G\"], + using lists_drop_emp[OF \bs \ lists G\<^sub>+\] + lists_gen_to_hull[OF \bs \ lists G\<^sub>+\, THEN nonsimp[of bs "\G\"], unfolded \concat bs = b\, OF \b \B \G\\] by simp thus "b \B G" by (simp add: all) qed lemma simp_el_hull: assumes "b \B G" shows "b \B \G\" - using simp_elI[of b "\G\"] unfolding self_gen + using simp_elI[of b "\G\"] unfolding self_gen using assms gen_in simp_el_indecomp[OF \b \B G\] by auto lemma concat_tl_basis: "x # xs \ lists \ G \ concat xs \ \G\" unfolding hull_concat_lists by auto -text\The basis generates the hull\ +text\The basis generates the hull\ lemma set_concat_len: assumes "us \ lists G\<^sub>+" "1 < \<^bold>|us\<^bold>|" "u \ set us" shows "\<^bold>|u\<^bold>| < \<^bold>|concat us\<^bold>|" proof- obtain x y where "us = x \ [u] \ y" and "x \ y \ \" - using split_list_long[OF \1 < \<^bold>|us\<^bold>|\ \u \ set us\]. + using split_list_long[OF \1 < \<^bold>|us\<^bold>|\ \u \ set us\]. hence "x \ y \ lists G\<^sub>+" - using \us \ lists G\<^sub>+\ by auto - hence "\<^bold>|concat (x \ y)\<^bold>| \ 0" - using \x \ y \ \\ in_lists_conv_set by force + using \us \ lists G\<^sub>+\ by auto + hence "\<^bold>|concat (x \ y)\<^bold>| \ 0" + using \x \ y \ \\ in_lists_conv_set by force hence "\<^bold>|concat us\<^bold>| = \<^bold>|u\<^bold>| + \<^bold>|concat x\<^bold>| + \<^bold>|concat y\<^bold>|" using length_append \us = x \ [u] \ y\ by simp thus ?thesis using \\<^bold>|concat (x \ y)\<^bold>| \ 0\ by auto -qed +qed lemma non_simp_dec: assumes "w \ \ G" "w \ \" "w \ G" obtains us where "us \ lists G\<^sub>+" "1 < \<^bold>|us\<^bold>|" "concat us = w" - using \w \ \\ \w \ G\ \w \ \ G\ nonsing_concat_len basisI[of w G, unfolded simple_element.simps] + using \w \ \\ \w \ G\ \w \ \ G\ nonsing_concat_len basisI[of w G, unfolded simple_element.simps] by blast lemma basis_gen: "w \ G \ w \ \\ G\" proof (induct "length w" arbitrary: w rule: less_induct) case less {assume "w \ \ G" "w \ \" "w \ G" obtain us where "us \ lists G\<^sub>+" "1 < \<^bold>|us\<^bold>|" "concat us = w" - using non_simp_dec[OF \w \ \ G\ \w \ \\ \w \ G\] by blast + using non_simp_dec[OF \w \ \ G\ \w \ \\ \w \ G\] by blast have "u \ set us \ u \ \\ G\" for u - using lists_drop_emp[OF \us \ lists G\<^sub>+\] - set_concat_len[OF \us \ lists G\<^sub>+\ \1 < \<^bold>|us\<^bold>|\, THEN less[unfolded \concat us = w\[symmetric], of u]] + using lists_drop_emp[OF \us \ lists G\<^sub>+\] + set_concat_len[OF \us \ lists G\<^sub>+\ \1 < \<^bold>|us\<^bold>|\, THEN less[unfolded \concat us = w\[symmetric], of u]] unfolding in_lists_conv_set[of us G] by blast from subsetI[of "set us", OF this] - have ?case + have ?case using concat_in_hull[of us "\\ G\", unfolded self_gen \concat us = w\] by blast } thus ?case - by auto + by auto qed theorem basis_gen_hull: "\\ G\ = \G\" proof show "\\ G\ \ \G\" unfolding hull_concat_lists by auto - show "\G\ \ \\ G\" - proof + show "\G\ \ \\ G\" + proof fix x show "x \ \G\ \ x \ \\ G\" proof (induct rule: hull.induct) show "\w1 w2. w1 \ G \ w2 \ \\ G\ \ w1 \ w2 \ \\ G\" - using hull_closed[of _ "\ G"] basis_gen[of _ G] by blast + using hull_closed[of _ "\ G"] basis_gen[of _ G] by blast qed auto qed qed lemma basis_gen_hull': "\\ \G\\ = \G\" using basis_gen_hull self_gen by blast theorem basis_of_hull: "\ G = \ \G\" proof show "\ G \ \ \G\" using basisD basisI simp_el_hull by blast show "\ \G\ \ \ G" using basisD basisI simp_el_hull' by blast qed text\The basis is the smallest generating set.\ theorem "\S\ = \G\ \ \ G \ S" - by (metis basis_of_hull basis_sub) + by (metis basis_of_hull basis_sub) text\An arbitrary set between basis and the hull is generating...\ lemma gen_sets: assumes "\ G \ S" and "S \ \G\" shows "\S\ = \G\" using image_mono[OF lists_mono[of S "\G\"], of concat, OF \S \ \G\\] image_mono[OF lists_mono[of "\ G" S], of concat, OF \\ G \ S\] - unfolding sym[OF hull_concat_lists] basis_gen_hull - using subset_antisym[of "\S\" "\G\"] self_gen by auto + unfolding sym[OF hull_concat_lists] basis_gen_hull + using subset_antisym[of "\S\" "\G\"] self_gen by auto text\... and has the same basis\ lemma basis_sets: "\ G \ S \ S \ \G\ \ \ G = \ S" by (metis basis_of_hull gen_sets) text\Any nonempty composed element has a decomposition into basis elements with many useful properties\ lemma non_simp_fac: assumes "w \ \" and "w \ \G\" and "w \ \ G" - obtains us where "1 < \<^bold>|us\<^bold>|" and "us \ \" and "us \ lists \ G" and - "hd us \ \" and "hd us \ \G\" and - "concat(tl us) \ \" and "concat(tl us) \ \G\" and + obtains us where "1 < \<^bold>|us\<^bold>|" and "us \ \" and "us \ lists \ G" and + "hd us \ \" and "hd us \ \G\" and + "concat(tl us) \ \" and "concat(tl us) \ \G\" and "w = hd us \ concat(tl us)" proof- obtain us where "us \ lists \ G" and "concat us = w" using \w \ \G\\ dec_dom'[of w "\ G"] decI[of w "\ G"] unfolding basis_gen_hull by blast hence "us \ \" using \w \ \\ concat.simps(1) by blast from lists_hd[OF this \us \ lists \ G\, THEN emp_not_basis] lists_hd[OF this \us \ lists \ G\, THEN gen_in[of "hd us" "\ G", unfolded basis_gen_hull]] have "hd us \ \" and "hd us \ \G\". have "1 < \<^bold>|us\<^bold>|" - using \w \ \ G\ lists_hd[OF \us \ \\ \us \ lists \ G\] \w \ \\ \w \ \G\\ + using \w \ \ G\ lists_hd[OF \us \ \\ \us \ lists \ G\] \w \ \\ \w \ \G\\ concat_len_one[of us, unfolded \concat us = w\] nonsing_concat_len[of us, unfolded \concat us = w\] by blast from nemp_concat_hull[OF long_list_tl[OF this], of "\ G", unfolded basis_drop_emp basis_gen_hull, OF tl_lists[OF \us \ lists \ G\]] have "concat (tl us) \ \G\" and "concat(tl us) \ \". have "w = hd us \ concat(tl us)" using \us \ \\ \us \ lists \ G\ \concat us = w\ concat.simps(2)[of "hd us" "tl us"] list.collapse[of us] by argo - from that[OF \1 < \<^bold>|us\<^bold>|\ \us \ \\ \us \ lists \ G\ \hd us \ \\ \hd us \ \G\\ \concat (tl us) \ \\ \concat (tl us) \ \G\\ this] + from that[OF \1 < \<^bold>|us\<^bold>|\ \us \ \\ \us \ lists \ G\ \hd us \ \\ \hd us \ \G\\ \concat (tl us) \ \\ \concat (tl us) \ \G\\ this] show thesis. qed -lemma basis_dec: "p \ \G\ \ s \ \G\ \ p \ s \ \ G \ p = \ \ s = \" - using basis_no_fact[of p G s] by simp +lemma basis_dec: "p \ \G\ \ s \ \G\ \ p \ s \ \ G \ p = \ \ s = \" + using basis_no_fact[of p G s] by simp lemma non_simp_fac': "w \ \ G \ w \ \ \ w \ \G\ \ \us. us \ lists G\<^sub>+ \ w = concat us \ \<^bold>|us\<^bold>| \ 1" by (metis basisI concat_len_one decI' dec_dom' decI dec_nemp lists_hd nemp_elem_setI simple_element.elims(3)) lemma emp_gen_iff: "G\<^sub>+ = {} \ \G\ = {\}" proof assume "G\<^sub>+ = {}" show "\G\ = {\}" using hull_drop_one[of G, unfolded \G\<^sub>+ = {}\ emp_gen_set]. next assume "\G\ = {\}" thus"G\<^sub>+ = {}" by blast qed lemma emp_basis_iff: "\ G = {} \ G\<^sub>+ = {}" - using emp_gen_iff[of "\ G", unfolded basis_gen_hull basis_drop_emp, folded emp_gen_iff]. + using emp_gen_iff[of "\ G", unfolded basis_gen_hull basis_drop_emp, folded emp_gen_iff]. section "Code" -text\A basis freely generating its hull is called a \emph{code}. By definition, +text\A basis freely generating its hull is called a \emph{code}. By definition, this means that generated elements have unique factorizations into the elements of the code.\ locale code = fixes \ assumes \_is_code: "xs \ lists \ \ ys \ lists \ \ concat xs = concat ys \ xs = ys" begin -lemma emp_not_in_code: "\ \ \" +lemma emp_not_in_code: "\ \ \" proof assume "\ \ \" hence "[] \ lists \" and "[\] \ lists \" and "concat [] = concat [\]" and "[] \ [\]" by simp+ - thus False + thus False using \_is_code by blast qed -lemma code_simple: "c \ \ \ c \B \" - unfolding simple_element.simps -proof - fix c assume "c \ \" +lemma code_simple: "c \ \ \ c \B \" + unfolding simple_element.simps +proof + fix c assume "c \ \" hence "[c] \ lists \" by simp show "\us. us \ lists \\<^sub>+ \ concat us = c \ \<^bold>|us\<^bold>| = 1" proof fix us {assume "us \ lists \\<^sub>+" "concat us = c" hence "us \ lists \" by blast - hence "us = [c]" + hence "us = [c]" using \concat us = c\ \c \ \\ \_is_code[of "[c]", OF \[c] \ lists \\ \us \ lists \\] emp_not_in_code by auto} thus "us \ lists \\<^sub>+ \ concat us = c \ \<^bold>|us\<^bold>| = 1" - using sing_len[of c] by fastforce + using sing_len[of c] by fastforce qed qed lemma code_is_basis: "\ \ = \" using code_simple basisdef[of \] basis_sub by blast lemma code_unique_dec: "us \ lists \ \ Dec \ (concat us) = us" - using dec_dom'[of "concat us" \, THEN \_is_code, of us] + using dec_dom'[of "concat us" \, THEN \_is_code, of us] decI[of "concat us" \] hull_concat_lists[of \] image_eqI[of "concat us" concat us "lists \"] by argo lemma code_unique_ref: "us \ lists \\\ \ refine \ us = decompose \ (concat us)" proof- assume "us \ lists \\\" hence "concat (refine \ us) = concat us" using ref by fastforce hence eq: "concat (refine \ us) = concat (decompose \ (concat us))" - using decI[OF hull_closed_lists[OF \us \ lists \\\\]] by auto + using decI[OF hull_closed_lists[OF \us \ lists \\\\]] by auto have dec: "Dec \ (concat us) \ lists \" using \us \ lists \\\\ dec_dom' hull_closed_lists by blast have ref: "Ref \ us \ lists \" using lists_drop_emp[OF ref_in[OF \us \ lists \\\\]]. show ?thesis using \_is_code[OF ref dec eq]. qed -lemma code_dec_morph: assumes "x \ \\\" "y \ \\\" +lemma code_dec_morph: assumes "x \ \\\" "y \ \\\" shows "(Dec \ x) \ (Dec \ y) = Dec \ (x\y)" proof- have eq: "(Dec \ x) \ (Dec \ y) = Dec \ (concat ((Dec \ x) \ (Dec \ y)))" using dec_dom'[OF \x \ \\\\] dec_dom'[OF \y \ \\\\] code.code_unique_dec[OF code_axioms, of "(Dec \ x) \ (Dec \ y)", unfolded append_in_lists_conv, symmetric] - by blast + by blast moreover have "concat ((Dec \ x) \ (Dec \ y)) = (x \ y)" using concat_morph[of "Dec \ x" "Dec \ y", symmetric] unfolding decI[OF \x \ \\\\] decI[OF \y \ \\\\]. ultimately show "(Dec \ x) \ (Dec \ y) = Dec \ (x\y)" by argo qed lemma code_el_dec: "c \ \ \ decompose \ c = [c]" using code_unique_dec[of "[c]"] by auto lemma code_ref_list: "us \ lists \ \ refine \ us = us" proof (induct us) case (Cons a us) then show ?case using code_el_dec by simp qed simp -lemma code_ref_gen: assumes "G \ \\\" "u \ \G\" +lemma code_ref_gen: assumes "G \ \\\" "u \ \G\" shows "Dec \ u \ \decompose \ ` G\" proof- have "refine \ (Dec G u) = Dec \ u" using dec_dom'[OF \u \ \G\\] \G \ \\\\ code_unique_ref[of "Dec G u", unfolded decI[OF \u \ \G\\]] by blast from ref_gen[of "Dec G u" G, OF dec_dom'[OF \u \ \G\\], of \, unfolded this, OF \G \ \\\\] show ?thesis. qed end \ \end context code\ section \Binary code\ -text\We pay a special attention to two element codes. -In particular, we show that two words form a code if and only if they do not commute. This means that two +text\We pay a special attention to two element codes. +In particular, we show that two words form a code if and only if they do not commute. This means that two words either commute, or do not satisfy any nontrivial relation. \ -locale binary_code = +locale binary_code = fixes u\<^sub>0 u\<^sub>1 - assumes non_comm: "u\<^sub>0 \ u\<^sub>1 \ u\<^sub>1 \ u\<^sub>0" + assumes non_comm: "u\<^sub>0 \ u\<^sub>1 \ u\<^sub>1 \ u\<^sub>0" begin -lemma bin_fst_nemp: "u\<^sub>0 \ \" - using non_comm by auto +lemma bin_fst_nemp: "u\<^sub>0 \ \" + using non_comm by auto -text\A crucial property of two element codes is the constant decoding delay given by the word $\alpha$, -which is a prefix of any generating word (sufficiently long), while the letter +text\A crucial property of two element codes is the constant decoding delay given by the word $\alpha$, +which is a prefix of any generating word (sufficiently long), while the letter immediately after this common prefix indicates the first element of the decomposition. \ definition \ where bin_lcp_def [simp]: "\ = u\<^sub>0\u\<^sub>1 \\<^sub>p u\<^sub>1\u\<^sub>0" definition c\<^sub>0 where fst_mismatch_def: "c\<^sub>0 = (u\<^sub>0\u\<^sub>1)!\<^bold>|\\<^bold>|" definition c\<^sub>1 where snd_mismatch_def: "c\<^sub>1 = (u\<^sub>1\u\<^sub>0)!\<^bold>|\\<^bold>|" lemma bin_mismatch_neq: "c\<^sub>0 \ c\<^sub>1" - unfolding fst_mismatch_def snd_mismatch_def bin_lcp_def + unfolding fst_mismatch_def snd_mismatch_def bin_lcp_def using non_comm lcp_mismatch' pref_comp_eq[of "u\<^sub>0 \ u\<^sub>1" "u\<^sub>1 \ u\<^sub>0", OF _ swap_len] unfolding prefix_comparable_def - by blast + by blast lemma bin_lcp_pref_fst_snd: "\ \p u\<^sub>0 \ u\<^sub>1" and bin_lcp_pref_snd_fst: "\ \p u\<^sub>1 \ u\<^sub>0" unfolding bin_lcp_def using longest_common_prefix_prefix1 longest_common_prefix_prefix2. lemma bin_lcp_short: "\<^bold>|\\<^bold>| < \<^bold>|u\<^sub>0\<^bold>| + \<^bold>|u\<^sub>1\<^bold>|" proof- have "\ u\<^sub>0\u\<^sub>1 \p u\<^sub>1\u\<^sub>0" using comm_ruler non_comm by blast from lcp_len'[OF this, folded bin_lcp_def, unfolded length_append] show "\<^bold>|\\<^bold>| < \<^bold>|u\<^sub>0\<^bold>| + \<^bold>|u\<^sub>1\<^bold>|". qed lemma bin_lcp_fst_mismatch_pref: "\ \ [c\<^sub>0] \p u\<^sub>0 \ \" proof- have "\ \ [c\<^sub>0] \p u\<^sub>0 \ u\<^sub>1" using append_one_prefix[of \ "u\<^sub>0 \ u\<^sub>1", folded fst_mismatch_def, OF bin_lcp_pref_fst_snd, unfolded length_append, OF bin_lcp_short]. hence "\ \ [c\<^sub>0] \p u\<^sub>0 \ (u\<^sub>1 \ u\<^sub>0)" using pref_prolong by blast from pref_prod_pref_short[OF this bin_lcp_pref_snd_fst, unfolded length_append sing_len] show "\ \ [c\<^sub>0] \p u\<^sub>0 \ \" using nemp_len[OF bin_fst_nemp] by linarith qed -lemma not_fst_snd_pref: "\ u\<^sub>0 \ u\<^sub>1 \p \" - using bin_lcp_short[folded length_append[of u\<^sub>0 u\<^sub>1]] prefix_order.order.antisym[OF bin_lcp_pref_fst_snd] by fastforce +lemma not_fst_snd_pref: "\ u\<^sub>0 \ u\<^sub>1 \p \" + using bin_lcp_short[folded length_append[of u\<^sub>0 u\<^sub>1]] prefix_order.antisym[OF bin_lcp_pref_fst_snd] by fastforce lemma bin_lcp_fst_mismatch_pref': "\ \ [c\<^sub>0] \p u\<^sub>0 \ u\<^sub>1" using strict_prefixI[OF bin_lcp_pref_fst_snd, THEN add_nth_pref, folded fst_mismatch_def] not_fst_snd_pref self_pref[of \] by fastforce interpretation symcode: binary_code u\<^sub>1 u\<^sub>0 rewrites "symcode.c\<^sub>0 = c\<^sub>1" and "symcode.c\<^sub>1 = c\<^sub>0" and "symcode.\ = \" proof- show "binary_code u\<^sub>1 u\<^sub>0" - unfolding binary_code_def using non_comm by simp + unfolding binary_code_def using non_comm by simp show "binary_code.\ u\<^sub>1 u\<^sub>0 = \" by (simp add: \binary_code u\<^sub>1 u\<^sub>0\ binary_code.bin_lcp_def lcp_sym) show "binary_code.c\<^sub>0 u\<^sub>1 u\<^sub>0 = c\<^sub>1" by (simp add: \binary_code u\<^sub>1 u\<^sub>0\ \binary_code.\ u\<^sub>1 u\<^sub>0 = \\ binary_code.fst_mismatch_def snd_mismatch_def) show "binary_code.c\<^sub>1 u\<^sub>1 u\<^sub>0 = c\<^sub>0" by (simp add: \binary_code u\<^sub>1 u\<^sub>0\ \binary_code.\ u\<^sub>1 u\<^sub>0 = \\ binary_code.snd_mismatch_def fst_mismatch_def) qed lemmas bin_snd_nemp = symcode.bin_fst_nemp and bin_snd_mismatch = symcode.bin_lcp_fst_mismatch_pref lemma bin_lcp_fst_lcp: "\ \p u\<^sub>0 \ \" and bin_lcp_snd_lcp: "\ \p u\<^sub>1 \ \" using bin_lcp_fst_mismatch_pref bin_snd_mismatch by auto lemma bin_all_nemp: "ws \ lists {u\<^sub>0,u\<^sub>1} \ concat ws = \ \ ws = \" using bin_fst_nemp bin_snd_nemp by(induct ws, simp, auto) lemma bin_lcp_all_lcp: "ws \ lists {u\<^sub>0,u\<^sub>1} \ \ \p concat ws \ \" proof(induct ws rule: rev_induct, simp) case (snoc x xs) have x_or: "x = u\<^sub>0 \ x = u\<^sub>1" using \xs \ [x] \ lists {u\<^sub>0, u\<^sub>1}\ by simp have "xs \ lists {u\<^sub>0, u\<^sub>1}" using \xs \ [x] \ lists {u\<^sub>0, u\<^sub>1}\ by auto - from pref_prolong[OF snoc.hyps, OF this, of "x\\", unfolded lassoc] + from pref_prolong[OF snoc.hyps, OF this, of "x\\", unfolded lassoc] show ?case using bin_lcp_fst_lcp bin_lcp_snd_lcp disjE[OF x_or] by auto qed -lemma bin_code_alpha: assumes "us \ lists {u\<^sub>0,u\<^sub>1}" and "vs \ lists {u\<^sub>0,u\<^sub>1}" and "hd us \ hd vs" +lemma bin_code_alpha: assumes "us \ lists {u\<^sub>0,u\<^sub>1}" and "vs \ lists {u\<^sub>0,u\<^sub>1}" and "hd us \ hd vs" shows "concat us \ \ \\<^sub>p concat vs \ \ = \" using assms proof (induct us vs rule: list_induct2', simp) case (2 x xs) show ?case using bin_lcp_all_lcp[OF \x # xs \ lists {u\<^sub>0, u\<^sub>1}\, folded lcp_pref_conv, unfolded lcp_sym[of \]] by simp next case (3 y ys) show ?case using bin_lcp_all_lcp[OF \y # ys \ lists {u\<^sub>0, u\<^sub>1}\, folded lcp_pref_conv] by simp next case (4 x xs y ys) interpret i: binary_code x y - using "4.prems"(1) "4.prems"(2) "4.prems"(3) non_comm binary_code.intro by auto + using "4.prems"(1) "4.prems"(2) "4.prems"(3) non_comm binary_code.intro by auto have alph: "{u\<^sub>0,u\<^sub>1} = {x,y}" using "4.prems"(1) "4.prems"(2) "4.prems"(3) by auto from disjE[OF this[unfolded doubleton_eq_iff]] have \id: "i.\ = \" unfolding bin_lcp_def i.bin_lcp_def using lcp_sym by auto have c0: "i.\ \ [i.c\<^sub>0] \p x \ concat xs \ i.\" - using i.bin_lcp_all_lcp[of xs] \x # xs \ lists {u\<^sub>0, u\<^sub>1}\[unfolded Cons_in_lists_iff alph] + using i.bin_lcp_all_lcp[of xs] \x # xs \ lists {u\<^sub>0, u\<^sub>1}\[unfolded Cons_in_lists_iff alph] pref_prolong[OF i.bin_lcp_fst_mismatch_pref] by blast have c1: "i.\ \ [i.c\<^sub>1] \p y \ concat ys \ i.\" - using i.bin_lcp_all_lcp[of ys] \y # ys \ lists {u\<^sub>0, u\<^sub>1}\[unfolded Cons_in_lists_iff alph] + using i.bin_lcp_all_lcp[of ys] \y # ys \ lists {u\<^sub>0, u\<^sub>1}\[unfolded Cons_in_lists_iff alph] pref_prolong[OF i.bin_snd_mismatch] by blast have "i.\\[i.c\<^sub>0] \\<^sub>p i.\\[i.c\<^sub>1] = i.\" by (simp add: i.bin_mismatch_neq lcp_first_mismatch') from lcp_rulers[OF c0 c1, unfolded this, unfolded \id] - show ?case - unfolding concat.simps(2) rassoc pref_cancel_conv using i.bin_mismatch_neq by simp + show ?case + unfolding concat.simps(2) rassoc pref_cancel_conv using i.bin_mismatch_neq by simp qed theorem bin_code: assumes "us \ lists {u\<^sub>0,u\<^sub>1}" and "vs \ lists {u\<^sub>0,u\<^sub>1}" and "concat us = concat vs" shows "us = vs" using assms proof (induct us vs rule: list_induct2', simp) case (2 x xs) then show ?case - using bin_fst_nemp bin_snd_nemp by auto + using bin_fst_nemp bin_snd_nemp by auto next case (3 y ys) - then show ?case - using bin_fst_nemp bin_snd_nemp by auto + then show ?case + using bin_fst_nemp bin_snd_nemp by auto next case (4 x xs y ys) - then show ?case + then show ?case proof(cases "x = y") - assume "x = y" thus "x # xs = y # ys" - using "4.hyps" \concat (x # xs) = concat (y # ys)\[unfolded concat.simps(2) \x = y\, unfolded cancel] + assume "x = y" thus "x # xs = y # ys" + using "4.hyps" \concat (x # xs) = concat (y # ys)\[unfolded concat.simps(2) \x = y\, unfolded cancel] \y # ys \ lists {u\<^sub>0, u\<^sub>1}\[unfolded Cons_in_lists_iff] \x # xs \ lists {u\<^sub>0, u\<^sub>1}\[unfolded Cons_in_lists_iff] - by simp + by simp next assume "x \ y" have "concat(y # ys) = \" using bin_code_alpha[OF \x # xs \ lists {u\<^sub>0, u\<^sub>1}\ \y # ys \ lists {u\<^sub>0, u\<^sub>1}\, unfolded list.sel(1) \concat (x # xs) = concat (y # ys)\, OF \x \ y\] by simp from bin_all_nemp[OF \y # ys \ lists {u\<^sub>0, u\<^sub>1}\ this] have False by simp - thus "x # xs = y # ys" by blast + thus "x # xs = y # ys" by blast qed qed end lemmas no_comm_bin_code = binary_code.bin_code[unfolded binary_code_def] theorem bin_code_code: assumes "u \ v \ v \ u" shows "code {u, v}" - unfolding code_def using no_comm_bin_code[OF assms] by blast + unfolding code_def using no_comm_bin_code[OF assms] by blast section \Free hull\ text\While not every set $G$ of generators is a code, there is a unique minimal free monoid containing it, called the \emph{free hull} of $G$. It can be defined inductively using the property known as the \emph{stability condition}. \ inductive_set free_hull :: "'a list set \ 'a list set" ("\_\\<^sub>F") for G where "\ \ \G\\<^sub>F" | free_gen_in: "w \ G \ w \ \G\\<^sub>F" | "w1 \ \G\\<^sub>F \ w2 \ \G\\<^sub>F \ w1 \ w2 \ \G\\<^sub>F" | "p \ \G\\<^sub>F \ q \ \G\\<^sub>F \ p \ w \ \G\\<^sub>F \ w \ q \ \G\\<^sub>F \ w \ \G\\<^sub>F" \ \the stability condition\ lemmas [intro] = free_hull.intros text\The defined set indeed is a hull.\ lemma free_hull_hull: "\\G\\<^sub>F\ = \G\\<^sub>F" proof show "\G\\<^sub>F \ \\G\\<^sub>F\" by (simp add: genset_sub) show "\\G\\<^sub>F\ \ \G\\<^sub>F" proof - fix x assume "x \ \\G\\<^sub>F\" + fix x assume "x \ \\G\\<^sub>F\" thus "x \ \G\\<^sub>F" proof (rule hull.induct) show " \ \ \G\\<^sub>F" by (simp add: free_hull.intros(1)) show "\w1 w2. w1 \ \G\\<^sub>F \ w2 \ \\G\\<^sub>F\ \ w2 \ \G\\<^sub>F \ w1 \ w2 \ \G\\<^sub>F" by (simp add: free_hull.intros(3)) qed qed qed text\The free hull is always (non-strictly) larger than the hull.\ lemma hull_in_free_hull: "\G\ \ \G\\<^sub>F" proof fix x assume "x \ \G\" - then show "x \ \G\\<^sub>F" - using free_hull.intros(3) + then show "x \ \G\\<^sub>F" + using free_hull.intros(3) hull_induct[of x G "\ x. x \ \G\\<^sub>F", OF \x \ \G\\ free_hull.intros(1)[of G] free_hull.intros(2)] by auto qed text\On the other hand, it can be proved that the \emph{free basis}, defined as the basis of the free hull, has a (non-strictly) smaller cardinality than the ordinary basis.\ definition free_basis :: "'a list set \ 'a list set" ("\\<^sub>F _" [54] 55) where "free_basis G \ \ \G\\<^sub>F" lemma basis_gen_hull_free: "\\\<^sub>F G\ = \G\\<^sub>F" unfolding free_basis_def using basis_gen_hull free_hull_hull by blast lemma genset_sub_free: "G \ \G\\<^sub>F" by (simp add: free_hull.free_gen_in subsetI) text -\We have developed two points of view on freeness: +\We have developed two points of view on freeness: \<^item> being a free hull, that is, to satisfy the stability condition; \<^item> being generated by a code.\ - + text\We now show their equivalence\ text\First, basis of a free hull is a code.\ lemma free_basis_code: "code (\\<^sub>F G)" proof - fix xs ys + fix xs ys show "xs \ lists (\\<^sub>F G) \ ys \ lists (\\<^sub>F G) \ concat xs = concat ys \ xs = ys" proof(induction xs ys rule: list_induct2', simp) case (2 x xs) - show ?case - using listsE[OF \x # xs \ lists (\\<^sub>F G)\, of "x \ \\<^sub>F G", unfolded free_basis_def, THEN emp_not_basis] - concat.simps(2)[of x xs, unfolded \concat (x # xs) = concat \\[unfolded concat.simps(1)], symmetric, unfolded append_is_Nil_conv[of x "concat xs"]] + show ?case + using listsE[OF \x # xs \ lists (\\<^sub>F G)\, of "x \ \\<^sub>F G", unfolded free_basis_def, THEN emp_not_basis] + concat.simps(2)[of x xs, unfolded \concat (x # xs) = concat \\[unfolded concat.simps(1)], symmetric, unfolded append_is_Nil_conv[of x "concat xs"]] by blast next case (3 y ys) - show ?case - using listsE[OF \y # ys \ lists (\\<^sub>F G)\, of "y \ \\<^sub>F G", unfolded free_basis_def, THEN emp_not_basis] + show ?case + using listsE[OF \y # ys \ lists (\\<^sub>F G)\, of "y \ \\<^sub>F G", unfolded free_basis_def, THEN emp_not_basis] concat.simps(2)[of y ys, unfolded \concat \ = concat (y # ys)\[unfolded concat.simps(1),symmetric],symmetric, unfolded append_is_Nil_conv[of y "concat ys"]] - by blast + by blast next case (4 x xs y ys) have "\<^bold>|x\<^bold>| = \<^bold>|y\<^bold>|" proof(rule ccontr) assume "\<^bold>|x\<^bold>| \ \<^bold>|y\<^bold>|" - have "x \ concat xs = y \ concat ys" + have "x \ concat xs = y \ concat ys" using \concat (x # xs) = concat (y # ys)\ by simp then obtain t where or: "x = y \ t \ t \ concat xs = concat ys \ x \ t = y \ concat xs = t \ concat ys" using append_eq_append_conv2[of x "concat xs" y "concat ys"] by blast hence "t \ \" using \\<^bold>|x\<^bold>| \ \<^bold>|y\<^bold>|\ by auto have "x \ \\<^sub>F G" and "y \ \\<^sub>F G" using listsE[OF \x # xs \ lists (\\<^sub>F G)\, of "x \ \\<^sub>F G" ] listsE[OF \y # ys \ lists (\\<^sub>F G)\, of "y \ \\<^sub>F G" ] by blast+ hence "x \ \" and "y \ \" unfolding free_basis_def using emp_not_basis by blast+ have "x \ \G\\<^sub>F" and "y \ \G\\<^sub>F" - using basis_sub[of "\G\\<^sub>F", unfolded free_basis_def[symmetric] ] \x # xs \ lists (\\<^sub>F G)\ + using basis_sub[of "\G\\<^sub>F", unfolded free_basis_def[symmetric] ] \x # xs \ lists (\\<^sub>F G)\ \y # ys \ lists (\\<^sub>F G)\ by auto have "concat xs \ \G\\<^sub>F" and "concat ys \ \G\\<^sub>F" - using concat_tl_basis[OF \x # xs \ lists (\\<^sub>F G)\[unfolded free_basis_def]] - concat_tl_basis[OF \y # ys \ lists (\\<^sub>F G)\[unfolded free_basis_def]] unfolding free_hull_hull. - have "t \ \G\\<^sub>F" + using concat_tl_basis[OF \x # xs \ lists (\\<^sub>F G)\[unfolded free_basis_def]] + concat_tl_basis[OF \y # ys \ lists (\\<^sub>F G)\[unfolded free_basis_def]] unfolding free_hull_hull. + have "t \ \G\\<^sub>F" using or free_hull.intros(4) \x \ \G\\<^sub>F\ \y \ \G\\<^sub>F\ \concat xs \ \G\\<^sub>F\ \concat ys \ \G\\<^sub>F\ by metis thus False - using or basis_dec[of x "\G\\<^sub>F" t, unfolded free_hull_hull, OF \x \ \G\\<^sub>F\ \t \ \G\\<^sub>F\] - basis_dec[of y "\G\\<^sub>F" t, unfolded free_hull_hull, OF \y \ \G\\<^sub>F\ \t \ \G\\<^sub>F\] + using or basis_dec[of x "\G\\<^sub>F" t, unfolded free_hull_hull, OF \x \ \G\\<^sub>F\ \t \ \G\\<^sub>F\] + basis_dec[of y "\G\\<^sub>F" t, unfolded free_hull_hull, OF \y \ \G\\<^sub>F\ \t \ \G\\<^sub>F\] using \t \ \\ \x \ \\ \y \ \\ \x \ \\<^sub>F G\ \y \ \\<^sub>F G\ unfolding free_basis_def by auto qed thus "x # xs = y # ys" using "4.IH" \x # xs \ lists (\\<^sub>F G)\ \y # ys \ lists (\\<^sub>F G)\ \concat (x # xs) = concat (y # ys)\ by auto next qed qed text\Second, a code generates its free hull.\ lemma code_gen_free_hull: "code C \ \C\\<^sub>F = \C\" proof - assume "code C" + assume "code C" show "\C\ \ \C\\<^sub>F" - using hull_mon[of C "\C\\<^sub>F"] - free_hull.free_gen_in[of _ C] subsetI[of C "\C\\<^sub>F"] + using hull_mon[of C "\C\\<^sub>F"] + free_hull.free_gen_in[of _ C] subsetI[of C "\C\\<^sub>F"] unfolding free_hull_hull[of C] by auto show "\C\\<^sub>F \ \C\" proof - fix x assume "x \ \C\\<^sub>F" + fix x assume "x \ \C\\<^sub>F" have "\ \ \C\" by (simp add: hull.intros(1)) show "x \ \C\" - proof(rule free_hull.induct[of x C],simp add: \x \ \C\\<^sub>F\, (simp add: hull.intros hull_closed)+, + proof(rule free_hull.induct[of x C],simp add: \x \ \C\\<^sub>F\, (simp add: hull.intros hull_closed)+, simp add: gen_in, simp add: hull_closed) fix p q w assume "p \ \C\" "q \ \C\" "p \ w \ \C\" "w \ q \ \C\" - have eq: "(Dec C p) \ (Dec C w \ q) = (Dec C p \ w) \ (Dec C q)" - using code.code_dec_morph[OF \code C\ \p \ \C\\ \w \ q \ \C\\, unfolded lassoc] - unfolding code.code_dec_morph[OF \code C\ \p \ w \ \C\\ \q \ \C\\, symmetric]. + have eq: "(Dec C p) \ (Dec C w \ q) = (Dec C p \ w) \ (Dec C q)" + using code.code_dec_morph[OF \code C\ \p \ \C\\ \w \ q \ \C\\, unfolded lassoc] + unfolding code.code_dec_morph[OF \code C\ \p \ w \ \C\\ \q \ \C\\, symmetric]. have "Dec C p \ Dec C p \ w" - using eqd_comp[OF eq]. - hence "Dec C p \p Dec C p \ w" + using eqd_comp[OF eq]. + hence "Dec C p \p Dec C p \ w" using \p \ w \ \C\\ \p \ \C\\ concat_morph decI prefD pref_antisym triv_pref unfolding prefix_comparable_def by metis then obtain ts where "(Dec C p) \ ts = Dec C p \ w" using lq_pref by blast - hence "ts \ lists C" - using append_in_lists_conv[of "Dec C p" ts C] dec_dom'[OF \p \ w \ \C\\] + hence "ts \ lists C" + using append_in_lists_conv[of "Dec C p" ts C] dec_dom'[OF \p \ w \ \C\\] unfolding \(Dec C p) \ ts = Dec C p \ w\ by blast hence "concat ts = w" - using concat_morph[of "Dec C p" ts] + using concat_morph[of "Dec C p" ts] unfolding \(Dec C p) \ ts = Dec C p \ w\ decI[OF \p \ w \ \C\\] decI[OF \p \ \C\\] by auto thus "w \ \C\" using \ts \ lists C\ by auto qed qed qed text\That is, a code is its own free basis\ lemma code_free_basis: assumes "code C" shows "C = \\<^sub>F C" - using basis_of_hull[of C, unfolded code_gen_free_hull[OF assms, symmetric] - code.code_is_basis[OF assms]] + using basis_of_hull[of C, unfolded code_gen_free_hull[OF assms, symmetric] + code.code_is_basis[OF assms]] unfolding free_basis_def. -text\Moreover, the free hull of G is the smallest code-generated hull containing G. +text\Moreover, the free hull of G is the smallest code-generated hull containing G. In other words, the term free hull is appropriate.\ text\First, several intuitive monotonicity and closure results.\ lemma free_hull_mono: "G \ H \ \G\\<^sub>F \ \H\\<^sub>F" proof assume "G \ H" fix x assume "x \ \G\\<^sub>F" have el: "\ w. w \ G \ w \ \H\\<^sub>F" using \G \ H\ free_hull.free_gen_in by auto show "x \ \H\\<^sub>F" - proof (rule free_hull.induct[of x G], simp add: \x \ \G\\<^sub>F\, simp add: free_hull.intros(1), + proof (rule free_hull.induct[of x G], simp add: \x \ \G\\<^sub>F\, simp add: free_hull.intros(1), simp add: el, simp add: free_hull.intros(3)) show "\p q w. p \ \H\\<^sub>F \ q \ \H\\<^sub>F \ p \ w \ \H\\<^sub>F \ w \ q \ \H\\<^sub>F \ w \ \H\\<^sub>F" - using free_hull.intros(4) by auto + using free_hull.intros(4) by auto qed qed lemma free_hull_idem: "\\G\\<^sub>F\\<^sub>F = \G\\<^sub>F" proof - show "\\G\\<^sub>F\\<^sub>F \ \G\\<^sub>F" - proof + show "\\G\\<^sub>F\\<^sub>F \ \G\\<^sub>F" + proof fix x assume "x \ \\G\\<^sub>F\\<^sub>F" show "x \ \G\\<^sub>F" - proof (rule free_hull.induct[of x "\G\\<^sub>F"], simp add: \x \ \\G\\<^sub>F\\<^sub>F\, + proof (rule free_hull.induct[of x "\G\\<^sub>F"], simp add: \x \ \\G\\<^sub>F\\<^sub>F\, simp add: free_hull.intros(1), simp add: free_hull.intros(2), simp add: free_hull.intros(3)) show "\p q w. p \ \G\\<^sub>F \ q \ \G\\<^sub>F \ p \ w \ \G\\<^sub>F \ w \ q \ \G\\<^sub>F \ w \ \G\\<^sub>F" - using free_hull.intros(4) by auto + using free_hull.intros(4) by auto qed qed next show "\G\\<^sub>F \ \\G\\<^sub>F\\<^sub>F" - using free_hull_hull hull_in_free_hull by auto + using free_hull_hull hull_in_free_hull by auto qed lemma hull_gen_free_hull: "\\G\\\<^sub>F = \G\\<^sub>F" proof show " \\G\\\<^sub>F \ \G\\<^sub>F" using free_hull_idem free_hull_mono hull_in_free_hull by metis next show "\G\\<^sub>F \ \\G\\\<^sub>F" - by (simp add: free_hull_mono genset_sub) + by (simp add: free_hull_mono genset_sub) qed text \Code is also the free basis of its hull.\ lemma code_free_basis_hull: "code C \ C = \\<^sub>F \C\" unfolding free_basis_def using code_free_basis[unfolded free_basis_def] - unfolding hull_gen_free_hull. + unfolding hull_gen_free_hull. text\The minimality of the free hull easily follows.\ theorem free_hull_min: assumes "code C" and "G \ \C\" shows "\G\\<^sub>F \ \C\" - using free_hull_mono[OF \G \ \C\\] unfolding hull_gen_free_hull - unfolding code_gen_free_hull[OF \code C\]. + using free_hull_mono[OF \G \ \C\\] unfolding hull_gen_free_hull + unfolding code_gen_free_hull[OF \code C\]. theorem free_hull_inter: "\G\\<^sub>F = \ {M. G \ M \ M = \M\\<^sub>F}" proof have "X \ {M. G \ M \ M = \M\\<^sub>F} \ \G\\<^sub>F \ X" for X unfolding mem_Collect_eq[of _ "\ M. G \ M \ M = \M\\<^sub>F"] - using free_hull_mono[of G X] by simp - from Inter_greatest[of "{M. G \ M \ M = \M\\<^sub>F}", OF this] - show "\G\\<^sub>F \ \ {M. G \ M \ M = \M\\<^sub>F}" + using free_hull_mono[of G X] by simp + from Inter_greatest[of "{M. G \ M \ M = \M\\<^sub>F}", OF this] + show "\G\\<^sub>F \ \ {M. G \ M \ M = \M\\<^sub>F}" by blast next show " \ {M. G \ M \ M = \M\\<^sub>F} \ \G\\<^sub>F" - by (simp add: Inter_lower free_hull_idem genset_sub_free) + by (simp add: Inter_lower free_hull_idem genset_sub_free) qed text\Decomposition into the free basis is a morphism.\ -lemma free_basis_dec_morph: "u \ \G\\<^sub>F \ v \ \G\\<^sub>F \ +lemma free_basis_dec_morph: "u \ \G\\<^sub>F \ v \ \G\\<^sub>F \ Dec (\\<^sub>F G) (u \ v) = (Dec (\\<^sub>F G) u) \ (Dec (\\<^sub>F G) v)" - using code.code_dec_morph[OF free_basis_code, of u G v, symmetric, + using code.code_dec_morph[OF free_basis_code, of u G v, symmetric, unfolded basis_gen_hull_free[of G]]. section \Lists as the free hull of singletons\ -text\A crucial property of free monoids of words is that they can be seen as lists over the free basis, +text\A crucial property of free monoids of words is that they can be seen as lists over the free basis, instead as lists over the original alphabet.\ abbreviation sings where "sings B \ {[b] | b. b \ B}" -lemma sings_image: "sings B = (\ x. [x]) ` B" +lemma sings_image: "sings B = (\ x. [x]) ` B" using Setcompr_eq_image. -lemma lists_sing_map_concat_ident: "xs \ lists (sings B) \ xs = map (\ x. [x]) (concat xs)" +lemma lists_sing_map_concat_ident: "xs \ lists (sings B) \ xs = map (\ x. [x]) (concat xs)" by (induct xs, simp, auto) lemma code_sings: "code (sings B)" proof - fix xs ys assume xs: "xs \ lists (sings B)" and ys: "ys \ lists (sings B)" - and eq: "concat xs = concat ys" + fix xs ys assume xs: "xs \ lists (sings B)" and ys: "ys \ lists (sings B)" + and eq: "concat xs = concat ys" from lists_sing_map_concat_ident[OF xs, unfolded eq] show "xs = ys" unfolding lists_sing_map_concat_ident[OF ys, symmetric]. qed lemma sings_gen_lists: "\sings B\ = lists B" unfolding hull_concat_lists proof(intro equalityI subsetI, standard) fix xs show "xs \ concat ` lists (sings B) \ \x\set xs. x \ B" - by force + by force assume "xs \ lists B" hence "map (\x. x # \) xs \ lists (sings B)" by force - from imageI[OF this, of concat] + from imageI[OF this, of concat] show "xs \ concat ` lists (sings B)" - unfolding concat_map_sing_ident[of xs]. -qed + unfolding concat_map_sing_ident[of xs]. +qed lemma "sings B = \\<^sub>F (lists B)" using code_free_basis_hull[OF code_sings, of B, unfolded sings_gen_lists]. lemma map_sings: "xs \ lists B \ map (\x. x # \) xs \ lists (sings B)" by (induct xs) auto lemma dec_sings: "xs \ lists B \ Dec (sings B) xs = map (\ x. [x]) xs" using code.code_unique_dec[OF code_sings, of "map (\ x. [x]) xs" B, OF map_sings] unfolding concat_map_sing_ident. end \ No newline at end of file