Theory CoWBasic

(*  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 "HOL-Eisbach.Eisbach_Tools"
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.
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  ("|_|") ― ‹note that it's bold |›
notation (latex output)
  length  ("latex‹\\ensuremath{\\left| ›_latex‹\\right|}›")
notation longest_common_prefix  (infixr "p" 61) ― ‹provided by Sublist.thy›

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 ("ε")


named_theorems emp_simps
lemmas [emp_simps] = append_Nil2 append_Nil list.map(1) list.size(3)

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  ("p")

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)
  where "(prefix_comparable u v)  u ≤p v  v ≤p u"

lemma pref_compI1: "u ≤p v   u  v"
  unfolding prefix_comparable_def..

lemma pref_compI2: "v ≤p u  u  v"
  unfolding prefix_comparable_def..

lemma pref_compE [elim]: assumes "u  v" obtains "u ≤p v" | "v ≤p u"
  using assms unfolding prefix_comparable_def..

lemma pref_compI[intro]: "u ≤p v  v ≤p u  u  v"
  unfolding prefix_comparable_def
  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  ("np" 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" 50)
notation (latex output) strict_prefix  ("<p")
lemmas [simp] = strict_prefix_def

interpretation lcp: semilattice_order "(∧p)" prefix strict_prefix
proof
  fix a b c :: "'a list"
  show "(a p b) p c = a p b p c"
    by (rule prefix_order.antisym)
    (meson longest_common_prefix_max_prefix longest_common_prefix_prefix1 longest_common_prefix_prefix2 prefix_order.trans)+
  show "a p b = b p a"
    by (simp add: longest_common_prefix_max_prefix longest_common_prefix_prefix1 longest_common_prefix_prefix2 prefix_order.antisym)
  show "a p a = a"
    by (simp add: longest_common_prefix_max_prefix longest_common_prefix_prefix1 prefix_order.eq_iff)
  show "a ≤p b = (a = a p b)"
    by (metis longest_common_prefix_max_prefix longest_common_prefix_prefix1 longest_common_prefix_prefix2 prefix_order.dual_order.eq_iff)
  thus "(a <p b) = (a = a p b  a  b)"
    by simp
qed

lemmas sprefI = strict_prefixI

lemma sprefI1[intro]: "v = u  z  z  ε  u <p v"
  by simp

lemma sprefI1'[intro]: "u  z = v  z  ε  u <p v"
  by force

lemma sprefI2[intro]: "u ≤p v  |u| < |v|  u <p v"
  by force

lemma sprefD: "u <p v  u ≤p v  u  v"
  by auto

lemmas sprefD1[dest] = prefix_order.strict_implies_order and
       sprefD2 = prefix_order.less_imp_neq

lemmas sprefE[elim?] = strict_prefixE

lemma spref_exE[elim?]: assumes "u <p v" obtains z where "u  z = v" and "z  ε"
   using assms unfolding strict_prefix_def prefix_def by blast

subsection ‹Suffix›

notation suffix (infixl "≤s" 50)
notation (latex output) suffix ("s")

lemma sufI[intro]: "p  s = w  s ≤s w"
  by (auto simp add: suffix_def)

lemma sufD[elim]: "u ≤s v   z. z  u = v"
  by (auto simp add: suffix_def)


notation strict_suffix (infixl "<s" 50)
notation (latex output) strict_suffix  ("<s")
lemmas [simp] = strict_suffix_def

lemmas [intro] = suffix_order.le_neq_trans

lemmas ssufI = suffix_order.le_neq_trans

lemma ssufI1[intro]: "u  v = w  u  ε  v <s w"
  by blast

lemma ssufI2[intro]: "u ≤s v  length u < length v  u <s v"
  by blast

lemma ssufE: "u <s v  (u ≤s v  u  v  thesis)  thesis"
  by auto

lemma ssufI3[intro]: "u  v = w  u ≤np w  v <s w"
  unfolding nonempty_prefix_def by blast

lemma ssufD[elim]: "u <s v  u ≤s v  u  v"
  by auto

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 "s" 50)
  where "(suffix_comparable u v)  (rev u)  (rev v)"

lemma suf_compI1[intro]: "u ≤s v  u s v"
  by (simp add: pref_compI suffix_comparable_def suffix_to_prefix)

lemma suf_compI2[intro]: "v ≤s u  u s v"
  by (simp add: pref_compI suffix_comparable_def suffix_to_prefix)

definition nonempty_suffix (infixl "≤ns" 60) where nonempty_suffix_def[simp]:  "u ≤ns v  u  ε  u ≤s v"
notation (latex output) nonempty_suffix  ("ns" 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" 50)
notation (latex output) sublist ("f")
lemmas fac_def = sublist_def


notation strict_sublist (infixl "<f" 50)
notation (latex output) strict_sublist ("<f")
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. pus = v)"
notation (latex output) nonempty_factor ("nf")

lemmas facI = sublist_appendI

lemma facI': "a  u  b = w  u ≤f w"
  by auto

lemma facE[elim]: assumes "u ≤f v"
  obtains p s where "v = p  u  s"
  using assms unfolding fac_def
  by blast

lemma facE'[elim]: assumes "u ≤f v"
  obtains p s where "p  u  s = v"
  using assms unfolding fac_def
  by blast

section "Various elementary lemmas"

lemmas drop_all_iff = drop_eq_Nil ― ‹backward compatibility with Isabelle 2021›

lemma exE2: " x y. P x y  ( x y. P x y  thesis)  thesis"
  by auto

lemmas concat_morph = concat_append

lemmas cancel = same_append_eq and
  cancel_right = append_same_eq

lemmas disjI = verit_and_neg(3)

lemma rev_in_conv: "rev u  A  u  rev ` A"
  by force

