Theory KMP

theory KMP
imports Coinductive_Language
theory KMP imports
  "../Coinductive_Languages/Coinductive_Language"
begin

primrec match where
  "match L [] = 𝔬 L"
| "match L (x # xs) = (𝔬 L ∨ match (𝔡 L x) xs)"

lemma match_alt: "match L w = (∃p s. w = p @ s ∧ in_language L p)"
  apply (induct w arbitrary: L)
  apply (auto simp: Cons_eq_append_conv intro: exI[of _ "[]"] exI[of _ "_ # _"])
  apply (metis in_language.simps(2))
  done

corec tab where
  "tab xs failure =
    Lang (xs = []) (λc. if xs = [] ∨ hd xs ≠ c then failure c else tab (tl xs) (𝔡 (failure c)))"

lemma tab_simps [simp]:
  "tab [] failure = Lang True failure"
  "tab (x # xs) failure = Lang False (λc. if x ≠ c then failure c else tab xs (𝔡 (failure x)))"
by(subst tab.code; simp add: fun_eq_iff; fail)+

lemma mk_table_sel [simp]:
  "𝔬 (tab xs failure) = (xs = [])"
  "𝔡 (tab xs failure) = (λc. if xs = [] ∨ hd xs ≠ c then failure c else tab (tl xs) (𝔡 (failure c)))"
  by (subst tab.code; simp)+

definition aux :: "'a list ⇒ ('a ⇒ 'a language) ⇒ 'a ⇒ 'a language"
where "aux xs failure = 𝔡 (tab xs failure)"

friend_of_corec aux
where "aux xs failure c = Lang
   (if xs = [] ∨ hd xs ≠ c then 𝔬 (failure c) else tl xs = [])
   (if xs = [] ∨ hd xs ≠ c then 𝔡 (failure c) else (aux (tl xs) (𝔡 (failure c))))"
subgoal unfolding aux_def mk_table_sel by(subst tab.code; simp)
apply (auto simp: rel_fun_def rel_prod.simps split: prod.split; metis fst_conv snd_conv)
done

context fixes xs :: "'a list" begin

corec table :: "'a language"
where "table = Lang (xs = []) (aux xs (λ_. table))"

lemma table: "table = tab xs (λ_. table)"
by(subst (1) table.code)(rule language.expand; simp add: aux_def)

end

definition is_substring_of :: "'a list ⇒ 'a list ⇒ bool"
where "is_substring_of xs ys = match (table xs) ys"

export_code is_substring_of in Haskell module_name KMP file code

end