Theory Myhill_1
theory Myhill_1
imports Folds
"HOL-Library.While_Combinator"
begin
section ‹First direction of MN: ‹finite partition ⇒ regular language››
notation
conc (infixr ‹⋅› 100) and
star (‹_⋆› [101] 102)
lemma Pair_Collect [simp]:
shows "(x, y) ∈ {(x, y). P x y} ⟷ P x y"
by simp
text ‹Myhill-Nerode relation›
definition
str_eq :: "'a lang ⇒ ('a list × 'a list) set" (‹≈_› [100] 100)
where
"≈A ≡ {(x, y). (∀z. x @ z ∈ A ⟷ y @ z ∈ A)}"
abbreviation
str_eq_applied :: "'a list ⇒ 'a lang ⇒ 'a list ⇒ bool" (‹_ ≈_ _›)
where
"x ≈A y ≡ (x, y) ∈ ≈A"
lemma str_eq_conv_Derivs:
"str_eq A = {(u,v). Derivs u A = Derivs v A}"
by (auto simp: str_eq_def Derivs_def)
definition
finals :: "'a lang ⇒ 'a lang set"
where
"finals A ≡ {≈A `` {s} | s . s ∈ A}"
lemma lang_is_union_of_finals:
shows "A = ⋃(finals A)"
unfolding finals_def
unfolding Image_def
unfolding str_eq_def
by (auto) (metis append_Nil2)
lemma finals_in_partitions:
shows "finals A ⊆ (UNIV // ≈A)"
unfolding finals_def quotient_def
by auto
subsection ‹Equational systems›
text ‹The two kinds of terms in the rhs of equations.›