lemmas map_rev_involution = list.map_comp[of rev rev, unfolded rev_involution' list.map_id]

lemma map_rev_lists_rev:  "map rev ` (lists (rev ` A)) = lists A"
  unfolding lists_image[of rev] image_comp
  by (simp add: rev_involution')

lemma inj_on_map_lists: assumes "inj_on f A"
  shows "inj_on (map f) (lists A)"
proof
  fix xs ys
  assume "xs  lists A" and "ys  lists A" and "map f xs = map f ys"
  have "x = y" if "x  set xs" and "y  set ys" and  "f x =  f y"  for x y
    using in_listsD[OF xs  lists A, rule_format, OF x  set xs]
          in_listsD[OF ys  lists A, rule_format, OF y  set ys]
         inj_on f A[unfolded inj_on_def, rule_format, OF _ _ f x =  f y] by blast
  from list.inj_map_strong[OF this  map f xs = map f ys]
  show "xs = ys".
qed

lemma bij_lists: "bij_betw f X Y  bij_betw (map f) (lists X) (lists Y)"
  unfolding bij_betw_def using inj_on_map_lists lists_image by metis

lemma concat_sing': "concat [r] = r"
  by simp

lemma concat_sing: assumes "s = [a]" shows "concat s = a"
  using concat_sing' unfolding s = [a].

lemma rev_sing: "rev [x] = [x]"
  by simp

lemma hd_word: "a#ws = [a]  ws"
  by simp

lemma pref_singE: assumes "p ≤p [a]" obtains "p = ε" | "p = [a]"
  using assms unfolding  prefix_Cons by fastforce

lemma map_hd:  "map f (a#v) = [f a]  (map f v)"
  by simp

lemma hd_tl: "w  ε  [hd w]  tl w = w"
  by simp

lemma hd_tlE: assumes "w  ε"
  obtains a w' where "w = a#w'"
  using exE2[OF assms[unfolded neq_Nil_conv]].

lemma hd_tl_lenE: assumes "0 < |w|"
  obtains a w' where "w = a#w'"
  using exE2[OF assms[unfolded length_greater_0_conv neq_Nil_conv]].

lemma hd_tl_longE: assumes "Suc 0 < |w|"
  obtains a w' where "w = a#w'" and "w'  ε" and "hd w = a" and "tl w = w'"
proof-
  obtain a w' where "w = a#w'"
    using  hd_tl_lenE[OF Suc_lessD[OF  assms]].
  hence "w'  ε" and  "hd w = a" and "tl w = w'" using assms by auto
  from that[OF w = a#w' this] show thesis.
qed

lemma hd_pref: "w  ε  [hd w] ≤p w"
  using hd_tl
  by blast

lemma add_nth: assumes "n < |w|" 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': assumes "w  ε" shows "[w!0] ≤p w"
  using hd_pref[OF w  ε, folded hd_conv_nth[OF w  ε, symmetric]].

lemma sub_lists_mono: "A  B  x  lists A  x  lists B"
  by auto

lemma lists_hd_in_set[simp]: "us  ε  us  lists Q  hd us  Q"
  by fastforce

lemma lists_last_in_set[simp]: "us  ε  us  lists Q  last us  Q"
  by fastforce

lemma replicate_in_lists: "replicate k z  lists {z}"
  by (induction k) auto

lemma tl_in_lists: assumes "us  lists A" shows "tl us  lists A"
  using suffix_lists[OF suffix_tl assms].

lemmas lists_butlast = tl_in_lists[reversed]

lemma long_list_tl: assumes "1 < |us|" shows "tl us  ε"
proof
  assume "tl us = ε"
  from assms
  have "us  ε" and "|us| = Suc |tl us|" and "|us|  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  |u|  drop n (take n u  w) = w"
  by simp

lemma split_list_long: assumes "1 < |us|" and "u  set us"
  obtains xs ys where "us = xs  [u]  ys" and "xsysε"
proof-
  obtain xs ys where "us = xs  [u]  ys"
    using split_list_first[OF u  set us] by auto
  hence "xsysε"
    using 1 < |us| 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"
 by (induct xs, simp_all add: subset_iff)

lemma concat_map_sing_ident: "concat (map (λ x. [x]) xs) = xs"
  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  ε], symmetric].

lemma spref_butlast_pref: assumes "u ≤p v" and "u  v" shows "u ≤p butlast v"
  using butlast_append prefixE[OF u ≤p v] u  v append_Nil2 prefixI by metis

lemma last_concat: "xs  ε  last xs  ε  last (concat xs) = last (last xs)"
  using concat_butlast_last last_appendR by metis

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 suf_drop_conv: "u ≤s w  drop (|w| - |u|) w = u"
  using suffix_take append_take_drop_id same_append_eq suffix_drop by metis

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

lemma fin_bin: "finite {x,y}"
  by simp

lemma rev_rev_image_eq [reversal_rule]: "rev ` rev ` X = X"
  by (simp add: image_comp)

lemma last_take_conv_nth: assumes "n < length xs" shows "last (take (Suc n) xs) = xs!n"
  unfolding take_Suc_conv_app_nth[OF assms] by simp

lemma inj_map_inv: "inj_on f A  u  lists A  u = map (the_inv_into A f) (map f u)"
  by (induct u,  simp, simp add: the_inv_into_f_f)

lemma last_sing[simp]: "last [c] = c"
   by simp

lemma hd_hdE: "(u = ε  thesis)  (u = [hd u]  thesis)  (u = [hd u, hd (tl u)]  tl (tl u)  thesis)  thesis"
  using Cons_eq_appendI[of "hd u" "[hd (tl u)]" _ "tl u" "tl (tl u)"] hd_tl[of u] hd_tl[of "tl u"] hd_word
   by fastforce

lemma same_sing_pref: "u  [a] ≤p v  u  [b] ≤p v  a = b"
  using prefix_same_cases by fastforce

lemma compow_Suc: "(f^^(Suc k)) w = f ((f^^k) w)"
  by simp

lemma compow_Suc': "(f^^(Suc k)) w = (f^^k) (f w)"
  by (simp add: funpow_swap1)

subsection ‹General facts›

lemma two_elem_sub: "x  A  y  A  {x,y}  A"
  by simp

thm fun.inj_map[THEN injD]

lemma inj_comp: assumes "inj (f :: 'a list  'b list)" shows "(g  w = h w  (f  g) w = (f  h) w)"
  by (rule iffI, simp) (use injD[OF assms] in fastforce)

lemma inj_comp_eq: assumes "inj (f :: 'a list  'b list)" shows "(g = h  f  g = f  h)"
  by (rule, fast)  (use fun.inj_map[OF assms, unfolded inj_on_def] in fast)

lemma two_elem_cases[elim!]: assumes "u  {x, y}" obtains (fst) "u = x" | (snd) "u = y"
  using assms by blast

lemma two_elem_cases2[elim]: assumes "u  {x, y}" "v  {x,y}" "u  v"
  shows "(u = x  v = y  thesis)  (u = y  v = x  thesis)  thesis"
  using assms by blast

lemma two_elemP: "u  {x, y}  P x  P y  P u"
  by blast

lemma pairs_extensional: "( r s. P r s  ( a b. Q a b  r = fa a  s = fb b))   {(r,s). P r s} = {(fa a, fb b) | a b. Q a b}"
  by auto

lemma pairs_extensional': "( r s. P r s  ( t. Q t  r = fa t  s = fb t))   {(r,s). P r s} = {(fa t, fb t) | t. Q t}"
  by auto

lemma doubleton_subset_cases: assumes "A  {x,y}"
  obtains  "A = {}" | "A = {x}" | "A = {y}" | "A = {x,y}"
  using assms by blast

subsection ‹Map injective function›

lemma map_pref_conv [reversal_rule]: assumes "inj f" shows "map f u ≤p map f v  u ≤p v"
  using map_mono_prefix[of "map f u" "map f v" "inv f"] map_mono_prefix
  unfolding map_map inv_o_cancel[OF inj f] list.map_id..

lemma map_suf_conv [reversal_rule]: assumes "inj f" shows "map f u ≤s map f v  u ≤s v"
  using map_mono_suffix[of "map f u" "map f v" "inv f"] map_mono_suffix
  unfolding map_map inv_o_cancel[OF inj f] list.map_id..

lemma map_fac_conv [reversal_rule]: assumes "inj f" shows "map f u ≤f map f v  u ≤f v"
  using map_mono_sublist[of "map f u" "map f v" "inv f"] map_mono_sublist
  unfolding map_map inv_o_cancel[OF inj f] list.map_id..

lemma map_lcp_conv: assumes "inj f" shows "(map f xs) p (map f ys) = map f (xs p ys)"
proof (induct xs ys rule: list_induct2')
  case (4 x xs y ys)
  then show ?case
  proof (cases "x = y")
    assume "x = y"
    thus ?case
      using "4.hyps" by simp
  next
    assume "x  y"
    hence "f x  f y"
      using inj_eq[OF inj f] by simp
    thus ?case using x  y by simp
  qed
qed simp_all

subsection ‹Orderings on lists: prefix, suffix, factor›

lemmas self_pref = prefix_order.refl and
       pref_antisym = prefix_order.antisym and
       pref_trans = prefix_order.trans and
       spref_trans = prefix_order.less_trans and
       suf_trans = suffix_order.trans and
       fac_trans[intro] = sublist_order.order.trans

subsection "On the empty word"

lemma nemp_elem_setI[intro]: "u  S  u  ε  u  S - {ε}"
  by simp


lemma emp_concat_emp: "us  lists (S - {ε})  concat us = ε  us = ε"
  using DiffD2 by auto

lemma take_nemp: "w  ε  0 < n  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

lemma pref_of_emp: "u  v = ε  u = ε"
  using append_is_Nil_conv by simp

lemma suf_of_emp: "u  v = ε  v = ε"
  using append_is_Nil_conv by simp

lemma nemp_comm: "(u  ε  v  ε  u  v  u  v = v  u)  u  v = v  u"
  by force

lemma non_triv_comm [intro]: "(u  ε  v  ε  u  v  u  v = v  u)  u  v = v  u"
  by force

lemma split_list': "a  set ws  p s. ws = p  [a]  s"
  using split_list by fastforce

lemma split_listE: assumes "a  set w"
  obtains p s where "w = p  [a]  s"
  using exE2[OF split_list'[OF assms]].

subsection ‹Counting letters›

declare count_list_rev [reversal_rule]

lemma count_list_map_conv [reversal_rule]:
  assumes "inj f" shows "count_list (map f ws) (f a) = count_list ws a"
  by (induction ws) (simp_all add: inj_eq[OF assms])

subsection "Set inspection method"

text‹This section defines a simple method that splits a goal into subgoals by enumerating
  all possibilites for @{term "x"} in a premise such as @{term "x  {a,b,c}"}.
  See the demonstrations below.›

method set_inspection = (
    (unfold insert_iff),
    (elim disjE emptyE),
    (simp_all only: singleton_iff refl True_implies_equals)
    )

lemma "u  {x,y}  P u"
  apply(set_inspection)
  oops

lemma "u. u  {x,y}  u = x  u = y"
  by(set_inspection, simp_all)


section "Length and its properties"

lemmas lenarg = arg_cong[of _ _ length] and
       lenmorph = length_append

lemma lenarg_not: "|u|  |v|  u  v"
  using size_neq_size_imp_neq.

lemma len_less_neq: "|u| < |v|  u  v"
  by blast

lemmas len_nemp_conv = length_greater_0_conv

lemma npos_len: "|u|  0  u = ε"
  by simp

lemma nemp_pos_len: "w  ε  0 < |w|"
  by blast

lemma nemp_le_len: "r  ε  1  |r|"
  by (simp add: leI)

lemma swap_len: "|u  v| = |v  u|"
  by simp

lemma len_after_drop: "p + q  |w|   q  |drop p w|"
  by simp

lemma short_take_append: "n  |u| take n (u  v) = take n u"
  by simp

lemma sing_word: "|us| = 1  [hd us] = us"
  by (cases us) simp+

lemma sing_word_concat: assumes "|us| = 1" shows "[concat us] = us"
  unfolding concat_sing[OF sing_word[OF |us| = 1, symmetric]] using sing_word[OF |us| = 1].

lemma len_one_concat_in: "ws  lists A  |ws| = 1  concat ws  A"
  using Cons_in_lists_iff sing_word_concat by metis

lemma concat_nemp:  "concat us  ε  us  ε"
  using concat.simps(1) by blast

lemma sing_len: "|[a]| = 1"
  by simp

lemmas pref_len = prefix_length_le and
       suf_len = suffix_length_le

lemmas spref_len = prefix_length_less and
       ssuf_len = suffix_length_less

lemma pref_len': "|u|  |u  z|"
  by auto

lemma suf_len': "|u|  |z  u|"
  by auto

lemma fac_len: "u ≤f v  |u|  |v|"
  by auto

lemma fac_len': "|w|  |u  w  v|"
  by simp

lemma fac_len_eq: "u ≤f v  |u| = |v|  u = v"
  unfolding fac_def using lenmorph npos_len by fastforce

thm length_take

lemma len_take1: "|take p w|  p"
  by simp

lemma len_take2: "|take p w|  |w|"
  by simp

lemma drop_len: "|u  w|  |u  v  w|"
  by simp

lemma drop_pref: "drop |u| (u  w) = w"
  by simp

lemma take_len: "p  |w|  |take p w| = p"
  using  min_absorb2[of p "|w|", folded length_take[of p w]].

lemma conj_len: "p  x = x  s  |p| = |s|"
  using lenmorph[of p x] lenmorph[of x s] add.commute add_left_imp_eq
  by auto

lemma take_nemp_len: "u  ε  r  ε  take |r| u  ε"
  by simp

lemma nemp_len: "u  ε  |u|  0"
  by simp

lemma emp_len: "w = ε  |w| = 0"
  by simp

lemma take_self: "take |w| w = w"
  using take_all[of w "|w|", OF order.refl].

lemma len_le_concat: "ε   set ws  |ws|  |concat ws|"
proof (induct ws)
  case (Cons a ws)
  hence "1  |a|"
    using list.set_intros(1)[of a ws] nemp_le_len[of a] by blast
  then show ?case
    unfolding   concat.simps(2)  unfolding  lenmorph hd_word[of a ws] sing_len
    using Cons.hyps Cons.prems by simp
qed simp

lemma eq_len_iff: assumes eq: "x  y = u  v" shows "|x|  |u|  |v|  |y|"
  using lenarg[OF eq] unfolding lenmorph by auto

lemma eq_len_iff_less: assumes eq: "x  y = u  v" shows "|x| < |u|  |v| < |y|"
  using lenarg[OF eq] unfolding lenmorph by auto

lemma Suc_len_nemp: "|w| = Suc n  w  ε"
  by force

lemma same_sufix_nil: assumes "z  u ≤p u" shows "z = ε"
  using prefix_length_le[OF assms] unfolding lenmorph by simp

lemma count_list_gr_0_iff: "0 < count_list u a  a  set u"
  by (intro iffI, use count_notin[folded not_gr0, of a u] in blast) (induction u, auto)

lemma mid_fac_ex: assumes "2  |w|"
  shows "w = [hd w](butlast (tl w))[last w]"
  using long_list_tl[OF 2  |w|[folded One_less_Two_le_iff]] append_butlast_last_id[of "tl w"] len_nemp_conv[of w]
  by (simp add: last_tl tl_Nil)

section "List inspection method"

text‹In this section we define a proof method, named list\_inspection, which splits the goal into subgoals which enumerate possibilities
  on lists with fixed length and given alphabet.
  More specifically, it looks for a premise of the form  such as @{term "|w| = 2  w  lists {x,y,z}"} or @{term "|w|  2  w  lists {x,y,z}"}
  and substitutes the goal by the goals listing all possibilities for the word @{term w}, see demonstrations
  after the method definition.›

context
begin

text‹First, we define an elementary lemma used for unfolding the premise.
Since it is very specific, we keep it private.›

private lemma hd_tl_len_list_iff:  "|w| = Suc n  w  lists A  hd w  A   w = hd w # tl w  |tl w| = n  tl w  lists A" (is "?L = ?R")
proof
  show "?L  ?R"
  proof (elim conjE)
    assume "|w| = Suc n" and "w  lists A"
    note Suc_len_nemp[OF |w| = Suc n]
    from lists_hd_in_set[OF w  ε w  lists A] list.collapse[OF w  ε] tl_in_lists[OF w  lists A]
    show "hd w  A  w = hd w # tl w  |tl w| = n  tl w  lists A"
      using |w| = Suc n by simp
  qed
next
  show "?R  ?L"
    using Cons_in_lists_iff[of "hd w" "tl w"] length_Cons[of "hd w" "tl w"] by presburger
qed

text‹We define a list of lemmas used for the main unfolding step.›

private lemmas len_list_word_dec =
    numeral_nat hd_tl_len_list_iff
    insert_iff empty_iff simp_thms length_0_conv

text‹The method itself accepts an argument called `add`, which is supplied to the method
 simp\_all to solve some simple cases, and thus lower the number of produced goals on the fly.›

method list_inspection = (
    ((match premises in len[thin]: "|w|  k" and list[thin]: "w  lists A"  for w k A 
        insert conjI[OF len list])+)?,
    (unfold numeral_nat le_Suc_eq le_0_eq), ― ‹unfold numeral and e.g. @{term "k  2"}
    (unfold conj_ac(1)[of "w  lists A" "|w|  k" for w A k]
          conj_disj_distribR[where ?R = "w  lists A" for w A])?,
    ((match premises in len[thin]: "|w| = k" and list[thin]: "w  lists A"  for w k A 
        insert conjI[OF len list])+)?,
        ― ‹transform into the conjuction such as @{term "length w = 2  w  lists {x,y,z}"}
    (unfold conj_ac(1)[of "w  lists A" "|w| = k" for w A k] len_list_word_dec), ― ‹unfold w›
    (elim disjE conjE), ― ‹split into all cases›
    (simp_all only: singleton_iff lists.Nil list.sel refl True_implies_equals)?, ― ‹replace w everywhere›
    (simp_all only: empty_iff lists.Nil bool_simps)? ― ‹solve simple cases›
)

subsubsection "List inspection demonstrations"

text‹The required premise in the form of conjuction.
First, inequality bound on the length, second, equality bound.›

lemma "|w| = 2  w  lists {x,y,z}  P w"
  apply(list_inspection)
  oops

text‹The required premise as 2 separate assumptions.›
lemma "|w|  2  w  lists {x,y,z}  P w"
  apply(list_inspection)
  oops


lemma "w ≤p w  |w|  2  w  lists {a,b}  hd w = a  w  ε   w = [a, b]  w = [a, a]  w = [a]"
  by list_inspection

lemma "w ≤p w  |w| = 2  w  lists {a,b}  hd w = a   w = [a, b]  w = [a, a]"
  by list_inspection

lemma "w ≤p w  |w| = 2  w  lists {a,b}  hd w = a   w = [a, b]  w = [a, a]"
  by list_inspection

lemma "w ≤p w  w  lists {a,b}  |w| = 2  hd w = a   w = [a, b]  w = [a, a]"
  by list_inspection

end
section "Prefix and prefix comparability properties"

lemmas pref_emp = prefix_bot.extremum_uniqueI

lemma triv_pref: "r ≤p r  s"
  using prefI[OF refl].

lemma triv_spref: "s  ε  r <p 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

lemma spref_cancel: "z  u <p z  v  u <p v"
  by simp

lemma spref_cancel': "u <p v  z  u <p z  v"
  by simp

lemmas pref_cancel_conv = same_prefix_prefix and
       suf_cancel_conv = same_suffix_suffix ― ‹provided by Sublist.thy›

lemma pref_cancel_hd_conv: "a # u ≤p a # v  u ≤p v"
  by simp

lemma spref_cancel_conv: "z  x <p z  y  x <p y"
  by auto

lemma spref_snoc_iff [simp]: "u <p v  [a]  u ≤p v"
  by (auto simp only: strict_prefix_def prefix_snoc) simp

lemma spref_two_lettersE: assumes "p <p [a,b]" obtains "p = ε" | "p = [a]"
  using assms pref_singE[of p a thesis]
  unfolding hd_word[of a "[b]"] spref_snoc_iff by fastforce

lemmas pref_ext[intro] = prefix_prefix ― ‹provided by Sublist.thy›

lemmas pref_extD = append_prefixD and
       suf_extD =  suffix_appendD

lemma spref_extD: "x  y <p z  x <p z"
  using prefix_order.le_less_trans[OF triv_pref].

lemma spref_ext: "r <p u  r <p u  v"
  by force

lemma pref_ext_nemp: "r ≤p u  v  ε  r <p u  v"
  by auto

lemma pref_take: "p ≤p w  take |p| w = p"
  unfolding prefix_def  by force

lemma pref_take_conv: "take (|r|) w = r  r ≤p w"
  using pref_take[of r w] take_is_prefix[of "|r|" 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 <p w  take |p| w = p"
  by (elim spref_exE) force

lemma pref_same_len: "u ≤p v  |u| = |v|  u = v"
  by (fastforce elim: prefixE)

lemma pref_same_len': "u  z ≤p v  w  |u| = |v|  u = v"
   by (fastforce elim: prefixE)

lemma pref_comp_eq: "u  v  |u| = |v|  u = v"
  using pref_same_len by fastforce

lemma ruler_eq_len: "u ≤p w  v ≤p w  |u| = |v|  u = v"
  by (fastforce simp add: prefix_def)

lemma pref_prod_eq: "u ≤p v  z  |u| = |v|  u = v"
  by (fastforce simp add: prefix_def)

lemmas  pref_comm_eq = pref_same_len[OF _ swap_len] and
        pref_comm_eq' = pref_prod_eq[OF _ swap_len, unfolded rassoc]

lemma pref_comm_eq_conv: "u  v ≤p v  u  u  v = v  u"
  using pref_comm_eq self_pref by metis

lemma add_nth_pref: assumes "u <p w" shows "u  [w!|u|] ≤p w"
  using add_nth[OF prefix_length_less[OF u <p w], unfolded spref_take[OF u <p w]].

lemma index_pref: "|u|  |w|  ( i < |u|.  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]
    take_is_prefix[of "|u|" w] by auto

lemma pref_index: assumes "u ≤p w" "i < |u|" shows "u!i = w!i"
  using nth_take[OF i < |u|, 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]] unfolding prefix_def by blast

subsection "Prefix comparability"

lemma pref_comp_sym[sym]: "u  v  v  u"
  by blast

lemma not_pref_comp_sym[sym]: "¬ u  v  ¬ v  u"
  by blast

lemma pref_comp_sym_iff: "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]

lemma ruler_eq: "u  x = v  y  u ≤p v  v ≤p u"
  by (metis prefI prefix_same_cases)

lemma ruler_eq': "u  x = v  y  u ≤p v  v <p u"
  using ruler_eq prefix_order.le_less by blast

lemmas ruler_eqE = ruler_eq[THEN disjE] and
       ruler_eqE' = ruler_eq'[THEN disjE] and
       ruler_pref = ruler[OF append_prefixD triv_pref] and
       ruler'[THEN pref_comp_eq]
lemmas ruler_prefE = ruler_pref[THEN disjE]

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 vz  w ≤p v  v ≤p w"
  using ruler by blast

lemma ruler_pref'': "w ≤p vz  w  v"
  unfolding prefix_comparable_def using ruler_pref'.

lemma pref_cancel_right: assumes "u  z ≤p v  z" shows "u ≤p v"
proof-
  have "|u|  |v|"
    using prefix_length_le[OF assms] by force
  from ruler_le[of u "v  z" v, OF pref_extD[OF assms] triv_pref this]
  show "u ≤p v".
qed

lemma pref_prod_pref_short: "u ≤p z  w  v ≤p w  |u|  |z  v|  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 zuw" shows "u ≤p zu"
  using pref_prod_pref[of  u z "u  w", OF u ≤p zuw triv_pref].

lemma pref_prod_long: "u ≤p v  w  |v|  |u|  v ≤p u"
  using ruler_le[OF triv_pref].

lemmas pref_prod_long_ext = pref_prod_long[OF append_prefixD]

lemma pref_prod_long_less: assumes "u ≤p v  w" and  "|v| < |u|" shows  "v <p u"
  using sprefI2[OF  pref_prod_long[OF u ≤p v  w less_imp_le[OF |v| < |u|]] |v| < |u|].

lemma pref_keeps_per_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 "ru"] by blast

lemma pref_keeps_per_root': "u <p r  u  v ≤p u  v <p r  v"
  using pref_keeps_per_root by auto

lemma per_root_pref_sing: "w <p r  w  u  [a] ≤p w  u  [a] ≤p r  u"
  using append_assoc pref_keeps_per_root' spref_snoc_iff by metis

lemma pref_prolong:  "w ≤p z  r  r ≤p s  w ≤p z  s"
  using pref_trans[OF _ pref_cancel'].

lemma spref__pref_prolong:  "w <p z  r  r ≤p s  w <p z  s"
  using prefix_order.less_le_trans[OF _ pref_cancel'].

lemma pref_spref_prolong:  "w ≤p z  r  r <p s  w <p z  s"
  using prefix_order.le_less_trans[OF _ spref_cancel'].

lemma spref_spref_prolong:  "w <p z  r  r <p s  w <p z  s"
  using prefix_order.less_trans[OF _ spref_cancel'].

lemmas pref_shorten = pref_trans[OF pref_cancel']

lemma pref_prolong': "u ≤p w  z  v  u ≤p z  u ≤p w  v  u"
  using ruler_le[OF _ pref_cancel' le_trans[OF suf_len' suf_len']].

lemma pref_prolong_per_root: "u ≤p r  s  s ≤p r  s  u ≤p r  u"
  using pref_prolong[of u r s "r  s", THEN pref_prod_pref].

thm pref_compE
lemma pref_prolong_comp: "u ≤p w  z  v  u  z  u ≤p w  v  u"
  using pref_prolong' pref_prolong by (elim pref_compE)

lemma pref_prod_le[intro]: "u ≤p v  w  |u|  |v|  u ≤p v"
  using ruler_le[OF _ triv_pref].

lemma prod_pref_prod_le: "uv ≤p xy  |u|  |x|  u ≤p x"
  using  pref_prod_le[OF append_prefixD].

lemma pref_prod_less: "u ≤p v  w  |u| < |v|  u <p v"
  using pref_prod_le[OF _ less_imp_le, THEN sprefI2].

lemma eq_le_pref[elim]: "x  y = u  v  |x|  |u|  x ≤p u"
  using pref_prod_le[OF prefI].

lemma eq_less_pref: "x  y = u  v  |x| < |u|  x <p u"
  using pref_prod_less[OF prefI].

lemma eq_less_suf: assumes "x  y = u  v" shows "|x| < |u|  v <s y"
  using eq_less_pref[reversed, folded strict_suffix_to_prefix, OF x  y = u  v[symmetric]]
  unfolding eq_len_iff_less[OF x  y = u  v].

lemma pref_prod_cancel: assumes "u ≤p pwq" and "|p|  |u|" and "|u|  |pw|"
  obtains r where "p  r = u" and "r ≤p w"
proof-
  obtain r where [symmetric]: "u = p  r" using pref_prod_long[OF u ≤p pwq  |p|  |u|]..
  moreover have "r ≤p w"
    using pref_prod_le[OF u ≤p pwq[unfolded lassoc] |u|  |pw|]
    unfolding p  r = u[symmetric] by simp
  ultimately show thesis..
qed

lemma pref_prod_cancel': assumes "u ≤p pwq" and "|p| < |u|" and "|u|  |pw|"
  obtains r where "p  r = u" and "r ≤p w" and "r  ε"
proof-
  obtain r where "p  r = u" and "r ≤p w"
    using pref_prod_cancel[OF u ≤p pwq less_imp_le[OF |p| < |u|] |u|  |pw|].
  moreover have "r  ε" using  p  r = u less_imp_neq[OF |p| < |u|] by fastforce
  ultimately show thesis..
qed

lemma non_comp_parallel: "¬ u  v  u  v"
  unfolding prefix_comparable_def parallel_def de_Morgan_disj..

lemma comp_refl: "u  u"
  unfolding prefix_comparable_def
  by simp

lemma incomp_cancel: "¬ pu  pv  ¬ u  v"
  unfolding prefix_comparable_def
  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 comm_comp_eq: "r  s  s  r  r  s = s  r"
  using comm_ruler by blast

lemma pref_share_take: "u ≤p v  q  |u|  take q u = take q v"
  by (auto simp add: prefix_def)

lemma pref_prod_longer: "u ≤p z  w  v ≤p w  |z  v|  |u|   z  v ≤p u"
  using ruler_le[OF pref_cancel'].

lemma pref_comp_not_pref: "u  v  ¬ v ≤p u  u <p v"
  by auto

lemma pref_comp_not_spref: "u  v  ¬ u <p v  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

lemmas last_no_split = prefix_snoc

lemma last_no_split': "u <p w  w ≤p u  [a]  w = u  [a]"
  unfolding prefix_order.less_le_not_le last_no_split by blast

lemma comp_shorter: "v  w  |v|  |w|  v ≤p w"
  unfolding prefix_comparable_def
  by (auto simp add: prefix_def)

lemma comp_shorter_conv: "|u|  |v|  u  v  u ≤p v"
  using comp_shorter by auto

lemma pref_comp_len_trans: "w ≤p v  u  v  |w|  |u|  w ≤p u"
  using ruler_le pref_trans by (elim pref_compE)

lemma comp_cancel: "z  w1  z  w2  w1  w2"
  unfolding prefix_comparable_def
  using pref_cancel by auto

lemma emp_pref: "ε ≤p u"
  by simp

lemma emp_spref:  "u  ε  ε <p u"
  by simp

lemma long_pref: "u ≤p v  |v|  |u|  u = v"
  by (auto simp add: prefix_def)

lemma not_comp_ext: "¬ w1   w2  ¬ w1  z  w2  z'"
  using contrapos_nn[OF _ ruler_comp[OF triv_pref triv_pref]].

lemma mismatch_incopm: "|u| = |v|  x  y  ¬ u  [x]  v  [y]"
  by (auto simp add: prefix_def)

lemma comp_prefs_comp: "u  z  v  w  u  v"
  using ruler_comp[OF triv_pref triv_pref].

lemma comp_hd_eq: "u  v  u  ε  v  ε  hd u = hd v"
  unfolding prefix_comparable_def
  by (auto simp add: prefix_def)

lemma pref_hd_eq': "p ≤p u  p ≤p v  p  ε   hd u = hd v"
  by (auto simp add: prefix_def)

lemma pref_hd_eq: "u ≤p v  u  ε   hd u = hd v"
  by (auto simp add: prefix_def)

lemma sing_pref_hd: "[a] ≤p v  hd v = a"
  by (auto simp add: prefix_def)

lemma suf_last_eq: "p ≤s u  p ≤s v  p  ε   last u = last v"
  by (auto simp add: suffix_def)

lemma comp_hd_eq': "u  r  v  s  u  ε  v  ε  hd u = hd v"
using comp_hd_eq[OF comp_prefs_comp].

subsection ‹Minimal and maximal prefix with a given property›

lemma le_take_pref: 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]]
  by force

lemma min_pref: assumes  "u ≤p w" and "P u"
  obtains v where "v ≤p w" and "P v" and  " y. y ≤p w  P y  v ≤p y"
  using assms
proof(induction "|u|" arbitrary: u rule: less_induct)
  case (less u')
  then show ?case
  proof (cases " y. y ≤p w  P y  u' ≤p y", blast)
    assume "¬ (y. y ≤p w  P y  u' ≤p y)"
    then obtain x where "x ≤p w" and "P x" and " ¬ u' ≤p x"
      by blast
    have "|x| < |u'|"
      using prefix_length_less[OF pref_comp_not_pref[OF ruler'[OF x ≤p w u' ≤p w]¬ u' ≤p x]].
    from less.hyps[OF this _ x ≤p w P x] that
    show thesis by blast
  qed
qed


lemma min_pref': assumes  "u ≤p w" and "P u"
  obtains v where "v ≤p w" and "P v" and  " y. y ≤p v  P y  y = v"
proof-
  from min_pref[of _  _ P, OF assms]
  obtain v where "v ≤p w" and "P v" and min: "y. y ≤p w  P y  v ≤p y" by blast
  have "y = v" if "y ≤p v" and "P y" for y
    using min[OF pref_trans[OF y ≤p v v ≤p w] P y] y ≤p v by force
  from that[OF v ≤p w P v this]
  show thesis.
qed

lemma max_pref: assumes  "u ≤p w" and "P u"
  obtains v where "v ≤p w" and "P v" and  " y. y ≤p w  P y  y ≤p v"
  using assms
proof(induction "|w|-|u|" arbitrary: u rule: less_induct)
  case (less u')
  then show ?case
  proof (cases " y. y ≤p w  P y  y ≤p u'", blast)
    assume "¬ (y. y ≤p w  P y   y ≤p  u')"
    then obtain x where "x ≤p w" and "P x" and "¬ x ≤p u'" and "u'  w"
      by blast
    from ruler'[OF x ≤p w u' ≤p w]
    have "|u'| < |x|"
      using  comp_shorter[OF x  u'] ¬ x ≤p u' by fastforce
    hence "|w| - |x| < |w| - |u'|"
      using  x ≤p w  u'  w diff_less_mono2 leI[THEN long_pref[OF u' ≤p w]] by blast
    from less.hyps[OF this _  x ≤p w P x] that
    show thesis by blast
  qed
qed

section "Suffix and suffix comparability  properties"

lemmas suf_emp = suffix_bot.extremum_uniqueI

lemma triv_suf: "u ≤s v  u"
  by (simp add: suffix_def)

lemma emp_ssuf: "u  ε  ε <s u"
  by simp

lemma suf_cancel: "uv ≤s wv  u ≤s w"
  by simp

lemma suf_cancel': "u ≤s w  uv ≤s wv"
  by simp

lemma ssuf_cancel_conv: "x  z <s y  z  x <s y"
  by auto

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 <p w  rev s <s rev w"
  using strict_suffix_to_prefix[of "rev s" "rev w", unfolded rev_rev_ident, symmetric].

lemma nsuf_rev_pref_iff: "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] =
  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 sufE = prefixE[reversed] and
       prefE = prefixE and
       ssuf_exE = spref_exE[reversed]

lemmas suf_prod_long_ext  = pref_prod_long_ext[reversed]

lemmas suf_prolong_per_root = pref_prolong_per_root[reversed]

lemmas suf_ext = suffix_appendI ― ‹provided by Sublist.thy›

lemmas ssuf_ext = spref_ext[reversed] and
  ssuf_extD = spref_extD[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] and
  strict_suffixE' = strict_prefixE'[reversed] and
  ssuf_tl_suf  = spref_butlast_pref[reversed]


lemma ssuf_Cons_iff [simp]: "u <s a # v  u ≤s v"
  by (auto simp only: strict_suffix_def suffix_Cons) (simp add: suffix_def)

lemma ssuf_induct [case_names ssuf]:
  assumes "u. (v. v <s u  P v)  P u"
  shows "P u"
proof (induction u rule: list.induct[of "λu. v. v ≤s u  P v", rule_format, OF _ _ triv_suf],
       use assms suffix_bot.extremum_strict in blast)
qed (metis assms ssuf_Cons_iff suffix_Cons)

subsection "Suffix comparability"

lemma eq_le_suf[elim]: assumes "x  y = u  v" "|x|  |u|" shows "v ≤s y"
  using eq_le_pref[reversed, OF assms(1)[symmetric]]
  lenarg[OF x  y = u  v, unfolded lenmorph] |x|  |u| by linarith

lemmas eq_le_suf'[elim] = eq_le_pref[reversed]

lemma eq_le_suf''[elim]: assumes "v  u = y  x" "|x|  |u|" shows "x ≤s u"
  using eq_le_suf'[OF assms(1)[symmetric] assms(2)].

lemma pref_comp_rev_suf_comp[reversal_rule]: "(rev w) s (rev v)  w  v"
  unfolding suffix_comparable_def by simp

lemma suf_comp_rev_pref_comp[reversal_rule]: "(rev w)  (rev v)  w s v"
  unfolding suffix_comparable_def 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_eq_len = ruler_eq_len[reversed] and
  suf_ruler_comp = ruler_comp[reversed] and
  ruler_suf = ruler_pref[reversed] and
  ruler_suf' = ruler_pref'[reversed] and
  ruler_suf'' = ruler_pref''[reversed] and
  suf_prod_le = pref_prod_le[reversed] and
  prod_suf_prod_le = prod_pref_prod_le[reversed] and
  suf_prod_eq = pref_prod_eq[reversed] and
  suf_prod_less = pref_prod_less[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_long_less = pref_prod_long_less[reversed] and
  suf_prod_longer = pref_prod_longer[reversed] and
  suf_keeps_root = pref_keeps_per_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
    suf_comp_cancel = comp_cancel[reversed] and
  suf_not_comp_ext = not_comp_ext[reversed] and
  mismatch_suf_incopm = mismatch_incopm[reversed] and
  suf_comp_sym[sym] = pref_comp_sym[reversed] and
  suf_comp_refl = comp_refl[reversed]

lemma suf_comp_or: "u s v  u ≤s v  v ≤s u"
  unfolding suffix_comparable_def prefix_comparable_def suf_rev_pref_iff..

lemma comm_comp_eq_conv: "r  s  s  r  r  s = s  r"
  using pref_comp_eq[OF _ swap_len] comp_refl by metis

lemma comm_comp_eq_conv_suf: "r  s s s  r  r  s = s  r"
  using pref_comp_eq[reversed, OF _ swap_len, of r s] suf_comp_refl[of "r  s"] by argo

lemma suf_comp_last_eq: assumes "u s v" "u  ε" "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 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"   ("(_¯>)(_)" [75,74] 74)
  where  "left_quotient u v  = drop |u| v"
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"  ("(_)(<¯_) " [76,77] 76)
  where "right_quotient u v  = rev ((rev v)¯>(rev u))"
notation (latex output) right_quotient ("latex‹\\ensuremath{ {›_ latex‹} \\cdot {› _ latex‹}^{-1}}›")

lemmas lq_def = left_quotient_def and
       rq_def = right_quotient_def

text‹Priorities of these operations are as follows:›

lemma "u<¯v<¯w = (u<¯v)<¯w"
  by simp

lemma "u¯>v¯>w = u¯>(v¯>w)"
  by simp

lemma "u¯>v<¯w = u¯>(v<¯w)"
  by simp

lemma "r  u¯>w<¯v  s = r  (u¯>w<¯v)  s"
  by simp

lemma rq_rev_lq[reversal_rule]: "(rev v)<¯(rev u) = rev (u¯>v)"
  unfolding right_quotient_def
  by simp

lemma lq_rev_rq[reversal_rule]: "(rev v)¯>rev u = rev (u<¯v)"
  unfolding right_quotient_def
  by simp

subsection ‹Left Quotient›

lemma lqI:  "u  z = v  u¯>v = z"
  unfolding left_quotient_def
  by force

lemma lq_triv[simp]:  "u¯>(u  z) = z"
  using lqI[OF refl].

lemma lq_triv'[simp]:  "u  u¯>(u  z) = u z"
  by simp

lemma append_lq: assumes "uv ≤p w" shows "(uv)¯>w = v¯>(u¯>w)"
  using lq_triv[of "uv"] lq_triv[of "v"] lq_triv[of "u" "v_"] assms[unfolded prefix_def]
  by force

lemma lq_self[simp]: "u¯>u = ε"
  unfolding left_quotient_def
  by simp

lemma lq_emp[simp]: "ε¯>u = u"
  unfolding left_quotient_def
  by simp

lemma lq_pref[simp]: "u ≤p v  u  (u¯>v) = v"
  unfolding left_quotient_def prefix_def
  by fastforce

lemma lq_pref_conv: "|u|  |v|  u ≤p v  u  u¯>v = v"
  using lq_pref by blast

lemma lq_len:  "|u¯>v| = |v| - |u|"
    unfolding left_quotient_def using length_drop.

lemmas lcp_lq = lq_pref[OF longest_common_prefix_prefix1] lq_pref[OF longest_common_prefix_prefix2]

lemma lq_pref_cancel: "u ≤p v  v  r = u  s   (u¯>v)  r = s"
  unfolding prefix_def
  by force

lemma lq_the: assumes "u ≤p v"
  shows "(THE z. u  z = v) = (u¯>v)"
proof-
  have "uz = v  z = (u¯>v)" for z
    by fastforce
  from the_equality[of "λz. uz=v", OF lq_pref this, OF assms]
  show ?thesis.
qed

lemma lq_same_len: "|u| = |v|  u¯>v = ε"
  unfolding left_quotient_def by simp

lemma lq_assoc: "|u|  |v|  (u¯>v)¯>w = v¯>(u  w)"
  unfolding  left_quotient_def using prefix_length_le by auto

lemma lq_assoc': "(u  w)¯>v = w¯>(u¯>v)"
  unfolding left_quotient_def lenmorph
  by (simp add: add.commute)

lemma lq_reassoc: "u ≤p v  (u¯>v)w = u¯>(vw)"
  unfolding prefix_def
  by force

lemma lq_trans: "u ≤p v  v ≤p w  (u¯>v)  (v¯>w) = u¯>w"
  by (simp add: lq_reassoc)

lemma lq_rq_reassoc_suf: assumes "u ≤p z" "u ≤s w" shows "wu¯>z = w<¯u  z"
  using rassoc[of "w<¯u"  u  "u¯>z", unfolded lq_pref[OF u ≤p z] lq_pref[reversed, OF u ≤s w]].

lemma lq_ne: "p ≤p up  u  ε  p¯>(up)  ε"
  using lq_pref[of p "u  p"] by fastforce

lemma lq_spref: "u <p v  u¯>v  ε"
  using lq_pref by (auto simp add: prefix_def)

lemma lq_suf_suf: "r ≤p s  (r¯>s) ≤s s"
  by (auto simp add: prefix_def)

lemma lq_short_len: "r ≤p s  |r| +  |r¯>s| = |s|"
  by (auto simp add: prefix_def)

lemma pref_lq: "v ≤p w  u¯>v ≤p u¯>w"
  unfolding left_quotient_def prefix_def
  using drop_append by blast

lemma spref_lq: "u ≤p v  v <p w  u¯>v <p u¯>w"
  by (auto simp add: prefix_def)

lemma pref_gcd_lq: assumes "u ≤p v" shows "(gcd |u| |u¯>v|) = gcd |u| |v|"
  using gcd_add2[of "|u|" "|u¯>v|", unfolded lq_short_len[OF assms], symmetric].

lemma conjug_lq: "x  z = z  y  y = z¯>(x  z)"
  by simp

lemma conjug_emp_emp: "p ≤p u  p  p¯>(u  p) = ε  u = ε"
  using lq_ne by blast


lemma hd_lq_conv_nth: assumes "u <p v" shows "hd(u¯>v) = v!|u|"
  using prefix_length_less[OF assms, THEN hd_drop_conv_nth] unfolding lq_def.

lemma concat_morph_lq: "us ≤p ws  concat (us¯>ws) = (concat us)¯>(concat ws)"
  by (auto simp add: prefix_def)


lemma pref_cancel_lq: assumes "u ≤p x  y"
  shows "x¯>u ≤p y"
  using pref_lq[OF u ≤p x  y, of x, unfolded lq_triv].

lemma pref_cancel_lq_ext: assumes "u  v ≤p x  y" and  "|x|  |u|" shows "x¯>u  v ≤p y"
proof-
  note pref_prod_long[OF append_prefixD, OF u  v ≤p x  y |x|  |u|]
  from  pref_cancel_lq[OF u  v ≤p x  y]
  show "x¯>u  v ≤p y"
    unfolding lq_reassoc[OF x ≤p u] using |x|  |u| by force
qed

lemma pref_cancel_lq_ext': assumes "u  v ≤p x  y" and  "|u|  |x|" shows "v ≤p u¯>x  y"
  using pref_lq[OF u  v ≤p x  y, of u]
  unfolding lq_triv lq_reassoc[OF pref_prod_le[OF append_prefixD[OF u  v ≤p x  y] |u|  |x|]].

lemma empty_lq_eq: "r ≤p z  r¯>z = ε  r = z"
  unfolding prefix_def by force

lemma le_if_then_lq: "|u|  |v|  (if |v|  |u| then v¯>u  else u¯>v) = u¯>v"
  by (cases "|u| = |v|", simp_all add: lq_same_len)

lemma append_comp_lq: "u  v  w  v  u¯>w"
proof (elim pref_compE)
  assume "u  v ≤p w"
  from pref_drop[OF this, of "|u|", unfolded drop_pref]
  show "v  u¯>w"
    unfolding left_quotient_def by (rule pref_compI1)
next
  assume "w ≤p u  v"
  from pref_drop[OF this, of "|u|", unfolded drop_pref]
  show "v  u¯>w"
    unfolding left_quotient_def by (rule pref_compI2)
qed

subsection "Right quotient"

lemmas rqI = lqI[reversed] and
  rq_triv[simp] = lq_triv[reversed] and
  rq_triv'[simp] = lq_triv'[reversed] and
  rq_self[simp] = lq_self[reversed] and
  rq_emp[simp] = lq_emp[reversed] and
  rq_suf[simp] = 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_def[reversed] and
  empty_rq_eq = empty_lq_eq[reversed] and
  append_rq = append_lq[reversed] and
  rq_same_len = lq_same_len[reversed] and
  rq_assoc = lq_assoc[reversed] and
  rq_assoc' = lq_assoc'[reversed] and
  le_if_then_rq =  le_if_then_lq[reversed] and
  append_comp_rq = append_comp_lq[reversed]

subsection ‹Left and right quotients combined›

lemma pref_lq_rq_id:  "p ≤p w  w<¯(p¯>w) = p"
  unfolding prefix_def
  using rq_triv[of p "p¯>w"] by force

lemmas suf_rq_lq_id = pref_lq_rq_id[reversed]

lemma rev_lq': "r ≤p s  rev (r¯>s) = (rev s)<¯(rev r)"
  by (simp add: rq_rev_lq)

lemma pref_rq_suf_lq: "s ≤s u  r ≤p (u<¯s)  s ≤s (r¯>u)"
  using lq_reassoc[of r "u<¯s" s] rq_suf[of s u] triv_suf[of s "r¯>u<¯s"]
  by presburger

lemmas suf_lq_pref_rq = pref_rq_suf_lq[reversed]

lemma "ws = v  v<¯s = w" using rqI.

lemma lq_rq_assoc: "s ≤s u  r ≤p (u<¯s)  (r¯>u)<¯s = r¯>(u<¯s)"
  using lq_reassoc[of r "u<¯s" s] rq_suf[of s u] rqI[of "r¯>u<¯s" s "r¯>u"]
  by argo

lemmas rq_lq_assoc = lq_rq_assoc[reversed]

lemma lq_prod: "u ≤p vu  u ≤p w   u¯>(vu)u¯>w = u¯>(vw)"
  using lq_reassoc[of u "v  u" "u¯>w"] lq_rq_reassoc_suf[of u w "v  u", unfolded rq_triv[of v u]]
  by (simp add: suffix_def)

lemmas rq_prod = lq_prod[reversed]

lemma pref_suf_mid: assumes "pws = p'vs'" and "p ≤p p'" and "s ≤s s'"
  shows "v ≤f w"
proof-
  have "pws  = (p  p¯>p')  v  (s'<¯s   s)"
    using pws = p'vs'
    unfolding lq_pref[OF p ≤p p'] rq_suf[OF s ≤s s'].
  thus ?thesis
    by simp
qed

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.
›

thm append_eq_conv_conj[folded left_quotient_def]
lemma eqd: "x  y = u  v  |x|  |u|   t. x  t = u  t  v = y"
  by (simp add: append_eq_conv_conj)

lemma eqdE: assumes "x  y = u  v" and "|x|  |u|"
  obtains t where "x  t = u" and "t  v = y"
  using eqd[OF assms] by blast

lemma eqd_lessE: assumes "x  y = u  v" and "|x| < |u|"
  obtains t where "x  t = u" and "t  v = y" and "t  ε"
  using eqdE[OF assms(1) less_imp_le[OF assms(2)]] assms(2)
  using append.right_neutral less_not_refl by metis

lemma eqdE': assumes "x  y = u  v" and "|v|  |y|"
  obtains t where "x  t = u" and "t  v = y"
  using eqdE[OF assms(1)] lenarg[OF assms(1), unfolded lenmorph] assms(2)
  by auto

thm long_pref

lemma eqd_pref_suf_iff: assumes "x  y = u  v" shows "x ≤p u  v ≤s y"
  by (rule linorder_le_cases[of "|x|" "|u|"], use eqd[OF assms] in blast)
  (use eqd[OF assms[symmetric]] in fastforce)


lemma eqd_spref_ssuf_iff: assumes "x  y = u  v" shows "x <p u  v <s y"
  using eqd_pref_suf_iff[OF assms] assms by force

lemma eqd_pref: "x  y = u  v  |x|  |u|  x  (x¯>u) = u  (x¯>u)  v = y"
  using eqd lq_triv by blast

lemma eqd_pref1: "x  y = u  v  |x|  |u|  x  (x¯>u) = u"
  using eqd_pref by blast

lemma eqd_pref2: "x  y = u  v  |x|  |u|  (x¯>u)  v = y"
  using eqd_pref by blast

lemma eqd_eq: assumes "x  y = u  v" "|x| = |u|" shows "x = u" "y = v"
  using assms by simp_all

lemma eqd_eq_suf: "x  y = u  v  |y| = |v|  x = u  y = v"
  by simp

lemma eqd_comp: assumes "x  y = u  v" shows "x  u"
  using le_cases[of "|x|" "|u|" "x  u"]
    eqd_pref1[of x y u v, THEN prefI[of x "x¯>u" u], OF assms]
    eqd_pref1[of u v x y, THEN prefI[of u "u¯>x" x], OF assms[symmetric]] by auto

― ‹not equal to eqd\_pref1[reversed]›
lemma eqd_suf1: "x  y = u  v  |x|  |u|  (y<¯v)v = y"
  using eqd_pref2 rq_triv by blast

― ‹not equal to eqd\_pref2[reversed]›
lemma eqd_suf2: assumes "x  y = u  v" "|x|  |u|" shows "x  (y<¯v) = u"
  using rq_reassoc[OF sufI[OF eqd_suf1[OF x  y = u  v |x|  |u|]], 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 "|x|  |u|"
  shows "(y<¯v)v = y  x  (y<¯v) = u"
  using eqd_suf1[OF assms] eqd_suf2[OF assms] by blast

lemma eqd_exchange_aux:
  assumes "u  v = x  y" and "u  v' = x  y'" and "u'  v = x'  y" and "|u|  |x|"
  shows "u'  v' = x'  y'"
  using eqd[OF u  v = x  y |u|  |x|] eqd[OF u  v' = x  y' |u|  |x|] u'  v = x'  y by force

lemma eqd_exchange:
  assumes "u  v = x  y" and "u  v' = x  y'" and "u'  v = x'  y"
  shows "u'  v' = x'  y'"
  using eqd_exchange_aux[OF assms]  eqd_exchange_aux[OF assms[symmetric], symmetric] by force

hide_fact eqd_exchange_aux

section ‹Longest common prefix›

lemmas lcp_simps = longest_common_prefix.simps ― ‹provided by Sublist.thy›

lemmas lcp_sym = lcp.commute

― ‹provided by Sublist.thy›
lemmas lcp_pref = longest_common_prefix_prefix1
lemmas lcp_pref' = longest_common_prefix_prefix2
lemmas pref_pref_lcp[intro] = longest_common_prefix_max_prefix

lemma lcp_pref_ext: "u ≤p v  u ≤p (u  w) p (v  z)"
  using longest_common_prefix_max_prefix prefix_prefix triv_pref by metis

lemma pref_non_pref_lcp_pref: assumes "u ≤p w" and "¬ u ≤p z" shows "w p z <p u"
proof-
  note ruler'[OF u ≤p w lcp_pref, of z, unfolded prefix_comparable_def]
  with pref_trans[of u "w p z", OF _ lcp_pref'] ¬ u ≤p z
  show "w p z <p u"
    by auto
qed

lemmas lcp_take = pref_take[OF lcp_pref] and
       lcp_take' = pref_take[OF lcp_pref']

lemma lcp_take_eq: "take (|u p v|) u = take (|u p v|) v"
  unfolding lcp_take lcp_take'..

lemma lcp_pref_conv: "u p v = u  u ≤p v"
  unfolding prefix_order.eq_iff[of "u p v" u]
  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 lcp_pref_conv': "u p v = v  v ≤p u"
  using lcp_pref_conv[of v u, unfolded lcp_sym[of v]].

lemmas lcp_left_idemp[simp] = lcp_pref[folded lcp_pref_conv'] and
       lcp_right_idemp[simp] = lcp_pref'[folded lcp_pref_conv] and
       lcp_left_idemp'[simp] = lcp_pref'[folded lcp_pref_conv'] and
       lcp_right_idemp'[simp] = lcp_pref[folded lcp_pref_conv]

lemma lcp_per_root:  "r  s p s  r ≤p r  (r  s p s  r)"
  using  pref_prod_pref[OF pref_prolong[OF lcp_pref triv_pref] lcp_pref'].

lemma lcp_per_root':  "r  s p s  r ≤p s  (r  s p s  r)"
  using lcp_per_root[of s r, unfolded lcp_sym[of "s  r"]].

lemma pref_lcp_pref: "w ≤p u p v  w ≤p u"
  using lcp_pref pref_trans by blast

lemma pref_lcp_pref': "w ≤p u p v  w ≤p v"
  using pref_lcp_pref[of w v u, unfolded lcp_sym[of v u]].

lemmas lcp_self = lcp.idem

lemma lcp_eq_len: "|u| = |u p v|  u = u p v"
  using  long_pref[OF lcp_pref, of u v] by auto

lemma lcp_len: "|u|  |u p v|  u ≤p v"
  using long_pref[OF lcp_pref, of u v] unfolding lcp_pref_conv[symmetric].

lemma lcp_len': "¬ u ≤p v  |u p v| < |u|"
  using not_le_imp_less[OF contrapos_nn[OF _ lcp_len]].

lemma incomp_lcp_len: "¬ u  v  |u p v| < min |u| |v|"
  using lcp_len'[of u v] lcp_len'[of v u] unfolding lcp_sym[of v] min_less_iff_conj by blast

lemma lcp_ext_right_conv: "¬ r  r'  (r  u) p (r'  v) = r p r'"
  unfolding prefix_comparable_def
  by (induct r r' rule: list_induct2') simp_all

lemma lcp_ext_right [case_names comp non_comp]: obtains  "r  r'" | "(r  u) p (r'  v) = r p r'"
  using lcp_ext_right_conv by blast

lemma lcp_same_len: "|u| = |v|  u  v  u  w p v  w' = u p v"
  using pref_comp_eq by (cases rule: lcp_ext_right) (elim notE)

lemma lcp_mismatch: "|u p v| < |u|  |u p v| < |v|  u! |u p v|  v! |u p v|"
  by (induct u v rule: list_induct2') auto

lemma lcp_mismatch': "¬ u  v  u! |u p v|  v! |u p v|"
  using incomp_lcp_len lcp_mismatch unfolding min_less_iff_conj..

lemma lcp_mismatchE: assumes "¬ us  vs"
  obtains us' vs'
  where "(us p vs)  us' = us" and "(us p vs)  vs' = vs" and
    "us'  ε" and "vs'  ε" and "hd us'  hd vs'"
proof -
  obtain us' vs' where us: "(us p vs)  us' = us" and vs: "(us p vs)  vs' = vs"
    using prefixE[OF lcp_pref prefixE[OF lcp_pref']]
    unfolding eq_commute[of "xy" for x y].
  with ¬ us  vs have "us'  ε" and "vs'  ε"
    unfolding prefix_comparable_def lcp_pref_conv[symmetric] lcp_sym[of vs]
    by fastforce+
  hence "us! |us p vs| = hd us'" and "vs! |us p vs| = hd vs'"
    using hd_lq_conv_nth[OF triv_spref, symmetric] unfolding lq_triv
    unfolding arg_cong[OF us[symmetric], of nth] arg_cong[OF vs[symmetric], of nth]
    by blast+
  from lcp_mismatch'[OF ¬ us  vs, unfolded this]
  have "hd us'  hd vs'".
  from that[OF us vs us'  ε vs'  ε this]
  show thesis.
qed

lemma lcp_mismatch_lq: assumes "¬ u  v"
  shows
  "(u p v)¯>u  ε" and
  "(u p v)¯>v  ε" and
  "hd ((u p v)¯>u)  hd ((u p v)¯>v)"
proof-
  from lcp_mismatchE[OF assms]
  obtain su sv where "(u p v)  su = u" and
      "(u p v)  sv = v" and "su  ε" and "sv  ε" and "hd su  hd sv".
  thus "(u p v)¯>u  ε" and "(u p v)¯>v  ε" and "hd ((u p v)¯>u)  hd ((u p v)¯>v)"
    using lqI[OF (u p v)  su = u] lqI[OF (u p v)  sv = v] by blast+
qed

lemma lcp_ext_left: "(z  u) p (z  v) = z  (u p v)"
  by (induct z) auto

lemma lcp_first_letters: "u!0  v!0  u p v = ε"
  by (induct u v rule: list_induct2') auto

lemma lcp_first_mismatch: "a  b  w  [a]  u p w  [b]  v  = w"
  by (simp add: lcp_ext_left)

lemma lcp_first_mismatch': "a  b  u  [a] p u  [b] = u"
  using lcp_first_mismatch[of a b u ε ε] by simp

lemma lcp_mismatch_eq_len: assumes "|u| = |v|" "x  y" shows "u  [x] p v  [y] = u p v"
  using lcp_self lcp_first_mismatch'[OF x  y] lcp_same_len[OF |u| = |v|]
  by (cases "u = v") auto

lemma lcp_first_mismatch_pref: assumes "p  [a] ≤p u" and "p  [b] ≤p v" and "a  b"
  shows "u p v = p"
  using assms(1-2) lcp_first_mismatch[OF a  b]
  unfolding  prefix_def rassoc by blast

lemma lcp_append_monotone: "u p x ≤p (u  v) p (x  y)"
  by (simp add: lcp.mono)

lemma lcp_distinct_hd: "hd u  hd v  u p v = ε"
  using pref_hd_eq'[OF lcp_pref lcp_pref'] by blast

lemma nemp_lcp_distinct_hd: assumes "u  ε" and "v  ε" and "u p v = ε"
  shows "hd u  hd v"
proof
  assume "hd u = hd v"
  from lcp_ext_left[of "[hd u]" "tl u" "tl v",
       unfolded hd_tl[OF u  ε] hd_tl[OF v  ε, folded this]]
  show False
    using u p v = ε by simp
qed

lemma lcp_lenI: assumes "i < min |u| |v|" and "take i u = take i v" and "u!i  v!i"
  shows "i = |u p v|"
proof-
  have u: "take i u  [u ! i]  drop (Suc i) u = u"
    using i < min |u| |v| id_take_nth_drop[of i u] by simp
  have v: "take i u  [v ! i]  drop (Suc i) v = v"
    using i < min |u| |v|
    unfolding take i u = take i v using id_take_nth_drop[of i v] by force
  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 p v = take i u".
  thus ?thesis
    using i < min |u| |v| by auto
qed

lemma lcp_prefs: "|u  w p v  w'| < |u|  |u  w p v  w'| < |v|  u p v = u  w p v  w'"
  by (induct u v rule: list_induct2') auto

lemma lcp_extend_eq: assumes "u ≤p v" and "u' ≤p v'" and
              "|v p v'|  |u|" and "|v p v'|  |u'|"
            shows "u p u' = v p v'"
proof-
  consider "|v p v'| = |u|" | "|v p v'| = |u'|" | "|v p v'| < |u|  |v p v'| < |u'|"
    using assms(3-4) by force
  thus ?thesis
  proof (cases)
    assume "|v p v'| = |u|"
    from ruler_eq_len[OF longest_common_prefix_prefix1 u ≤p v this]
    have "u ≤p u'"
      using  prefix_length_prefix[OF longest_common_prefix_prefix2 assms(2,4)] by blast
    thus ?thesis
      unfolding v p v' = u lcp_pref_conv.
  next
    assume "|v p v'| = |u'|"
    from ruler_eq_len[OF longest_common_prefix_prefix2 u' ≤p v' this]
    have "u' ≤p u"
      using  prefix_length_prefix[OF longest_common_prefix_prefix1 assms(1,3)] by blast
    thus ?thesis
      unfolding v p v' = u'  lcp_pref_conv'.
  next
    assume "|v p v'| < |u|  |v p v'| < |u'|"
    thus ?thesis
      using  lcp_prefs[of u "u¯>v" u' "u'¯>v'", unfolded lq_pref[OF u ≤p v] lq_pref[OF u' ≤p v']]
      by blast
  qed
qed

lemma long_lcp_same: assumes "¬ (u p v ≤p w)" shows "u p w = v p w"
proof-
  have "v p w ≤p u"
    using ruler[OF lcp_pref' lcp_pref', of u v w] assms unfolding lcp_sym[of v] by force
  have  "u p w ≤p v"
    using ruler[OF lcp_pref lcp_pref, of u v w] assms by force
  show ?thesis
    unfolding prefix_order.eq_iff
    using v p w ≤p u u p w ≤p v by force
qed

lemma long_lcp_sameE: obtains "u p v ≤p w" | "u p w = v p w"
  using long_lcp_same by blast

lemma ruler_spref_lcp: assumes "u p w <p v p w"
  shows "u p v = u p w"
proof-
  have "¬ v p w ≤p u"
    using prefix_order.leD[of "v p w" "u p w"] assms by force
  from long_lcp_same[OF this]
  show ?thesis
    unfolding lcp_sym[of u].
qed

subsection "Longest common prefix and prefix comparability"
find_theorems name:ruler
lemma lexord_cancel_right: "(u  z, v  w)  lexord r  ¬ u  v  (u,v)  lexord r"
  unfolding prefix_comparable_def
  by (induction rule: list_induct2') auto

lemma lcp_rulersE: assumes "r ≤p s" and "r' ≤p s'" obtains "r  r'" | "s p s' = r p r'"
  by (cases rule: lcp_ext_right[of _ _ _ "r¯>s" "r'¯>s'"]) (assumption, simp only: assms lq_pref)

lemma lcp_rulers: "r ≤p s  r' ≤p s'  (r  r'   s p s' = r p r')"
  by (cases rule: lcp_ext_right[of _ _ _ "r¯>s" "r'¯>s'"], blast) (meson lcp_rulersE)

lemma lcp_rulers': "w ≤p r  w' ≤p s  ¬ w  w'  (r p s) = w p w'"
  using lcp_rulers by blast

lemma lcp_ruler: "r  w1  r  w2  ¬ w1  w2  r ≤p w1 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_compI1[OF pref_trans] ruler' by (elim pref_compE)

lemma comp_monotone': "w  r   w p w'  r"
  using comp_monotone[OF _ lcp_pref].

lemma double_ruler_aux: assumes "w  r" and "w'  r'" and "¬ r  r'" and "|w|  |w'|"
  shows "w p w' = take |w| (r p r')"
proof-
  have pref1: "w p w' ≤p r p r'"
    using comp_monotone'[OF w'  r']  lcp_ruler[OF comp_monotone'[OF w  r] _ ¬ r  r']
    unfolding lcp_sym[of w'] by simp
  show ?thesis
  proof (cases)
    assume "w  w'"
    hence "w p w' = w"
      using |w|  |w'|
      by (simp add: comp_shorter lcp.absorb1)
    show ?thesis
      using pref_take[OF pref1, symmetric] unfolding w p w' = w.
  next
    assume "¬ w  w'"
    hence pref2: "r p r' ≤p w p w'"
      using comp_monotone'[OF w'  r'[symmetric]]  lcp_ruler[OF comp_monotone'[OF w  r[symmetric]] _ ¬ w  w']
      unfolding lcp_sym[of r'] by simp
    hence "w p w' = r p r'"
      using pref1 pref_antisym by blast
    then show ?thesis
      using lcp_take len_take2 take_all_iff by metis
  qed
qed

lemma double_ruler: assumes "w  r" and "w'  r'" and "¬ r  r'"
  shows "w p w' = take (min |w| |w'|) (r p r')"
  by (cases "|w|" "|w'|" rule: le_cases)
  (use double_ruler_aux[OF assms] double_ruler_aux[OF assms(2,1) assms(3)[symmetric], unfolded lcp_sym[of r'] lcp_sym[of w']]
  in linarith)+

hide_fact double_ruler_aux

lemmas pref_lcp_iff = lcp.bounded_iff

lemma pref_comp_ruler: assumes "w  u  [x]" and "w  v  [y]" and "x  y" and "|u| = |v|"
  shows "w ≤p u  w ≤p v"
  using double_ruler[OF w  u  [x] w  v  [y] mismatch_incopm[OF |u| = |v| x  y]]
   take_is_prefix  lcp_self lcp_mismatch_eq_len[OF |u| = |v| x  y] pref_lcp_iff by metis

lemma  comp_per_partes:
  shows "(u  w  v  u¯>w)  u  v  w"
proof
  assume "u  v  w"
  from comp_monotone[OF _ triv_pref, OF this] append_comp_lq[OF this]
  show  "u  w  v  u¯>w"
    by blast
next
  assume c2: "u  w  v  u¯>w"
  hence "u  v  u  u¯>w"
    unfolding comp_cancel by blast
  show "u  v  w"
    by (rule pref_compE[OF conjunct1[OF c2]], use u  v  u  u¯>w in force,blast)
qed

lemmas  scomp_per_partes = comp_per_partes[reversed]


subsection ‹Longest common suffix›

definition longest_common_suffix ("_ s _ " [61,62] 64)
  where
    "longest_common_suffix u v  rev (rev u p rev v)"

lemma lcs_lcp [reversal_rule]: "rev u p rev v = rev (u s v)"
  unfolding longest_common_suffix_def rev_rev_ident..

lemmas lcs_simp = lcp_simps[reversed] and
       lcs_sym = lcp_sym[reversed] and
       lcs_suf = lcp_pref[reversed] and
       lcs_suf' = lcp_pref'[reversed] and
       suf_suf_lcs = pref_pref_lcp[reversed] and
       suf_non_suf_lcs_suf = pref_non_pref_lcp_pref[reversed] and
       lcs_drop_eq = lcp_take_eq[reversed] and
       lcs_take = lcp_take[reversed] and
       lcs_take' = lcp_take'[reversed] and
       lcs_suf_conv = lcp_pref_conv[reversed] and
       lcs_suf_conv' = lcp_pref_conv'[reversed] and
       lcs_per_root = lcp_per_root[reversed] and
       lcs_per_root' = lcp_per_root'[reversed] and
       suf_lcs_suf = pref_lcp_pref[reversed] and
       suf_lcs_suf' = pref_lcp_pref'[reversed] and
       lcs_self[simp] = lcp_self[reversed] and
       lcs_eq_len = lcp_eq_len[reversed] and
       lcs_len = lcp_len[reversed] and
       lcs_len' = lcp_len'[reversed] and
       suf_incomp_lcs_len = incomp_lcp_len[reversed] and
       lcs_ext_left_conv = lcp_ext_right_conv[reversed] and
       lcs_ext_left [case_names comp non_comp] = lcp_ext_right[reversed] and
       lcs_same_len = lcp_same_len[reversed] and
       lcs_mismatch = lcp_mismatch[reversed] and
       lcs_mismatch' = lcp_mismatch'[reversed] and
       lcs_mismatchE = lcp_mismatchE[reversed] and
       lcs_mismatch_rq = lcp_mismatch_lq[reversed] and
       lcs_ext_right = lcp_ext_left[reversed] and
       lcs_first_mismatch = lcp_first_mismatch[reversed, unfolded rassoc] and
       lcs_first_mismatch' = lcp_first_mismatch'[reversed, unfolded rassoc] and
       lcs_mismatch_eq_len = lcp_mismatch_eq_len[reversed] and
       lcs_first_mismatch_suf = lcp_first_mismatch_pref[reversed] and
       lcs_rulers = lcp_rulers[reversed] and
       lcs_rulers' = lcp_rulers'[reversed] and
       suf_suf_lcs' = lcp.mono[reversed] and
       lcs_distinct_last = lcp_distinct_hd[reversed] and
       lcs_lenI = lcp_lenI[reversed] and
       lcs_sufs = lcp_prefs[reversed]

lemmas lcs_ruler = lcp_ruler[reversed] and
       suf_comp_monotone = comp_monotone[reversed] and
       suf_comp_monotone' = comp_monotone'[reversed] and
       double_ruler_suf = double_ruler[reversed] and
       suf_lcs_iff = pref_lcp_iff[reversed] and
       suf_comp_ruler = pref_comp_ruler[reversed]

section "Mismatch"

text ‹The first pair of letters on which two words/lists disagree›

function 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))"
  using shuffles.cases by blast+
termination
  by (relation "measure (λ (t,s). length t)", simp_all)

text ‹Alternatively, mismatch pair may be defined using the longest common prefix as follows.›

lemma mismatch_pair_lcp: "mismatch_pair u v = (u!|upv|,v!|upv|)"
  by (induction u v rule: mismatch_pair.induct) simp_all

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 (pu) (pv)"
  unfolding mismatch_pair_lcp by (simp add: lcp_ext_left)

lemma mismatch_ext_right: assumes  "¬ u  v"
  shows "mismatch_pair u v = mismatch_pair (uz) (vw)"
proof-
  have less1: "|u p v| < |u|" and less2: "|v p u| < |v|"
    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

lemma mismatchI: "¬ u  v  i < min |u| |v|  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

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"
  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.
(Due to the definition of lexord which is designed for irreflexive relations.)›

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 |u| |v|" 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 |u| |v| 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"

lemmas [simp] = sublist_Cons_right

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 add: prefix_def fac_def)

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
  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.order.eq_iff by blast

lemma fac_triv: "pxq = x  p = ε"
  using long_pref[OF prefI suf_len'] unfolding append_self_conv2 rassoc.

lemma fac_triv': "pxq = x  q = ε"
  using fac_triv[reversed] unfolding rassoc.

lemmas
  suf_fac = suffix_imp_sublist and
  pref_fac = prefix_imp_sublist

lemma fac_ConsE: assumes "u ≤f (a#v)"
  obtains "u ≤p (a#v)" | "u ≤f v"
  using assms unfolding sublist_Cons_right
  by blast

lemmas
  fac_snocE = fac_ConsE[reversed]

lemma fac_elim_suf: assumes "f ≤f ms" "¬ f ≤f s"
  shows "f ≤f m(take (|f|-1) s)"
  using assms
proof(induction s rule:rev_induct)
  case (snoc s ss)
  have "¬ f ≤f ss"
    using ¬ f ≤f ss  [s][unfolded sublist_append] by blast

  show ?case
  proof(cases)
    assume  "f ≤f m  ss"
    hence "f ≤f m  take (|f| - 1) ss"
      using ¬ f ≤f ss snoc.IH by blast
    then show ?thesis
      unfolding take_append lassoc using append_assoc sublist_append by metis
  next
    assume "¬ f ≤f m  ss"
    hence "f ≤s m  ss  [s]"
      using  snoc.prems(1)[unfolded lassoc sublist_snoc, unfolded rassoc] by blast
    from suf_prod_le[OF this, THEN suffix_imp_sublist] ¬ f ≤f ss  [s]
    have "|ss  [s]| < |f|"
      by linarith
    from this Suc_less_iff_Suc_le length_append_singleton[of ss s]
    show ?thesis
      using snoc.prems(1) take_all_iff by metis
  qed
qed auto

lemmas fac_elim_pref = fac_elim_suf[reversed]

lemma fac_elim: assumes "f ≤f pms" and  "¬ f ≤f p" and "¬ f ≤f s"
  shows "f ≤f (drop (|p| - (|f| - 1)) p)  m  (take (|f|-1) s)"
  using  fac_elim_suf[OF fac_elim_pref[OF f ≤f pms, unfolded lassoc], unfolded rassoc, OF assms(2-3)].

lemma fac_ext_pref: "u ≤f w  u ≤f p  w"
  by (meson sublist_append)

lemma fac_ext_suf: "u ≤f w  u ≤f w  s"
  by (meson sublist_append)

lemma fac_ext: "u ≤f w  u ≤f p  w  s"
  by (meson fac_ext_pref fac_ext_suf)

lemma fac_ext_hd:"u ≤f w  u ≤f a#w"
  by (metis sublist_Cons_right)

lemma card_switch_fac: assumes "2  card (set ws)"
  obtains c d where "c  d" and  "[c,d] ≤f ws"
  using assms
proof (induct ws, force)
  case (Cons a ws)
  then show ?case
  proof (cases)
    assume "2  card (set ws)"
    from Cons.hyps[OF _ this] Cons.prems(1) fac_ext_hd
    show thesis by metis
  next
    assume "¬ 2  card (set ws)"
    have "ws  ε"
      using 2  card (set (a # ws))  by force
    hence "a = hd ws  set (a # ws) = set ws"
      using hd_Cons_tl[OF ws  ε] by force
    hence "a  hd ws"
      using 2  card (set (a # ws)) ¬ 2  card (set ws)  by force
    from Cons.prems(1)[OF this]
    show thesis
      using  Cons_eq_appendI[OF _ hd_tl[OF ws  ε, symmetric]] sublist_append_rightI by blast
  qed
qed

lemma fac_overlap_len: assumes "u ≤f x  y  z" and "|u|  |y|"
  shows "u ≤f x  y  u ≤f y  z"
proof-
  obtain s p where eq: "x  y  z = p  u  s"
    using u ≤f x  y  z unfolding fac_def by blast
  show ?thesis
  proof (rule le_cases)
    assume "|p|  |x|"
    from add_le_mono[OF this |u|  |y|]
    have "|p  u|  |x  y|"
      unfolding lenmorph.
    from eq_le_pref[OF eq[symmetric, unfolded lassoc] this]
    have "u ≤f x  y"
      using fac_pref by blast
    thus ?thesis by blast
  next
    assume "|x|  |p|"
    from eqd[OF eq this]
    show "u ≤f x  y  u ≤f y  z"
      unfolding fac_def by metis
  qed
qed

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.
›


primrec list_power :: "'a list  nat  'a list"  (infixr "@" 80)
  where
    pow_0: "u @ 0 = ε"
  | pow_Suc: "u @ Suc n = u  u @ n"

term power.power








context
begin

interpretation monoid_mult "ε" "append"
  rewrites "power u n = u@n"
proof-
  show "class.monoid_mult ε (⋅)"
    by (unfold_locales, simp_all)
  show "power.power ε (⋅) u n = u @ n"
    unfolding power.power_def list_power_def by blast
qed


― ‹inherited power properties›

lemma emp_pow_emp[simp]: "r = ε  r@n = ε"
  by simp

lemma pow_pos:"0 < k  a@k = a  a@(k-1)"
  by (simp add: power_eq_if)

lemma pow_pos':"0 < k  a@k = a@(k-1)  a"
  using power_minus_mult by metis

lemma pow_diff: "k < n  a@(n - k) = a  a@(n-k-1)"
  by (rule pow_pos) simp

lemma pow_diff': "k < n  a@(n - k) = a@(n-k-1)  a"
  by (rule pow_pos') simp

lemmas pow_zero = power.power_0 and
  pow_one = power_Suc0_right and
  pow_1 = power_one_right and
  emp_pow[emp_simps] = power_one and
  pow_two[simp] = power2_eq_square and
  pow_Suc = power_Suc and
  pow_Suc' = power_Suc2 and
  pow_comm = power_commutes and
  add_exps = power_add and
  pow_eq_if_list = power_eq_if and
  pow_mult = power_mult and
   comm_add_exp = power_commuting_commutes

lemma pow_rev_emp_conv[reversal_rule]: "power.power (rev ε) (⋅) = (@)"
      unfolding power.power_def list_power_def by simp

lemma pow_rev_map_rev_emp_conv [reversal_rule]: "power.power (rev (map rev  ε)) (⋅) = (@)"
    unfolding power.power_def list_power_def by simp
end

named_theorems exp_simps
lemmas [exp_simps]  = pow_zero pow_one emp_pow
                    numeral_nat less_eq_Suc_le neq0_conv pow_mult[symmetric]

named_theorems cow_simps
lemmas [cow_simps] = emp_simps exp_simps

― ‹more power properties›

lemma sing_Cons_to_pow: "[a, a] = [a] @ Suc (Suc 0)" "a # [a] @ k = [a] @ Suc k"
  by simp_all

lemma zero_exp: "n = 0  r@n = ε"
  by simp

lemma nemp_pow: "t@m  ε  0 < m"
  using zero_exp by blast

lemma pow_nemp_pos[intro]: assumes "u = t@m" "u  ε" shows "0 < m"
  using nemp_pow[OF u  ε[unfolded u = t@m]].



lemma nemp_exp_pos[intro]: "w  ε  r@k = w  0 < k"
  using nemp_pow by blast

lemma nemp_exp_pos'[intro]: "w  ε  w = r@k  0 < k"
  using nemp_pow by blast

lemma nemp_pow_nemp[intro]: "t@m  ε  t  ε"
  using emp_pow by auto

lemma sing_pow_nth:"i < m  ([a]@m) ! i = a"
  by (induct i m rule: diff_induct) auto

lemma pow_is_concat_replicate: "u@n = concat (replicate n u)"
  by (induct n) auto

lemma pow_slide: "u  (v  u)@n   v = (u  v)@(Suc n)"
  by (induct n) simp+

lemma hd_pow: assumes "0 < n" shows "hd(u@n) = hd u"
  unfolding pow_pos[OF 0 < n] using  hd_append2 by (cases "u = ε", simp_all)

lemma pop_pow: "m  k u@m  u@(k-m) =  u@k"
  using le_add_diff_inverse add_exps  by metis

lemma pop_pow_cancel: "u@k  v = u@m  w  m  k  u@(k-m)  v = w"
  using  lassoc pop_pow[of m k u] same_append_eq[of "u@m" "u@(k-m)v" w, unfolded lassoc] by argo

lemma pows_comm: "t@k  t@m = t@m  t@k"
  unfolding add_exps[symmetric] add.commute[of k]..

lemma comm_add_exps: assumes "r  u = u  r" shows "r@m  u@k = u@k  r@m"
  using comm_add_exp[OF comm_add_exp[OF assms, symmetric], symmetric].

lemma rev_pow: "rev (x@m) = (rev x)@m"
  by (induct m, simp, simp add: pow_comm)

lemma pows_comp: "x@i  x@j"
  unfolding prefix_comparable_def using ruler_eqE[OF pows_comm, of x i j] by blast

lemmas pows_suf_comp = pows_comp[reversed, folded rev_pow suffix_comparable_def]

lemmas [reversal_rule] = rev_pow[symmetric]

lemmas pow_eq_if_list' = pow_eq_if_list[reversed] and
  pop_pow_one' = pow_pos[reversed] and
  pop_pow' = pop_pow[reversed] and
  pop_pow_cancel' = pop_pow_cancel[reversed]

lemma pow_len:  "|u@k| = k * |u|"
  by (induct k) simp+

lemma pow_set: "set (w@Suc k) = set w"
  by (induction k, simp_all)

lemma eq_pow_exp[simp]: assumes "u  ε" shows "u@k = u@m  k = m"
proof
  assume "k = m" thus "u@k = u@m" by simp
next
  assume "u@k = u@m"
  from lenarg[OF this, unfolded pow_len mult_cancel2]
  show "k = m"
    using u  ε[folded length_0_conv] by blast
qed

lemma emp_pow_pos_emp [intro]: assumes "v@j = ε" "0 < j" shows "v = ε"
  using pow_pos[OF 0 < j, of v, unfolded v@j = ε] by blast

lemma nemp_emp_pow: assumes "u  ε" shows "u@m = ε  m = 0"
  using  eq_pow_exp[OF assms, of m 0, unfolded pow_zero].

lemma nemp_pow_nemp_pos_conv: assumes "u  ε" shows "u@m  ε  0 < m"
  unfolding nemp_emp_pow[OF assms] by blast

lemma nemp_Suc_pow_nemp: "u  ε  u@Suc k  ε"
  by simp

lemma nonzero_pow_emp: "0 < m  u@m = ε   u = ε"
  by (cases "u = ε", simp)
  (use nemp_emp_pow[of u m] in blast)

lemma pow_eq_eq:
  assumes "u@k = v@k" and "0 < k"
  shows "u = v"
proof-
  have "|u| = |v|"
    using lenarg[OF u@k = v@k, unfolded pow_len] 0 < k by simp
  from eqd_eq[of u "u@(k-1)" v "v@(k-1)", OF _ this]
  show ?thesis
    using u@k = v@k unfolding pow_pos[OF 0 < k] by blast
qed

lemma Suc_pow_eq_eq[elim]: "u@Suc k = v@Suc k  u = v"
  using pow_eq_eq by blast

lemma map_pow[simp]: "map f (r@k) = (map f r)@k"
  by (induct k, simp_all)

lemmas [reversal_rule] = map_pow[symmetric]

lemma concat_pow[simp]: "concat (r@k) = (concat r)@k"
  by (induct k, simp_all)

lemma concat_sing_pow[simp]: "concat ([a]@k) = a@k"
  unfolding concat_pow concat_sing'..

lemma sing_pow_empty: "[a]@n = ε  n = 0"
  using nemp_emp_pow[OF list.simps(3), of _ ε].

lemma sing_pow_lists: "a  A  [a]@n  lists A"
  by (induct n, auto)

lemma long_pow: "r  ε  m  |r@m|"
  unfolding pow_len[of r m] using nemp_le_len[of r] by simp

lemma long_pow_exp': "r  ε  m < |r@(Suc m)|"
  using Suc_le_lessD long_pow by blast

lemma long_pow_expE: assumes "r  ε" obtains n where  "m  |r@Suc n|"
  using long_pow_exp'[OF r  ε] nat_less_le by blast

lemma pref_pow_ext: "x ≤p r@k  x ≤p r@Suc k"
  using pref_trans[OF _ prefI[OF pow_Suc'[symmetric]]].

lemma pref_pow_ext': "u ≤p r@k  u ≤p r  r@k"
  using pref_pow_ext[unfolded pow_Suc].

lemma pref_pow_root_ext: "x ≤p r@k  r  x ≤p r@Suc k"
  by simp

lemma pref_prod_root: "u ≤p r@k  u ≤p r  u"
  using pref_pow_ext'[THEN pref_prod_pref].

lemma le_exps_pref:  "k  l  r@k ≤p r@l"
  using leI pop_pow[of k l r] by blast

lemma pref_exp_le: assumes "u  ε" "u@m ≤p u@n" shows "m  n"
  using mult_cancel_le[OF nemp_len[OF u  ε], of m n]
    prefix_length_le[OF u@m ≤p u@n, unfolded pow_len[of u m] pow_len[of u n]]
  by blast

lemma sing_exp_pref_iff: assumes "a  b"
  shows "[a]@i ≤p [a]@k[b]  w  i  k"
proof
  assume "i  k"
  thus "[a]@i ≤p [a]@k[b]  w"
    using pref_ext[OF le_exps_pref[OF i  k]] by blast
next
  have "¬ [a]@i ≤p [a]@k[b]  w" if "¬ i  k"
  proof (rule notI)
  assume "[a]@i ≤p [a]@k[b]  w"
  hence "k < i" and "0 < i - k" using ¬ i  k by force+
  from pop_pow[OF less_imp_le, OF this(1)]
  have "[a]@k  [a]@(i - k) = [a]@i".
  from [a]@i ≤p [a]@k[b]  w[folded this, unfolded pref_cancel_conv
       pow_pos[OF 0 < i - k]]
  show False
    using a  b by simp
 qed
 thus "[a] @ i ≤p [a] @ k  [b]  w  i  k"
   by blast
qed

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 = le_exps_pref[reversed] and
  suf_exp_le = pref_exp_le[reversed] and
  sing_exp_suf_iff = sing_exp_pref_iff[reversed]

lemma comm_common_power: assumes "r  u = u  r" shows "r@|u| = u@|r|"
  using eqd_eq[OF comm_add_exps[OF r  u = u  r], of "|u|" "|r|"]
  unfolding pow_len by fastforce

lemma one_generated_list_power: "u  lists {x}  k. concat u = x@k"
  by(induction u rule: lists.induct, unfold concat.simps(1), use pow_zero[of x, symmetric] in fast,
        unfold concat.simps(2))
  (use pow_Suc[symmetric, of x] singletonD in metis)

lemma pow_lists: assumes "0 < k" shows "u@k  lists B  u  lists B"
  unfolding pow_Suc[of u "k-1", unfolded Suc_minus_pos[OF 0 < k]] by simp

lemma concat_morph_power: "xs  lists B  xs = ts@k  concat ts@k = concat xs"
  by (induct k arbitrary: xs ts) simp_all


lemma per_exp_pref: "u ≤p r  u  u ≤p r@k  u"
proof(induct k)
  case (Suc k) show ?case
    unfolding pow_Suc rassoc
    using Suc.hyps Suc.prems pref_prolong by blast
qed simp

lemmas
    per_exp_suf = per_exp_pref[reversed]

lemma hd_sing_pow: "k  0  hd ([a]@k) = a"
  by (induction k) simp+


lemma sing_pref_comp_mismatch:
  assumes "b  a" and "c  a" and "[a]@k  [b]  [a]@l  [c]"
  shows "k = l  b = c"
proof
  show "k = l"
    using assms
  proof (induction k l rule: diff_induct)
    show " b  a  c  a  [a] @ x  [b]  [a] @ 0  [c]  x = 0" for x
      by (rule ccontr, elim not0_SucE) fastforce
  qed (simp add:prefix_comparable_def)+
  show "b = c"
    using assms(3) unfolding k = l by auto
qed

lemma sing_pref_comp_lcp: assumes "r  s" and "a  b" and "a  c"
  shows  "[a]@r  [b]  u p [a]@s  [c]  v = [a]@(min r s)"
proof-
  have "r  s   [a]@r  [b]  u p [a]@s  [c]  v = [a]@(min r s)"
  proof (rule diff_induct[of "λ r s. r  s  [a]@r  [b]  u p [a]@s  [c]  v = [a]@(min r s)"])
    have "[a] @ Suc (x - 1)  [b]  u p [c]  v = [a] @ min x 0" if "x  0" for x
      unfolding pow_Suc  min_0R exp_simps rassoc by (simp add: a  c)
    thus "x  0  [a] @ x  [b]  u p [a] @ 0  [c]  v = [a] @ min x 0" for x  by force
    show "0  Suc y  [a] @ 0  [b]  u p [a] @ Suc y  [c]  v = [a] @ min 0 (Suc y)" for y
      unfolding pow_Suc  min_0L exp_simps rassoc using a  b by auto
    show "x  y  [a] @ x  [b]  u p [a] @ y  [c]  v = [a] @ min x y 
           Suc x  Suc y  [a] @ Suc x  [b]  u p [a] @ Suc y  [c]  v = [a] @ min (Suc x) (Suc y)" for x y
      unfolding pow_Suc rassoc min_Suc_Suc by simp
  qed
  with assms
  show ?thesis by blast
qed

lemmas sing_suf_comp_mismatch = sing_pref_comp_mismatch[reversed]

lemma exp_pref_cancel: assumes "t@m  y = t@k" shows "y = t@(k - m)"
  using lqI[of "t@m" "t@(k-m)" "t@k"]  unfolding lqI[OF t@m  y = t@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@k"] pow_zero[of t]
    pref_antisym[of "t@m" "t@k", OF prefI[OF  t@m  y = t@k] le_exps_pref[of k m t]]
  by presburger

lemmas exp_suf_cancel = exp_pref_cancel[reversed]

lemma index_pow_mod: "i < |r@k|  (r@k)!i = r!(i mod |r|)"
proof(induction k)
  have aux:  "|r@(Suc l)| = |r@l| + |r|" for l
    by simp
  have aux1: "|(r@l)|  i  i < |r@l| + |r|   i mod |r| = i -  |r@l|" for l
    unfolding pow_len[of r l] using less_diff_conv2[of "l * |r|" i "|r|", unfolded add.commute[of "|r|"  "l * |r|"]]
      get_mod[of "i - l * |r|" "|r|" l] le_add_diff_inverse[of "l*|r|" i] by argo
  case (Suc k)
  show ?case
    unfolding aux sym[OF pow_Suc'[symmetric]] nth_append le_mod_geq
    using aux1[ OF _ Suc.prems[unfolded aux]]
      Suc.IH pow_Suc'[symmetric] Suc.prems[unfolded aux] leI[of i "|r @ k|"] by presburger
qed auto

lemma sing_pow_len [simp]: "|[r]@l| = l"
  by (induct l) auto

lemma take_sing_pow: "k  l  take k ([r]@l) = [r]@k"
proof (induct k)
  case (Suc k)
  have "k < |[r]@l|" using Suc_le_lessD[OF Suc k  l] unfolding sing_pow_len.
  from take_Suc_conv_app_nth[OF this]
  show ?case
    unfolding Suc.hyps[OF Suc_leD[OF Suc k  l]] pow_Suc'
    unfolding sing_pow_nth[OF Suc_le_lessD[OF Suc k  l]].
qed simp

lemma concat_take_sing: assumes "k  l" shows "concat (take k ([r]@l)) = r@k"
  unfolding take_sing_pow[OF k  l] using concat_sing_pow.

lemma unique_letter_word: assumes "c. c  set w  c = a" shows "w = [a]@|w|"
  using assms proof (induction w)
  case (Cons b w)
  have "[a] @ |w| = w" using Cons.IH[OF Cons.prems[OF list.set_intros(2)]]..
  then show "b # w = [a] @ |b # w|"
    unfolding Cons.prems[OF list.set_intros(1)] by auto
qed simp

lemma card_set_le_1_imp_hd_pow: assumes "card (set u)  1" shows "[hd u] @ |u| = u"
proof (cases "u = ε")
  assume "u  ε"
  then have "card (set u) = 1" using card (set u)  1
    unfolding le_less less_one card_0_eq[OF finite_set] set_empty by blast
  then have "set u = {hd u}" using hd_in_set[OF u  ε]
  by (elim card_1_singletonE) simp
  then show "[hd u]@|u| = u"
    by (intro unique_letter_word[symmetric]) blast
qed simp

lemma unique_letter_wordE'[elim]: assumes "( c. c  set w  c = a)" obtains k where "w = [a]@k"
  using unique_letter_word assms by metis

lemma unique_letter_wordE''[elim]: assumes "set w  {a}" obtains k where "w = [a] @ k"
  using assms unique_letter_word[of w a] by blast

lemma unique_letter_wordE[elim]: assumes "set w = {a}" obtains k where "w = [a]@Suc k"
proof-
  have "w  ε" using assms by force
  obtain l where "w = [a]@l"
    using unique_letter_wordE''[of w a thesis] assms by force
  with w  ε
  have "l  0"
    by blast
  show thesis
    using that[of "l-1"] unfolding w = [a]@l Suc_minus[OF l  0] by blast
qed

lemma conjug_pow: "x  z = z  y  x@k  z = z  y@k"
  by (induct k) fastforce+

lemma lq_conjug_pow: assumes "p ≤p x  p" shows "p¯>(x@k  p) = (p¯>(x  p))@k"
  using lqI[OF sym[OF conjug_pow[of x p  "p¯>(x  p)", OF sym[OF lq_pref[OF p ≤p x  p]], of k]]].

lemmas rq_conjug_pow = lq_conjug_pow[reversed]

lemma pow_pref_root_one: assumes "0 < k" and "r  ε" and "r@k ≤p r"
  shows  "k = 1"
  unfolding eq_pow_exp[OF r  ε, of k 1, symmetric] pow_1
  using r@k ≤p r triv_pref[of r "r@(k-1)", folded pow_pos[OF 0 < k]] by auto

lemma count_list_pow: "count_list (w@k) a = k * (count_list w a)"
  by (induction k, simp, simp)

lemma comp_pows_pref: assumes  "v  ε" and "(u  v)@k  u ≤p (u  v)@m" shows "k  m"
  using pref_exp_le[OF _  pref_extD[OF assms(2)]] assms(1) by blast

lemma comp_pows_pref': assumes  "v  ε" and "(u  v)@k ≤p (u  v)@m  u" shows "k  m"
proof(rule ccontr)
  assume "¬ k  m"
  hence "Suc m  k" by simp
  from le_exps_pref[OF this, unfolded pow_Suc']
  have "(u  v)@m  (u  v) ≤p (u  v)@k".
  from pref_trans[OF this assms(2)] v  ε
  show False by auto
qed

lemma comp_pows_not_pref: "¬ (u  v)@k  u ≤p (u  v)@m  m  k"
  by (induction k m rule: diff_induct) auto

lemma comp_pows_spref: "u@k <p u@m  k < m"
  by (induction k m rule: diff_induct) auto

lemma comp_pows_spref_ext: "(u  v)@k  u <p (u  v)@m  k < m"
  by (induction k m rule: diff_induct) auto

lemma comp_pows_pref_zero:"(u  v)@k <p u  k = 0"
  by (induct k) auto

lemma comp_pows_spref': "(u  v)@k <p (u  v)@m  u  k < Suc m"
  by (induction k m rule: diff_induct, simp_all add: comp_pows_pref_zero)

lemmas comp_pows_suf = comp_pows_pref[reversed] and
       comp_pows_suf' =  comp_pows_pref'[reversed] and
       comp_pows_not_suf = comp_pows_not_pref[reversed] and
       comp_pows_ssuf = comp_pows_spref[reversed] and
       comp_pows_ssuf_ext = comp_pows_spref_ext[reversed] and
       comp_pows_suf_zero = comp_pows_pref_zero[reversed] and
       comp_pows_ssuf' = comp_pows_spref'[reversed]

subsection Comparison

― ‹Lemmas allowing to compare complicated terms with powers›

named_theorems shifts
lemma shift_pow[shifts]: "(uv)@ku = u(vu)@k"
  using conjug_pow[OF rassoc].
  lemma[shifts]: "(u  v)@k  u  z = u  (v  u)@k  z"
  by (simp add: shift_pow)
lemma[shifts]: "u@k  u  z = u  u@k  z"
  by (simp add: conjug_pow)
lemma[shifts]: "r@k ≤p r  r@k"
  by (simp add: pow_comm[symmetric])
lemma [shifts]: "r@k ≤p r  r@k  z"
  unfolding lassoc pow_comm[symmetric] unfolding rassoc by blast
lemma [shifts]: "(r  q)@k ≤p r  q   (r  q)@k  z"
  unfolding lassoc pow_comm[symmetric] unfolding rassoc by simp
lemma [shifts]: "(r  q)@k ≤p r  q   (r  q)@k"
  unfolding lassoc pow_comm[symmetric] unfolding rassoc by simp
lemma[shifts]: "r@k  u ≤p r  r@k  v  u ≤p r  v"
  unfolding lassoc pow_comm[symmetric] unfolding rassoc pref_cancel_conv..
lemma[shifts]: "u  u@k  z = u@k  w  u  z = w"
   unfolding lassoc pow_comm[symmetric] unfolding rassoc cancel..
lemma[shifts]: "(r  q)@k  u ≤p r  q   (r  q)@k  v  u ≤p r  q  v"
  unfolding lassoc pow_comm[symmetric] unfolding rassoc pref_cancel_conv..
lemma[shifts]: "(r  q)@k  u = r  q   (r  q)@k  v  u = r  q  v"
  unfolding lassoc pow_comm[symmetric] unfolding rassoc cancel..
lemma[shifts]: "r  q   (r  q)@k  v = (r  q)@k  u  r  q  v = u"
  unfolding lassoc pow_comm[symmetric] unfolding rassoc cancel..
lemma shifts_spec [shifts]: "(u@k  v)@l  u  u@k  z = u@k  (v  u@k)@l  u  z"
  unfolding lassoc cancel_right unfolding rassoc pow_comm[symmetric]
  unfolding lassoc cancel_right shift_pow..
lemmas [shifts] = shifts_spec[of "r  q", unfolded rassoc] for r q
lemmas [shifts] = shifts_spec[of "r  q" _ _ _ ε , unfolded rassoc emp_simps] for r q
lemmas [shifts] = shifts_spec[of "r  q" _ "r  q", unfolded rassoc] for r q
lemmas [shifts] = shifts_spec[of "r  q" _ "r  q" _ ε , unfolded rassoc emp_simps] for r q
lemma[shifts]: "(u  (v  u)@k)@j  (u  v)@k = (u  v)@k  (u  (u  v)@k)@j"
  by (metis shift_pow)
lemma[shifts]: "(u  (v  u)@k  z)@j  (u  v)@k = (u  v)@k  (u  z  (u  v)@k)@j"
  by (simp add: conjug_pow)
lemmas[shifts] = pow_comm cancel rassoc pow_Suc pref_cancel_conv suf_cancel_conv add_exps cancel_right numeral_nat pow_zero emp_simps
lemmas[shifts] = less_eq_Suc_le
lemmas[shifts] =  neq0_conv
lemma shifts_hd_hd [shifts]: "a#b#v = [a]  b#v"
  using hd_word.
lemmas [shifts] =  shifts_hd_hd[of _ _ ε]
lemma[shifts]: "n  k  x@k = x@(n + (k -n))"
  by simp
lemma[shifts]: "n < k  x@k = x@(n + (k -n))"
  by simp
lemmas[shifts] = cancel cancel_right pref_cancel_conv suf_cancel_conv triv_pref
lemmas[shifts] = pow_diff

lemmas shifts_rev = shifts[reversed]

lemmas shift_simps = shifts shifts[reversed]

method comparison = ((simp only: shifts; fail) | (simp only: shifts_rev; fail))

section ‹Rotation›

lemma rotate_root_self: "rotate |r| (r@k) = r@k"
proof (cases "r = ε")
  assume "r  ε"
  show ?thesis
  proof (cases k)
    fix pred
    assume k: "k = Suc pred"
    show ?thesis
      unfolding k pow_Suc rotate_append pow_comm..
  qed simp
qed simp

lemma rotate_pow_self: "rotate (l*|u|) (u@k) = u@k"
proof(induct l)
  case (Suc l)
  show ?case
    unfolding mult_Suc rotate_rotate[symmetric] Suc.hyps
    using rotate_root_self.
qed simp

lemma rotate_pow_mod:  "rotate n (u@k) = rotate (n mod |u|) (u@k)"
  using rotate_rotate[of "n mod |u|" "n div |u| * |u|" "u@k", symmetric]
  unfolding rotate_pow_self[of "n div |u|" u k] div_mult_mod_eq[of n "|u|", unfolded add.commute[of "n div |u| * |u|" "n mod |u|"]].

lemma rotate_conj_pow: "rotate |u| ((uv)@k) = (vu)@k"
 by (induct k, simp, simp add: rotate_append shift_pow)

lemma rotate_pow_comm: "rotate n (u@k) = (rotate n u)@k"
proof (cases "u = ε")
  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 |u|) u" "drop (n mod |u|) u" k, unfolded append_take_drop_id[of "n mod |u|" u]]
    unfolding  mod_le_divisor[of "|u|" n, THEN take_len, OF uε[unfolded length_greater_0_conv[symmetric]]].
qed simp

lemmas rotate_pow_comm_two = rotate_pow_comm[of _ _ 2, unfolded pow_two]

lemma rotate_back: "rotate (|u| - n mod |u|) (rotate n u) = u"
proof  (cases "u = ε")
  assume "u  ε"
  show ?thesis
  unfolding rotate_conv_mod[of n u] rotate_rotate[of "|u| - n mod |u|" "n mod |u|" u]
  le_add_diff_inverse2[OF mod_le_divisor, OF nemp_pos_len[OF u  ε]]
  by simp
qed simp


lemma rotate_backE: obtains m where "rotate m (rotate n u) = u"
  using rotate_back by blast

lemma rotate_back': assumes "rotate m w = rotate n w"
  shows "rotate (m-n) w = w"
proof (cases)
  assume "n  m"
  from rotate_backE obtain k where   "rotate k (rotate n w) = w".
  hence nk: "rotate n (rotate k w) = w"
    unfolding rotate_rotate add.commute[of _ k].
  have mn: "rotate m (rotate k w) = (rotate n (rotate k w))"
    unfolding rotate_rotate add.commute[of _ k] unfolding rotate_rotate[symmetric] assms..
  have "rotate (m - n) (rotate n (rotate k w)) = rotate m (rotate k w)"
    unfolding rotate_rotate using n  m by simp
  from this[unfolded mn nk]
  show ?thesis.
qed simp

lemma rotate_class_rotate': "(n. rotate n w = u)  (n. rotate n (rotate l w) = u)"
proof
  obtain m where rot_m: "rotate m (rotate l w) = w" using rotate_backE.
  assume "n. rotate n w = u"
  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
        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
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_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
    unfolding prefix_comparable_def by blast
  from lexord_mismatch[OF this irrefl r]
  show ?thesis.
qed


section ‹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 concat_take_is_prefix: "concat(take n ws) ≤p concat ws"
  using concat_morph[of "take n ws" "drop n ws",symmetric, unfolded append_take_drop_id[of n ws], THEN prefI].

lemma concat_take_Suc: assumes "j < |ws|" shows "concat(take j ws)  ws!j = concat(take (Suc j) ws)"
  unfolding take_Suc_conv_app_nth[OF j < |ws|]
  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 <p concat ws"
  obtains j r where "j < |ws|" and "r <p ws!j" and "concat (take j ws)  r = u"
proof-
  have "|ws|  0"
    using assms by auto
  then obtain l where "Suc l = |ws|"
    using Suc_pred by blast
  let ?P = "λ j. u <p concat(take (Suc j) ws)"
  have "?P l"
    using assms  Suc l = |ws| by auto
  define j where "j = (LEAST j. ?P j)" ― ‹smallest j such that concat (take (Suc j) ws) <p u›
  have "u <p concat(take (Suc j) ws)"
    using  LeastI[of ?P, OF ?P l] unfolding sym[OF j_def].
  have  "j < |ws|"
    using Least_le[of ?P, OF ?P l] Suc l = |ws| unfolding sym[OF j_def]
    by auto
  have "concat(take j ws) ≤p u"
    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 (cases "j = 0", simp) force
  from prefixE[OF this]
  obtain r where "u = concat(take j ws)  r".
  from u <p concat (take (Suc j) ws)[unfolded this]
  have "r <p ws!j"
    unfolding concat_take_Suc[OF j < |ws|, symmetric]  spref_cancel_conv.
  show thesis
    using that[OF j < |ws| r <p ws!j u = concat(take j ws)  r[symmetric]].
qed

thm prefI

lemma pref_mod_pow: assumes "u ≤p w@l" and "w  ε"
  obtains k z where "k  l" and "z <p w" and "w@kz = u"
proof (cases "u = w@l")
  assume "u  w@l"
  from sprefI[OF u ≤p w@l this]
  have "u <p w @ l".
  have "w@l = concat ([w]@l)"
    by simp
  from pref_mod_list[of u "[w]@l", unfolded sing_pow_len concat_sing_pow, OF u <p w@l]
  obtain j r where "j < l" "r <p ([w] @ l) ! j" "concat (take j ([w] @ l))  r = u".
  hence "j  l" and "r <p w" and "w@j  r = u"
    unfolding sing_pow_nth[OF j < l] concat_take_sing[OF less_imp_le[OF j < l]] by auto
  from that[OF this]
  show thesis.
qed (use emp_spref assms in blast)

lemma pref_mod_pow': assumes "u <p w@l"
  obtains k z where "k < l" and "z <p w" and "w@kz = u"
proof-
  have "w  ε" using assms by force
  from pref_mod_pow[OF sprefD1[OF assms] this]
  obtain k z where "k  l" "z <p w" "w @ k  z = u".
  note spref_extD[OF u <p w@l[folded w @ k  z = u]]
  have "k < l"
    using comp_pows_spref[OF w @ k <p w @ l].
  from that[OF this z <p w w @ k  z = u]
  show thesis.
qed

lemma split_pow: assumes "u  v = w@k" "0 < k" "v  ε"
  obtains p s i j where "w = p  s" "s  ε" "u = (p  s)@i  p" "v = (s  p)@j  s" "k = i + j + 1"
proof-
  have "u <p w@k"
    using assms(1,3) by blast
  from pref_mod_pow'[OF this]
  obtain ku p where "ku < k" "p <p w" "w @ ku  p = u".
  from spref_exE[OF this(2)]
  obtain s where "p  s = w" "s  ε".
  obtain kv where "k = Suc(ku + kv)"
    using  less_imp_Suc_add[OF ku < k] by blast
  from u  v = w@k[folded this[symmetric] p  s = w w @ ku  p = u, unfolded rassoc pow_Suc']
  have "v = s  w@kv"
    unfolding shifts unfolding lassoc shift_pow[symmetric] unfolding rassoc cancel p  s = w.
  show thesis
    using that[OF p  s = w[symmetric] s  ε w @ ku  p = u[folded ps = w, symmetric]
          v = s  w@kv[folded ps = w,folded shift_pow] k = Suc(ku + kv)[unfolded Suc_eq_plus1]].
qed










lemma del_emp_concat: "concat us = concat (filter (λx. x  ε) us)"
  by (induct us) simp+

lemma lists_minus: "us  lists (C - A)  us  lists C"
  by blast

lemma lists_minus': "us  lists C  (filter (λx. x  ε) us)  lists (C - {ε})"
  by (simp add: in_lists_conv_set)

lemma pref_concat_pref: "us ≤p ws  concat us ≤p concat ws"
  by (auto simp add: prefix_def)

lemmas suf_concat_suf = pref_concat_pref[reversed]

lemma concat_mono_fac: "us ≤f ws  concat us ≤f concat ws"
  using  concat_morph facE facI' by metis

lemma ruler_concat_less: assumes "us ≤p ws" and "vs ≤p ws" and "|concat us| < |concat vs|"
  shows "us <p vs"
  using ruler[OF us ≤p ws vs ≤p ws] pref_concat_pref[of vs us, THEN prefix_length_le] |concat us| < |concat vs|
  by force

lemma concat_take_mono_strict: assumes "concat (take i ws) <p concat (take j ws)"
  shows "take i ws <p take j ws"
  using ruler_concat_less[OF _ _ prefix_length_less, OF take_is_prefix take_is_prefix assms].

lemma take_pp_less: assumes "take k ws <p take n ws" shows "k < n"
  using  conjunct2[OF sprefD[OF assms]]
    leI[of k n, THEN[2] le_take_pref[of n k ws, THEN[2] pref_antisym[of "take k ws" "take n ws"]], OF conjunct1[OF sprefD[OF assms]]]
  by blast

lemma concat_pp_less: assumes "concat (take k ws) <p concat (take n ws)" shows "k < n"
  using le_take_pref[of n k ws, THEN pref_concat_pref] conjunct1[OF sprefD[OF assms]]
    conjunct2[OF sprefD[OF assms]] pref_antisym[of "concat(take k ws)" "concat(take n ws)"]
  by fastforce

lemma take_le_take: "j  k  take j (take k xs) = take j xs"
proof (rule disjE[OF le_less_linear, of k "|xs|"])
  assume "j  k" and "k  |xs|"
  show ?thesis
    using pref_share_take[OF take_is_prefix, of j k xs, unfolded take_len[OF k  |xs|], OF j  k].
qed simp

lemma concat_interval: assumes "concat (take k vs) = concat (take j vs)  s" shows "concat (drop j (take k vs)) = s"
proof (rule disjE[OF le_less_linear, of k j])
  note eq1 = assms[folded  arg_cong[OF takedrop[of j "take k vs"], of concat, unfolded concat_morph]]
  assume "j < k"
  from eq1[unfolded take_le_take[OF less_imp_le[OF this]]]
  show ?thesis
    unfolding cancel.
next
  note eq1 = assms[folded  arg_cong[OF takedrop[of j "take k vs"], of concat, unfolded concat_morph]]
  assume "k  j"
  from pref_concat_pref[OF le_take_pref, OF this, of vs, unfolded assms]
  have "s = ε"
    by force
  from drop_all[OF le_trans[OF len_take1 k  j], of vs]
  have "concat (drop j (take k vs)) = ε"
    using concat.simps(1) by force
  with s = ε
  show ?thesis by blast
qed

lemma bin_lists_count_zero': assumes "ws  lists {x,y}" and "count_list ws y = 0"
  shows "ws  lists {x}"
  using assms
proof (induct ws)
  case (Cons a ws)
  have "a  y"
    using count_list (a # ws) y = 0 count_list.simps(2) by force
  hence "count_list ws y = 0"
    using count_list (a # ws) y = 0 count_list.simps(2) by force
  from Cons.hyps(3)[OF this]
  show ?case
    using a  {x,y}  a  y by auto
qed simp

lemma bin_lists_count_zero: assumes "ws  lists {x,y}" and "count_list ws x = 0"
  shows "ws  lists {y}"
  using assms unfolding insert_commute[of x y "{}"] using  bin_lists_count_zero' by metis

lemma count_in: "count_list ws a  0  a  set ws"
  using count_notin[of a ws] by fast

lemma count_in_conv: "count_list w a  0   a  set w"
  by (induct w, auto)

lemma two_in_set_concat_len: assumes "u  v" and "{u,v}  set ws"
  shows "|u| + |v|  |concat ws|"
proof-
  let ?ws = "filter (λ x. x  {u,v}) ws"
  have set: "set ?ws = {u,v}"
    using {u,v}  set ws by auto
  have "|concat ?ws|  |concat ws|"
    unfolding length_concat  using sum_list_filter_le_nat by blast
  have sum: "sum (λ x. count_list ?ws x * |x|) {u,v} = (count_list ?ws u) * |u| + (count_list ?ws v)*|v|"
    using assms by simp
  have "count_list ?ws u  0" and "count_list ?ws v  0"
    unfolding count_in_conv using assms by simp_all
  hence "|u| + |v|  |concat ?ws|"
    unfolding length_concat sum_list_map_eq_sum_count set sum
    using add_le_mono quotient_smaller by presburger
  thus ?thesis
    using |concat ?ws|  |concat ws| by linarith
qed


section ‹Root›

definition root :: "'a list  'a list  bool" ("_  _*" [51,51] 60 )
  where  "u  r* =  ( k. r@k = u)"
notation (latex output) root ("_  _*")

abbreviation not_root :: "['a list, 'a list]  bool"  ("_  _*" [51,51] 60 )
  where "u  r*  ¬ (u  r*)"

text‹Empty word has all roots, including the empty root.›

lemma emp_all_roots [simp]: "ε  r*"
  unfolding root_def using pow_0 by blast

lemma emp_all_roots'[elim]: "u = ε  u  r*"
  using emp_all_roots by blast

lemma rootI: "r@k  r*"
  using root_def by auto

lemma self_root: "u  u*"
  using rootI[of u "Suc 0"] by simp

lemma rootE[elim]: assumes "u  r*" obtains k where "r@k = u"
  using assms root_def by blast

lemma root_exp: "x  r*  x = r@(|x| div |r|)"
proof (rule iffI, cases "r = ε", force)
  assume "x  r*" and "r  ε"
  then obtain k where "r@k = x"
    unfolding root_def by blast
  from lenarg[OF this, unfolded pow_len]
  have "k = |x| div |r|"
    using nonzero_mult_div_cancel_right[OF nemp_len[OF r  ε], of k]  by auto
  from r@k = x[unfolded this, symmetric]
  show "x = r @ (|x| div |r|)".
qed (use root_def in metis)

lemma root_nemp_expE: assumes "w  r*" and "w  ε"
  obtains k where "r@k = w" "0 < k"
  using assms(1) assms(2) nemp_exp_pos root_exp by metis

lemma root_rev_iff[reversal_rule]: "rev u  rev t*  u  t*"
  unfolding root_def[reversed] using root_def..

lemma per_root_pref: "w  ε  w  r*  r ≤p w"
  using root_nemp_expE pow_pos triv_pref by metis

lemmas per_root_suf =  per_root_pref[reversed]

lemma per_exp_eq: "u ≤p ru  |u| = k*|r|  u  r*"
  using per_exp_pref[THEN pref_prod_eq] unfolding pow_len root_def by blast

lemma take_root: assumes "0 < k" shows "r = take |r| (r@k)"
  unfolding pow_pos[OF assms] by force

lemma root_nemp: "u  ε  u  r*  r  ε"
  unfolding root_def using emp_pow by auto

lemma root_shorter: assumes "u  ε" "u  r*" "u  r" shows "|r| < |u|"
proof (rule not_le_imp_less)
  from root_nemp_expE[OF u  r* u  ε]
  obtain k where "r@k = u" and "0 < k".
  from take_root[OF 0 < k, of r, unfolded r @ k = u]
  show "¬ |u|  |r|"
    using u  r by force
qed

lemma root_shorter_eq: "u  ε  u  r*  |r|  |u|"
  using root_shorter le_eq_less_or_eq by auto

lemma root_trans[trans]: "v  u*; u  t*  v  t*"
  by (metis root_def pow_mult)

lemma root_pow_root[intro]: "v  u*  v@n  u*"
  using rootI root_trans by blast

lemma root_len: "u  q*  k. |u| = k*|q|"
  unfolding root_def using pow_len by auto

lemma root_len_dvd: "u  t*  |t| dvd |u|"
  using root_len root_def by force

lemma common_root_len_gcd: "u  t*  v  t*  |t| dvd (gcd |u| |v|)"
  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 by blast
next
  assume "z  w  z*" thus "w  z*"
    unfolding root_def
    using exp_pref_cancel[of z 1 w, unfolded pow_1] by metis
qed

lemma add_roots[intro]: "w  z*  w'  z*  w  w'  z*"
  unfolding root_def using add_exps by blast

lemma concat_sing_list_pow: "ws  lists {u}  |ws| = k  concat ws = u@k"
proof(induct k arbitrary: ws)
  case (Suc k)
  have "ws  ε"
    using  list.size(3) nat.distinct(2)[of k, folded |ws| = Suc k] by blast
  from hd_Cons_tl[OF this]
  have "ws = hd ws # tl ws"  and "|tl ws| = k"
    using |ws| = Suc k by simp+
  then show ?case
    unfolding  pow_Suc hd_concat_tl[OF ws  ε, symmetric]
    using Suc.hyps[OF tl_in_lists[OF ws  lists {u}] |tl ws| = k]
      Nitpick.size_list_simp(2) lists_hd_in_set[of "ws" "{u}"] ws  lists{u} by blast
qed simp

lemma concat_sing_list_pow': "ws  lists{u}  concat ws = u@|ws|"
  by (simp add: concat_sing_list_pow)

lemma root_pref_cancel[elim]: assumes "xy  t*" and  "x  t*" shows "y  t*"
proof-
  obtain n m where "t@m = x  y" and "t@n = x"
    using xy  t*[unfolded root_def] x  t*[unfolded root_def] by blast
  from exp_pref_cancel[of t n y m, unfolded this]
  show "y  t*"
    using rootI by auto
qed

lemma root_suf_cancel [elim]: "u  v  r*  v  r*  u  r*"
  using exp_suf_cancel[of u r] unfolding root_def by metis

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 k m. x = t@k  y = t@m)"
  using  comm_append_are_replicate[of x y, folded pow_is_concat_replicate] pows_comm by auto

corollary comm_root:  "x  y = y  x     ( t. x  t*  y  t*)"
  unfolding root_def comm by fast

lemma comm_rootI:  "x  t*  y  t*  x  y = y  x"
  using comm_root  by blast

lemma commE[elim]: assumes  "x  y = y  x"
  obtains  t m k where "x = t@k" and "y = t@m" and "t  ε"
proof-
  from assms[unfolded comm]
  obtain t k m where "x = t@k" and "y = t@m"
    by blast
  from that[OF this]
  show thesis
  proof (cases "x  ε  y  ε")
    assume "x  ε  y  ε"
    thus thesis
      unfolding x = t@k y = t@m using t  ε  thesis
      by fastforce
  next
    assume "¬ (x  ε  y  ε)"
    hence "x = ε" "y = ε"
      by blast+
    from that[of "[undefined]" 0 0, unfolded this]
    show thesis
      by simp
  qed
qed

lemma comm_nemp_eqE: assumes "u  v = v  u" "u  ε" "v  ε"
  obtains k m where  "u@k = v@m" "0 < k" "0 < m"
proof-
  from commE[OF u  v = v  u]
  obtain t m k where "u = t@k" and "v = t@m".
  hence "0 < m" "0 < k"
    using u  ε v  ε by blast+
  have "u@m = v@k"
    unfolding u = t@k v = t@m pow_mult[symmetric]
    by (simp add: mult.commute)
  from that[OF this 0 < m 0 < k]
  show thesis.
qed

lemma comm_prod[intro]: assumes "ru = ur" and "rv = vr"
  shows "r(uv) = (uv)r"
  unfolding lassoc ru = ur unfolding rassoc rv = vr..

lemma LS_comm:
  assumes "y @ k  x = z @ l"
      and "z  y = y  z"
    shows "x  y = y  x"
proof -
  from z  y = y  z
  have "(y @ k  x)  y = y  (y @ k  x)"
    unfolding y @ k  x = z @ l by (fact comm_add_exp)
  then show "x  y = y  x"
    unfolding lassoc pow_comm[symmetric] unfolding rassoc cancel.
qed

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
word, which can be called its period word. Both these points of view are expressed by the following definition.
›

subsection "Periodic root"


lemma "u <p r  u  u ≤p r  u  r  ε"
  by simp

lemma per_rootI[intro]: "u ≤p r  u  r  ε  u <p r  u"
  by simp

lemma per_rootI'[intro]: assumes "u ≤p r@k" and  "r  ε" shows  "u <p r  u"
  using per_rootI[OF  pref_prod_pref[OF pref_pow_ext'[OF u ≤p r@k] u ≤p r@k] rε].

lemma per_root_nemp[dest]: "u <p r  u  r  ε"
  by simp

text ‹Empty word is not a periodic root but it has all nonempty periodic roots.›



text ‹Any nonempty word is its own periodic root.›

lemmas root_self = triv_spref

text‹"Short roots are prefixes"›

lemma "w <p r  u  |r|  |w|   r ≤p w"
  using pref_prod_long[OF sprefD1].

text ‹Periodic words are prefixes of the power of the root, which motivates the notation›

lemma pref_pow_ext_take: assumes "u ≤p r@k" shows "u ≤p take |r| u  r@k"
proof (rule le_cases[of "|u|" "|r|"])
  assume "|r|  |u|"
  show "u ≤p take |r| u  r @ k"
    unfolding pref_take[OF pref_prod_long[OF pref_pow_ext'[OF u ≤p r@k] |r|  |u|]]  using pref_pow_ext'[OF u ≤p r@k].
qed simp

lemma pref_pow_take: assumes "u ≤p r@k" shows "u ≤p take |r| u  u"
  using pref_prod_pref[of u "take |r| u" "r@k", OF pref_pow_ext_take u ≤p r@k, OF u ≤p r@k].


lemma per_root_powE: assumes "u <p r  u"
  obtains k where "u <p r@k" and "0 < k"
  using pref_prod_less[OF per_exp_pref[OF sprefD1]
  long_pow_exp'[OF per_root_nemp], OF assms assms] by blast

thm per_rootI per_rootI'

lemma per_root_powE': assumes "x <p r  x"
  obtains k where "x ≤p r@k" and "0 < k"
  using per_root_powE[OF assms] sprefD1 by metis

lemma per_root_modE' [elim]: assumes "u <p r  u"
  obtains p where "p <p r" and "r@(|u| div |r|)  p = u"
proof-
  have "r  ε"
    using assms by blast
  obtain m where "u <p r@m"
    using per_root_powE[OF u <p r  u].
  from pref_mod_pow[OF sprefD1[OF this] per_root_nemp[OF assms]]
  obtain k z where "k  m" and "z <p r" and "r @ k  z = u".
  have "k = (|u| div |r|)"
    using lenarg[OF r @ k  z = u, unfolded lenmorph pow_len]
    get_div[OF prefix_length_less[OF z <p r]] by metis
  thus ?thesis
    using that r @ k  z = u z <p r by blast
qed

lemma per_root_modE [elim]: assumes "u <p r  u"
  obtains n p s where "p  s = r" and "r@n  p = u" and "s  ε"
  using per_root_modE'[OF assms] spref_exE by metis

    lemma nemp_per_root_conv: "r  ε  u <p r  u  u ≤p r  u"
   by force



lemma root_ruler: assumes "w <p uw" "v <p uv"
  shows "w  v"
proof-
  obtain k l where "w <p u@k" "v <p u@l"
    using assms per_root_powE by metis
  moreover have "u@k  u@l"
    using conjug_pow eqd_comp by metis
  ultimately show ?thesis
    by (rule ruler_comp[OF sprefD1 sprefD1])
qed

lemmas same_len_nemp_root_eq = root_ruler[THEN pref_comp_eq]







lemma per_root_add_exp: assumes "u <p r  u" "0 < m" shows "u <p r@m  u"
  using 0 < m
proof (induct m)
    case (Suc m)
    then show ?case
      unfolding pow_Suc rassoc
      using spref_trans[OF u <p r  u, of "r  r @ m  u"] u <p r  u
      unfolding spref_cancel_conv by (cases "m = 0") simp_all
 qed simp

theorem per_root_pow_conv: "x <p r  x  ( k. x ≤p r@k)  r  ε"
  by (rule iffI) (use per_root_powE' per_root_nemp in metis, use per_rootI' in blast)

lemma per_root_exp': assumes "x ≤p r@k" shows "x ≤p r@|x|"
proof(cases "r = ε")
  assume "r  ε"
  have "|x|  |r@|x||"
    unfolding pow_len using nemp_le_len[OF r  ε] by force
  with pref_ext[OF x ≤p r@k, of "r@|x|", unfolded pows_comm[of r k]]
  show ?thesis
    by blast
qed (use assms in force)

lemma per_root_exp: assumes "x <p r  x" shows "x ≤p r@|x|"
proof-
  obtain k where "x ≤p r@k"
    using x <p r  x unfolding per_root_pow_conv by blast
  from per_root_exp'[OF this]
  show "x ≤p r@|x|".
qed

lemma per_root_drop_exp: "u <p (r@m)  u   u <p r  u"
  unfolding per_root_pow_conv unfolding pow_mult[symmetric]
  using emp_pow by blast

lemma per_root_exp_conv: "u <p (r@Suc m)  u   u <p r  u"
  by (rule iffI) (use per_root_drop_exp in blast, use per_root_add_exp in blast)

lemma pref_drop_exp: assumes "x ≤p z  x@m" shows "x ≤p z  x"
  using assms pow_comm pref_prod_pref pref_prolong triv_pref by metis

lemma per_root_drop_exp': "x ≤p r@(Suc k)  x@m   x ≤p  r  x"
  using nemp_Suc_pow_nemp per_rootI per_root_drop_exp pref_drop_exp sprefD by metis

lemma per_drop_exp': "0 < k  x ≤p r@k  x  x ≤p  r  x"
  using  nonzero_pow_emp per_rootI per_root_drop_exp sprefD by metis

lemmas per_drop_exp_rev = per_drop_exp'[reversed]

corollary comm_drop_exp: assumes "0 < m" and "u  r@m = r@m'  u" shows "r  u = u  r"
proof
  assume "r  ε" "u  ε"
  hence "m = m'"
    using lenarg[OF u  r@m = r@m'  u] unfolding lenmorph pow_len
    by auto
  have "ur ≤p ur@m"
    unfolding pow_pos[OF 0 < m] by simp
    have "ur ≤p r@m'  u  r"
      using pref_ext[of "u  r" "r@m  u" r, unfolded rassoc m = m', OF  ur ≤p ur@m[unfolded u  r@m = r@m'  u]].
    hence "ur ≤p r(ur)"
      using per_root_drop_exp[of "ur" r m'] 0 < m[unfolded m = m'] per_drop_exp' by blast
    from comm_ruler[OF self_pref[of "r  u"], of "r  u  r", OF this]
    show "r  u = u  r"
      unfolding prefix_comparable_def
      by force
qed

lemma comm_drop_exp': assumes "u@k  v = v  u@k'" "0 < k'" shows "u  v = v  u"
  using  comm_drop_exp[OF 0 < k' assms(1)[symmetric]].

lemma comm_drop_exps[elim]: assumes "u@m  v@k = v@k  u@m" and "0 < m" and "0 < k" shows "u  v = v  u"
  using comm_drop_exp[OF 0 < k u@m  v@k = v@k  u@m] comm_drop_exp[OF 0 < m, of v u m] by blast

lemma comm_pow_roots:
  assumes "0 < m" and "0 < k"
  shows "u@m  v@k = v@k  u@m  u  v = v  u"
  by (rule, use comm_drop_exps[OF _ assms] in blast)
  (use comm_add_exps in blast)

corollary pow_comm_comm: assumes "x@j = y@k" and "0 < j" shows "xy = yx"
  using  comm_drop_exp[OF 0 < j, of y x j, unfolded x@j = y@k, OF pow_comm[symmetric]].


lemma pow_comm_comm': assumes comm: "u@(Suc k) = v@(Suc l)" shows "u  v = v  u"
  using comm pow_comm_comm by blast

lemma comm_trans: assumes uv: "uv =  vu" and vw: "wv = vw" 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@(|v| * |w|) = w@(|v| * |u|)"
      unfolding pow_mult comm_common_power[OF uv] comm_common_power[OF vw]
      unfolding pow_mult[symmetric] mult.commute[of "|u|"]..
    obtain k l where k: "|v| * |w| = Suc k" and l: "|v| * |u| = 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

lemma root_comm_root: assumes "x ≤p u  x" and "v  u = u  v" and "u  ε"
  shows "x ≤p v  x"
  using per_rootI[OF x ≤p ux u  ε] per_exp_pref  commE[OF v  u = u  v] per_drop_exp'
   assms(1) assms(3) nemp_pow by metis


lemma drop_per_pref: assumes "w <p u  w" shows "drop |u| w ≤p w"
  using pref_drop[OF sprefD1[OF w <p u  w], of "|u|", unfolded drop_pref[of u w]].

lemma per_root_trans[intro]: assumes "w <p u  w" and "u  t*" shows "w <p t  w"
  using per_root_drop_exp rootE[OF u  t*] w <p u  w by metis

lemma per_root_trans'[intro]: "w ≤p u  w  u  r*  u  ε  w ≤p r  w"
  using per_root_trans sprefD1 per_rootI by metis

lemmas per_root_trans_suf'[intro] = per_root_trans'[reversed]

text‹Note that
@{term "w <p u  w  u <p t  u  w <p t  w"}
does not hold.
›

lemma per_root_same_prefix:"w <p r  w  w' ≤p r  w'   w  w'"
  using root_ruler by auto

lemma take_after_drop:  "|u| + q  |w|  w <p u  w  take q (drop |u| w) = take q w"
  using pref_share_take[OF drop_per_pref[of w u] len_after_drop[of "|u|" 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: "|u| + |v|  |w|"
  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
  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_root: assumes "w <p u  w" and  "w <p v  w" and "|u| + |v|  |w|" shows "uv = vu"
  using two_pers[OF sprefD1[OF assms(1)] sprefD1[OF assms(2)] assms(3)].

subsection ‹Maximal root-prefix›

lemma max_root_mismatch: assumes "u  [a] <p r  u  [a]" and "u  [b] ≤p w" and "a  b"
  shows "w p r  w = u"
proof (rule lcp_first_mismatch_pref[OF u  [b] ≤p w _ a  b[symmetric]])
  have "u  [a] ≤p r  u"
    using assms(1)[unfolded lassoc spref_snoc_iff].
  thus "u  [a] ≤p r  w"
    using append_prefixD[OF u  [b] ≤p w] pref_prolong by blast
qed


lemma max_pref_per_root:  "u p r  u ≤p r  (u p r  u)"
  by (rule pref_prod_pref[of _ _ u]) force+

lemma max_pref_pref:
  assumes "r  ε"
  shows "u p r  u ≤p r@|u p r  u|"
proof-
  have "u p r  u <p r  (u p r  u)"
    using assms max_pref_per_root by auto
  from per_root_exp[OF this]
  show ?thesis.
qed


lemma max_pref_lcp_root_pow: assumes "r  ε" and "|u p r  u|  k"
  shows "u p r  u = u p r@k" (is "?max = u p r@k")
proof (rule pref_antisym)
  from max_pref_pref[OF assms(1)] le_exps_pref[OF assms(2)]
  have "?max ≤p r@k"
    using pref_trans by blast
  thus "?max ≤p u p r@k"
    by force
  show "u p r@k ≤p ?max"
  proof (rule lcp.boundedI, force)
    show "u p r @ k ≤p r  u"
    proof (rule pref_prolong)
      show "u p r @ k ≤p r  (u p r @ k)"
        using  lcp.cobounded2  by (rule pref_prod_root[of "u p r@k"])
      show "u p r @ k ≤p u"
        using lcp.cobounded1.
    qed
  qed
qed

lemma max_pref_shorter_lcp: assumes "u p r  u <p v p r  v"
  shows "u p v = u p r  u"
proof (cases)
  assume "r = ε"
  then show ?thesis
  using assms by (clarify, unfold emp_simps lcp.idem) (use lcp.absorb3 in blast)
next
  let ?u = "u p r  u" and ?v = "v p r  v"
  assume "r  ε"
  from max_pref_lcp_root_pow[OF this]
  obtain k where "?u = u p r@k" and "?v = v p r@k"
    using pref_len' suf_len' by meson
  from ruler_spref_lcp[OF assms[unfolded this], folded ?u = u p r@k]
  show "u p v = u p r  u".
qed


find_theorems "?u p ?r  ?u"


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,
it is secondary, and less convenient for reasoning.›

definition period :: "'a list  nat  bool"
  where [simp]: "period w n  w <p (take n w)  w"

lemma period_I': "w  ε  0 < n  w ≤p (take n w)  w  period w n"
  unfolding period_def  by fastforce

lemma periodI[intro]: "w  ε  w <p r  w  period w |r|"
  by (elim period_I'[of _ "|r|", OF _ nemp_pos_len])
     (blast, use pref_pow_take per_root_powE' in metis)

text‹The numeric definition respects the following convention about empty words and empty periods.›

lemma emp_no_period: "¬ period ε n"
  by simp

lemma "¬ period w 0"
  by simp



lemma per_nemp: "period w n   w  ε"
  by simp

lemma per_not_zero: "period w n   0 < n"
  by simp

lemma per_pref: "period w n   w ≤p take n w  w"
  by simp

text‹A nonempty word has all "long" periods›

lemma all_long_pers: " w  ε; |w|  n   period w n"
  by simp

lemma len_is_per: "w  ε  period w |w|"
  by simp

text‹The standard numeric definition of a period uses indeces.›

lemma period_indeces: assumes "period w n" and "i + n < |w|" 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 < |w|]]].
  also have "... = w ! (i + n)"
    using pref_index[OF per_pref[OF period w n] i + n < |w|, symmetric] unfolding add.commute[of n i].
  finally show ?thesis.
qed

lemma indeces_period:
  assumes  "w  ε" and "0 < n" and  forall: " i. i + n < |w|  w!i = w!(i+n)"
  shows "period w n"
proof-
  have "|w|  |take n w  w|"
    by auto
  {fix j assume "j < |w|"
    have "w ! j = (take n w  w) ! j"
    proof (cases "j < |take n w|")
      assume "j < |take n w|" show "w ! j = (take n w  w) ! j"
        using pref_index[OF take_is_prefix j < |take n w|, symmetric]
        unfolding pref_index[OF triv_pref j < |take n w|, of w].
    next
      assume "¬ j < |take n w|"
      from leI[OF this] j < |w|
      have "|take n w| = n"
        by force
      hence "j = (j - n) + n" and "(j - n) + n < |w|"
        using  leI[OF ¬ j < |take n w|] j < |w| 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 |take n w| = n]]
      show "w ! j = (take n w  w) ! j"
        using j = (j - n) + n by simp
    qed}
  with index_pref[OF |w|  |take n w  w|]
  have "w ≤p take n w  w" by blast
  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 period_rev: assumes "period w p" shows "period (rev w) p"
proof (rule indeces_period[of "rev w" p, OF _ per_not_zero[OF assms]])
  show "rev w  ε"
    using assms[unfolded period_def] by force
next
  fix i assume "i + p < |rev w|"
  from this[unfolded length_rev] add_lessD1
  have "i < |w|" and "i + p < |w|" by blast+
  have e: "|w| - Suc (i + p) + p = |w| - Suc i" using i + p < |rev w| by simp
  have "|w| - Suc (i + p) + p < |w|"
    using i + p < |w| Suc_diff_Suc i < |w|
      diff_less_Suc e less_irrefl_nat not_less_less_Suc_eq by metis
  from period_indeces[OF assms this] rev_nth[OF i  < |w|, folded e] rev_nth[OF i + p < |w|]
  show "rev w ! i = rev w !(i+p)" by presburger
qed

lemma period_rev_conv [reversal_rule]: "period (rev w) n  period w n"
  using period_rev period_rev[of "rev w"] unfolding rev_rev_ident by (intro iffI)

lemma period_fac: assumes "period (uwv) p" and "w  ε"
  shows "period w p"
proof (rule indeces_period)
  show "0 < p" using per_not_zero[OF period (uwv) p].
  fix i assume "i + p < |w|"
  hence "|u| + i + p  < |uwv|"
    by simp
  from period_indeces[OF period (uwv) p this]
  have "(uwv)!(|u| + i) = (uwv)! (|u| + (i + p))"
    by (simp add: add.assoc)
  thus "w!i = w!(i+p)"
    using nth_append_length_plus[of u "wv" i, unfolded lassoc] i + p < |w| add_lessD1[OF i + p < |w|]
      nth_append[of w v] by auto
qed (simp add: w  ε)

lemma period_fac': "period v p  u ≤f v  u  ε  period u p"
  by (elim facE, hypsubst, rule period_fac)

lemma pow_per[intro]: assumes "y  ε" and "0 < k" shows "period (y@k) |y|"
  using period_I'[OF _ nemp_pos_len[OF y  ε] pref_pow_ext_take, OF _ self_pref]
  assms by blast

lemma per_fac: assumes "w  ε" and "w ≤f y@k" shows "period w |y|"
proof-
  have "y  ε"
    using assms  by force
  have "0 < k"
    using assms nemp_exp_pos sublist_Nil_right by metis
  from pow_per[OF y  ε this] period_fac  facE[OF w ≤f y@k] w  ε
  show "period w |y|" by metis
qed

text‹The numeric definition is equivalent to being prefix of a power.›

theorem period_pref: "period w n  (k r. w ≤np r@k  |r| = n)" (is "_  ?R")
proof(cases "w = ε")
  assume "w  ε"
  show "period w n  ?R"
  proof
    assume "period w n"
    consider (short) "|w|  n" |  (long) "n < |w|"
      by linarith
    then show ?R
    proof(cases)
      assume "|w|  n"
      from le_add_diff_inverse[OF this]
      obtain z where "|w  z| = n"
        unfolding lenmorph using exE[OF Ex_list_of_length[of "n - |w|"]] by metis
      thus ?R
        using  pow_1 npI'[OF w  ε] by metis
    next
      assume "n < |w|"
      then show ?R
        unfolding nonempty_prefix_def
        using w  ε take_len[OF less_imp_le[OF n < |w|]]
        per_root_powE[OF period w n[unfolded period_def]]
        sprefD1 by metis
    qed
  next
    assume ?R
    then obtain k r where "w ≤np r@k" and "n = |r|" by blast
    have "w ≤p take n w  w"
      using  pref_pow_take[OF npD[OF w ≤np r @ k], folded n = |r|].
    have "n  0"
      unfolding length_0_conv[of r, folded n = |r|] using w ≤np r @ k by force
    hence "take n w  ε"
      unfolding n = |r| using w  ε by simp
    thus "period w n"
      unfolding period_def using w ≤p take n w  w by blast
  qed
qed simp

text ‹Two more characterizations of a period›

theorem per_shift: assumes "w  ε" "0 < n"
  shows "period w n  drop n w ≤p w"
proof
  assume "period w n" show "drop n w ≤p w"
    using drop_per_pref[OF period w n[unfolded period_def]]
      append_take_drop_id[of n w, unfolded append_eq_conv_conj] by argo
next
  assume "drop n w ≤p w"
  show "period w n"
    using conjI[OF pref_cancel'[OF drop n w ≤p w, of "take n w"] take_nemp[OF w  ε 0 < n]]
    unfolding  append_take_drop_id  by force
qed

lemma rotate_per_root: assumes "w  ε" and "0 < n" and "w = rotate n w"
  shows "period w n"
proof (cases "|w|  n")
  assume "|w|  n"
  from all_long_pers[OF w  ε, OF this]
  show "period w n".
next
  assume not: "¬ |w|  n"
  have "drop (n mod |w|) w ≤p w"
    using prefI[OF rotate_drop_take[symmetric, of n w]]
    unfolding w = rotate n w[symmetric].
  from per_shift[OF w  ε 0 < n] this[unfolded mod_less[OF not[unfolded not_le]]]
  show "period w n"..
qed

subsubsection ‹Various lemmas on periods›

lemma period_drop: assumes "period w p" and "p < |w|"
  shows "period (drop p w) p"
  using period_fac[of "take p w" "drop p w" ε p] p < |w| period w p
  unfolding append_take_drop_id drop_eq_Nil not_le append_Nil2 by blast

lemma ext_per_left: assumes "period w p" and  "p  |w|"
  shows "period (take p w  w) p"
proof-
  have f: "take p (take p w  w) = take p w"
    using p  |w| by simp
  show ?thesis
    using  period w p pref_cancel'[of w "take p w  w" "take p w" ]
    unfolding f period_def
    by blast
qed

lemma ext_per_left_power: "period w p  p  |w|  period ((take p w)@k  w) p"
proof (induction k)
  case (Suc k)
  show ?case
    using ext_per_left[OF Suc.IH[OF period w p p  |w|]] p  |w|
    unfolding pref_share_take[OF per_exp_pref[OF per_pref[OF period w p]] p  |w|,symmetric]
      lassoc pow_Suc[symmetric] by fastforce
qed auto

lemma take_several_pers: assumes "period w n" and "m*n  |w|"
  shows "(take n w)@m = take (m*n) w"
proof (cases "m = 0")
  assume "m  0"
  have "|(take n w)@m| = m*n"
    unfolding pow_len nat_prod_le[OF m  0 m*n  |w|, THEN take_len] by blast
  have "(take n w)@m ≤p w"
    using  period w n[unfolded period_def]
        ruler_le[of "take n w@m" "take n w@m  w" w, OF triv_pref] m * n  |w|[folded |take n w@m| = m * n]
        per_exp_pref sprefD by metis
  show ?thesis
    using pref_take[OF take n w@m ≤p w, unfolded  |take n w@m| = m * n, symmetric].
qed simp

lemma per_div: assumes "n dvd |w|" and "period w n"
  shows "(take n w)@(|w| div n) = w"
  using take_several_pers[OF period w n div_times_less_eq_dividend] unfolding dvd_div_mult_self[OF n dvd |w|] take_self.

lemma per_mult: assumes "period w n" and "0 < m" shows "period w (m*n)"
proof (cases "m*n  |w|")
  have "w  ε" using per_nemp[OF period w n].
  assume "¬ m * n  |w|" thus "period w (m*n)"
    using all_long_pers[of  w "m * n", OF w  ε] by linarith
next
  assume "m * n  |w|"
  show "period w (m*n)"
    using  period w n
    unfolding period_def
    using per_root_add_exp[of w "take n w"] 0 < m
     take_several_pers[OF period w n m*n  |w|, symmetric]
    by presburger
qed


theorem  two_periods:
  assumes "period w p" "period w q"  "p + q  |w|"
  shows "period w (gcd p q)"
proof-
  obtain t where "take p w  t*" "take q w  t*"
    using two_pers_root[OF period w p[unfolded period_def] period w q[unfolded period_def],
        unfolded take_len[OF add_leD1[OF p + q  |w|]] take_len[OF add_leD2[OF p + q  |w|]],
        OF p + q  |w|, unfolded comm_root[of "take p w" "take q w"]] by blast
  hence "w <p t  w"
    using period w p period_def per_root_trans by blast
  have "period w |t|"
    using  periodI[OF  per_nemp[OF period w p] w <p t  w].
  have "|t| 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  |w|]] take_len[OF add_leD2[OF p + q  |w|]].
  from dvd_div_eq_0_iff[OF this]
  have "0 < gcd p q div |t|"
    using per_not_zero[OF period w p] unfolding  gcd_nat.eq_neutr_iff by blast
  from per_mult[OF period w |t| this]
  show ?thesis
    unfolding dvd_div_mult_self[OF |t| dvd (gcd p q)].
qed

lemma index_mod_per_root: assumes "r  ε" and i: " i < |w|. w!i = r!(i mod |r|)" shows  "w <p r  w"
proof-
  have "i < |w|  (r  w) ! i = r ! (i mod |r|)" 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  using r  ε by auto
qed

lemma index_pref_pow_mod: "w ≤p r@k  i < |w|   w!i = r!(i mod |r| )"
  using  index_pow_mod[of i r k] less_le_trans[of i "|w|" "|r@k|"] prefix_length_le[of w "r@k"] pref_index[of w "r@k" i] by argo

lemma index_per_root_mod: "w <p r  w  i < |w|   w!i = r!(i mod |r|)"
  using index_pref_pow_mod[of w r _ i] per_root_powE' by metis

lemma root_divisor: assumes "period w k" and  "k dvd |w|" shows "w  (take k w)*"
  using rootI[of "take k w" "(|w| div k)"] unfolding
    take_several_pers[OF period w k, of "|w| div k", unfolded dvd_div_mult_self[OF k dvd |w|] take_self, OF , OF order_refl].

lemma per_pref': assumes "u  ε" and "period v k" and  "u ≤p v" shows "period u k"
proof-
  { assume "k  |u|"
    have "take k v = take k u"
      using  pref_share_take[OF u ≤p v k  |u|]  by auto
    hence "take k v  ε"
      using period 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  period v k
      unfolding period_def   take k v = take k u
      using pref_trans[OF u ≤p v, of "take k u  v"]
      by blast
    hence "u ≤p take k u  u"
      using u ≤p v pref_prod_pref by blast
    hence ?thesis
      using take k u  ε period_def by blast
  }
  thus ?thesis
    using u  ε all_long_pers nat_le_linear by blast
qed

subsection "Period: overview"
notepad
begin
  fix w r::"'a list"
  fix n::nat
  assume "w  ε" "r  ε" "0 < n"
  have "¬ w <p ε  w"
    by simp
  have "¬ ε <p ε  ε"
    by simp
  have "ε <p r  ε"
    using r  ε by blast
  have "¬ period w 0"
    by simp
  have "¬ period ε 0"
    by simp
  have "¬ period ε n"
    by simp
end

subsection ‹Singleton and its power›

primrec letter_pref_exp :: "'a list  'a  nat" where
  "letter_pref_exp ε a = 0" |
  "letter_pref_exp (b # xs) a = (if b  a then 0 else Suc (letter_pref_exp xs a))"

definition letter_suf_exp :: "'a list  'a  nat" where
  "letter_suf_exp w a = letter_pref_exp (rev w) a"

lemma concat_len_one: assumes "|us| = 1" shows "concat us = hd us"
  using concat_sing[OF sing_word[OF |us| = 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]@k = c # w" unfolding root_def by blast
  thus "c = a  w  [a]*"
  proof (cases "0 < k")
    assume "[a] @ k = c # w" and "0 < k"
    from eqd_eq[of "[a]", OF this(1)[unfolded hd_word[of _ w] pow_pos[OF 0 < k]]]
    show ?thesis
      unfolding root_def by auto
  qed simp
qed

lemma pref_sing_pow: assumes "w ≤p [a]@m" shows "w = [a]@(|w|)"
proof-
  have  "[a]@m = [a]@(|w|)[a]@(m-|w|)"
    using pop_pow[OF prefix_length_le[OF assms, unfolded sing_pow_len], of "[a]", symmetric].
  show ?thesis
    using  eqd_eq(1)[of w "w¯>[a]@m" "[a]@(|w|)""[a]@(m-|w|)",
          unfolded lq_pref[OF assms] sing_pow_len,
          OF [a]@m = [a]@(|w|)[a]@(m-|w|) refl].
qed

lemma sing_pow_palindrom: assumes "w = [a]@k" shows "rev w = w"
  using rev_pow[of "[a]" "|w|", unfolded rev_sing]
  unfolding pref_sing_pow[of w a k, unfolded assms[unfolded root_def, symmetric], OF self_pref, symmetric].

lemma suf_sing_power: assumes "w ≤s [a]@m" shows "w  [a]*"
  using sing_pow_palindrom[of "rev w" a "|rev w|", unfolded rev_rev_ident]
    pref_sing_pow[of "rev w" a m, OF w ≤s [a]@m[unfolded suffix_to_prefix rev_pow rev_rev_ident rev_sing]]
    rootI[of "[a]" "|rev w|"]  by auto

lemma sing_fac_pow: assumes "w  [a]*" and  "v ≤f w" shows "v  [a]*"
proof-
  obtain k where "w = [a]@k" using w  [a]*[unfolded root_def] by blast
  obtain p where "p ≤p w" and "v ≤s p"
    using fac_pref_suf[OF v ≤f w] by blast
  hence "v ≤s [a]@ |p|"
    using pref_sing_pow[OF p ≤p w[unfolded w = [a]@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

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)
    case (Cons c w)
    then show ?case
      by (simp add: sing_pow_hd_tl)
  qed simp
next
  assume "w  [a]*"
  then show ?All
  proof (induct w)
    case (Cons c w)
    then show ?case
      unfolding sing_pow_hd_tl by simp
  qed simp
qed

lemma sing_fac_set: "[a] ≤f x  a  set x"
  by fastforce

lemma set_sing_pow_hd [simp]: assumes "0 < k" shows "a  set ([a]@k)"
  using assms gr0_conv_Suc by force

lemma neq_set_not_root: "a  b  b  set x  x  [a]*"
  using all_set_sing_pow by metis

lemma sing_pow_set_Suc[simp]: "set ([a]@Suc k) = {a}"
  by (induct k, simp_all)

lemma sing_pow_set[simp]: assumes "0 < k" shows "set ([a]@k) = {a}"
  using sing_pow_set_Suc[of _ "k-1", unfolded Suc_minus_pos[OF assms]].

lemma sing_pow_set_sub: "set ([a]@k)  {a}"
  by (induct k, simp_all)

lemma unique_letter_fac_expE: assumes "w ≤f [a]@k"
  obtains m where "w = [a]@m"
  using unique_letter_wordE''[OF subset_trans[OF set_mono_sublist[OF assms] sing_pow_set_sub]] by blast


lemma neq_in_set_not_pow: "a  b  b  set x  x  [a]@k"
  by (cases "0 < k", use sing_pow_set singleton_iff in metis) force

lemma sing_pow_card_set_Suc: assumes "c = [a]@Suc k" shows "card(set c) = 1"
proof-
  have "card {a} = 1" by simp
  from this[folded sing_pow_set_Suc[of a k]]
  show "card(set c) = 1"
    unfolding assms.
qed

lemma sing_pow_card_set:  assumes "k  0" and "c = [a]@k" shows "card(set c) = 1"
  using sing_pow_card_set_Suc[of c a "k - 1", unfolded Suc_minus[OF k  0], OF c = [a]@k].

lemma sing_pow_set': "u  [a]*  u  ε  set u = {a}"
  unfolding all_set_sing_pow[symmetric]
  using lists_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 root_sing_set_iff: "u  [a]*  set u  {a}"
  by (rule, use sing_pow_set'[of u a, folded set_empty2] in force, use all_set_sing_pow[of u a] in force)

lemma letter_pref_exp_hd: "u  ε  hd u = a  letter_pref_exp u a  0"
  by (induct u, auto)


lemma letter_pref_exp_pref: "[a]@(letter_pref_exp w a) ≤p w "
  by(induct w, simp, simp)

lemma letter_pref_exp_Suc: "¬ [a]@(Suc (letter_pref_exp w a)) ≤p w "
  by (induct w, simp_all add: prefix_def)

lemma takeWhile_letter_pref_exp: "takeWhile (λx. x = a) w =[a]@(letter_pref_exp w a)"
  by (induct w, simp, simp)

lemma concat_takeWhile_sing: "concat (takeWhile (λ x. x = u) ws) = u@|takeWhile (λ x. x = u) ws|"
  unfolding takeWhile_letter_pref_exp  concat_sing_pow  sing_pow_len ..

lemma dropWhile_distinct: assumes "w  [a]@(letter_pref_exp w a)"
  shows "[a]@(letter_pref_exp w a)[hd (dropWhile (λx. x = a) w)] ≤p w"
proof-
  have nemp: "dropWhile (λx. x = a) w  ε"
    using takeWhile_dropWhile_id[of "λx. x = a" w, unfolded  takeWhile_letter_pref_exp] w  [a]@(letter_pref_exp w a)
    by force
  from takeWhile_dropWhile_id[of "λx. x = a" w, unfolded takeWhile_letter_pref_exp]
  have "[a]@(letter_pref_exp w a)[hd (dropWhile (λx. x = a) w)] tl (dropWhile (λx. x = a) w) = w"
    unfolding hd_tl[OF nemp].
  thus ?thesis
    unfolding lassoc using triv_pref by blast
qed

lemma letter_pref_exp_mismatch: "u = [a]@letter_pref_exp u a  v  v  ε  hd v  a"
  using hd_pref letter_pref_exp_Suc[unfolded pow_Suc'] same_prefix_prefix by metis

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]@|w|"
  by(induct w,  auto)

lemma dropWhile_sing_pow: "dropWhile (λ x. x = a) w = ε  w = [a]@|w|"
  by(induct w,  auto)

lemma nemp_takeWhile_hd: "us  ε  hd (takeWhile (λ a. a = hd us) us) = hd us"
  by (simp add: pref_hd_eq takeWhile_eq_Nil_iff takeWhile_is_prefix)

lemma nemp_takeWhile_last: "us  ε  last (takeWhile (λ a. a = hd us) us) = hd us"
proof (induct us)
  case (Cons a us)
  then show ?case
    by (simp add: takeWhile_eq_Nil_iff) blast
qed simp

lemma card_set_decompose: assumes "1 < card (set us)"
  shows "takeWhile (λ a. a = hd us) us  ε" and "dropWhile (λ a. a = hd us) us  ε" and
         "set (takeWhile (λ a. a = hd us) us) = {hd us}" and
         "last (takeWhile (λ a. a = hd us) us)  hd(dropWhile (λ a. a = hd us) us)"
proof-
  have "us  ε"
    using assms by force
  thus "takeWhile (λa. a = hd us) us  ε"
    by (simp add: takeWhile_eq_Nil_iff)
  from sing_pow_set'[OF takeWhile_sing_root this]
  show set: "set (takeWhile (λ a. a = hd us) us) = {hd us}".
  show "dropWhile (λa. a = hd us) us  ε"
  proof (rule notI)
    assume "dropWhile (λa. a = hd us) us = ε"
    from set[unfolded takeWhile_dropWhile_id[of "(λa. a = hd us)" us, unfolded this emp_simps]]
    show False
      using assms by force
  qed
  from hd_dropWhile[OF this]
  show "last (takeWhile (λ a. a = hd us) us)  hd(dropWhile (λ a. a = hd us) us)"
    unfolding nemp_takeWhile_last[OF us  ε] by simp
qed

lemma distinct_letter_in: assumes "w  [a]*"
  obtains m b q where "[a]@m  [b]  q = w" and "b  a"
proof-
  have "dropWhile (λ x. x = a) w  ε"
    unfolding dropWhile_sing_pow using assms rootI[of "[a]" "|w|"] 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"
    using hd_dropWhile[OF dropWhile (λx. x = a) w  ε].
  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]@m  [b]  q = w" and "b  hd w" and "m  0"
proof-
  obtain m b q where a1: "[hd w]@m  [b]  q = w" and a2: "b  hd w"
    using distinct_letter_in[OF assms].
  have "m  0"
  proof (rule notI)
    assume "m = 0"
    note a1[unfolded this pow_zero emp_simps, folded hd_word]
    thus False using a2 by force
  qed
  from that[OF a1 a2 this]
  show thesis.
qed

lemma distinct_letter_in_hd': assumes "w  [hd w]*"
  obtains m b q where  "[hd w]@Suc m  [b]  q = w" and "b  hd w"
using distinct_letter_in_hd[OF assms] Suc_minus by metis

lemma distinct_letter_in_suf: assumes "w  [a]*"
  obtains m b where "[b]  [a]@m  ≤s w" and "b  a"
  using distinct_letter_in[reversed, unfolded rassoc, OF assms]
  unfolding suffix_def by metis

lemma sing_pow_exp: assumes "w  [a]*" shows "w = [a]@|w|"
proof-
  obtain k where "[a] @ 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 < |w|" shows "w ! i = a"
  using  sing_pow_nth[OF i < |w|, 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]@m  [b] ≤p z" and  "a  b"
  shows "w  z p z  w = [a]@m"
proof-
  obtain k z' where "w = [a]@k" "z = [a]@m  [b]  z'" "k  0"
    using w  [a]* [a]@m  [b] ≤p z w  ε nemp_pow[of "[a]"]
    unfolding root_def
    by (auto simp add: prefix_def)
  hence eq1: "w  z = [a]@m  ([a]@k  [b]  z')" and eq2: "z  w = [a]@m  ([b]  z' [a]@k)"
    by (simp add: w = [a]@k z = [a]@m  [b]  z' pows_comm)+
  have "hd ([a]@k  [b]  z') = a"
    using hd_append2[OF w  ε, of "[b]z'",
        unfolded w = (a # ε)@k hd_sing_pow[OF k  0, of a]].
  moreover have "hd([b]  z' [a]@k) = b"
    by simp
  ultimately have "[a]@k  [b]  z' p [b]  z' [a]@k = ε"
    by (simp add: a  b lcp_distinct_hd)
  thus ?thesis using eq1 eq2 lcp_ext_left[of "[a]@m" "[a]@k  [b]  z'" "[b]  z' [a]@k"]
    by simp
qed

lemma per_one: assumes "w <p [a]  w" shows "w  [a]*"
proof-
  have "w <p (a # ε) @ n  0 < n   w  [a]*" for n
    using pref_sing_pow[of w a] sprefD1 rootI[of "[a]" "|w|"] by metis
  with  rootI per_root_powE[OF assms]
  show ?thesis
    by blast
qed

lemma per_one': "w  [a]*  w <p [a]  w"
  using comm_root  self_root triv_spref[OF not_Cons_self2]  by blast

lemma per_sing_one: assumes "w  ε" "w <p [a]  w" shows "period w 1"
  using periodI[OF w  ε w <p [a]  w] 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"
  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 <p w"
  unfolding border_def  by simp

lemma borderD_suf: "x ≤b w  x ≤s w"
  unfolding border_def by blast

lemma borderD_ssuf: "x ≤b w  x <s w"
  unfolding border_def by blast

lemma borderD_nemp: "x ≤b w  x  ε"
  using border_def by blast

lemma borderD_neq: "x ≤b w  x  w"
  unfolding border_def by blast

lemma borderedI: "u ≤b w  bordered w"
  unfolding bordered_def by fast

lemma border_lq_nemp: assumes "x ≤b w" shows "x¯>w  ε"
  using assms borderD_spref lq_spref by blast

lemma border_rq_nemp: assumes "x ≤b w" shows "w<¯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.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<¯x)  x"
  unfolding border_def using rq_suf_suf ruler' by metis

lemmas border_lq_suf_comp = border_lq_comp[reversed]

subsection "The shortest border"

lemma border_len:  assumes "x ≤b w"
  shows "1 < |w|" and "0 < |x|" and "|x| < |w|"
proof-
  show "|x| < |w|"
    using assms prefix_length_less unfolding border_def strict_prefix_def
    by blast
  show "0 < |x|"
    using assms unfolding border_def by blast
  thus "1 < |w|"
    using assms |x| < |w| unfolding border_def
    by linarith
qed

lemma borders_compare: assumes "x ≤b w" and "x' ≤b w" and "|x'| < |x|"
  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 |x'| < |x|]]
    suf_ruler_le[OF borderD_suf[OF x' ≤b w] borderD_suf[OF x ≤b w] less_imp_le_nat[OF |x'| < |x|]]
    borderD_nemp[OF x' ≤b w] |x'| < |x|
    borderI by blast

lemma unbordered_border:
  "bordered w    x. x ≤b w  ¬ bordered x"
proof (induction "|w|" arbitrary: w rule: less_induct)
  case less
  obtain x' where "x' ≤b w"
    using bordered_def less.prems by blast
  show ?case
  proof (cases "bordered x'")
    assume "¬ bordered x'"
    thus ?case
      using x' ≤b w by blast
  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
  qed
qed

lemma unbordered_border_shortest: "x ≤b w  ¬ bordered x   y ≤b w  |x|  |y|"
  using bordered_def[of x] borders_compare[of x w y] not_le_imp_less[of "|x|" "|y|"] by blast

lemma long_border_bordered: assumes long: "|w| < |x| + |x|" and border: "x ≤b w"
  shows "(w<¯x)¯>x ≤b x"
proof-
  define p s where "p = w<¯x" and "s = x¯>w"
  hence eq: "px = xs"
    using assms unfolding border_def using lq_pref[of x w] rq_suf[of x w]  by simp
  have "|p| < |x|"
    using lenarg[OF p_def] long unfolding rq_len by linarith
  have px: "p  p¯>x = x" and sx: "p¯>x  s = x"
    using eqd_pref[OF eq less_imp_le, OF |p| < |x|] by blast+
  have "p¯>x  ε"
    using |p| < |x| px by fastforce
  have "p  ε"
    using border_rq_nemp[OF border] p_def
    by presburger
  have "p¯>x  x"
    using p  ε px by force
  have "(p¯>x) ≤b x"
    unfolding border_def
    using eqd_pref[OF eq less_imp_le, OF |p| < |x|] p¯>x  ε p¯>x  x by blast
  thus ?thesis
    unfolding p_def.
qed

thm long_border_bordered[reversed]

lemma border_short_dec: assumes border: "x ≤b w" and short: "|x| + |x|  |w|"
  shows "x  x¯>(w<¯x)  x = w"
proof-
  have eq: "xx¯>w = w<¯xx"
    using lq_pref[OF borderD_pref[OF border]] rq_suf[OF borderD_suf[OF border]] by simp
  have "|x|  |w<¯x|"
    using short unfolding rq_len by linarith
  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"
  obtains u v where "uvu = w" and "u  ε"
proof-
  obtain u where "u ≤b w" and "¬ bordered u"
    using unbordered_border[OF assms] by blast
  have "|u| + |u|  |w|"
    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

lemma emp_not_bordered: "¬ bordered ε"
  by simp

lemma bordered_nemp: "bordered w  w  ε"
  using emp_not_bordered by blast

lemma sing_not_bordered: "¬ bordered [a]"
  using bordered_dec[of "[a]" False] append_eq_Cons_conv[of _ _ a ε] suf_nemp by fast

subsection "Relation to period and conjugation"

lemma border_conjug_eq: "x ≤b w  (w<¯x)  w = w  (x¯>w)"
  using lq_rq_reassoc_suf[OF borderD_pref borderD_suf, symmetric] by blast

lemma border_per_root: "x ≤b w  w ≤p (w<¯x)  w"
  using border_conjug_eq by blast

lemma per_root_border: assumes "|r| < |w|" and "r  ε" and "w ≤p r  w"
  shows "r¯>w ≤b w"
proof
  have "|r|  |w|" and "r ≤p w"
    using less_imp_le[OF |r| < |w|] pref_prod_long[OF w ≤p r  w] by blast+
  show "r¯>w ≤p w"
    using pref_lq[OF w ≤p r  w, of r] unfolding lq_triv.
  show "r¯>w ≤s w"
    using r ≤p w by (auto simp add: prefix_def)
  show "r¯>w  w"
    using r ≤p w r  ε unfolding prefix_def  by fastforce
  show "r¯>w  ε"
    using lq_pref[OF r ≤p w] |r| < |w| by force
qed

lemma pref_suf_neq_per: assumes "x ≤p w" and "x ≤s w" and "x  w" shows "period w (|w|-|x|)"
proof-
  have "(w<¯x)x = w"
    using  rq_suf[OF x ≤s w].
  have "x(x¯>w) = w"
    using  lq_pref[OF x ≤p w].
  have take: "w<¯x = take (|w|-|x|) w"
    using rq_take.
  have nemp: "take (|w|-|x|) w  ε"
    using x ≤p w x  w unfolding prefix_def by auto
  have "w ≤p take (|w|-|x|) w  w"
    using triv_pref[of w "x¯>w"]
    unfolding lassoc[of "w<¯x" x "x¯>w", unfolded x  x¯>w = w w<¯x  x = w, symmetric] take.
  thus "period w (|w|-|x|)"
    unfolding period_def   using nemp by blast
qed

lemma border_per: "x ≤b w  period w (|w|-|x|)"
  unfolding  border_def using pref_suf_neq_per by blast

lemma per_border: assumes "n < |w|" and "period w n"
  shows "take (|w| - n) w  ≤b w"
proof-
  have eq: "take (|w| - n) w = drop n w"
    using pref_take[OF period w n[unfolded
   per_shift[OF per_nemp[OF period w n] per_not_zero[OF period w n]]], unfolded length_drop].
  have "take (|w| - n) w  ε"
    using n < |w| take_eq_Nil by fastforce
  moreover have "take (|w| - n) w  w"
    using  per_not_zero[OF period w n] n < |w| unfolding take_all_iff[of "|w|-n" w] by fastforce
  ultimately show ?thesis
    unfolding border_def using take_is_prefix[of "|w|-n" w] suffix_drop[of n w, folded eq] by blast
qed

section ‹The longest border and the shortest period›

subsection ‹The longest border›

definition max_borderP :: "'a list  'a list  bool" where
  "max_borderP u w = (u ≤p w  u ≤s w  (u = w  w = ε)  ( v. v ≤b w   v ≤p u))"

lemma max_borderP_emp_emp: "max_borderP ε ε"
  unfolding max_borderP_def by simp

lemma max_borderP_exE: obtains u where "max_borderP u w"
proof-
  define P where "P = (λ x. x ≤p w  x ≤s w  (x = w  w = ε))"
  have "P ε"
    unfolding P_def by blast
  obtain v where "v ≤p w" and "P v" and "(y. y ≤p w  P y  y ≤p v)"
    using max_pref[of ε w P thesis, OF prefix_bot.extremum P ε] by blast
  hence "max_borderP v w"
    unfolding  max_borderP_def border_def P_def by presburger
  from that[OF this]
  show thesis.
qed

lemma max_borderP_of_nemp: "max_borderP u ε  u = ε"
  by (metis max_borderP_def suffix_bot.extremum_unique)

lemma max_borderP_D_neq: "w  ε  max_borderP u w  u  w"
  by (simp add: max_borderP_def)

lemma max_borderP_D_pref: "max_borderP u w  u ≤p w"
  by (simp add: max_borderP_def)

lemma max_borderP_D_suf: "max_borderP u w  u ≤s w"
  by (simp add: max_borderP_def)

lemma max_borderP_D_max: "max_borderP u w  v ≤b w   v ≤p u"
  by (simp add: max_borderP_def)

lemma  max_borderP_D_max': "max_borderP u w  v ≤b w  v ≤s u"
  unfolding max_borderP_def using borderD_suf  suf_pref_eq suffix_same_cases by metis

lemma unbordered_max_border_emp:  "¬ bordered w  max_borderP u w  u = ε"
  unfolding max_borderP_def bordered_def border_def by blast

lemma bordered_max_border_nemp:  "bordered w  max_borderP u w  u  ε"
  unfolding max_borderP_def bordered_def border_def using prefix_Nil by blast

lemma max_borderP_border: "max_borderP u w  u  ε  u ≤b w"
  unfolding max_borderP_def border_def by blast

lemma max_borderP_rev: "max_borderP (rev u) (rev w)  max_borderP u w"
proof-
  assume "max_borderP (rev u) (rev w)"
  from this[unfolded max_borderP_def rev_is_rev_conv, folded pref_rev_suf_iff suf_rev_pref_iff]
  have "u = w  w = ε" and "u ≤p w" and "u ≤s w" and allv: "v ≤b rev w  v ≤p rev u" for v
    by blast+
  show "max_borderP u w"
  proof (unfold max_borderP_def, intro conjI, simp_all only: u ≤p w u ≤s w)
    show "u = w  w = ε" by fact
    show "v. v ≤b w  v ≤p u"
  proof (rule allI, rule impI)
      fix v assume "v ≤b w"
      show "v ≤p u"
        using max_borderP (rev u) (rev w) v ≤b w border_rev_conv max_borderP_D_max' pref_rev_suf_iff by metis
    qed
  qed
qed

lemma max_borderP_rev_conv: "max_borderP (rev u) (rev w)  max_borderP u w"
  using max_borderP_rev max_borderP_rev[of "rev u" "rev w", unfolded rev_rev_ident] by blast

(* TODO zkusit use argmax?
  SH: nasledujici jednoduche dukazy ukazuji, ze se tim asi nic neziska *)
term arg_max
definition max_border :: "'a list  'a list" where
  "max_border w = (THE u. (max_borderP u w))"

lemma max_border_unique: assumes "max_borderP u w" "max_borderP v w"
  shows "u = v"
  using max_borderP_D_max[OF max_borderP u w, OF max_borderP_border[OF max_borderP v w]]
        max_borderP_D_max[OF max_borderP v w, OF max_borderP_border[OF max_borderP u w]]
  by force

lemma max_border_ex: "max_borderP (max_border w) w"
proof (rule max_borderP_exE[of w])
  fix u assume "max_borderP u w"
  with max_border_unique[OF this]
  show ?thesis
    unfolding max_border_def
    by (elim theI[of "λ x. max_borderP x w"]) simp
qed

lemma max_borderP_max_border: "max_borderP u w  max_border w = u"
  using max_border_unique[OF max_border_ex].

lemma max_border_len_rev: "|max_border u| =  |max_border (rev u)|"
  by (cases "u = ε", simp, metis length_rev max_borderP_max_border max_borderP_rev_conv max_border_ex)

lemma max_border_border: assumes "bordered w" shows "max_border w ≤b w"
  using max_border_ex bordered_max_border_nemp[OF assms, of "max_border w"]
  unfolding max_borderP_def border_def by blast

theorem max_border_border':  "max_border w  ε  max_border w ≤b w"
  using max_borderP_border max_border_ex  by blast

lemma max_border_sing_emp: "max_border [a] = ε"
  using max_border_ex[THEN unbordered_max_border_emp[OF sing_not_bordered]] by fast

lemma max_border_suf: "max_border w ≤s w"
  using max_borderP_D_suf max_border_ex by auto

lemma max_border_nemp_neq: "w  ε  max_border w  w"
  by (simp add: max_borderP_D_neq max_border_ex)

lemma max_borderI: assumes "u  w" and "u ≤p w" and "u ≤s w" and " v. v ≤b w  v ≤p u"
  shows "max_border w = u"
  using assms max_border_ex
  by (intro max_borderP_max_border, unfold max_borderP_def border_def, blast)

lemma max_border_less_len: assumes "w  ε" shows "|max_border w| < |w|"
  using assms border_len(3) leI list.size(3) max_border_border' npos_len by metis

theorem max_border_max_pref: assumes "u ≤b w" shows "u ≤p max_border w"
  using  max_borderP_D_max[OF max_border_ex u ≤b w].

theorem max_border_max_suf: assumes "u ≤b w" shows "u ≤s max_border w"
  using  max_borderP_D_max'[OF max_border_ex u ≤b w].

lemma bordered_max_bord_nemp_conv[code]: "bordered w  max_border w  ε"
  using bordered_max_border_nemp max_border_ex unbordered_max_border_emp by blast

lemma max_bord_take: "max_border w = take |max_border w| w"
proof (cases "bordered w")
  assume "bordered w"
  from borderD_pref[OF max_border_border[OF this]]
  show "max_border w = take |max_border w| w"
    by (simp add: pref_take)
next
  assume "¬ bordered w"
  hence "max_border w = ε"
    using bordered_max_bord_nemp_conv by blast
  thus "max_border w = take |max_border w| w"
    by simp
qed


subsection ‹The shortest period›

(* TODO define min_period first, then use it here
  SH: prazdne slovo bude mit nedefinovanou periodu
*)
definition min_period_root :: "'a list  'a list" ("π") where
  "min_period_root w = take (LEAST n. period w n) w"

definition min_period :: "'a list  nat" where
  "min_period w = |π w|"

lemma min_per_emp[simp]: "π ε = ε"
  unfolding min_period_root_def by simp

lemma min_per_zero[simp]: "min_period ε = 0"
  by (simp add: min_period_def)

lemma min_per_per: "w  ε  period w (min_period w)"
  unfolding min_period_def min_period_root_def
  using len_is_per LeastI_ex period_def periodI by metis

lemma min_per_pos: "w  ε  0 < min_period w"
  using min_per_per by auto

lemma min_per_len:  "min_period w  |w|"
  unfolding min_period_def min_period_root_def using len_is_per Least_le by simp

lemmas min_per_root_len = min_per_len[unfolded min_period_def]

lemma min_per_sing: "min_period [a] = 1"
  using min_per_pos[of "[a]"] min_per_len[of "[a]"] by simp

lemma min_per_root_per_root: assumes "w  ε" shows "w <p (π w)  w"
  using LeastI_ex assms len_is_per period_def unfolding min_period_root_def by metis

lemma min_per_pref: "π w ≤p w"
  unfolding  min_period_root_def using take_is_prefix by blast

lemma min_per_nemp: "w  ε  π w  ε"
  using min_per_root_per_root by blast

lemma min_per_min: assumes "w <p r  w" shows "π w ≤p r"
proof (cases "w = ε")
  assume "w  ε"
  have "period w |π w|"
    using w  ε min_per_root_per_root periodI by blast
  have "period w |r|"
    using w  ε assms periodI by blast
  from Least_le[of "λ n. period w n", OF this]
  have "|π w|  |r|"
    unfolding  min_period_root_def using dual_order.trans len_take1 by metis
  with pref_trans[OF  min_per_pref sprefD1[OF w <p r  w]]
  show "π w ≤p r"
    using pref_prod_le by blast
qed simp

lemma lq_min_per_pref:  "π w¯>w ≤p w"
  unfolding same_prefix_prefix[of "π w" _ w, symmetric]  lq_pref[OF min_per_pref] using sprefD1[OF min_per_root_per_root]
  by (cases "w = ε", simp)

lemma max_bord_emp: "max_border ε = ε"
  by (simp add: max_borderP_of_nemp max_border_ex)

theorem min_per_max_border: "π w  max_border w = w"
proof (cases "w = ε")
  assume "w  ε"
  have "max_border w = (π w)¯>w"
  proof (intro max_borderI)
    show "π w¯>w  w"
      using  min_per_nemp[OF w  ε]  lq_pref[OF min_per_pref]  append_self_conv2 by metis
    show "π w¯>w ≤s w"
      using lq_suf_suf[OF min_per_pref].
    show "π w¯>w ≤p w"
      using lq_min_per_pref by blast
    show "v. v ≤b w  v ≤p π w¯>w"
    proof (rule allI, rule impI)
      fix v assume "v ≤b w"
      have "w <p (w<¯v)  w"
        using per_border v ≤b w border_per_root[OF v ≤b w] border_rq_nemp[OF v ≤b w]   by blast
      from min_per_min[OF this]
      have "π w ≤p w<¯v".
      from pref_rq_suf_lq[OF borderD_suf[OF v ≤b w] this]
      have "v ≤s π w¯>w".
      from  suf_pref_eq[OF this] ruler[OF borderD_pref[OF v ≤b w] π w¯>w ≤p w]
      show "v ≤p π w¯>w"
        by blast
    qed
  qed
  thus ?thesis
    using lq_pref[OF min_per_pref, of w] by simp
qed (simp add:  max_bord_emp)

lemma min_per_len_diff: "min_period w = |w| - |max_border w|"
  unfolding min_period_def  using lenarg[OF min_per_max_border,unfolded lenmorph,of w] by linarith

lemma min_per_root_take [code]: "π w = take (|w| - |max_border w|) w"
  using cancel_right max_border_suf min_per_max_border suffix_take by metis

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"
  where  "primitive u = ( r k. r@k = u  k = 1)"

lemma emp_not_prim[simp]: "¬ primitive ε"
  unfolding primitive_def
  by (metis pow_eq_if_list zero_neq_one)

lemma primI[intro]: "( r k. r@k = u  k = 1)  primitive u"
  by (simp add: primitive_def)

lemma prim_nemp: "primitive u  u  ε"
  by force

lemma prim_exp_one: "primitive u  r@k = u  k = 1"
  using primitive_def by blast

lemma pow_nemp_imprim[intro]: "2  k   ¬ primitive (u@k)"
  using prim_exp_one by fastforce

lemma pow_not_prim: "¬ primitive (u@Suc(Suc k))"
  using prim_exp_one by fastforce

lemma pow_non_prim: "k  1  ¬ primitive (w@k)"
  using prim_exp_one
  by auto

lemma prim_exp_eq: "primitive u  r@k = u  u = r"
  using prim_exp_one pow_1 by blast

lemma prim_per_div: assumes "primitive v" and "n  0" and "n  |v|" and "period v (gcd |v| n)"
  shows "n = |v|"
proof-
  have "gcd |v| n dvd |v|"
    by simp
  from  prim_exp_eq[OF primitive v per_div[OF this period v (gcd |v| n)]]
  have "gcd |v| n = |v|"
    using take_len[OF le_trans[OF gcd_le2_nat[OF n  0] n  |v|], of "|v|"]  by presburger
  from gcd_le2_nat[OF n  0, of "|v|", unfolded this] n  |v|
  show "n = |v|" by force
qed

lemma prim_triv_root: "primitive u  u  t*  t = u"
  using prim_exp_eq unfolding root_def
  unfolding primitive_def root_def by fastforce

lemma prim_comm_root[elim]: assumes "primitive r" and "u  r = r  u" shows "u  r*"
  using ur = ru[unfolded comm] prim_exp_eq[OF primitive r] rootI by metis

lemma prim_comm_exp[elim]: assumes "primitive r" and "ur = ru" obtains k where "r@k = u"
  using rootE[OF prim_comm_root[OF assms]].

lemma pow_prim_root: assumes "w@k = r@n" and "0 < n" "primitive r"
  shows "w  r*"
  using pow_comm_comm[OF w@k = r@n[symmetric] 0 < n] prim_comm_root[OF primitive r]
    by presburger

lemma prim_root_drop_exp[elim]: assumes "u@k  r*" and "0 < k" and  "primitive r"
  shows "u  r*"
  using pow_comm_comm[of u k r, OF _ 0 < k, THEN  prim_comm_root[OF primitive r]]
    u@k  r*[unfolded root_def] unfolding root_def by metis

lemma prim_card_set: assumes "primitive u" and "|u|  1" shows "1 < card (set u)"
  using |u|  1 primitive u pow_non_prim[OF |u|  1, of "[hd u]"]
  by (elim not_le_imp_less[OF contrapos_nn] card_set_le_1_imp_hd_pow[elim_format]) simp

lemma comm_not_prim: assumes "u  ε" "v  ε" "u  v = v  u" shows "¬ primitive (u  v)"
proof-
  obtain t k m where "u = t@k"  "v = t@m"
    using uv = vu[unfolded comm] by blast
  show ?thesis using pow_non_prim[of "k+m" "t"]
    unfolding u = t@k v = t@m add_exps[of t k m]
    using nemp_pow[OF u  ε[unfolded u = t@k]] nemp_pow[OF v  ε[unfolded v = t@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@k = rotate n w"
    obtain l where "(rotate l r)@k = w"
      using rotate_backE[of n w, folded r@k = rotate n w, unfolded rotate_pow_comm] by blast
    from prim_exp_one[OF primitive w this]
    show "k = 1".
  qed
next
  assume "primitive (rotate n w)"  show "primitive w"
  proof (rule primI)
    fix r k assume "r@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@k = w" and "w  r"
proof-
  from ¬ primitive w[unfolded primitive_def]
  obtain r k where "k  1" and "r@k = w"  by blast
  have "r  ε"
    using w  ε r@k = w emp_pow by blast
  have "k  0"
    using w  ε r@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 @ k = w]] by simp
  show thesis
    using that[OF r  ε _ r@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 < |w|"
  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"]
    unfolding take_len[OF less_imp_le_nat[OF n < |w|]] append_take_drop_id rotate n w = w.
  have "take n w  ε" "drop n w  ε"
    using 0 < n n < |w| by auto+
  from primitive w show False
    using comm_not_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 < |w|  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@l = w" and "w  r" by blast
  have "rotate |r| w = w"
    using rotate_root_self[of r l, unfolded r@l = w].
  moreover have "0 < |r|"
    by (simp add: r  ε)
  moreover have "|r| < |w|"
    unfolding pow_len[of r l, unfolded r@l = w]  using  1 < l 0 < |r|  by auto
  ultimately show False
    using assms(2) by blast
qed

corollary prim_iff_rotate: assumes "w  ε" shows
  "primitive w  ( n. 0 < n  n < |w|  rotate n w  w)"
  using no_rotate_prim[OF w  ε] prim_no_rotate by blast

lemma prim_sing: "primitive [a]"
  using prim_iff_rotate[of "[a]"] by fastforce

lemma sing_pow_conv [simp]: "[u] = t@k  t = [u]  k = 1"
  using pow_non_prim pow_1 prim_sing by metis

lemma prim_rev_iff[reversal_rule]: "primitive (rev u)  primitive u"
  unfolding primitive_def[reversed] using primitive_def..

lemma prim_map_prim: "primitive (map f ws)  primitive ws"
  unfolding primitive_def using map_pow  by metis

lemma inj_map_prim: assumes "inj_on f A" and "u  lists A" and
  "primitive u"
shows "primitive (map f u)"
  using prim_map_prim[of "the_inv_into A f" "map f u", folded inj_map_inv[OF assms(1-2)], OF assms(3)].

lemma prim_map_iff [reversal_rule]:
  assumes "inj f" shows "primitive (map f ws) = primitive (ws)"
  using inj_map_prim[of _ UNIV, unfolded lists_UNIV, OF inj f UNIV_I]
    prim_map_prim by (intro iffI)

lemma prim_concat_prim: "primitive (concat ws)  primitive ws"
  unfolding primitive_def using concat_pow by metis

lemma eq_append_not_prim: "x = y  ¬ primitive (x  y)"
  by (metis append_Nil2 comm_not_prim prim_nemp)

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_root :: "'a list  'a list" ("ρ") where
  "primitive_root x = (if x  ε then (THE r. primitive r  ( k. x = r@k)) else ε)"

definition primitive_root_exp :: "'a list  nat" ("eρ") where
 "primitive_root_exp x = (if x  ε then (THE k. x = (ρ x)@k) else 0)"



lemma primroot_emp[simp]: "ρ ε = ε"
  unfolding primitive_root_def by simp

lemma comm_prim: assumes "primitive r" and  "primitive s" and "rs = sr"
  shows "r = s"
  using rs = sr[unfolded comm] assms[unfolded primitive_def, rule_format] by metis

lemma primroot_ex: assumes "x  ε" shows " r k. primitive r  k  0  x = r@k"
  using x  ε
proof(induction "|x|" arbitrary: x rule: less_induct)
  case less
  then show " r k.  primitive r  k  0  x = r@k"
  proof (cases "primitive x")
    assume "¬ primitive x"
    from non_prim[OF this x  ε]
    obtain r l where "r  ε" and "1 < l" and "r@l = x" and "x  r" by blast
    from less.hyps[OF root_shorter[OF x  ε rootI[of r l, unfolded r@l = x] x  r]  r  ε]
    obtain k pr where "primitive pr" "k  0" "r = pr@k"
      by blast
    have "k*l  0"
      using 1 < l k  0 by force
    have "x = pr@(k*l)"
      using pow_mult[of pr k l, folded r = pr@k, unfolded r@l = x, symmetric].
    thus "r k. primitive r  k  0  x = r @ k"
      using primitive pr k*l  0 by fast
  next
    assume "primitive x"
    have "x = x@Suc 0"
      by simp
    thus " r k.  primitive r  k  0  x = r@k"
      using primitive x by force
  qed
qed

lemma primroot_exE: assumes"x  ε"
  obtains r k where "primitive r" and "0 < k" and "x = r@k"
  using assms  primroot_ex[OF x  ε] by blast

text‹Uniqueness of the primitive root follows from the following lemma›

lemma primroot_unique: assumes "u  ε" and "primitive r" and "u = r@k" shows "ρ u = r"
proof-
  have "0 < k"
    using u  ε u = r@k by blast
  have "s = r" if "primitive s" and "u = s@l" for s l
  proof-
    from pow_comm_comm[OF u = s@l[unfolded u = r@k] 0 < k]
    obtain t where "s  t*" and "r  t*"
      using comm_root 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
  hence "primitive s  (k. u = s @ k)  s = r" for s
    by presburger
  from the_equality[of "λ r. primitive r  (k. u = r @ k)" r, OF _ this]
  show "ρ u = r"
    using primitive r u = r@k  unfolding primitive_root_def if_P[OF u  ε] by blast
qed

lemma primroot_unique': assumes "0 < k" "primitive r" and "u = r@k" shows "ρ u = r"
  using  primroot_unique[OF _ assms(2,3)] using prim_nemp[OF primitive r] 0 < k unfolding u = r@k
  using nonzero_pow_emp by blast

lemma prim_self_root[intro]: "primitive x  ρ x  = x"
  using emp_not_prim primroot_unique pow_1 by metis

lemma primroot_exp_unique: assumes "u  ε" and "(ρ u)@k = u" shows "eρ u = k"
  unfolding primitive_root_exp_def if_P[OF u  ε]
proof (rule the_equality)
  show "u = (ρ u)@k" using (ρ u)@k = u[symmetric].
  have "ρ u  ε"
    using assms by force
  show "ka = k" if "u = ρ u @ ka" for ka
    using eq_pow_exp[OF ρ u  ε, of k ka, folded u = (ρ u)@k that] by blast
qed

lemma primroot_prim[intro]:  "x  ε  primitive (ρ x)"
  using primroot_unique primroot_ex by metis

text‹Existence and uniqueness of the primitive root justifies the function @{term primitive_root}: it indeed yields the primitive root of a nonempty word.›


lemma primroot_is_root[simp]: "x  (ρ x)*"
  by (cases "x = ε", force, unfold root_def) (use primroot_exE primroot_unique in metis)


lemma primroot_expE: obtains k where "(ρ x)@k = x" and "0 < k"
proof (cases "x = ε")
  assume "x  ε"
  with primroot_is_root[unfolded root_def] that
  show thesis by fastforce
qed auto

lemma primroot_exp_eq [simp]: "(ρ u)@(eρ u) = u"
  using primroot_expE[of u "ρ u @ eρ u = u"] primroot_exp_unique pow_0 primitive_root_exp_def by metis

lemma primroot_exp_len:
  shows "eρ w * |ρ w| = |w|"
  using lenarg[OF primroot_exp_eq] unfolding pow_len.

lemma primroot_exp_nemp [intro]: "u  ε  0 < eρ u"
  using  primroot_exp_eq nemp_pow by metis



lemma primroot_nemp[intro!]: "x  ε  ρ x  ε"
  using prim_nemp by blast

lemma primroot_idemp[simp]: "ρ (ρ x) = ρ x"
 by (cases "x = ε")  (simp only: primroot_emp, use prim_self_root in blast)

lemma prim_primroot_conv: assumes "w  ε" shows "primitive w  ρ w = w"
  using assms prim_self_root primroot_prim[OF w  ε] by metis

lemma not_prim_primroot_expE: assumes "¬ primitive w"
  obtains k where "ρ w @k = w" and "2  k"
  using primroot_exp_eq primroot_prim assms
proof (cases "w = ε")
  assume "w  ε"
  with primroot_exp_eq[of w]
  have "eρ w  1" "eρ w  0"
    using pow_zero pow_1 primroot_prim[OF w  ε] ¬ primitive w by force+
  with that[OF ρ w @ eρ w = w]
  show thesis by force
qed force


lemma not_prim_expE: assumes "¬ primitive x" and "x  ε"
  obtains r k where "primitive r" and "2  k" and "r@k = x"
  using not_prim_primroot_expE[OF ¬ primitive x] primroot_prim[OF x  ε] by metis

lemma primroot_of_root: assumes "u  ε" and "u  q*" shows "ρ q = ρ u"
proof-
  have "q  ε"
    using assms by force
  from primroot_unique[OF u  ε primroot_prim[OF this], symmetric]
  root_trans[OF u  q* primroot_is_root[of q]]
  show ?thesis
    unfolding root_def by blast
qed



lemma primroot_shorter_root: assumes "u  ε" and "u  q*" shows "|ρ u|  |q|"
  unfolding primroot_of_root[OF assms, symmetric]
  using root_nemp[OF assms] root_shorter_eq[of q, OF _ primroot_is_root] by blast


lemma primroot_len_le: "u  ε  |ρ u|  |u|"
  using primroot_expE primroot_shorter_root[OF _ self_root] by auto

lemma primroot_take: assumes "u  ε" shows "ρ u = (take ( |ρ u| ) u)"
proof-
  obtain k where "(ρ u)@k = u" and "0 < k"
    using primroot_expE by blast
  show "ρ u = (take ( |ρ u| ) u)"
    using take_root[of _ "(ρ u)", OF 0 < k, unfolded (ρ u)@k = u].
qed


lemma primroot_rotate_comm: assumes "w  ε" shows "ρ (rotate n w) = rotate n (ρ w)"
proof-
  obtain l where  "(ρ w)@l = w"
    using primroot_expE.
  hence "rotate n w  (rotate n (ρ w))*"
    using rotate_pow_comm root_def by metis
  have "rotate n w  ε"
    using assms by auto
  have "primitive (rotate n (ρ w))"
    using assms prim_rotate_conv by blast
  show  ?thesis
    using primroot_unique[OF rotate n w  ε primitive (rotate n (ρ w))]
    rootE[OF rotate n w  (rotate n (ρ w))*] by metis
qed

lemma primroot_rotate: "ρ w = r  ρ (rotate (k*|r|) w) = r" (is "?L  ?R")
proof(cases "w = ε")
  case False
  show ?thesis
    unfolding primroot_rotate_comm[OF w  ε, of "k*|r|"]
    using length_rotate[of "k*|r|" "ρ w"] mod_mult_self2_is_0[of k "|r|"]
      rotate_id[of "k*|r|" "ρ w"]
    by metis
qed (simp add: rotate_is_Nil_conv[of "k*|r|" w])

lemma primrootI[intro]: assumes pow: "u = r@(Suc k)" and "primitive r" shows "ρ u = r"
proof-
  have "u  ε"
    using pow primitive r prim_nemp by auto
  show "ρ u = r"
    using primroot_unique[OF u  ε primitive r u = r@(Suc k)].
qed

lemma primroot_pref: "ρ u ≤p u"
  by (cases "u = ε", use primroot_emp in blast)
     (simp add: per_root_pref[OF _ primroot_is_root])

lemma short_primroot: assumes "u  ε" "¬ primitive u" shows "|ρ u| < |u|"
  using primroot_prim[OF u  ε]  le_neq_implies_less pref_len primroot_pref
        long_pref assms by metis

lemma prim_primroot_cases: obtains "u = ε" | "primitive u" | "|ρ u| < |u|"
  using short_primroot by blast

text‹We also have the standard characterization of commutation for nonempty words.›

lemma comm_rootE: assumes  "x  y = y  x"
  obtains  t where "x   t*" and "y  t*" and "t  ε"
  using assms[unfolded comm_root]
  using emp_all_roots list.discI root_nemp by metis

theorem comm_primroots: assumes "u  ε" and "v  ε" shows "u  v = v  u  ρ u = ρ v"
proof
  assume "u  v = v  u"
  from comm_rootE[OF this]
  obtain t where "u  t*" and "v  t*".
  show "ρ u = ρ v"
    using primroot_of_root[OF v  ε v  t*, unfolded primroot_of_root[OF u  ε u  t*]].
next
  assume "ρ u = ρ v"
  from pows_comm[of "ρ u" "eρ u" "eρ v"]
  show "u  v = v  u"
    unfolding primroot_exp_eq unfolding ρ u = ρ v primroot_exp_eq.
qed

lemma comm_primroots': "u  ε  v  ε  u  v = v  u  ρ u = ρ v"
  by (simp add: comm_primroots)

lemma same_primroots_comm:  "ρ x = ρ y  x  y = y  x"
  using comm_primroots by blast

lemma pow_primroot: assumes "x  ε" shows "ρ (x@Suc k) = ρ x"
  using  comm_primroots'[OF nemp_Suc_pow_nemp, OF assms assms, of k, folded pow_Suc' pow_Suc] by blast

lemma comm_primroot_exp: assumes "v  ε" and "u  v = v  u"
  obtains n where "(ρ v)@n = u"
proof(cases)
  assume "u = ε" thus thesis using that pow_0 by blast
next
  assume "u  ε" thus thesis using that[OF primroot_expE] u  v = v  u[unfolded comm_primroots[OF u  ε v  ε]] by metis
qed

lemma comm_primrootE: assumes  "x  y = y  x"
  obtains t where "x  t*" and "y  t*" and "primitive t"
  using comm_primroots assms emp_all_roots prim_sing primroot_is_root primroot_prim by metis

lemma primE: obtains t where "primitive t"
  using comm_primrootE by metis

lemma comm_primrootE':  assumes  "x  y = y  x"
  obtains  t m k where "x = t@k" and "y = t@m" and "primitive t"
  using comm_primrootE[OF x  y = y  x, unfolded root_def] by metis

lemma comm_nemp_pows_posE: assumes  "x  y = y  x" and "x  ε" and "y  ε"
  obtains t k m where "x = t@k" and "y = t@m" and "0 < k" and "0 < m" and "primitive t"
proof-
  from comm_primrootE[OF x  y = y  x, unfolded root_def]
  obtain t k m where "t@k = x" "t@m = y" "primitive t"
    by metis
  note nemp_exp_pos[OF x  ε t@k = x] nemp_exp_pos[OF y  ε t@m = y]
  show thesis
    using that[OF t@k = x[symmetric] t@m = y[symmetric] 0 < k 0 < m primitive t].
qed

lemma comm_primroot_conv:  "u  v = v  u  u  ρ v = ρ v  u"
proof (cases "u = ε  v = ε")
  assume "¬ (u = ε  v = ε)"
  hence "u  ε" "v  ε"
    by blast+
  show ?thesis
    using comm_primroots[OF u  ε v  ε, folded
   comm_primroots[OF u  ε primroot_nemp[OF v  ε], unfolded primroot_idemp]].
qed force

lemma comm_primroot [simp, intro]: "u  ρ u = ρ u  u"
  using comm_primroot_conv by blast

lemma comp_primroot_conv': shows "u  v = v  u  ρ u  ρ v = ρ v  ρ u"
  using comm_primroot_conv[of u v] comm_primroot_conv[of "ρ v" u]
  unfolding eq_sym_conv[of "ρ v  u"] eq_sym_conv[of "ρ v  ρ u"] by blast

lemma per_root_primroot: "w <p r  w   w <p ρ r  w"
  using per_root_trans[OF _ primroot_is_root].

lemma primroot_per_root: "r  ε   r <p ρ r  r"
  by blast

lemma prim_comm_short_emp: assumes "primitive p" and "up=pu" and "|u| < |p|"
  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 |u| < |p|[folded this]
  show False
    using primroot_len_le[OF u  ε] by auto
qed

lemma primroot_rev[reversal_rule]: shows "ρ (rev u) = rev (ρ u)"
proof (cases "u = ε")
  assume "u  ε"
  hence "rev u  ε"
    by simp
  have "primitive (rev (ρ u))"
    using primroot_prim[OF u  ε] unfolding prim_rev_iff.
  have "rev u = (rev (ρ u))@eρ u"
    unfolding rev_pow[symmetric] primroot_exp_eq..
  from primroot_unique[OF rev u  ε primitive (rev (ρ u)) this]
  show ?thesis.
qed simp

lemmas primroot_suf = primroot_pref[reversed]

lemma per_le_prim_iff:
  assumes "u ≤p p  u" and "p  ε" and "2 * |p|  |u|"
  shows "primitive u  u  p  p  u"
proof
  have "|p| < |u|" using  2 * |p|  |u|
      nemp_len[OF p  ε] by linarith
  with p  ε
  show "primitive u  u  p  p  u"
    by (intro notI, elim notE) (rule prim_comm_short_emp[OF _ sym])
  show "u  p  p  u  primitive u"
  proof (elim swap[of "_ = _"], elim not_prim_primroot_expE)
    fix k z assume "2  k" and eq: "z @ k = u"
    from this(1) lenarg[OF this(2)] 2 * |p|  |u|
    have "|z| + |p|  |u|"
      by (elim at_least2_Suc) (simp only: pow_Suc lenmorph[of z])
    with u ≤p p  u have "z  p = p  z"
      by (rule two_pers[rotated 1]) (simp flip: eq pow_comm)
    from comm_add_exp[OF this, of k]
    show "u  p = p  u" unfolding eq.
  qed
qed

lemma per_root_mod_primE [elim]: assumes "u <p r  u"
  obtains n p s where "p  s = ρ r" and "(ps)@n  p = u" and "s  ε"
  using per_root_modE[OF per_root_primroot[OF assms]] primroot_prim[OF per_root_nemp[OF assms]]
   emp_not_prim by metis

subsection ‹Primitivity and the shortest period›

lemma min_per_primitive: assumes "w  ε" shows "primitive (π w)"
proof-
  have "ρ(π w)  ε"
    using assms min_per_nemp primroot_nemp by blast
  obtain k where "π w = (ρ (π w))@k"
    using  primroot_expE by metis
  from rootI[of "ρ (π w)" k, folded this]
  have "w <p (ρ (π w))  w"
    using  min_per_root_per_root[OF assms, THEN per_root_trans]  by presburger
  from pow_pref_root_one[OF _ ρ(π w)  ε, of k, folded  π w = (ρ (π w))@k, OF _ min_per_min[OF this]]
  have "k = 1"
    using  π w = (ρ (π w))@k min_per_nemp[OF w  ε] pow_zero[of "ρ (π w)"] by fastforce
  show "primitive (π w)"
    using primroot_prim[OF ρ (π w)  ε, folded π w = (ρ (π w))@k[unfolded k = 1 One_nat_def pow_one]].
qed

lemma min_per_short_primroot: assumes "w  ε" and "(ρ w)@k = w" and "k  1"
  shows "π w = ρ w"
proof-
  have "k  0"
    using assms pow_zero by blast
  with k  1  have "2  k"
    by fastforce
           have "w <p (ρ w)  w"
    using assms(1) assms(2) per_root_drop_exp root_self by metis
  have "w <p (π w)  w"
    using assms(1) min_per_root_per_root by blast
  have "π w ≤p ρ w"
    using min_per_min[OF w <p (ρ w)  w].
  from  prefix_length_le[OF this]
  have "|π w| + |ρ w|  |w|"
    unfolding lenarg[OF (ρ w)@k =w, unfolded pow_len, symmetric] using
     mult_le_mono1[OF 2  k, of "|ρ w|"] unfolding one_add_one[symmetric] distrib_right mult_1
    by simp
  from two_pers_root[OF w <p (π w)  w w <p (ρ w)  w this]
  have "π w  ρ w = ρ w  π w".
  from this[unfolded comm_primroots[OF per_root_nemp[OF w <p (π w)  w] per_root_nemp[OF w <p (ρ w)  w]]]
  show "π w = ρ w"
    unfolding prim_self_root[of "ρ w", OF primroot_prim[OF w  ε]]
      prim_self_root[of "π w", OF min_per_primitive[OF w  ε]].
qed


lemma primitive_iff_per: "primitive w  w  ε  (π w = w  π w  w  w  π w)"
proof
  assume "primitive w"
  hence "w  ε" by fastforce
  show "w  ε  (π w = w  π w  w  w  π w)"
  proof (rule conjI)
    show "π w = w  π w  w  w  π w"
      using comm_prim [OF min_per_primitive[OF w  ε] primitive w]
      by (intro verit_or_neg(1))
  qed fact
next
  assume asm: "w  ε  (π w = w  π w  w  w  π w)"
  have "w  ε" and imp: "π w  w = w  π w  π w = w"
    using asm by blast+
  obtain k where "(ρ w)@k = w" "0 < k"
    using primroot_expE.
  show "primitive w"
  proof-
    from imp[unfolded min_per_short_primroot[OF w  ε (ρ w)@k = w]]
    have "ρ w = w"
      using pow_comm[symmetric, of "ρ w" k, unfolded ρ w @k = w]
       ρ w @ k = w min_per_short_primroot[OF w  ε ρ w@k = w] pow_1 w  ε by metis
    thus "primitive w"
      using prim_primroot_conv[OF w  ε] by simp
  qed
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 (infix "" 51)
  where "u  v  r s. r  s = u  s  r = v"

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 conjugE_nemp[elim]:
  assumes "u  v" and "u  ε"
  obtains r s where "r  s = u" and "s  r = v" and "s  ε"
  using assms unfolding conjugate_def
proof (cases "u = v")
  assume "u  v"
  obtain r s where "r  s = u" and "s  r = v" using conjugE[OF u  v].
  hence "s  ε" using u  v by force
  thus thesis using that[OF r  s = u s  r = v] by blast
qed (simp add: that[OF _ _ u  ε])

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_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 rotate_conjug: "w  rotate n w"
  using conjug_rotate_iff by blast

lemma conjug_rotate_iff_le:
  shows "u  v  ( n  |u| - 1. v = rotate n u)"
proof
  show "n  |u| - 1 . v = rotate n u  u  v"
    using conjug_rotate_iff by blast
next
  assume "u  v"
  thus " n  |u| - 1. v = rotate n u"
  proof (cases "u = ε")
    assume "u  ε"
    obtain r s where "r  s = u" and  "s  r = v" and "s  ε"
      using conjugE_nemp[OF u  v u  ε].
    hence "v = rotate |r| u"
      using rotate_append[of r s] by argo
    moreover have "|r|  |u| - 1"
      using lenarg[OF r  s = u, unfolded lenmorph] nemp_len[OF s  ε] by linarith
    ultimately  show "n  |u| - 1. v = rotate n u"
      by blast
  qed auto
qed

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 conjug_refl: "u  u"
  by standard+

lemma conjug_sym[sym]: "u  v  v  u"
  by (elim conjugE, intro conjugI) assumption

lemma conjug_swap: "u  v  v  u"
  by blast

lemma conjug_nemp_iff: "u  v  u = ε  v = ε"
  by (elim conjugE1, intro iffI) simp+

lemma conjug_len: "u  v   |u| = |v|"
  by (elim conjugE, hypsubst, rule swap_len)

lemma pow_conjug:
  assumes eq: "t@i  r  u = t@k" and t: "r  s = t"
  shows "u  t@i  r = (s  r)@k"
proof -
  have "t@i  r  u  t@i  r = t@i  t@k  r" unfolding eq[unfolded lassoc] lassoc append_same_eq pows_comm..
  also have "  = t@i  r  (s  r)@k" unfolding conjug_pow[OF rassoc, symmetric] t..
  finally show "u  t@i  r = (s  r)@k" unfolding same_append_eq.
qed

lemma conjug_set: assumes "u  v" shows "set u = set v"
  using conjugE[OF u  v] set_append Un_commute by metis

lemma conjug_concat_conjug: "xs  ys  concat xs  concat ys"
  unfolding conjugate_def using concat_morph by metis

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)@k  u = z" and "v  ε"
proof -
  have "z ≤p x  z" using eq[symmetric]..
  from this and x  ε have "z <p x  z"..
  then obtain k u v where "x@k  u = z" and x: "u  v = x" and "v  ε"..
  have z: "(uv)@k  u = z" unfolding x x@k  u = z..
  have "z  y = (uv)  ((uv)@k  u)" unfolding z unfolding x eq..
  also have " = (uv)@k  u  (v  u)" unfolding lassoc pow_comm[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 "xz = zy" and "x  ε"
  shows " u v k. u  v = x  v  u  = y  (u  v)@k  u = z"
  using assms by blast

lemma conjug_eq_primrootE' [elim, consumes 2]:
  assumes eq: "x  z = z  y" and "x  ε"
  obtains r s i n where
    "(r  s)@i = x" and
    "(s  r)@i = y" and
    "(r  s)@n  r = z" and
    "s  ε" and "0 < i" and "primitive (r  s)"
     proof -
  obtain i where "(ρ x)@i = x" "0 < i"
    using primroot_expE by blast
  have "z <p x  z" using prefI[OF x  z = z  y[symmetric]] x  ε..
  from per_root_primroot[OF this]
  have "z <p (ρ x)  z".
  from per_root_modE[OF this]
  obtain n r s where "r  s = ρ x" "ρ x @ n  r = z" "s  ε".
  have x: "(rs)@i = x" unfolding r  s = ρ x (ρ x)@i = x..
  have z: "(rs)@n  r = z" unfolding r  s = ρ x using  (ρ x)@n  r = z.
  have y [symmetric]: "y = (sr)@i"
    using eq[symmetric, folded x z, unfolded lassoc pows_comm[of _ i], unfolded rassoc cancel,
          unfolded shift_pow cancel].
  from x  ε have "primitive (r  s)" unfolding r  s = ρ x..
  from that[OF x y z s  ε 0 < i this]
  show thesis.
qed

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 pow_conjug_conjug_conv: assumes "0 < k" shows "u@k  v@k  u  v"
proof
  assume "u @ k  v @ k"
  obtain r s where "r  s = u@k" and "s  r = v@k"
    using conjugE[OF u@k  v@k].
  hence "v@k = (rotate |r| u)@k"
    using rotate_append rotate_pow_comm by metis
  hence "v = rotate |r| u"
    using pow_eq_eq[OF _ 0 < k] by blast
  thus "u  v"
    using rotate_conjug by blast
next
  assume "u  v"
  obtain r s where "u = r  s" and "v = s  r"
    using conjugE[OF u  v] by metis
  have "u@k  r = r  v@k"
    unfolding u = r  s v = s  r shift_pow..
  thus "u@k  v@k"
    using conjugI1 by blast
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

text‹Of course, conjugacy is an equivalence relation.›
lemma conjug_equiv: "equivp (∼)"
  by (simp add: conjug_refl conjug_sym conjug_trans equivpI reflpI sympI transpI)

lemma rotate_fac_pref: assumes "u ≤f w"
  obtains w' where "w'  w" and "u ≤p w'"
proof-
  from facE[OF u ≤f w]
  obtain p s where "w = p  u  s".
  from that[OF conjugI'[of "u  s" p, unfolded rassoc, folded this] triv_pref]
  show thesis.
qed

lemma rotate_into_pos_sq: assumes "sp ≤f ww" and  "|s|  |w|" and "|p|  |w|"
obtains w' where "w  w'" "p ≤p w'" "s ≤s w'"
proof-
  obtain pw where "pwsp ≤p ww"
    by (meson assms(1) fac_pref)
  hence "pw  s ≤p w w"
    unfolding lassoc prefix_def by force

  hence "take |pw  s| (w  w) = pw  s"
    using pref_take by blast

  have "p ≤p drop |pw  s| (w  w)"
    using pref_drop[OF pwsp ≤p ww[unfolded lassoc]] drop_pref  by metis

  let ?w = "rotate |pw  s| w"

  have "|?w| = |w|" by auto

  have "rotate |pw  s| (w  w) = ?w  ?w"
    using rotate_pow_comm_two.

  hence eq: "?w  ?w = (drop |pw  s| (w  w))  take |pw  s| (w  w)"
    by (metis pw  s ≤p w  w append_take_drop_id pref_take rotate_append)

  have "p ≤p ?w"
    using pref_prod_le[OF _ |p|  |w|[folded |?w| = |w|]]
          prefix_prefix[OF p ≤p drop |pw  s| (w  w), of "take |pw  s| (w  w)", folded eq].

  have "s ≤s ?w"
    using pref_prod_le[reversed, OF _ |s|  |w|[folded |?w| = |w|], of ?w]
    unfolding eq take |pw  s| (w  w) = pw  s lassoc by blast

  show thesis
    using that[OF rotate_conjug p ≤p ?w s ≤s ?w].
qed

lemma rotate_into_pref_sq: assumes "p ≤f ww" and  "|p|  |w|"
obtains w' where "w  w'" "p ≤p w'"
  using rotate_into_pos_sq[of ε, unfolded emp_simps, OF p ≤f ww _ |p|  |w|] by auto

lemmas rotate_into_suf_sq = rotate_into_pref_sq[reversed]

lemma rotate_into_pos: assumes "sp ≤f w"
  obtains w' where "w  w'" "p ≤p w'" "s ≤s w'"
proof(rule rotate_into_pos_sq)
  show "sp ≤f ww"
    using s  p ≤f w by blast
  show "|s|  |w|"
    using order.trans[OF pref_len' fac_len[OF s  p ≤f w] ].
  show "|p|  |w|"
    using order.trans[OF suf_len' fac_len[OF s  p ≤f w]].
qed

lemma rotate_into_pos_conjug: assumes "w  v" and "sp ≤f v"
  obtains w' where "w  w'" "p ≤p w'" "s ≤s w'"
  using assms conjug_trans rotate_into_pos by metis

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)@i = v" and u: "(s  r)@i = u" and prim': "primitive (r  s)" and "0 < i"..
  have "r  s = v" using v unfolding prim_exp_one[OF prim u] pow_1.
  show "primitive v" using prim' unfolding r  s = v.
qed

lemma conjug_prim_iff: assumes "u  v" shows "primitive u = primitive v"
  using prim_conjug[OF _ u  v] prim_conjug[OF _ conjug_sym[OF u  v]]..

lemmas conjug_prim_iff' =  conjug_prim_iff[OF conjugI']

lemmas conjug_concat_prim_iff = conjug_concat_conjug[THEN conjug_prim_iff]

lemma conjug_eq_primrootE [elim, consumes 2]:
  assumes eq: "x  z = z  y" and "x  ε"
  obtains r s i n where
    "(r  s)@i = x" and
    "(s  r)@i = y" and
    "(r  s)@n  r = z" and
    "s  ε" and "0 < i" and "primitive (r  s)"
     and "primitive (s  r)"
  using conjug_eq_primrootE'[OF assms] conjug_prim_iff' by metis


lemma conjug_primrootsE: assumes "ρ p  ρ q"
  obtains r s k l where "p = (r  s)@k" and "q = (s  r)@l"  and "primitive (rs)"
proof(cases)
  assume "p = ε  q = ε"
  obtain w::"'a list" where "primitive w"
    by blast
  from that[of w ε 0 0, unfolded emp_simps]
  show ?thesis
    by (simp add: p = ε  q = ε primitive w)
next
  assume "¬ (p = ε  q = ε)"
  hence "primitive (ρ p)"
    using assms conjug_prim_iff by auto
  from conjugE[OF ρ p  ρ q]
  obtain r s where
    "r  s = ρ p" and
    "s  r = ρ q".
  from that[of r s "eρ p" "eρ q", unfolded this, OF _ _ primitive (ρ p)]
  show ?thesis
    using primroot_exp_eq[symmetric]
    by blast
qed

lemma root_conjug: "u ≤p r  u  u¯>(ru)  r"
  using conjugI1 conjug_sym lq_pref by metis

lemmas conjug_prim_iff_pref = conjug_prim_iff[OF root_conjug]

lemma conjug_primroot_word:
  assumes conjug: "u  t = t  v"
  shows "(ρ u)  t = t  (ρ v)"
proof (cases "u = ε")
  assume "u  ε"
  from u  t = t  v u  ε obtain r s i n where
    u: "(r  s)@i = u" and v: "(s  r)@i = v" and prim: "primitive (r  s)"
    and "(r  s)@n  r = t" and "0 < i"..
  have rs: "ρ u = r  s" and sr: "ρ v = s  r"
    using prim_conjug[OF prim conjugI'] u v 0 < i prim
     primroot_unique' by meson+
  show ?thesis
    unfolding (r  s)@n  r = t[symmetric] rs sr
    by comparison
next
  assume "u = ε"
  hence "v = ε"
    using assms by force
  show ?thesis
    unfolding u = ε v = ε by simp
qed

lemma conjug_primroot:
  assumes "u  v"
  shows "ρ u  ρ v"
proof(cases)
  assume "u = ε" with u  v show "ρ u  ρ v"
    using conjug_nemp_iff by blast
next
  assume "u  ε"
  from u  v obtain t where "u  t = t  v"..
  from conjug_primroot_word[OF this]
  show "ρ u  ρ v"
    by (simp add: conjugI1)
qed

lemma conjug_primroots_nemp:  assumes "x  y  y  x" and "r  s = ρ (x  y)" and  "s  r = ρ (y  x)"
  shows "r  ε" and "s  ε"
proof-
  have "x  y  ε" and "y  x  ε"
    using assms(1) by force+
  have "r  ε  s  ε"
  proof (rule contrapos_np[OF assms(1)])
    assume "¬ (r  ε  s  ε)"
    hence "ρ (x  y) = ρ (y  x)"
      using assms(2-3) by force
    with comm_primroots[symmetric, OF x  y  ε y  x  ε]
    show "x  y = y  x"
      using eqd_eq[OF _ swap_len] by meson
  qed
  thus "r  ε" and "s  ε"
    by blast+
qed

lemma conjugE_primrootsE[elim]: assumes "x  y  y  x"
  obtains r s where "r  s = ρ (x  y)" and "s  r = ρ (y  x)" and "r  ε" and "s  ε"
proof-
  have "ρ (x  y)  ε"
    using assms by force
  from conjugE_nemp[OF conjug_primroot[OF conjugI', of x y] this] conjug_primroots_nemp[OF assms] that
  show thesis
    by auto
qed

lemma conjug_add_exp: "u  v   u@k  v@k"
  by (elim conjugE1, intro conjugI1, rule conjug_pow)

lemma conjug_primroot_iff:
  assumes nemp:"u  ε" and len: "|u| = |v|"
  shows "ρ u   ρ v  u  v"
proof
  show "u  v  ρ u   ρ v" using conjug_primroot.
  assume conjug: "ρ u   ρ v"
  have "v  ε" using nemp_len[OF nemp] unfolding len length_0_conv.
  with nemp obtain k l where roots: "(ρ u)@k = u" "(ρ v)@l = v"
    using primroot_exp_eq by blast

  have "|(ρ u)@k| = |(ρ v)@l|" 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 two_conjugs_aux: assumes "uv = xy" and "vu = yx" and "u  ε" and  "u  x" and "|u|  |x|"
  obtains r s k l m n where
    "u = (s  r)@k  s" and  "v = (r  s)@l  r" and
    "x = (s  r)@m  s" and  "y = (r  s)@n  r" and
    "primitive (r  s)" and "primitive (s  r)"
proof-
  have "|u| < |x|"
    using u  x eqd_eq(1)[OF uv = xy] le_neq_implies_less[OF |u|  |x|] by blast
  hence "x  ε"
    by force
  from eqd_lessE[OF uv = xy |u| < |x|]
  obtain t where "u  t = x" "t  y = v" "t  ε".
  from vu = yx[folded this(1-2)]
  obtain exp where "y  u = (ρ t)@exp"
    using comm_primroot_exp[OF t  ε, of "y  u"] unfolding rassoc by metis
  hence "0 < exp"
    using u  ε by blast
  from split_pow[OF y  u = (ρ t)@exp this u  ε]
  obtain r s n k  where  "u = (s  r)@k  s" "y = (r  s)@n  r" "r  s = ρ t"
    by metis
  have "primitive (r  s)"
    unfolding r  s = ρ t using t  ε by blast
  hence "primitive (s  r)"
    using conjug_prim_iff' by blast
  define e where "e = eρ t"
  have t: "t = (rs)@e"
    unfolding r  s = ρ t e_def by simp
  have eq1: "t  (r  s) @ n  r = (r  s) @ (eρ t + n)  r"
    unfolding add_exps r  s = ρ t primroot_exp_eq rassoc..
  have eq2: "((s  r) @ k  s)  t = (s  r) @ (k + e)  s"
    unfolding t by comparison
  show thesis
    using that[OF u = (s  r)@k  s _ _  y = (r  s)@n  r primitive (r  s) primitive (s  r),
          folded u  t = x t  y = v, unfolded u = (s  r)@k  s y = (r  s)@n  r, OF eq1 eq2].
qed

lemma two_conjugs: assumes "uv = xy" and "vu = yx" and "u  ε" and "x  ε" and "u  x"
  obtains r s k l m n where
    "u = (s  r)@k  s" and  "v = (r  s)@l  r" and
    "x = (s  r)@m  s" and  "y = (r  s)@n  r" and
    "primitive (r  s)" and "primitive (s  r)"
  by (rule le_cases[of "|u|" "|x|"],
  use two_conjugs_aux[OF assms(1-3,5)] in metis)
  (use two_conjugs_aux[OF assms(1-2)[symmetric] assms(4) assms(5)[symmetric]] in metis)

lemma fac_pow_pref_conjug:
  assumes "u ≤f t@k"
  obtains t' where "t  t'" and "u ≤p t'@k"
proof (cases "t = ε")
  assume "t  ε"
  obtain p q where eq: "p  u  q = t@k" using facE'[OF assms].
  obtain i r where "i  k" and "r <p t" and p: "t@i  r = p"
    using pref_mod_pow[OF prefI[OF eq] t ε].
  from r <p t obtain s where t: "r  s = t"..
  have eq': "t@i  r  (u  q) = t@k" using eq unfolding lassoc p.
  have  "u ≤p (s  r)@k" using pow_conjug[OF eq' t] unfolding rassoc..
  with conjugI'[of r s] show thesis unfolding t..
qed (use assms in auto)

lemmas fac_pow_suf_conjug = fac_pow_pref_conjug[reversed]

lemma fac_pow_len_conjug[intro]: assumes  "|u| = |v|" and "u ≤f v@k" shows "v  u"
proof-
  obtain t where "v  t" and "u ≤p t@k"
    using fac_pow_pref_conjug[OF u ≤f v @ k].
  have "u = t"
    using pref_prod_eq[OF pref_prod_root[OF u ≤p t@k] conjug_len[OF v  t,folded |u| = |v|]].
  from v  t[folded this]
  show "v  u".
qed

lemma conjug_fac_sq:
  "u  v  u ≤f v  v"
  by (elim conjugE, unfold eq_commute[of "_  _"]) (intro facI', simp)

lemma conjug_fac_pow_conv: assumes "|u| = |v|" and "2  k"
  shows "u  v  u ≤f v@k"
proof
  assume "u  v"
  have f:  "v  v ≤f v @k"
    using 2  k  unfolding pow_two[symmetric] using le_exps_pref by blast
  from fac_trans[OF conjug_fac_sq[OF u  v] this]
  show "u ≤f v @ k".
next
  show " u ≤f v @ k  u  v"
    using fac_pow_len_conjug[OF |u| = |v|, THEN conjug_sym].
qed

lemma conjug_fac_Suc: assumes "t  v"
  shows "t@k ≤f v@Suc k"
proof-
  obtain r s where "v = r  s" and "t = s  r"
    using t  v by blast
  show ?thesis
    unfolding v = r  s t = s  r
    unfolding pow_slide[of r s k, symmetric]
    by force
qed

lemma fac_pow_conjug: assumes "u ≤f v@k" and  "t  v"
  shows "u ≤f t@Suc k"
proof-
  obtain r s where "v = r  s" and "t = s  r"
    using t  v by blast
  have "s  v@k  r = t@Suc k"
    unfolding v = r  s t = s  r shift_pow pow_Suc rassoc..
  from facI[of "v@k" s r, unfolded this]
  show "u ≤f t@Suc k"
    using u ≤f v@k by blast
qed

lemma border_conjug: "x ≤b w  w<¯x  x¯>w"
  using border_conjug_eq conjugI1 by blast

lemma count_list_conjug: assumes "u  v" shows "count_list u a = count_list v a"
proof-
  from  conjugE[OF u  v]
  obtain r s where "r  s = u" "s  r = v".
  show "count_list u a = count_list v a"
    unfolding r  s = u[symmetric] s  r = v[symmetric] count_list_append by presburger
qed

lemma conjug_in_lists: "us  vs  vs  lists A  us  lists A"
  unfolding conjugate_def  by auto

lemma conjug_in_lists': "us  vs  us  lists A  vs  lists A"
  unfolding conjugate_def  by auto

lemma conjug_in_lists_iff: "us  vs  us  lists A  vs  lists A"
  unfolding conjugate_def  by auto


lemma prim_conjug_unique: assumes "primitive (u  v)" and "u  v = r  s" and "v  u = s  r" and "u  v  v  u"
  shows "u = r" and "v = s"
proof-
  have "u = r" if "primitive (u  v)" and "u  v = r  s" and "v  u = s  r" and "u  v   v  u" and  "|v|  |s|" for u v r s :: "'a list"
  proof-
    from eqdE[OF v  u = s  r |v|  |s|]
    obtain t where "v  t = s" "t  r = u".
    have "t  (r  v) = (r  v)  t"
      unfolding lassoc t  r = u unfolding rassoc v  t = s by fact
    from comm_not_prim[OF _ _ this, unfolded lassoc t  r = u]
    have "t = ε"
      using primitive (u  v) u  v  v  u by blast
    thus "u = r"
      using t  r = u by force
  qed
  from this[OF assms]
  this[OF primitive (u  v)[unfolded u  v = r  s] assms(2-3)[symmetric] assms(4)[unfolded u  v = r  s v  u = s  r]]
  show "u = r"
    by fastforce
  thus "v = s"
    using u  v = r  s by fast
qed

lemma prim_conjugE[elim, consumes 3]:  assumes "(u  v)  z = z  (v  u)" and "primitive (u  v)" and "v  ε"
  obtains k where "(u  v)@k  u = z"
proof-
  from conjug_eqE[OF assms(1) prim_nemp[OF assms(2)]]
  obtain x y m where "x  y = u  v" and "y  x = v  u" and "(x  y)@m  x = z" and "y  ε".
  from prim_conjug_unique[OF primitive (u  v) x  y = u  v[symmetric] y  x = v  u[symmetric]]
  consider "u  v = v  u" | "u = x  v = y" by blast
  thus thesis
  proof (cases)
    assume "u  v = v  u"
    from comm_not_prim[OF _ v  ε this] primitive (u  v)
    have "u = ε" by blast
    from (u  v)  z = z  (v  u)[symmetric] primitive (u  v)
    obtain k where "z = (u  v)@k  u"
      unfolding u = ε emp_simps  by blast
    from that[OF this[symmetric]]
    show thesis.
  next
    assume "u = x  v = y"
    with (x  y)@m  x = z that
    show thesis by blast
  qed
qed

lemma prim_conjugE'[elim, consumes 3]: assumes "(r  s)  z = z  (s  r)" and "primitive (r  s)" and "z  ε"
  obtains k where "(r  s)@k  r = z"
proof (cases s = ε)
  assume "s = ε"
  from assms(1-2)[unfolded this emp_simps]
  have  "primitive r" and "z  r = r  z"  by force+
  from prim_comm_exp[OF this]
  obtain k where "z = r@k" "0 < k"
    using nemp_exp_pos[OF z  ε] by metis
  have "r@(k-1)r = z"
    unfolding pow_pos'[OF 0 < k, of r, folded z = r@k]..
  from that[unfolded s = ε emp_simps, OF this]
  show thesis.
qed (use prim_conjugE[OF assms(1-2)] in blast)

lemma conjug_primroots_unique:  assumes "x  y  y  x" and
       "r  s = ρ (x  y)" and  "s  r = ρ (y  x)" and
       "r'  s' = ρ (x  y)" and  "s'  r' = ρ (y  x)"
     shows "r = r'" and "s = s'"
proof-
  have "x  y  ε" and "y  x  ε" and "x  ε" and "y  ε" and "(x  y)  (y  x)  (y  x)  (x  y)"
    using x  y  y  x eqd_eq[OF _ swap_len] by blast+
  show "r = r'"
  proof (rule prim_conjug_unique(1))
    from primroot_prim[OF x  y  ε, folded r  s = ρ (x  y)]
    show "primitive (r  s)".
    from r  s = ρ (x  y)[folded r'  s' = ρ (x  y)] s  r = ρ (y  x)[folded s'  r' = ρ (y  x)]
    show "r  s = r'  s'" and "s  r = s'  r'".
    show "r  s  s  r"
      unfolding r  s = ρ (x  y)  s  r = ρ (y  x)
      using same_primroots_comm (x  y)  (y  x)  (y  x)  (x  y) by blast
  qed
  thus "s = s'"
    using r  s = ρ (x  y)[folded r'  s' = ρ (x  y)] by blast
qed

lemma prim_conjug_pref: assumes "primitive (s  r)" and "u  r  s ≤p (s  r)@n" and "r  ε"
  obtains n where "(s  r)@n  s = u"
proof-
  have "u  r  s ≤p (s  r  u)  r  s"
    using pref_prod_root[OF u  r  s ≤p (s  r)@n] unfolding rassoc.
  from pref_prod_eq[OF this, unfolded lenmorph]
  have "(s  r)  u = u  (r  s)"
    unfolding rassoc by force
  from prim_conjugE[OF this primitive (s  r) r  ε]
  show thesis
    using that.
qed

lemma fac_per_conjug: assumes "period w n" and  "v ≤f w" and "|v| = n"
  shows "v  take n w"
proof-
  have "|take n w| = |v|"
    using fac_len[OF v ≤f w] |v| = n take_len by blast
  from per_root_powE'[OF period w n[unfolded period_def]]
  obtain k where "w ≤p take n w @ k".
  from fac_pow_len_conjug[OF |take n w| = |v|[symmetric], THEN conjug_sym]
       fac_trans[OF  v ≤f w pref_fac, OF this]
  show ?thesis.
qed

lemma fac_pers_conjug: assumes "period w n" and  "v ≤f w" and "|v| = n" and "u ≤f w" and "|u| = n"
  shows "v  u"
  using  conjug_trans[OF fac_per_conjug[OF period w n v ≤f w |v| = n]
      conjug_sym[OF fac_per_conjug[OF period w n u ≤f w |u| = n]]].

lemma conjug_pow_powE: assumes "w  r@k" obtains s where "w = s@k"
proof-
  obtain u v where "w = u  v" and "v  u = r@k"
    using assms by blast
  have "w = (v¯>(rv))@k"
    unfolding w = u  v lq_conjug_pow[OF pref_prod_root, OF prefI[OF v  u = r @ k], symmetric] v  u = r @ k[symmetric]
    by simp
  from that[OF this]
  show thesis.
qed

lemma find_second_letter:  assumes "a  b" and  "set ws = {a,b}"
  shows "dropWhile (λ c. c = a) ws  ε" and "hd (dropWhile (λ c. c = a) ws) = b"
proof-
  let ?a = "(λ c. c = a)"

  define wsb where "wsb = dropWhile ?a ws  takeWhile ?a ws"
  have "wsb  ws"
    unfolding wsb_def using takeWhile_dropWhile_id[of ?a ws] conjugI' by blast
  hence "set wsb = {a,b}"
    using set ws = {a,b} by (simp add: conjug_set)

  have "takeWhile ?a ws  ws"
    unfolding takeWhile_eq_all_conv using set ws = {a,b} a  b by simp
  thus "dropWhile ?a ws  ε" by simp
  from hd_dropWhile[OF this] set_dropWhileD[OF hd_in_set[OF this], unfolded set ws = {a,b}]
  show "hd (dropWhile ?a ws) = b"
    by blast
qed

lemma fac_conjuq_sq:
  assumes "u  v" and "|w|  |u|" and "w ≤f u  u"
  shows "w ≤f v  v"
proof -
  have assm_le: "w ≤f s  r  s  r"
    if "p  w  q = r  s  r  s" and "|r|  |p|" for w s r p q :: "'a list"
  proof -
    obtain p' where "r  p' = p"
      using p  w  q = r  s  r  s |r|  |p| unfolding rassoc by (rule eqdE[OF sym])
    show "w ≤f s  r  s  r"
      using p  w  q = r  s  r  s
      by (intro facI'[of p' _ "q  r"]) (simp flip: r  p' = p)
  qed
  obtain r s where "r  s = u" and "s  r = v" using u  v..
  obtain p q where "p  w  q = u  u" using w ≤f u  u ..
  from lenarg[OF this] |w|  |u|
  have "|r|  |p|  |s|  |q|"
    unfolding r  s = u[symmetric] lenmorph by linarith
  then show "w ≤f v  v"
    using p  w  q = u  u unfolding r  s = u[symmetric] s  r = v[symmetric]
    by (elim disjE) (simp only: assm_le rassoc, simp only: assm_le[reversed] lassoc)
qed

lemma fac_conjuq_sq_iff:
  assumes "u  v" shows "|w|  |u|  w ≤f u  u  w ≤f v  v"
  using fac_conjuq_sq[OF u  v] fac_conjuq_sq[OF u  v[symmetric]]
  unfolding conjug_len[OF u  v[symmetric]]..

lemma map_conjug:
  "u  v  map f u  map f v"
  by (elim conjugE, unfold eq_commute[of "_  _"]) auto

lemma map_conjug_iff [reversal_rule]:
  assumes "inj f" shows "map f u  map f v  u  v"
  using map_conjug map_conjug[of "map f u" "map f v" "inv f"]
  unfolding map_map inv_o_cancel[OF inj f] list.map_id by (intro iffI)

lemma card_conjug: assumes "w  ε"
  shows "card (Collect (conjugate w)) = |ρ w|"
proof-
  define f where "f = (λn. rotate n w)"

  have "ρ w  ε"
    by (simp add: assms primroot_nemp)
  obtain k where "(ρ w)@k = w"
    using primroot_expE
    by blast
  have "f`{0..<|ρ w|} = {w'. w  w'}"
    unfolding set_eq_iff
    unfolding mem_Collect_eq conjug_rotate_iff image_iff
    unfolding atLeast0LessThan
    unfolding f_def
    using lessThan_iff rotate_pow_mod[of _ "ρ w" k] mod_less_divisor[OF nemp_pos_len[OF ρ w  ε]]
    unfolding (ρ w)@k = w
    by meson

  have "inj_on f {0..<|ρ w|}"
  proof (rule inj_onI)
    fix x y
    assume "x  {0..<|ρ w|}" "y  {0..<|ρ w|}" "f x = f y"
    hence roxy: "rotate x (ρ w) = rotate y (ρ w)"
      unfolding f_def
      by (metis assms primroot_rotate_comm)
    show "x = y"
      using prim_no_rotate[OF primroot_prim[OF w  ε]] rotate_back'[OF roxy] rotate_back'[OF roxy[symmetric]] x  {0..<|ρ w|} y  {0..<|ρ w|}
      unfolding atLeast0LessThan lessThan_iff
      using bot_nat_0.not_eq_extremum less_imp_diff_less nat_le_linear zero_diff_eq by metis
  qed
  from card_image[OF this]
  show ?thesis
    unfolding f ` {0..<|ρ w|} = {w'. w  w'}
    unfolding atLeast0LessThan card_lessThan.
qed

lemma finite_Bex_conjug: assumes "finite A"
  shows "finite {r. Bex A (conjugate r)}"
  unfolding finite_Collect_bex[OF finite A, of conjugate]
proof
  fix y
  assume "y  A"
  show "finite {r. r  y}"
  proof(cases "y = ε")
    case True
    then show ?thesis
      unfolding conjug_swap[of _ y]
      by (metis (mono_tags, opaque_lifting) y  A assms conjug_nemp_iff finite_subset mem_Collect_eq subset_eq)
  next
    case False
    then show ?thesis
      unfolding conjug_swap[of _ y]
      by (simp add: card_conjug card_ge_0_finite primroot_nemp)
  qed
qed

subsection ‹Enumerating conjugates›

definition bounded_conjug
  where "bounded_conjug w' w k  ( n  k. w = rotate n w')"

named_theorems bounded_conjug

lemma[bounded_conjug]: "bounded_conjug w' w 0  w = w'"
  unfolding bounded_conjug_def by auto

lemma[bounded_conjug]: "bounded_conjug w' w (Suc k)  bounded_conjug w' w k  w = rotate (Suc k) w'"
  unfolding bounded_conjug_def using le_SucE le_imp_less_Suc le_less by metis

lemma[bounded_conjug]: "w'  w  bounded_conjug w w' (|w|-1)"
  unfolding bounded_conjug_def conjug_swap[of w'] using conjug_rotate_iff_le.

lemma "w  [a,b,c]  w = [a,b,c]  w = [b,c,a]  w = [c,a,b]"
  by (simp add: bounded_conjug)

subsection ‹General lemmas using  conjugation›

lemma switch_fac: assumes "x  y" and  "set ws = {x,y}" shows "[x,y] ≤f ws  ws"
proof-
  let ?y = "(λ a. a = y)" and ?x = "(λ a. a = x)"
  have "ws  ε"
    using set ws = {x,y} by force

  define wsx where "wsx = dropWhile ?y ws  takeWhile ?y ws"
  have "wsx  ws"
    unfolding wsx_def using takeWhile_dropWhile_id[of ?y ws] conjugI' by blast
  have "set wsx = {x,y}"
    unfolding wsx_def using set ws = {x,y} conjugI' conjug_set takeWhile_dropWhile_id by metis
  from find_second_letter[OF  x  y[symmetric] set ws = {x,y}[unfolded insert_commute[of x]]]
  have  "dropWhile (λc. c = y) ws  ε" and "hd wsx = x"
    unfolding wsx_def using hd_append by simp_all
  hence "takeWhile ?x wsx  ε"
    unfolding wsx_def takeWhile_eq_Nil_iff by blast
  from root_nemp_expE[OF takeWhile_sing_root[of x wsx] this]
  obtain k where [symmetric]: "[x]@k = takeWhile ?x wsx" and "0 < k".
  note find_second_letter[OF x  y set wsx = {x,y}]
  have "wsx = [x]@(k - 1)  [x]  [hd (dropWhile ?x wsx)]  tl (dropWhile ?x wsx)"
    unfolding lassoc pow_pos'[OF 0 < k,symmetric] takeWhile ?x wsx = [x]@k[symmetric]
    unfolding rassoc hd_tl[OF dropWhile ?x wsx  ε] takeWhile_dropWhile_id..
  from this[unfolded hd (dropWhile ?x wsx) = y]
  have "[x,y] ≤f wsx" by (auto simp add: fac_def)
  thus "[x,y] ≤f ws  ws"
    using fac_trans[OF _ conjug_fac_sq[OF wsx  ws]] by blast
qed

lemma imprim_ext_pref_comm: assumes "¬ primitive (u  v)" and "¬ primitive (u  v  u)"
  shows "u  v = v  u"
using ¬ primitive (u  v) proof (elim not_prim_primroot_expE)
  fix z n assume "z @ n = u  v" and "2  n"
  have "2 * |z|  |u  v  u|"
    by (simp add: pow_len 2  n trans_le_add1 flip: z@n = u  v rassoc)
  moreover have "u  v  u ≤p z  u  v  u"
    by (intro pref_prod_root[of _ _ "n + n"]) (simp add: z @ n = u  v add_exps)
  ultimately have "(u  v  u)  z = z  u  v  u"
    using ¬ primitive (u  v  u)  per_le_prim_iff
    by (cases "z = ε") blast+
  from comm_add_exp[OF this[symmetric], of n]
  show "u  v = v  u"
    unfolding z @ n = u  v by simp
qed

lemma imprim_ext_suf_comm:
  "¬ primitive (u  v)  ¬ primitive (u  v  v)  u  v = v  u"
  by (intro imprim_ext_pref_comm[symmetric])
     (unfold conjug_prim_iff[OF conjugI', of v] rassoc)

lemma prim_xyky: assumes "2  k" and "¬ primitive ((x  y)@k  y)" shows "x  y = y  x"
proof-
  have "k  0" using 2  k by simp
  have "(x  y)@k = (x  y)@(k - 1)  x  y"
    unfolding rassoc pow_Suc'[symmetric] Suc_minus[OF k  0]..
  have "(x  y)@k  y = ((x  y)@(k -1)  x)  y  y"
    unfolding lassoc cancel_right unfolding rassoc pow_Suc'[symmetric] Suc_minus[OF k  0]..
  from imprim_ext_suf_comm[OF _ ¬ primitive ((x  y)@k  y)[unfolded this],
       unfolded rassoc pow_Suc'[symmetric] Suc_minus[OF k  0], OF pow_nemp_imprim[OF 2  k]]
  show "x  y = y  x"
    unfolding (x  y)@k = (x  y)@(k -1)  x  y shift_pow
     pow_Suc'[of "x  y", unfolded rassoc, symmetric] pow_Suc[of "y  x", unfolded rassoc, symmetric]
    using pow_eq_eq by blast
qed

lemma fac_pow_div: assumes "u ≤f w@l" "primitive w"
  shows "w@((|u| div |w|) - 1) ≤f u"
proof-
  obtain w' where
    "w  w'" and
    "u ≤p w' @ l"
    using fac_pow_pref_conjug[OF u ≤f w@l].

  note prim_nemp[OF primitive w]
  hence "w'  ε"
    using conjug_nemp_iff w  w' by blast

  obtain s where  "s <p w'" and  "w' @ (|u| div |w'|)  s = u"
    using per_root_modE'[OF per_rootI', OF u ≤p w' @ l w'  ε].

  have "w@((|u| div |w|) - 1) ≤f w' @ (|u| div |w'|)"
    unfolding conjug_len[OF w  w']
    using conjug_fac_Suc[OF w  w']
    by (cases "(|u| div |w'|) = 0", force)
    (use Suc_minus in metis)
  thus ?thesis
    using fac_ext_suf[of _ "w' @ (|u| div |w'|)" s, unfolded w' @ (|u| div |w'|)  s = u]
    by presburger
qed

section ‹Element of lists: a method for testing if a word is in lists A›

lemma append_in_lists[simp, intro]: "u  lists A  v  lists A  u  v  lists A"
  by simp

lemma pref_in_lists: "u ≤p v  v  lists A  u  lists A"
  by (auto simp add: prefix_def)

lemmas suf_in_lists = pref_in_lists[reversed]

lemma fac_in_lists: "ws  lists S  vs ≤f ws  vs  lists S"
  by force

lemma lq_in_lists: "v  lists A  u¯>v  lists A"
  unfolding left_quotient_def using fac_in_lists[OF _ sublist_drop].

lemmas rq_in_lists = lq_in_lists[reversed]

lemma take_in_lists: "w  lists A  take j w  lists A"
  using pref_in_lists[OF take_is_prefix].

lemma drop_in_lists: "w  lists A  drop j w  lists A"
  using suf_in_lists[OF suffix_drop].

lemma lcp_in_lists: "u  lists A   u p v  lists A"
  using pref_in_lists[OF lcp_pref].

lemma lcp_in_lists': "v  lists A   u p v  lists A"
  using pref_in_lists[OF lcp_pref'].

lemma append_in_lists_dest: "u  v  lists A  u  lists A"
  by simp

lemma append_in_lists_dest': "u  v  lists A  v  lists A"
  by simp

lemma pow_in_lists: "u  lists A  u@k  lists A"
  by (induct k) auto

lemma takeWhile_in_list: "u  lists A  takeWhile P u  lists A"
  using take_in_lists[of u _ "|takeWhile P u|", folded takeWhile_eq_take].

lemma rev_in_lists: "u  lists A  rev u  lists A"
  by auto

lemma append_in_lists_dest1: "u  v = w  w  lists A  u  lists A"
  by auto

lemma append_in_lists_dest2: "u  v = w  w  lists A  v  lists A"
  by auto

lemma pow_in_lists_dest1: "u  v = w@n  w  lists A  u  lists A"
  using append_in_lists_dest pow_in_lists by metis

lemma pow_in_lists_dest1_sym: "w@n = u  v  w  lists A  u  lists A"
  using append_in_lists_dest pow_in_lists by metis

lemma pow_in_lists_dest2: "u  v = w@n  w  lists A  v  lists A"
  using append_in_lists_dest' pow_in_lists by metis

lemma pow_in_lists_dest2_sym: "w@n = u  v  w  lists A  v  lists A"
  using append_in_lists_dest' pow_in_lists by metis

lemma per_in_lists: "w <p r  w  r  lists A  w  lists A"
  using  pow_in_lists[of r A] pref_in_lists per_root_pow_conv by metis

lemma nth_in_lists: "j < |w|  w  lists A  w ! j  A"
  using in_lists_conv_set nth_mem by force

method inlists =
 (insert method_facts, use nothing in ((elim suf_in_lists | elim pref_in_lists[elim_format] | rule lcp_in_lists | rule drop_in_lists |
    rule lq_in_lists | rule rq_in_lists |
     rule take_in_lists | intro lq_in_lists | rule nth_in_lists |
     rule append_in_lists | elim conjug_in_lists | rule pow_in_lists | rule takeWhile_in_list
   | elim append_in_lists_dest1 | elim append_in_lists_dest2
   | elim pow_in_lists_dest2 | elim pow_in_lists_dest2_sym
   | elim pow_in_lists_dest1 | elim pow_in_lists_dest1_sym)
   | (simp | fact))+)

section ‹Reversed mappings›

definition rev_map :: "('a list  'b list)  ('a list  'b list)" where
  "rev_map f = rev  f  rev"

lemma rev_map_idemp[simp]: "rev_map (rev_map f) = f"
  unfolding rev_map_def by auto

lemma rev_map_arg: "rev_map f u = rev (f (rev u))"
  by (simp add: rev_map_def)

lemma rev_map_arg': "rev ((rev_map f) w) = f (rev w)"
  by (simp add: rev_map_def)

lemmas rev_map_arg_rev[reversal_rule] = rev_map_arg[reversed add: rev_rev_ident]

lemma rev_map_sing: "rev_map f [a] =  rev (f [a])"
  unfolding rev_map_def by simp

lemma rev_maps_eq_iff[simp]: "rev_map g = rev_map h  g = h"
  using arg_cong[of "rev_map g" "rev_map h" rev_map, unfolded rev_map_idemp] by fast

lemma rev_map_funpow[reversal_rule]: "(rev_map (f::'a list 'a list)) ^^ k = rev_map  (f ^^ k)"
  unfolding funpow.simps rev_map_def
  by(induct k, simp+)

section ‹Overlapping powers, periods, prefixes and suffixes›

lemma pref_suf_overlapE: assumes "p ≤p w" and "s ≤s w" and "|w|  |p| + |s|"
  obtains p1 u s1 where "p1  u  s1 = w" and "p1  u = p" and "u  s1 = s"
proof-
  define u where "u = (w<¯s)¯>p"
  have "u ≤s p"
    unfolding u_def lq_def using suffix_drop.
  obtain p1 s1 where "p1  u = p" and "p  s1 = w"
    using  suffixE[OF u ≤s p] prefixE[OF p ≤p w] by metis
  note p  s1 = w[folded p1  u = p, unfolded rassoc]

  have "|s1|  |s|"
    using |w|  |p| + |s|[folded p  s1 = w, unfolded lenmorph] by force
  hence "s1 ≤s s"
    using p  s1 = w s ≤s w suf_prod_long by blast

  from rq_lq_assoc[OF rq_suf_suf[OF s ≤s w], of s1] u_def[folded rqI[OF p  s1 = w]]
  have "u = s<¯s1"
    using suf_rq_lq_id[OF s ≤s w] s1 ≤s s by presburger
  hence "u  s1 = s"
    using  rq_suf[OF s1 ≤s s] by blast

  from that[OF p1  u  s1 = w p1  u = p this]
  show thesis.
qed

lemma mid_sq: assumes "pxq=xx" shows "xp=px" and "xq=qx"
proof-
  have "(xp)xq = (px)qx"
    using assms by auto
  from eqd_eq[OF this]
  show "xp=px" and "xq=qx"
    by simp+
qed

lemma mid_sq': assumes "pxq=xx" shows "q  p = x" and "p  q = x"
proof-
  have "pqx = xx"
    using assms[unfolded  mid_sq(2)[OF assms]].
  thus "pq = x"  by auto
  from assms[folded this] this
  show "qp = x"  by auto
qed

lemma mid_sq_pref: "p  u ≤p u  u  p  u = u  p"
  using mid_sq(1)[symmetric] unfolding prefix_def rassoc by metis

lemmas mid_sq_suf = mid_sq_pref[reversed]

lemma mid_sq_pref_suf: assumes "pxq=xx" shows "p ≤p x" and "p ≤s x" and "q ≤p x" and "q ≤s x"
  using assms mid_sq'[OF assms] by blast+

lemma mid_pow: assumes "px@(Suc l)q = x@k"
  shows "xp=px" and "xq=qx"
proof-
  have "xpx@lxq = x(px@Suc l  q)"
    by comparison
  also have "... = (px@Suc l  q)  x"
    unfolding rassoc assms  by comparison
  also have "... =  pxx@lqx" by simp
  finally have eq: "xpx@lxq = pxx@lqx".

  have "(xp)x@lxq = (px)x@lqx"
    using eq unfolding rassoc.
  from eqd_comp[OF this]
  show "xp = px"
    using comm_ruler by blast

  have "(xpx@l)(xq) = (xpx@l)(qx)"
    using eq unfolding lassoc xp = px.
  from this[unfolded cancel]
  show "xq = qx".
qed

lemma root_suf_comm: assumes "x ≤p r  x" and  "r ≤s r  x" shows "r  x = x  r"
proof-
  have "r  x = x  x¯>(r  x)"
    using lq_pref[OF x ≤p r  x, symmetric].
  from this and conj_len[OF this]
  have "r = x¯>(r  x)"
    using lq_pref[OF x ≤p r  x] suf_ruler_eq_len[OF r ≤s r  x, of "x¯>(r  x)"] by blast
  from r  x = x  x¯>(r  x)[folded this]
  show "r  x = x  r".
qed

lemma pref_marker: assumes "w ≤p v  w" and "u  v ≤p w"
  shows "u  v = v  u"
  using append_prefixD[OF u  v ≤p w] comm_ruler[OF u  v ≤p w, of "v  w", unfolded same_prefix_prefix]
    w ≤p v  w by blast

lemma pref_marker_ext: assumes "|x|  |y|" and "v  ε" and "y  v ≤p x  v@k"
  obtains n where "y = x  (ρ v)@n"
proof-
  note pref_prod_long_ext[OF y  v ≤p x  v@k |x|  |y|]
  have "x¯>y  v ≤p v@k"
    using pref_cancel_lq_ext[OF y  v ≤p x  v@k |x|  |y|].
  from pref_marker[OF _ this]
  have "x¯>y  v = v  x¯>y"
    unfolding pow_comm[symmetric] by blast
  then obtain n where "x¯>y = (ρ v)@n"
    using v  ε
    using comm_primroots pow_zero primroot_expE by metis
  hence "y = x  (ρ v)@n"
    using x ≤p y by (auto simp add: prefix_def)
  from that[OF this] show thesis.
qed

lemma pref_marker_sq: "p  x ≤p x  x  p  x = x  p"
  using pref_marker same_prefix_prefix triv_pref by metis

lemmas suf_marker_sq = pref_marker_sq[reversed]

lemma pref_marker_conjug: assumes "w  ε" and "w  r  s ≤p s  (r  s)@m" and "primitive (r  s)"
  obtains n where "w = s  (r  s)@n"
proof-
  have "(r  w)  r  s ≤p (r  s)@Suc m"
    using w  r  s ≤p s  (r  s)@m by auto
  from pref_marker[OF _ this, folded pow_comm, OF triv_pref]
  have "(r  w)  r  s = (r  s)  r  w".
  from comm_primroots'[OF _ prim_nemp[OF primitive (r  s)] this, unfolded prim_self_root[OF primitive (r  s)]]
  have "ρ (r  w) = r  s"
    using w  ε by blast
  then obtain n where "r  w = (r  s)@n" "0 < n"
    using w  ε primroot_expE by metis
  thus thesis
    using pow_pos[OF 0 < n, of "r  s", folded r  w = (r  s)@n,
    unfolded rassoc cancel] that by force
qed

lemmas pref_marker_reversed = pref_marker[reversed]

lemma suf_marker_per_root: assumes "w ≤p v  w" and "p  v  u ≤p w"
  shows "u ≤p v  u"
proof-
  have "p  v = v  p"
    using pref_marker[OF w ≤p v  w, of p] p  v  u ≤p w by (auto simp add: prefix_def)
  from pref_trans[OF p  v  u ≤p w[unfolded lassoc this, unfolded rassoc] w ≤p v  w]
  have "p  u ≤p w"
    using pref_cancel by auto
  from  ruler_le[OF this p  v  u ≤p w]
  have "p  u ≤p p  v  u"
    by force
  thus ?thesis
    unfolding pref_cancel_conv.
qed

lemma suf_marker_per_root': assumes "w ≤p v  w" and "p  v  u ≤p w" and "v  ε"
  shows "u ≤p p  u"
proof-
  have "p  v = v  p"
    using pref_marker[OF w ≤p v  w, of p] p  v  u ≤p w by (fastforce simp add: prefix_def)
  from root_comm_root[OF suf_marker_per_root[OF w ≤p v  w p  v  u ≤p w] this v  ε]
  show "u ≤p p  u".
qed

lemma marker_fac_pref: assumes "u ≤f r@k" and  "r ≤p u" shows "u ≤p r@k"
  using assms
proof (cases "r = ε")
  assume "r  ε"
  have "|u|  |r@k|"
    using u ≤f r@k by force
  obtain u' where "r  u' = u"
    using r ≤p u by (auto simp add: prefix_def)
  obtain p s where "p  u  s = r@k"
    using u ≤f r@k by blast
  from suf_marker_per_root[of "r@k" r p "u'  s", folded pow_comm, OF triv_pref]
  have "u'  s ≤p r  (u'  s)"
    using p  u  s = r@k[folded r  u' = u, unfolded rassoc] by fastforce
  hence "u'  s ≤p r@k  (u'  s)"
    using per_exp_pref by blast
  hence "u ≤p (r@k  r)  (u'  s)"
    unfolding r  u' = u[symmetric] pow_Suc'[symmetric] pow_Suc rassoc
       by (auto simp add: prefix_def)
  thus "u ≤p r@k"
    unfolding rassoc using |u|  |r@k| by blast
qed simp

lemma marker_fac_pref_len: assumes "u ≤f r@k" and "t ≤p u" and "|t| = |r|"
  shows "u ≤p t@k"
proof-
  have "|u|  |r@k|"
    using u ≤f r@k by force
  hence "|u|  |t@k|"
    unfolding pow_len |t| = |r|.
  have "t ≤f r@k"
    using assms by blast
  hence "t  r"
    using |t| = |r| by (simp add: conjug_sym fac_pow_len_conjug)
  from fac_pow_conjug[OF u ≤f r@k this]
  have "u ≤p t@Suc k"
    using marker_fac_pref[OF _  t ≤p u] by blast
  thus "u ≤p t@k"
    using |u|  |t@k| unfolding pow_Suc' by blast
qed

lemma root_suf_comm': "x ≤p r  x  r ≤s x  r  x = x  r"
  using root_suf_comm suffix_appendI[of r x r] by blast

lemmas suf_root_pref_comm = root_suf_comm'[reversed]

lemma marker_pref_suf_fac:  assumes "u ≤p v" and "u ≤s v" and "v ≤f u@k"
  shows "u  v = v  u"
  using root_suf_comm'[OF pref_prod_root[OF  marker_fac_pref[OF v ≤f u@k u ≤p v]] u ≤s v].

lemma pref_suf_per_fac_comm:
  assumes "v ≤p u  v" and "v ≤s v  u" and "u ≤f v@k" shows "u  v = v  u"
      using marker_pref_suf_fac[OF _ _ u ≤f v@k] root_suf_comm[OF v ≤p u  v suf_ext] root_suf_comm[reversed, OF v ≤s v  u pref_ext]
      ruler_pref'[OF v ≤p u  v] ruler_suf'[OF v ≤s v  u] by argo

lemma mid_long_pow: assumes eq: "y@m = u  x@(Suc k)  v" and "|y|  |x@k|"
  shows "(u  v)  y = y  (u  v)" and  "(u  x@l  v)  y = y  (u  x@l  v)" and "(u¯>(yu))  x = x  (u¯>(yu))"
proof-
  have eq': "x x v  u = u¯>(uxxv)u" by simp
  let ?y = "u¯>(yu)"
  have "u ≤p y  u"
    using eq prefI pref_prod_root[of u y m,unfolded eq] by simp
  hence "?y  y"
    using root_conjug by blast
  from conjug_len[OF this]
  have "|?y|  |x@k|"
    using |y|  |x@k| by simp
  from lq_conjug_pow[OF u ≤p y  u, of m]
  have "?y@m = x@Suc kvu"
    unfolding eq eq' by simp
  hence "x@Suc k ≤p ?y  x@Suc k"
    using rassoc prefI pref_prod_root[of "x@Suc k" ?y m] by blast
  have "x @ Suc k ≤p x  x @ Suc k"
    using pref_pow_ext' by blast
  have com: "?y  x = x  ?y"
    using  |?y|  |x@k| two_pers[OF x@Suc k ≤p ?y  x@Suc k x @ Suc k ≤p x  x @ Suc k]
    unfolding pow_Suc' lenmorph by linarith
  thus "?y  x = x  ?y"
    by blast
  have "?y  x@Suc k = x@Suc k  ?y"
    using com comm_add_exp by metis
  from pow_comm[of ?y m, unfolded ?y @ m = x@(Suc k)  v  u, unfolded lassoc this, unfolded rassoc]
  have "x@Suc k  v  u  ?y = x@Suc k  ?y  v  u".
  hence "u  ?y  v  u = u  v  u  ?y" by simp
  thus "(u  v)   y = y  (u  v)"
    unfolding lassoc lq_pref[OF u ≤p y  u] by fastforce
  have "u  x@l  v   u  ?y =  u  (?y  x@l)  v  u"
    unfolding comm_add_exp[OF com[symmetric], of l, symmetric] rassoc cancel
    using u  ?y  v  u = u  v  u  ?y[unfolded cancel, symmetric].
  thus "(u  x@l  v)  y = y  (u  x@l  v)"
    unfolding lq_pref[OF u ≤p y  u] lassoc by blast
qed

lemma mid_pow_pref_suf': assumes  "sw@(Suc l)p ≤f w@k" shows "p ≤p w@k"  and "s ≤s w@k"
proof-
  obtain v u where dec: "v  s   w@(Suc l)  p  u = w@k"
    using facE'[OF assms, unfolded rassoc].
  hence "(v  s)  w = w  (v  s)" and "w  (p  u) = (p  u)  w"
    using mid_pow[of "v  s" w l "p  u" k] unfolding rassoc by presburger+
  have "|p|  |w@k|" and "|s|  |w@k|"
    using fac_len[OF assms] unfolding lenmorph by linarith+

  from per_exp_pref[of "p  u" w k, unfolded w  (p  u) = (p  u)  w, OF triv_pref]
  have "p ≤p w@k  (p  u)"
    using   prefix_order.trans[OF triv_pref[of p u]] by blast
  thus "p ≤p w@k"
    using |p|  |w @ k| pref_prod_le by blast

  from per_exp_suf[of "v  s" w k, unfolded (v  s)  w = w  (v  s), OF triv_suf]
  have "s ≤s (v  s)  w@k"
    using  suffix_order.trans[OF triv_suf[of s v], of "(v  s)  w@k"] by blast
  thus "s ≤s w@k"
    using |s|  |w @ k| suf_prod_le by blast
qed

lemma mid_pow_pref_suf: assumes  "swp ≤f w@k" shows "p ≤p w@k"  and "s ≤s w@k"
  using mid_pow_pref_suf'[of s w 0 p k, unfolded pow_one, OF assms].

lemma fac_marker_pref: "y  x ≤f y@k  x ≤p y  x"
  using mid_pow_pref_suf(1)[of ε, unfolded emp_simps, THEN pref_prod_root].

lemmas fac_marker_suf = fac_marker_pref[reversed]

lemma prim_overlap_sqE [consumes 2]:
  assumes prim: "primitive r" and eq: "p  r  q = r  r"
  obtains (pref_emp) "p = ε" | (suff_emp) "q = ε"
proof (cases "|p| = 0", blast)
  assume "|p|  0" and qemp: "q = ε  thesis"
  hence "|q| < |r|"
    using lenarg[OF eq] unfolding lenmorph by linarith
  have "q = ε"
    using prim_comm_short_emp[OF prim  mid_sq(2)[OF eq, symmetric] |q| < |r|].
  from qemp[OF this]
  show thesis.
qed

lemma prim_overlap_sqE' [consumes 2]:
  assumes prim: "primitive r" and eq: "p  r  q = r  r"
  obtains (pref_emp) "p = ε" | (suff_emp) "p = r"
  using append_Nil2 eq mid_sq'(2) prim prim_overlap_sqE by metis

lemma prim_overlap_sq:
  assumes prim: "primitive r" and eq: "p  r  q = r  r"
  shows "p = ε  q = ε"
  using prim_overlap_sqE[OF prim eq disjI1 disjI2].

lemma prim_overlap_sq':
  assumes prim: "primitive r" and pref: "p  r ≤p r  r" and len: "|p| < |r|"
  shows "p = ε"
  using mid_sq(1)[symmetric, THEN prim_comm_short_emp[OF prim _ len ]] pref
   by (auto simp add: prefix_def)

lemma prim_overlap_pow:
  assumes prim: "primitive r" and pref: "u  r ≤p r@k"
  obtains i where "u = r@i" and "i < k"
proof-
  obtain q where eq: "u  r @ Suc 0  q = r @ k"
    using pref by (auto simp add: prefix_def)
  from mid_pow(1)[OF this, symmetric]
  have "u  r = r  u".
  from prim_comm_exp[OF primitive r this]
  obtain i where "r@i = u".
  hence "|r @ Suc i|  |r @ k|"
    using pref by (auto simp add: prefix_def)
  from mult_cancel_le[OF nemp_len[OF prim_nemp[OF prim]] this[unfolded pow_len]]
  have "i < k"  by auto
  from that[OF r@i = u[symmetric] this]
  show thesis.
qed

lemma prim_overlap_pow':
  assumes prim: "primitive r" and pref: "u  r ≤p r@k" and less: "|u| < |r|"
  shows "u = ε"
proof-
  obtain i where "u = r@i"
    using prim_overlap_pow[OF prim pref] by force
  from less[unfolded pow_len[of r i, folded this]]
  have "i = 0" by force
  from u = r@i[unfolded this pow_zero]
  show "u = ε".
qed

lemma prim_sqs_overlap:
  assumes prim: "primitive r" and comp: "u  r  r  v  r  r"
    and len_u: "|u| < |v| + |r|" and len_v: "|v| < |u| + |r|"
  shows "u = v"
proof (cases rule: le_cases)
  have wlog_le: "u = v" if comp: "u  (r  r)  v  (r  r)" and len_v: "|v| < |u| + |r|"
    and "|u|  |v|" for u v
  proof -
    obtain w where v: "u  w = v"
      using comp_shorter[OF comp_prefs_comp[OF comp] |u|  |v|] by (auto simp add: prefix_def)
    have "|w| < |r|" using len_v unfolding v[symmetric] by simp
    have comp': "r  r  (w  r)  r" using comp unfolding v[symmetric] rassoc comp_cancel.
    moreover have "|w  r|  |r  r|" using less_imp_le_nat[OF |w| < |r|] by simp
    ultimately have pref: "w  r ≤p r  r"
      by (rule pref_comp_len_trans[OF triv_pref])
    from this |w| < |r| have "w = ε" by (rule prim_overlap_sq'[OF prim])
    show "u = v" using v unfolding w = ε append_Nil2.
  qed
  show "|u|  |v|  u = v" using wlog_le[OF comp len_v].
  show "|v|  |u|  u = v" using wlog_le[OF comp[symmetric] len_u, symmetric].
qed

lemma drop_pref_prim: assumes "Suc n < |w|" and "w ≤p drop (Suc n) (w  w)" and "primitive w"
  shows False
  using assms
proof (cases "w = ε")
  assume "w  ε"
  obtain s where "drop (Suc n) (w  w) = w  s"
    using prefD[OF w ≤p drop (Suc n) (w  w)] by blast
  note takedrop[of "Suc n" "w  w", unfolded this]
  from Suc n < |w| w  ε prim_overlap_sqE'[OF primitive w this]
  show False by auto
qed simp

lemma root_suf_conjug: assumes "primitive (s  r)" and "y ≤p (s  r)  y" and "y ≤s y  (r  s)" and "|s  r|  |y|"
             obtains l where "y = (s  r)@l  s"
proof-
  have "y  ε"
    using assms(1) assms(4) by force
  have "r  s ≤s y"
    using suf_prod_long[OF y ≤s y  (r  s) |s  r|  |y|[unfolded swap_len]].
  have "primitive (r  s)"
    using prim_conjug[OF primitive (s  r) conjugI'].
  have "r  y ≤p (r  s)  (r  y)"
    using y ≤p (s  r)  y by auto
  from prim_comm_exp[OF primitive (r  s) root_suf_comm'[OF this suf_ext[OF r  s ≤s y], symmetric]]
  obtain k where [symmetric]: "(r  s)@k = r  y" and "0 < k"
    using y  ε using nemp_exp_pos sufI suf_emp by metis
  hence "y = (s  r)@(k-1)  s"
    unfolding pow_pos[of _ "rs", OF 0 < k] rassoc cancel shift_pow by blast
  from that[OF this]
  show thesis.
qed

lemma pref_suf_pows_comm: assumes "x ≤p y@(Suc k)x@l" and  "y ≤s y@m  x@(Suc n)"
  shows "x  y = y  x"
  using root_suf_comm[OF per_root_drop_exp'[OF assms(1)] per_root_drop_exp'[reversed, OF assms(2)], symmetric].

lemma root_suf_pow_comm: assumes "x ≤p r  x" and  "r ≤s x@(Suc k)" shows "r  x = x  r"
  using  root_suf_comm[OF x ≤p r  x suf_prod_root[OF r ≤s x@(Suc k)]].

lemma suf_pow_short_suf: "r ≤s x@k  |x|  |r|  x ≤s r"
  using suf_prod_root[THEN suf_prod_long].

thm suf_pow_short_suf[reversed]

lemma sq_short_per: assumes "|u|  |v|" and "vv ≤p u(vv)"
  shows "uv = vu"
  using
    pref_marker[of "vv", OF vv ≤p u(vv)
    pref_prod_long[OF append_prefixD[OF vv ≤p u(vv)] |u|  |v|,
      THEN pref_cancel'], symmetric].

lemma fac_marker: assumes "w ≤p uw" and "uvu ≤f w"
  shows "u  v = v  u"
proof-
  obtain p s where "w = puvus"
    using uvu ≤f w[unfolded fac_def]
    by auto

  hence "puvu = upuv"
    using pref_marker[OF w ≤p uw, unfolded w = puvus, of "p  u  v"]
    by force

  thus "uv = vu"
    using eqd_eq[of "p  u" "v  u" "u  p" "u  v", unfolded rassoc, OF _ swap_len]
    by presburger
qed

lemma "4 = Suc(Suc(Suc(Suc 0)))"
  using [[simp_trace]] by simp


lemma xyxy_conj_yxxy: assumes "x  y  x  y  y  x  x  y"
  shows "x  y = y  x"
proof-
  have four: "x@4 = xxxx" for x :: "'a list"
    unfolding numeral_Bit0 by simp
  from conjug_fac_sq[OF assms[symmetric]]
  have "y  x  x  y ≤f (x  y)@4"
    unfolding four rassoc.
  from marker_fac_pref[reversed,
      OF this triv_suf[of "xy" "yx", unfolded rassoc]]
  have "y  x  x  y ≤s (x  y) @ 4".
  hence "y  x  x  y ≤s (xyxy)xyxy"
    unfolding four rassoc.
  from suf_prod_eq[OF this]
  show "x  y = y  x"
    by simp
qed


lemma per_glue: assumes "period u n" and "period v n" and "u ≤p w" and "v ≤s w" and
              "|w| + n  |u| + |v|"
            shows "period w n"
proof (rule indeces_period)
  show  "w  ε"
    using period u n u ≤p w by force
  show "0 < n"
    using period u n per_not_zero by metis
  fix i assume "i + n < |w|"
  show "w ! i = w ! (i + n)"
  proof (cases)
    assume "i + n < |u|"
    hence "w ! i = u ! i" and "w ! (i+n) = u ! (i+n)"
      using add_lessD1 u ≤p w pref_index by metis+
    thus "w ! i = w ! (i + n)"
      unfolding w ! i = u ! i w ! (i+n) = u ! (i+n)
      using period_indeces[OF period u n i + n < |u|]  by blast
  next
    assume "¬ i + n < |u|"
    obtain p where "w = p  v"
      using v ≤s w by (auto simp add: suffix_def)
    have "¬ i < |p|"
      using  ¬ i + n < |u| |w| + n  |u| + |v| unfolding lenarg[OF w = p  v, unfolded lenmorph]
      by auto
    hence "w!i = v!(i - |p|)" and "w!(i+n) = v!((i - |p|) + n)"
      unfolding w = p  v nth_append by simp_all
    have "i - |p| + n < |v|"
      using ¬ i < |p| i + n < |w| w = p  v by auto
    from period_indeces[OF period v n this]
    show "w ! i = w ! (i + n)"
      unfolding w!i = v!(i - |p|) w!(i+n) = v!(i - |p| + n).
  qed
qed

lemma per_glue_facs: assumes "u  z ≤f w@k" and "z  v ≤f w@m" and "|w|  |z|"
  obtains l where "u  z  v ≤f w@l"
  using assms
proof (cases "k = 0")
  assume "k  0"
  have "z ≤f w@k"
    using u  z ≤f w@k by blast
  have "z ≤f w@m"
    using z  v ≤f w@m by blast
  define t where "t = take |w| z"
  have "|t| = |w|" and "t ≤p z"
    unfolding t_def using |w|  |z| take_is_prefix by (force,blast)
  hence "w  t"
    using z ≤f w@m by blast
  from marker_fac_pref_len[OF z  v ≤f (w) @ m _ |t| = |w| ]
  have "z  v ≤p t@m"
    using  t ≤p z by force
  have "u  z ≤f t@Suc k"
    using  fac_pow_conjug[OF u  z ≤f w@k  w  t[symmetric]].
  with t ≤p z
  have "u ≤s t@Suc k"
    using mid_pow_pref_suf(2)[of u t "t¯>z" "Suc k"] lq_pref by metis
  have "(t@Suc k<¯u) (u  z  v)  (z  v)¯>(t@m) = t@Suc k  t@m"
    unfolding lassoc rq_suf[OF u ≤s t@Suc k] unfolding rassoc cancel  using  lq_pref[OF z  v ≤p t@m] unfolding rassoc.
  from facI[of "u  z  v" "t@Suc k<¯u" "(z  v)¯>(t@m)", unfolded this, folded add_exps]
  obtain l where "u  z  v ≤f t@l"
    by metis
  from that[OF fac_pow_conjug[OF this w  t]]
  show thesis.
qed simp

lemma per_fac_pow_fac: assumes "period w n" and "v ≤f w" and "|v| = n"
  obtains k where "w ≤f v@k"
proof-
  obtain m where "w ≤f (take n w)@m"
    using  per_root_powE[OF period w n[unfolded  period_def]] pref_fac sprefD1 by metis
  obtain r s where "r  s = v" and "s  r = take n w"
    using fac_per_conjug[OF assms, THEN conjugE].
  hence "r  (take n w)@m  s = v@Suc m"
    by (metis pow_slide)
  from that[OF fac_trans, OF w ≤f (take n w)@m] sublist_appendI[of "(take n w)@m" r s, unfolded this]
  show thesis
    by blast
qed

lemma refine_per: assumes "period w n" and "v ≤f w" and "n  |v|" and "period v k" and "k dvd n"
  shows "period w k"
proof-
  have "n  0"
    using period w n by auto
  have "w  ε"
    using period w n by auto
  have "v  ε"
    using period v k by auto
  have "|take n w| = n"
    using take_len[OF le_trans[OF n  |v| fac_len[OF v ≤f w]]].
  have "|take n v| = n"
    using take_len[OF n  |v|].
  have "period v n"
    using  period_fac'[OF period w n v ≤f w v  ε] by blast
  have "take n v ≤f w"
    using v ≤f w n  |v| sublist_order.dual_order.trans sublist_take by metis
  have "period (take n v) k"
    using period w n period v k per_not_zero per_pref' take_is_prefix take_nemp by metis
  have "k  n"
    using k dvd n n  0 by auto
  hence "take k (take n v) =  take k v"
    using take_le_take by blast
  hence "(take k v)@(n div k) = take n v"
    using  per_div[OF _ period (take n v) k, unfolded |take n v| = n, OF k dvd n] by presburger
  have "|take k v| = k"
    using  order.trans[OF k  n n  |v|, THEN take_len].
  obtain e where  "w ≤f (take n v)@e"
    using per_fac_pow_fac[OF period w n take n v ≤f w |take n v| = n].
  from per_fac[OF w  ε this[folded (take k v)@(n div k) = take n v, folded  pow_mult]]
  show ?thesis
    unfolding |take k v| = k by blast
qed

lemma xy_per_comp: assumes "xy ≤p qxy"
  and "q  ε" and "q  y"
shows "x  y"
proof(cases rule: pref_compE[OF q  y])
  assume "q ≤p y"
  have "xq = qx"
    using
      pref_cancel'[OF q ≤p y, of x, THEN pref_trans, OF x  y ≤p q  x  y]
    unfolding lassoc
    using ruler_eq_len[OF _ triv_pref swap_len]
    by blast
  thus ?thesis
    using assms(1) assms(2) pref_comp_sym root_comm_root
      ruler_pref'' same_prefix_prefix by metis
next
  assume "y ≤p q"
  then show ?thesis
    by (meson append_prefixD prefix_append ruler' assms)
qed

lemma prim_xyxyy: "x  y  y  x  primitive (x  y  x  y  y)"
proof (rule prim_conjug)
  show "y  x  y  x  y  x  y  x  y  y"
    by (intro conjugI1) simp
  show "x  y  y  x  primitive (y  x  y  x  y)"
    by (intro iffD2[OF per_le_prim_iff[of _ "y  x"]]) auto
qed

lemma prim_min_per_suf_eq: assumes "primitive x" and "π x ≤s x" shows "π x = x"
  using assms(1) min_per_root_per_root[OF prim_nemp[OF primitive x], unfolded  ] root_suf_comm'[OF _ π x ≤s x]
  unfolding primitive_iff_per  by blast

lemma primroot_code[code]: "ρ x = (if x  ε then (if π x ≤s x then π x else x) else Code.abort (STR ''Empty word has no primitive root.'') (λ_. (ρ x)))"
proof(cases "x = ε")
  assume "x  ε"
  thus ?thesis
    unfolding if_P[OF x  ε]
  proof(cases)
    assume "eρ x = 1"
    have "primitive x"
      using primroot_exp_eq[of x, unfolded eρ x = 1 exp_simps]
      unfolding prim_primroot_conv[OF x  ε].
    from prim_min_per_suf_eq[OF this] prim_self_root[OF this]
    show "ρ x = (if π x ≤s x then π x else x)"
      by argo
  next
    assume "eρ x  1"
    show "ρ x = (if π x ≤s x then π x else x)"
      using primroot_suf
      unfolding min_per_short_primroot[OF x  ε primroot_exp_eq eρ x  1]
      by auto
  qed
qed (simp add: primitive_root_def)

lemma per_lemma_pref_suf: assumes "w <p p  w" and "w <s w  q" and
  fw: "|p| + |q|  |w|"
obtains r s k l m where "p = (r  s)@k" and "q = (s  r)@l" and "w = (r  s)@m  r" and "primitive (rs)"
proof-
  let ?q = "(w  q)<¯w"
  have "w <p ?q  w"
    using ssufD1[OF w <s w  q] rq_suf[symmetric, THEN per_rootI[OF prefI rq_ssuf[OF w <s w  q]]]
    by argo
  have "q  ?q"
    by (meson assms(2) conjugI1 conjug_sym rq_suf suffix_order.less_imp_le)

  have nemps': "p  ε" "?q ε"
    using assms(1) w <p ?qw by fastforce+
  from two_pers[OF sprefD1[OF w <p p  w] sprefD1[OF w <p ?qw]] fw
  have "p  ?q = ?q  p"
    unfolding conjug_len[OF q  (w  q)<¯w]
    by blast
  then have "ρ p = ρ ?q" using comm_primroots[OF nemps'] by force
  hence [symmetric]: "ρ q  ρ p"
    using conjug_primroot[OF q  (w  q)<¯w]
    by argo
  from conjug_primrootsE[OF this]
  obtain r s k l where
    "p = (r  s) @ k" and
    "q = (s  r) @ l" and
    "primitive (r  s)".
  have "w ≤p (rs)w"
    using assms per_root_drop_exp  sprefD1 p = (r  s) @ k
    by meson
  have "w ≤s w(sr)"
    using assms(2) per_root_drop_exp[reversed] ssufD1 q = (s  r) @ l
    by meson
  have "|r  s|  |w|"
    using conjug_nemp_iff[OF q  ?q] dual_order.trans length_0_conv  nemps' primroot_len_le[OF nemps'(1)] fw
    unfolding primroot_unique[OF nemps'(1) primitive (r  s) p = (r  s) @ k]
    by linarith
  from root_suf_conjug[OF primitive (r  s) w ≤p (rs)w w ≤s w(sr) this]
  obtain m where "w = (r  s) @ m  r".
  from that[OF p = (r  s) @ k q = (s  r) @ l this primitive (r  s)]
  show ?thesis.
qed

lemma fac_two_conjug_primroot:
  assumes facs: "w ≤f p@k" "w ≤f q@l" and nemps: "p  ε" "q  ε" and len: "|p| + |q|  |w|"
  obtains r s m where "ρ p  r  s" and "ρ q  r  s" and "w = (r  s)@m  r" and "primitive (rs)"
proof -
  obtain p' where "w <p p'w" "p  p'" "p'  ε"
    using conjug_nemp_iff fac_pow_pref_conjug[OF facs(1)] nemps(1) per_rootI' by metis
  obtain q' where "w <s wq'" "q  q'" "q'  ε"
    using fac_pow_pref_conjug[reversed, OF w ≤f q@l] conjug_nemp_iff  nemps(2) per_rootI'[reversed] by metis
  from per_lemma_pref_suf[OF w <p p'w w <s wq']
  obtain r s k l m where
    "p' = (r  s) @ k" and
    "q' = (s  r) @ l" and
    "w = (r  s) @ m  r" and
    "primitive (r  s)"
    using len[unfolded conjug_len[OF p  p'] conjug_len[OF q  q']]
    by blast
  moreover have "ρ p' = rs"
    using p' = (r  s) @ k primitive (r  s) p'  ε primroot_unique by blast
  hence "ρ p  rs"
    using conjug_primroot[OF p  p']
    by simp
  moreover have "ρ q' = sr"
    using q' = (s  r) @ l primitive (r  s)[unfolded conjug_prim_iff'[of r]] q'  ε primroot_unique by blast
  hence "ρ q  sr"
    using conjug_primroot[OF q  q']  by simp
  hence "ρ q  rs"
    using conjug_trans[OF _ conjugI']
    by meson
  ultimately show ?thesis
    using that by blast
qed

corollary fac_two_conjug_primroot':
   assumes facs: "u ≤f r@k" "u ≤f s@l" and nemps: "r  ε" "s  ε" and len: "|r| + |s|  |u|"
   shows "ρ r  ρ s"
  using fac_two_conjug_primroot[OF assms] conjug_trans[OF _ conjug_sym[of "ρ s"]]
  by metis

lemma fac_two_conjug_primroot'':
  assumes facs: "u ≤f r@k" "u ≤f s@l" and "u  ε" and len: "|r| + |s|  |u|"
  shows "ρ r  ρ s"
proof -
  have nemps: "r  ε" "s  ε" using facs u  ε by auto
  show "conjugate (ρ r) (ρ s)" using fac_two_conjug_primroot'[OF facs nemps len].
qed

lemma  fac_two_prim_conjug:
  assumes "w ≤f u@n" "w ≤f v@m" "primitive u" "primitive v" "|u| + |v|  |w|"
  shows "u  v"
  using fac_two_conjug_primroot'[OF assms(1-2) _ _ assms(5)] prim_nemp[OF primitive u] prim_nemp[OF primitive v]
  unfolding prim_self_root[OF primitive u] prim_self_root[OF primitive v].

lemma fac_pow_conjug_primroot: assumes "u@k ≤f v@l" and "|u@k|  2*|v|" and "2  k" and "u  ε"
  shows "ρ u  ρ v"
proof(rule fac_two_conjug_primroot''[OF _ assms(1)], force)
  have "0 < k"
     using 2  k by linarith
  show "|u| + |v|  |u @ k|"
  proof(cases "|u|" "|v|" rule: le_cases)
    assume "|u|  |v|"
    thus ?thesis
      using assms(2) by linarith
  next
    assume "|v|  |u|"
    hence " |u| + |v|  2*|u|"
      by simp
    thus ?thesis
      unfolding pow_len
      using mult_le_mono1[OF 2  k] le_trans
      by blast
  qed
  show "u @ k  ε"
    using u  ε 0 < k by blast
qed

section ‹Testing primitivity›

text‹This section defines a proof method used to prove that a word is primitive.›

lemma primitive_iff [code]: "primitive w  ¬ w ≤f tl w  butlast w"
proof-
  have "¬ primitive w  w ≤f tl w  butlast w"
  proof
    assume "¬ primitive w"
    then obtain r k where "k  1" and "w = r@k"
      unfolding primitive_def by blast
    show "w ≤f tl w  butlast w"
    proof (cases "w = ε")
      assume "w  ε"
      from this[unfolded w = r@k]
      have  "0 < k"
        using nemp_pow by blast
      have "r  ε"
        using pow_zero r @ k  ε by force
      have "r@(k-1)  ε"
        unfolding nemp_emp_pow[OF r  ε, of "k-1"]
        using 0 < k k  1 by force
      have "r  w  r@(k-1) = w  w"
        unfolding w = r@k pows_comm[of r k "k - 1"]
        unfolding lassoc cancel_right pow_pos[OF 0 < k]..
      hence "[hd r]  tl r  w  butlast (r@(k-1))  [last (r@(k-1))] = [hd w]  tl w  butlast w  [last w]"
        unfolding hd_tl[reversed, OF r@(k-1)  ε] hd_tl[reversed, OF w  ε]
        unfolding lassoc hd_tl[OF r  ε] hd_tl[OF w  ε].
      hence "tl r  w  butlast (r@(k-1)) = tl w  butlast w"
        by force
      thus ?thesis
        unfolding fac_def by metis
    qed simp
  next
    assume "w ≤f tl w  butlast w"
    show "¬ primitive w"
    proof (cases "w = ε")
      assume "w  ε"
      from facE[OF w ≤f tl w  butlast w]
      obtain p s where "tl w  butlast w = p  w  s".
      have "[hd w]  (p  w  s)  [last w] = w  w"
        unfolding tl w  butlast w = p  w  s[symmetric]
        unfolding lassoc hd_tl[OF w  ε]
        unfolding rassoc hd_tl[reversed, OF w  ε]..
      from prim_overlap_sqE[of w "[hd w]  p" "s  [last w]" False, unfolded rassoc, OF _  this[unfolded rassoc]]
      show "¬ primitive w"
        by blast
    qed simp
  qed
  thus ?thesis by blast
qed

method primitivity_inspection =  (insert method_facts, use nothing in
    simp add: primitive_iff pow_pos)

― ‹This is out of scope of the method, and has to be proved separately›
lemma alternate_prim: assumes "x  y" shows "primitive ([x,y]@n[x])"
proof-
  consider "n = 0" | "n = 1" | "2  n" by linarith
  then show ?thesis
  proof(cases)
   assume "2  n"
   have pref: "[x, y] @ n  [x] ≤p [x, y]  [x, y] @ n  [x]"
     by comparison
   have neq: "([x, y] @ n  [x])  [x, y]  [x, y]  [x, y] @ n  [x]"
     using  x  y by force
   then show ?thesis
     using per_le_prim_iff[of "[x,y]@n[x]" "[x,y]", OF pref] 2  n
     unfolding lenmorph pow_len
     by fastforce
 qed (simp_all add: x  y primitive_iff)
qed








end