Theory MSO

Up to index of Isabelle/HOL/MSO

theory MSO
imports Regular_Exp RBT_Set
(* Author: Dmitriy Traytel *)

header {* Monadic Second-Order Logic on Finite Words *}

(*<*)
theory MSO
imports Regular_Exp "~~/src/HOL/Library/Product_ord" "~~/src/HOL/Library/RBT_Set"
begin

declare set.simps [code]
hide_const (open) RBT.map RBT.fold
(*>*)

subsection {* Library Functions *}

primrec map_index' :: "nat => (nat => 'a => 'b) => 'a list => 'b list" where
"map_index' n f [] = []"
| "map_index' n f (x#xs) = f n x # map_index' (Suc n) f xs"

lemma length_map_index'[simp]: "length (map_index' n f xs) = length xs"
by (induct xs arbitrary: n) auto

lemma map_index'_map_zip: "map_index' n f xs = map (split f) (zip [n ..< n + length xs] xs)"
proof (induct xs arbitrary: n)
case (Cons x xs)
hence "map_index' n f (x#xs) = f n x # map (split f) (zip [Suc n ..< n + length (x # xs)] xs)" by simp
also have "… = map (split f) (zip (n # [Suc n ..< n + length (x # xs)]) (x # xs))" by simp
also have "(n # [Suc n ..< n + length (x # xs)]) = [n ..< n + length (x # xs)]" by (induct xs) auto
finally show ?case by simp
qed simp

abbreviation "map_index ≡ map_index' 0"

lemmas map_index = map_index'_map_zip[of 0, simplified]

lemma take_map_index: "take p (map_index f xs) = map_index f (take p xs)"
unfolding map_index by (auto simp: min_def take_map take_zip)

lemma drop_map_index: "drop p (map_index f xs) = map_index' p f (drop p xs)"
unfolding map_index'_map_zip by (cases "p < length xs") (auto simp: drop_map drop_zip)

lemma map_map_index[simp]: "map g (map_index f xs) = map_index (λn x. g (f n x)) xs"
unfolding map_index by auto

lemma map_index_map[simp]: "map_index f (map g xs) = map_index (λn x. f n (g x)) xs"
unfolding map_index by (auto simp: map_zip_map2)

lemma set_map_index[simp]: "x ∈ set (map_index f xs) = (∃i < length xs. f i (xs ! i) = x)"
unfolding map_index by (auto simp: set_zip intro!: image_eqI[of _ "split f"])

lemma nth_map_index[simp]: "p < length xs ==> map_index f xs ! p = f p (xs ! p)"
unfolding map_index by auto

lemma map_index_cong:
"∀p < length xs. f p (xs ! p) = g p (xs ! p) ==> map_index f xs = map_index g xs"
unfolding map_index by (auto simp: set_zip)

lemma map_index_id: "map_index (curry snd) xs = xs"
unfolding map_index by auto

lemma map_index_no_index[simp]: "map_index (λn x. f x) xs = map f xs"
unfolding map_index by (induct xs rule: rev_induct) auto

lemma map_index_congL:
"∀p < length xs. f p (xs ! p) = xs ! p ==> map_index f xs = xs"
by (rule trans[OF map_index_cong map_index_id]) auto

lemma map_index'_is_NilD: "map_index' n f xs = [] ==> xs = []"
by (induct xs) auto

declare map_index'_is_NilD[of 0, dest!]

lemma map_index'_is_ConsD:
"map_index' n f xs = y # ys ==> ∃z zs. xs = z # zs ∧ f n z = y ∧ map_index' (n + 1) f zs = ys"
by (induct xs arbitrary: n) auto

lemma map_index'_eq_imp_length_eq: "map_index' n f xs = map_index' n g ys ==> length xs = length ys"
proof (induct ys arbitrary: xs n)
case (Cons y ys) thus ?case by (cases xs) auto
qed (auto dest!: map_index'_is_NilD)

lemmas map_index_eq_imp_length_eq = map_index'_eq_imp_length_eq[of 0]

lemma map_index'_comp[simp]: "map_index' n f (map_index' n g xs) = map_index' n (λn. f n o g n) xs"
by (induct xs arbitrary: n) auto

primrec combinatorial_product :: "'a list list => 'a list list" where
"combinatorial_product [] = [[]]"
| "combinatorial_product (xs # xss) = List.maps (λx. map (Cons x) (combinatorial_product xss)) xs"

abbreviation "bool_combinatorial_product n ≡ combinatorial_product (replicate n [True, False])"

lemma set_bool_combinatorial_product[simp]:
"bs ∈ set (bool_combinatorial_product n) <-> length bs = n"
proof (induct bs arbitrary: n)
case Nil thus ?case by (cases n) (auto simp: maps_def)
next
case (Cons b bs) thus ?case by (cases n) (auto simp: maps_def)
qed

fun insert_nth :: "nat => 'a => 'a list => 'a list" where
"insert_nth 0 x xs = x # xs"
| "insert_nth (Suc n) x (y # xs) = y # insert_nth n x xs"
| "insert_nth _ x [] = [x]"

lemma insert_nth_take_drop[simp]: "insert_nth n x xs = take n xs @ [x] @ drop n xs"
by (induct n x xs rule: insert_nth.induct) auto

lemma length_insert_nth: "length (insert_nth n x xs) = Suc (length xs)"
by (induct xs) auto

subsection {* Interpretations and Encodings *}

type_synonym 'a interp = "'a list × (nat + nat set) list"

abbreviation "enc_atom_bool I n ≡ map (λx. case x of Inl p => n = p | Inr P => n ∈ P) I"

abbreviation "enc_atom I n a ≡ (a, enc_atom_bool I n)"

fun enc :: "'a interp => ('a × bool list) list" where
"enc (w, I) = map_index (enc_atom I) w"

subsection {* Syntax and Semantics of MSO *}

datatype 'a formula =
Q 'a nat
| Less nat nat
| In nat nat
| Not "'a formula"
| Or "'a formula" "'a formula"
| And "'a formula" "'a formula"
| Exists "'a formula"
| EXISTS "'a formula"

definition "σ = (λΣ n. concat (map (λbs. map (λa. (a, bs)) Σ) (List.n_lists n [True, False])))"
definition "π = (λ(a, bs). (a, tl bs))"

locale MSO = project "set o (σ Σ)" π for Σ :: "'a :: linorder list"
begin

abbreviation "Σ_combinatorial_product n ≡
List.maps (λbools. map (λa. (a, bools)) Σ) (bool_combinatorial_product n)"


primrec FOV :: "'a formula => nat set" where
"FOV (Q a m) = {m}"
| "FOV (Less m1 m2) = {m1, m2}"
| "FOV (In m M) = {m}"
| "FOV (Not φ) = FOV φ"
| "FOV (Or φ1 φ2) = FOV φ1 ∪ FOV φ2"
| "FOV (And φ1 φ2) = FOV φ1 ∪ FOV φ2"
| "FOV (Exists φ) = (λx. x - 1) ` (FOV φ - {0})"
| "FOV (EXISTS φ) = (λx. x - 1) ` FOV φ"

primrec SOV :: "'a formula => nat set" where
"SOV (Q a m) = {}"
| "SOV (Less m1 m2) = {}"
| "SOV (In m M) = {M}"
| "SOV (Not φ) = SOV φ"
| "SOV (Or φ1 φ2) = SOV φ1 ∪ SOV φ2"
| "SOV (And φ1 φ2) = SOV φ1 ∪ SOV φ2"
| "SOV (Exists φ) = (λx. x - 1) ` SOV φ"
| "SOV (EXISTS φ) = (λx. x - 1) ` (SOV φ - {0})"

(* Normal form: quantified variables are used in the body *)
primrec pre_wf_formula :: "nat => 'a formula => bool" where
"pre_wf_formula n (Q a m) = (a ∈ set Σ ∧ m < n)"
| "pre_wf_formula n (Less m1 m2) = (m1 < n ∧ m2 < n)"
| "pre_wf_formula n (In m M) = (m < n ∧ M < n)"
| "pre_wf_formula n (Not φ) = pre_wf_formula n φ"
| "pre_wf_formula n (Or φ1 φ2) = (pre_wf_formula n φ1 ∧ pre_wf_formula n φ2)"
| "pre_wf_formula n (And φ1 φ2) = (pre_wf_formula n φ1 ∧ pre_wf_formula n φ2)"
| "pre_wf_formula n (Exists φ) = (pre_wf_formula (n + 1) φ ∧ 0 ∈ FOV φ ∧ 0 ∉ SOV φ)"
| "pre_wf_formula n (EXISTS φ) = (pre_wf_formula (n + 1) φ ∧ 0 ∉ FOV φ ∧ 0 ∈ SOV φ)"

abbreviation "closed ≡ pre_wf_formula 0"

abbreviation "wf_formula n φ ≡ pre_wf_formula n φ ∧ FOV φ ∩ SOV φ = {}"

lemma max_idx_vars: "pre_wf_formula n φ ==> ∀p ∈ FOV φ ∪ SOV φ. p < n"
by (induct φ arbitrary: n)
(auto split: split_if_asm, (metis Un_iff diff_Suc_less less_SucE less_imp_diff_less)+)

lemma finite_FOV: "finite (FOV φ)"
by (induct φ) (auto split: split_if_asm)

lemma finite_SOV: "finite (SOV φ)"
by (induct φ) (auto split: split_if_asm)

abbreviation "wf_interp w I ≡ (length w > 0 ∧
(∀a ∈ set w. a ∈ set Σ) ∧
(∀x ∈ set I. case x of Inl p => p < length w | Inr P => ∀p ∈ P. p < length w))"


fun wf_interp_for_formula :: "'a interp => 'a formula => bool" where
"wf_interp_for_formula (w, I) φ =
(wf_interp w I ∧
(∀n ∈ FOV φ. case I ! n of Inl _ => True | _ => False) ∧
(∀n ∈ SOV φ. case I ! n of Inl _ => False | _ => True))"


fun satisfies :: "'a interp => 'a formula => bool" (infix "\<Turnstile>" 50) where
"(w, I) \<Turnstile> Q a m = (w ! (case I ! m of Inl p => p) = a)"
| "(w, I) \<Turnstile> Less m1 m2 = ((case I ! m1 of Inl p => p) < (case I ! m2 of Inl p => p))"
| "(w, I) \<Turnstile> In m M = ((case I ! m of Inl p => p) ∈ (case I ! M of Inr P => P))"
| "(w, I) \<Turnstile> Not φ = (¬ (w, I) \<Turnstile> φ)"
| "(w, I) \<Turnstile> (Or φ1 φ2) = ((w, I) \<Turnstile> φ1 ∨ (w, I) \<Turnstile> φ2)"
| "(w, I) \<Turnstile> (And φ1 φ2) = ((w, I) \<Turnstile> φ1 ∧ (w, I) \<Turnstile> φ2)"
| "(w, I) \<Turnstile> (Exists φ) = (∃p. p ∈ {0 .. length w - 1} ∧ (w, Inl p # I) \<Turnstile> φ)"
| "(w, I) \<Turnstile> (EXISTS φ) = (∃P. P ⊆ {0 .. length w - 1} ∧ (w, Inr P # I) \<Turnstile> φ)"

definition langMSO :: "nat => 'a formula => ('a × bool list) list set" where
"langMSO n φ = {enc (w, I) | w I.
length I = n ∧ wf_interp_for_formula (w, I) φ ∧ satisfies (w, I) φ}"


subsection {* ENC *}

definition "arbitrary_except n pbs xs ≡
PLUS (map (λbs. PLUS (map (λx. Atom (x, fold (λ(p, b). insert_nth p b) pbs bs)) xs))
(bool_combinatorial_product (n - length pbs)))"


lemma length_fold_insert_nth:
"length (fold (λ(p, b). insert_nth p b) pbs bs) = length bs + length pbs"
by (induct pbs arbitrary: bs) auto

lemma invar_fold_insert_nth:
"[|∀x∈set pbs. p < fst x; p < length bs; bs ! p = b|] ==>
fold (λ(x, y). insert_nth x y) pbs bs ! p = b"

by (induct pbs arbitrary: bs) (auto simp: nth_append)

lemma nth_fold_insert_nth:
"[|sorted (map fst pbs); distinct (map fst pbs); ∀(p, b) ∈ set pbs. p < length bs + length pbs;
i < length pbs; pbs ! i = (p, b)|] ==>
fold (λ(p, b). insert_nth p b) pbs bs ! p = b"

proof (induct pbs arbitrary: bs i p b)
case (Cons pb pbs)
show ?case
proof (cases i)
case 0
with Cons.prems have "p < Suc (length bs)"
proof (induct pbs rule: rev_induct)
case (snoc pb' pbs)
then obtain p' b' where "pb' = (p', b')" by auto
with snoc.prems have "∀p ∈ fst ` set pbs. p < p'" "p' ≤ Suc (length bs + length pbs)"
by (auto simp: image_iff sorted_Cons sorted_append le_eq_less_or_eq)
with snoc.prems show ?case by (intro snoc(1)) (auto simp: sorted_Cons sorted_append)
qed auto
with 0 Cons.prems show ?thesis unfolding fold.simps o_apply
by (intro invar_fold_insert_nth) (auto simp: sorted_Cons image_iff le_eq_less_or_eq nth_append)
next
case (Suc n) with Cons.prems show ?thesis unfolding fold.simps
by (auto intro!: Cons(1) simp: sorted_Cons)
qed
qed simp

lemma wf_rexp_arbitrary_except:
"[|length pbs ≤ n; set xs ⊆ set Σ|] ==> wf n (arbitrary_except n pbs xs)"
by (auto simp: arbitrary_except_def wf_PLUS) (force simp: σ_def set_n_lists length_fold_insert_nth)

definition valid_ENC :: "nat => nat => ('a × bool list) rexp" where
"valid_ENC n p = (if n = 0 then Regular_Exp.Not Zero else
TIMES [
Star (arbitrary_except n [(p, False)] Σ),
arbitrary_except n [(p, True)] Σ,
Star (arbitrary_except n [(p, False)] Σ)])"


lemma wf_rexp_valid_ENC: "wf n (valid_ENC n p)"
unfolding valid_ENC_def by (auto intro!: wf_rexp_arbitrary_except)

definition ENC :: "nat => 'a formula => ('a × bool list) rexp" where
"ENC n φ = flatten INTERSECT (valid_ENC n ` FOV φ)"

lemma wf_rexp_ENC: "wf n (ENC n φ)"
unfolding ENC_def
by (intro iffD2[OF wf_flatten_INTERSECT]) (auto intro: finite_FOV simp: wf_rexp_valid_ENC)

lemma enc_atom_σ_eq: "i < length w ==>
(length I = n ∧ w ! i ∈ set Σ) <-> enc_atom I i (w ! i) ∈ set (σ Σ n)"

by (auto simp: σ_def set_n_lists intro!: exI[of _ "enc_atom_bool I i"] imageI)

lemmas enc_atom_σ = iffD1[OF enc_atom_σ_eq, OF _ conjI]

lemma enc_atom_bool_take_drop_True:
"[|r < length I; case I ! r of Inl p' => p = p' | Inr P => p ∈ P|] ==>
enc_atom_bool I p = take r (enc_atom_bool I p) @ True # drop (Suc r) (enc_atom_bool I p)"

by (auto intro: trans[OF id_take_nth_drop])

lemma enc_atom_bool_take_drop_True2:
"[|r < length I; case I ! r of Inl p' => p = p' | Inr P => p ∈ P;
s < length I; case I ! s of Inl p' => p = p' | Inr P => p ∈ P; r < s|] ==>
enc_atom_bool I p = take r (enc_atom_bool I p) @ True #
take (s - Suc r) (drop (Suc r) (enc_atom_bool I p)) @ True #
drop (Suc s) (enc_atom_bool I p)"

using id_take_nth_drop[of r "enc_atom_bool I p"]
id_take_nth_drop[of "s - r - 1" "drop (Suc r) (enc_atom_bool I p)"]
by auto

lemma enc_atom_bool_take_drop_False:
"[|r < length I; case I ! r of Inl p' => p ≠ p' | Inr P => p ∉ P|] ==>
enc_atom_bool I p = take r (enc_atom_bool I p) @ False # drop (Suc r) (enc_atom_bool I p)"

by (auto intro: trans[OF id_take_nth_drop] split: sum.splits)

lemma enc_atom_lang_arbitrary_except_True: "[|r < length I;
case I ! r of Inl p' => p = p' | Inr P => p ∈ P; length I = n; a ∈ set xs; set xs ⊆ set Σ|] ==>
[enc_atom I p a] ∈ lang n (arbitrary_except n [(r, True)] xs)"

unfolding arbitrary_except_def
by (auto intro!: bexI[of _ "(λJ. take r J @ drop (r + 1) J) (enc_atom_bool I p)"]
intro: enc_atom_bool_take_drop_True)

lemma enc_atom_lang_arbitrary_except_True2:"[|r < length I; s < length I; r < s;
case I ! r of Inl p' => p = p' | Inr P => p ∈ P;
case I ! s of Inl p' => p = p' | Inr P => p ∈ P; length I = n; a ∈ set Σ|] ==>
[enc_atom I p a] ∈ lang n (arbitrary_except n [(r, True), (s, True)] Σ)"

unfolding arbitrary_except_def
by (auto intro!: bexI[of _ "(λJ. take r J @ take (s - r - 1) (drop (r + 1) J) @ drop (s - r) (drop (r + 1) J)) (enc_atom_bool I p)"]
intro!: enc_atom_bool_take_drop_True2 simp: min_absorb2 take_Cons' drop_Cons')

lemma enc_atom_lang_arbitrary_except_False:"[|r < length I;
case I ! r of Inl p' => p ≠ p' | Inr P => p ∉ P; length I = n; a ∈ set xs; set xs ⊆ set Σ|] ==>
[enc_atom I p a] ∈ lang n (arbitrary_except n [(r, False)] xs)"

unfolding arbitrary_except_def
by (auto intro!: bexI[of _ "(λJ. take r J @ drop (r + 1) J) (enc_atom_bool I p)"]
intro: enc_atom_bool_take_drop_False)

lemma arbitrary_except:
"[|v ∈ lang n (arbitrary_except n [(r, b)] S); r < n; set S ⊆ set Σ|] ==>
∃x. v = [x] ∧ snd x ! r = b ∧ fst x ∈ set S"

unfolding arbitrary_except_def by (auto simp: nth_append)

lemma arbitrary_except2:
"[|v ∈ lang n (arbitrary_except n [(r, b), (s, b')] S); r < s; r < n; s < n; set S ⊆ set Σ|] ==>
∃x. v = [x] ∧ snd x ! r = b ∧ snd x ! s = b' ∧ fst x ∈ set S"

unfolding arbitrary_except_def by (auto simp: nth_append min_absorb2)

lemma star_arbitrary_except:
"[|v ∈ star (lang n (arbitrary_except n [(r, b)] Σ)); r < n; i < length v|] ==>
snd (v ! i) ! r = b"

proof (induct arbitrary: i rule: star_induct)
case (append u v) thus ?case by (cases i) (auto dest: arbitrary_except)
qed simp

definition "dec_word ≡ map fst"

definition "positions_in_row w i =
Option.these (set (map_index (λp a_bs. if nth (snd a_bs) i then Some p else None) w))"


definition "dec_interp n FO w ≡ map (λi.
if i ∈ FO
then Inl (the_elem (positions_in_row w i))
else Inr (positions_in_row w i)) [0..<n]"


lemma positions_in_row: "positions_in_row w i = {p. p < length w ∧ snd (w ! p) ! i}"
unfolding positions_in_row_def these_def by (auto intro!: image_eqI[of _ the])

lemma positions_in_row_unique: "∃!p. p < length w ∧ snd (w ! p) ! i ==>
the_elem (positions_in_row w i) = (THE p. p < length w ∧ snd (w ! p) ! i)"

by (rule the1I2) (auto simp: the_elem_def positions_in_row)

lemma positions_in_row_length: "∃!p. p < length w ∧ snd (w ! p) ! i ==>
the_elem (positions_in_row w i) < length w"

unfolding positions_in_row_unique by (rule the1I2) auto

lemma positions_in_row_nth: "∃!p. p < length w ∧ snd (w ! p) ! i ==>
snd (w ! the_elem (positions_in_row w i)) ! i"

unfolding positions_in_row_unique by (rule the1I2) auto

lemma dec_interp_Inl: "[|i ∈ FO; i < n|] ==> ∃p. dec_interp n FO x ! i = Inl p"
unfolding dec_interp_def using nth_map[of n "[0..<n]"] by auto

lemma dec_interp_not_Inr: "[|dec_interp n FO x ! i = Inr P; i ∈ FO; i < n|] ==> False"
unfolding dec_interp_def using nth_map[of n "[0..<n]"] by auto

lemma dec_interp_Inr: "[|i ∉ FO; i < n|] ==> ∃P. dec_interp n FO x ! i = Inr P"
unfolding dec_interp_def using nth_map[of n "[0..<n]"] by auto

lemma dec_interp_not_Inl: "[|dec_interp n FO x ! i = Inl p; i ∉ FO; i < n|] ==> False"
unfolding dec_interp_def using nth_map[of n "[0..<n]"] by auto

lemma Inl_dec_interp_length:
assumes "∀i ∈ FO. ∃!p. p < length w ∧ snd (w ! p) ! i"
shows "Inl p ∈ set (dec_interp n FO w) ==> p < length w"
unfolding dec_interp_def by (auto intro: positions_in_row_length[OF bspec[OF assms]])

lemma Inr_dec_interp_length: "[|Inr P ∈ set (dec_interp n FO w); p ∈ P|] ==> p < length w"
unfolding dec_interp_def by (auto simp: positions_in_row)

lemma enc_atom_dec:
"[|wf_word n w; ∀i ∈ FO. i < n --> (∃!p. p < length w ∧ snd (w ! p) ! i); p < length w|] ==>
enc_atom (dec_interp n FO w) p (fst (w ! p)) = w ! p"

unfolding wf_word dec_interp_def map_filter_def Let_def
apply (auto simp: comp_def intro!: trans[OF iffD2[OF Pair_eq] pair_collapse])
apply (intro nth_equalityI)
apply (auto simp add: σ_def set_n_lists dest!: nth_mem) []
apply (auto simp: positions_in_row)
apply (drule bspec)
apply assumption
apply (drule mp)
apply assumption
apply (frule positions_in_row_unique)
apply (erule notE)
apply (simp add: positions_in_row)
apply (erule the1I2)
apply simp

apply (drule bspec)
apply assumption
apply (drule mp)
apply assumption
apply (frule positions_in_row_unique)
apply (simp add: positions_in_row)
apply (rule the_equality[symmetric])
apply auto
done

lemma enc_dec:
"[|wf_word n w; ∀i ∈ FO. i < n --> (∃!p. p < length w ∧ snd (w ! p) ! i)|] ==>
enc (dec_word w, dec_interp n FO w) = w"

unfolding enc.simps dec_word_def
by (intro trans[OF map_index_map map_index_congL] allI impI enc_atom_dec) assumption+

lemma dec_word_enc: "dec_word (enc (w, I)) = w"
unfolding dec_word_def by auto

lemma enc_unique:
assumes "wf_interp w I" "i < length I"
shows "∃p. I ! i = Inl p ==> ∃!p. p < length (enc (w, I)) ∧ snd (enc (w, I) ! p) ! i"
proof (erule exE)
fix p assume "I ! i = Inl p"
with assms show ?thesis by (auto simp: map_index all_set_conv_all_nth intro!: exI[of _ p])
qed

lemma dec_interp_enc_Inl:
"[|dec_interp n FO (enc (w, I)) ! i = Inl p'; I ! i = Inl p; i ∈ FO; i < n; length I = n; p < length w; wf_interp w I|] ==>
p = p'"

unfolding dec_interp_def using nth_map[of _ "[0..<n]"] positions_in_row_unique[OF enc_unique]
by (auto intro: sym[OF the_equality])

lemma dec_interp_enc_Inr:
"[|dec_interp n FO (enc (w, I)) ! i = Inr P'; I ! i = Inr P; i ∉ FO; i < n; length I = n; ∀p ∈ P. p < length w|] ==>
P = P'"

unfolding dec_interp_def positions_in_row by auto

lemma lang_ENC:
assumes "wf_formula n φ"
shows "lang n (ENC n φ) - {[]} = {enc (w, I) | w I . length I = n ∧ wf_interp_for_formula (w, I) φ}"
(is "?L = ?R")
proof (cases "FOV φ = {}")
case True with assms show ?thesis
apply (auto simp: ENC_def intro!: enc_atom_σ)
apply (intro exI conjI)
apply (rule trans[OF sym[OF enc_dec[of _ _ "{}"]] enc.simps])
apply (auto simp add: wf_word dec_word_def dec_interp_def)
apply (auto simp: σ_def positions_in_row set_n_lists split: sum.splits)
apply (drule trans[OF sym nth_map])
apply (auto intro: bspec[OF max_idx_vars])
done
next
case False
hence nonempty: "valid_ENC n ` FOV φ ≠ {}" by simp
have finite: "finite (valid_ENC n ` FOV φ)" by (rule finite_imageI[OF finite_FOV])
from False assms(1) have "0 < n" by (cases n) (auto split: dest!: max_idx_vars)
with wf_rexp_valid_ENC have wf_rexp: "∀x ∈ valid_ENC n ` FOV φ. wf n x" by auto
{ fix r w I assume "r ∈ FOV φ" and *: "length I = n" "wf_interp_for_formula (w, I) φ"
then obtain p where p: "I ! r = Inl p" by (cases "I ! r") auto
moreover from `r ∈ FOV φ` assms *(1) have "r < length I" by (auto dest!: max_idx_vars)
ultimately have "Inl p ∈ set I" by (auto simp add: in_set_conv_nth)
with *(2) have "p < length w" by auto
with *(2) obtain a where a: "w ! p = a" "a ∈ set Σ" by auto
from `r < length I` p *(1) `a ∈ set Σ`
have "[enc_atom I p a] ∈ lang n (arbitrary_except n [(r, True)] Σ)"
by (intro enc_atom_lang_arbitrary_except_True[OF _ _ _ _ subset_refl]) auto
moreover
from `r < length I` p *(1) `a ∈ set Σ` *(2) `p < length w`
have "take p (enc (w, I)) ∈ star (lang n (arbitrary_except n [(r, False)] Σ))"
by (auto simp: in_set_conv_nth intro!: Ball_starI enc_atom_lang_arbitrary_except_False) auto
moreover
from `r < length I` p *(1) `a ∈ set Σ` *(2) `p < length w`
have "drop (Suc p) (enc (w, I)) ∈ star (lang n (arbitrary_except n [(r, False)] Σ))"
by (auto simp: in_set_conv_nth intro!: Ball_starI enc_atom_lang_arbitrary_except_False) auto
ultimately have "take p (enc (w, I)) @ [enc_atom I p a] @ drop (p + 1) (enc (w, I)) ∈
lang n (valid_ENC n r)"
using `0 < n` unfolding valid_ENC_def by (auto simp del: append.simps)
with `p < length w` a have "enc (w, I) ∈ lang n (valid_ENC n r)"
using id_take_nth_drop[of p "enc (w, I)"] by auto
}
hence "?R ⊆ ?L" using lang_flatten_INTERSECT[OF finite nonempty wf_rexp] by (auto simp add: ENC_def)
moreover
{ fix x assume "x ∈ (\<Inter>r ∈ valid_ENC n ` FOV φ. lang n r)"
hence r: "∀r ∈ FOV φ. x ∈ lang n (valid_ENC n r)" by blast
have "length (dec_interp n (FOV φ) x) = n" unfolding dec_interp_def by simp
moreover
{ fix r assume "r ∈ FOV φ"
with assms have "r < n" using max_idx_vars[of n φ] by auto
from `r ∈ FOV φ` r obtain u v w where uvw: "x = u @ v @ w"
"u ∈ star (lang n (arbitrary_except n [(r, False)] Σ))"
"v ∈ lang n (arbitrary_except n [(r, True)] Σ)"
"w ∈ star (lang n (arbitrary_except n [(r, False)] Σ))" using `0 < n` unfolding valid_ENC_def by fastforce
hence "length u < length x" "!!i. i < length x --> snd (x ! i) ! r <-> i = length u"
by (auto simp: nth_append nth_Cons' split: split_if_asm
dest!: arbitrary_except[OF _ `r < n`]
dest: star_arbitrary_except[OF _ `r < n`, of u]
elim!: iffD1[OF star_arbitrary_except[OF _ `r < n`, of w False]]) auto
hence "∃!p. p < length x ∧ snd (x ! p) ! r" by auto
} note * = this
have x_wf_word: "wf_word n x" using wf_lang_wf_word[OF wf_rexp_valid_ENC] False r by auto
with * have "x = enc (dec_word x, dec_interp n (FOV φ) x)" by (intro sym[OF enc_dec]) auto
moreover
from * have "wf_interp_for_formula (dec_word x, dec_interp n (FOV φ) x) φ"
using r False x_wf_word[unfolded wf_word, unfolded σ_def] assms
apply (auto simp: dec_word_def split: sum.splits)
apply fastforce
using Inl_dec_interp_length[OF ballI] apply blast
using Inr_dec_interp_length apply blast
using dec_interp_Inl[OF _ bspec[OF max_idx_vars], of _ "FOV φ" n φ x] apply force
using dec_interp_Inr[OF _ bspec[OF max_idx_vars], of _ "FOV φ" n φ x] apply fastforce
done
ultimately have "x ∈ {enc (w, I) | w I. length I = n ∧ wf_interp_for_formula (w, I) φ}" by blast
}
with False lang_flatten_INTERSECT[OF finite nonempty wf_rexp] have "?L ⊆ ?R" by (auto simp add: ENC_def)
ultimately show ?thesis by blast
qed

subsection {* Welldefinedness of enc wrt. Models *}

lemma enc_alt_def:
"enc (w, x # I) = map_index (λn (a, bs). (a, (case x of Inl p => n = p | Inr P => n ∈ P) # bs)) (enc (w, I))"
by (auto simp: comp_def)

lemma enc_extend_interp: "enc (w, I) = enc (w', I') ==> enc (w, x # I) = enc (w', x # I')"
unfolding enc_alt_def by auto

lemma wf_interp_for_formula_Exists:
"[|wf_formula (length I) (Exists φ); w ≠ []|]==>
wf_interp_for_formula (w, I) (Exists φ) <-> (∀p < length w. wf_interp_for_formula (w, Inl p # I) φ)"

apply (clarsimp split: sum.splits split_if_asm)
apply safe
apply (metis (hide_lams) DiffI Suc_pred gr0I nth_Cons_0 nth_Cons_Suc singleton_iff sum.simps(4))
apply (metis diff_Suc_Suc diff_zero nat.exhaust nth_Cons_Suc)
apply (metis length_greater_0_conv)
apply (metis length_greater_0_conv)
apply (metis length_greater_0_conv)
apply (metis diff_Suc_Suc gr0_implies_Suc length_greater_0_conv minus_nat.diff_0 nth_Cons_Suc)
apply (auto simp: nth_Cons' split: split_if_asm)
done

lemma wf_interp_for_formula_any_Inl: "wf_interp_for_formula (w, Inl p # I) φ ==>
∀p < length w. wf_interp_for_formula (w, Inl p # I) φ"

by (auto simp: nth_Cons' split: split_if_asm)

lemma wf_interp_for_formula_EXISTS:
"[|wf_formula (length I) (EXISTS φ); w ≠ []|]==>
wf_interp_for_formula (w, I) (EXISTS φ) <-> (∀P ⊆ {0 .. length w - 1}. wf_interp_for_formula (w, Inr P # I) φ)"

apply (clarsimp split: sum.splits split_if_asm)
apply safe
apply (cases w)
apply auto [2]
apply (metis Suc_pred gr0I nth_Cons_Suc)
apply (metis (hide_lams) DiffI Suc_pred gr0I nth_Cons_0 nth_Cons_Suc singleton_iff sum.simps(4))
unfolding sym[OF length_greater_0_conv] nth_Cons' One_nat_def
apply auto [2]
apply (metis empty_subsetI)
apply (metis empty_subsetI)
apply (metis empty_subsetI neq0_conv)
done

lemma wf_interp_for_formula_any_Inr: "wf_interp_for_formula (w, Inr P # I) φ ==>
∀P ⊆ {0 .. length w - 1}. wf_interp_for_formula (w, Inr P # I) φ"

by (cases w) (auto simp: nth_Cons' split: sum.splits split_if_asm)

lemma enc_word_length: "enc (w, I) = enc (w', I') ==> length w = length w'"
by (auto elim: map_index_eq_imp_length_eq)

lemma enc_length:
assumes "w ≠ []" "enc (w, I) = enc (w', I')"
shows "length I = length I'"
using assms
length_map[of "(λx. case x of Inl p => 0 = p | Inr P => 0 ∈ P)" I]
length_map[of "(λx. case x of Inl p => 0 = p | Inr P => 0 ∈ P)" I']
by (induct rule: list_induct2[OF enc_word_length[OF assms(2)]]) auto

lemma wf_interp_for_formula_Or:
"[|wf_formula n (Or φ1 φ2); length I = n|] ==> wf_interp_for_formula (w, I) (Or φ1 φ2) =
(wf_interp_for_formula (w, I) φ1 ∧ wf_interp_for_formula (w, I) φ2)"

by auto

lemma wf_interp_for_formula_And:
"[|wf_formula n (And φ1 φ2); length I = n|] ==> wf_interp_for_formula (w, I) (And φ1 φ2) =
(wf_interp_for_formula (w, I) φ1 ∧ wf_interp_for_formula (w, I) φ2)"

by auto

lemma max_idx_Suc: "pre_wf_formula n φ ==> pre_wf_formula (Suc n) φ"
by (induct φ arbitrary: n) auto

lemma wf_formula_Suc: "wf_formula n φ ==> wf_formula (Suc n) φ"
by (auto simp: max_idx_Suc)

lemma enc_wf_interp:
assumes "wf_formula (length I) φ" "wf_interp_for_formula (w, I) φ"
shows "wf_interp_for_formula (dec_word (enc (w, I)), dec_interp (length I) (FOV φ) (enc (w, I))) φ"
(is "wf_interp_for_formula (_, ?dec) φ")
unfolding dec_word_enc
proof -
{ fix i assume i: "i ∈ FOV φ"
with assms(2) have "∃p. I ! i = Inl p" by (cases "I ! i") auto
with i assms have "∃!p. p < length (enc (w, I)) ∧ snd (enc (w, I) ! p) ! i"
by (intro enc_unique[of w I i]) (auto intro!: bspec[OF max_idx_vars] split: sum.splits)
} note * = this
have "∀x ∈ set ?dec. sum_case (λp. p < length w) (λP. ∀p∈P. p < length w) x"
proof (intro ballI, split sum.split, safe)
fix p assume "Inl p ∈ set ?dec"
thus "p < length w" using Inl_dec_interp_length[OF ballI[OF *]] by auto
next
fix p P assume "Inr P ∈ set ?dec" "p ∈ P"
thus "p < length w" using Inr_dec_interp_length by fastforce
qed
thus "wf_interp_for_formula (w, ?dec) φ"
using assms
dec_interp_Inl[of _ "FOV φ" "length I" "enc (w, I)", OF _ bspec[OF max_idx_vars]]
dec_interp_Inr[of _ "FOV φ" "length I" "enc (w, I)", OF _ bspec[OF max_idx_vars]]
by (fastforce split: sum.splits)
qed

lemma enc_welldef: "[|enc (w, I) = enc (w', I'); wf_formula (length I) φ;
wf_interp_for_formula (w, I) φ; wf_interp_for_formula (w', I') φ|] ==>
satisfies (w, I) φ <-> satisfies (w', I') φ"

proof (induction φ arbitrary: I I')
case (Q a m)
let ?dec = "λw I. (dec_word (enc (w, I)), dec_interp (length I) (FOV (Q a m)) (enc (w, I)))"
from Q(2,3) have "satisfies (w, I) (Q a m) = satisfies (?dec w I) (Q a m)"
unfolding dec_word_enc
using dec_interp_enc_Inl[of "length I" "{m}" w I m]
by (auto intro: nth_mem dest: dec_interp_not_Inr split: sum.splits) (metis nth_mem)+
moreover
from Q(3) have *: "w ≠ []" by simp
from Q(2,4) have "satisfies (w', I') (Q a m) = satisfies (?dec w' I') (Q a m)"
unfolding dec_word_enc enc_length[OF * Q(1)]
using dec_interp_enc_Inl[of "length I'" "{m}" w' I' m]
by (auto dest: dec_interp_not_Inr split: sum.splits) (metis nth_mem)+
ultimately show ?case unfolding Q(1) enc_length[OF * Q(1)] ..
next
case (Less m n)
let ?dec = "λw I. (dec_word (enc (w, I)), dec_interp (length I) (FOV (Less m n)) (enc (w, I)))"
from Less(2,3) have "satisfies (w, I) (Less m n) = satisfies (?dec w I) (Less m n)"
unfolding dec_word_enc
using dec_interp_enc_Inl[of "length I" "{m, n}" w I m] dec_interp_enc_Inl[of "length I" "{m, n}" w I n]
by (fastforce intro: nth_mem dest: dec_interp_not_Inr split: sum.splits)
moreover
from Less(3) have *: "w ≠ []" by simp
from Less(2,4) have "satisfies (w', I') (Less m n) = satisfies (?dec w' I') (Less m n)"
unfolding dec_word_enc enc_length[OF * Less(1)]
using dec_interp_enc_Inl[of "length I'" "{m, n}" w' I' m] dec_interp_enc_Inl[of "length I'" "{m, n}" w' I' n]
by (fastforce intro: nth_mem dest: dec_interp_not_Inr split: sum.splits)
ultimately show ?case unfolding Less(1) enc_length[OF * Less(1)] ..
next
case (In m M)
let ?dec = "λw I. (dec_word (enc (w, I)), dec_interp (length I) (FOV (In m M)) (enc (w, I)))"
from In(2,3) have "satisfies (w, I) (In m M) = satisfies (?dec w I) (In m M)"
unfolding dec_word_enc
using dec_interp_enc_Inl[of "length I" "{m}" w I m] dec_interp_enc_Inr[of "length I" "{m}" w I M]
by (auto dest: dec_interp_not_Inr dec_interp_not_Inl split: sum.splits) (metis nth_mem)+
moreover
from In(3) have *: "w ≠ []" by simp
from In(2,4) have "satisfies (w', I') (In m M) = satisfies (?dec w' I') (In m M)"
unfolding dec_word_enc enc_length[OF * In(1)]
using dec_interp_enc_Inl[of "length I'" "{m}" w' I' m] dec_interp_enc_Inr[of "length I'" "{m}" w' I' M]
by (auto dest: dec_interp_not_Inr dec_interp_not_Inl split: sum.splits) (metis nth_mem)+
ultimately show ?case unfolding In(1) enc_length[OF * In(1)] ..
next
case (Or φ1 φ2) show ?case unfolding satisfies.simps(5)
proof (intro disj_cong)
from Or(3-6) show "satisfies (w, I) φ1 = satisfies (w', I') φ1"
by (intro Or(1)) auto
next
from Or(3-6) show "satisfies (w, I) φ2 = satisfies (w', I') φ2"
by (intro Or(2)) auto
qed
next
case (And φ1 φ2) show ?case unfolding satisfies.simps(6)
proof (intro conj_cong)
from And(3-6) show "satisfies (w, I) φ1 = satisfies (w', I') φ1"
by (intro And(1)) auto
next
from And(3-6) show "satisfies (w, I) φ2 = satisfies (w', I') φ2"
by (intro And(2)) auto
qed
next
case (Exists φ)
hence "w ≠ []" "w' ≠ []" by auto
hence length: "length w = length w'" "length I = length I'"
using enc_word_length[OF Exists.prems(1)] enc_length[OF _ Exists.prems(1)] by auto
show ?case
proof
assume "satisfies (w, I) (Exists φ)"
with Exists.prems(3) obtain p where "p < length w" "satisfies (w, Inl p # I) φ"
by (auto intro: ord_less_eq_trans[OF le_imp_less_Suc Suc_pred])
moreover
with Exists.prems have "satisfies (w', Inl p # I') φ"
apply (intro iffD1[OF Exists.IH[of "Inl p # I" "Inl p # I'"]])
apply (elim enc_extend_interp)
apply (auto split: sum.splits split_if_asm) []
apply (blast dest!: wf_interp_for_formula_Exists[OF _ `w ≠ []`])
apply (blast dest!: wf_interp_for_formula_Exists[OF _ `w' ≠ []`, of I', unfolded length[symmetric]])
apply assumption
done
ultimately show "satisfies (w', I') (Exists φ)" using length by (auto intro!: exI[of _ p])
next
assume "satisfies (w', I') (Exists φ)"
with Exists.prems(1,2,4) obtain p where "p < length w'" "satisfies (w', Inl p # I') φ"
by (auto intro: ord_less_eq_trans[OF le_imp_less_Suc Suc_pred])
moreover
with Exists.prems have "satisfies (w, Inl p # I) φ"
apply (intro iffD2[OF Exists.IH[of "Inl p # I" "Inl p # I'"]])
apply (elim enc_extend_interp)
apply (auto split: sum.splits split_if_asm) []
apply (blast dest!: wf_interp_for_formula_Exists[OF _ `w ≠ []`, of I, unfolded length(1)])
apply (blast dest!: wf_interp_for_formula_Exists[OF _ `w' ≠ []`, of I', unfolded length(2)[symmetric]])
apply assumption
done
ultimately show "satisfies (w, I) (Exists φ)" using length by (auto intro!: exI[of _ p])
qed
next
case (EXISTS φ)
hence "w ≠ []" "w' ≠ []" by auto
hence length: "length w = length w'" "length I = length I'"
using enc_word_length[OF EXISTS.prems(1)] enc_length[OF _ EXISTS.prems(1)] by auto
show ?case
proof
assume "satisfies (w, I) (EXISTS φ)"
then obtain P where "P ⊆ {0 .. length w - 1}" "satisfies (w, Inr P # I) φ" by auto
moreover
with EXISTS.prems have "satisfies (w', Inr P # I') φ"
apply (intro iffD1[OF EXISTS.IH[of "Inr P # I" "Inr P # I'"]])
apply (elim enc_extend_interp)
apply (auto split: sum.splits split_if_asm) []
apply (blast dest!: wf_interp_for_formula_EXISTS[OF _ `w ≠ []`])
apply (blast dest!: wf_interp_for_formula_EXISTS[OF _ `w' ≠ []`, of I', unfolded length[symmetric]])
apply assumption
done
ultimately show "satisfies (w', I') (EXISTS φ)" using length by (auto intro!: exI[of _ P])
next
assume "satisfies (w', I') (EXISTS φ)"
then obtain P where P: "P ⊆ {0 .. length w' - 1}" "satisfies (w', Inr P # I') φ" by auto
moreover
with EXISTS.prems have "satisfies (w, Inr P # I) φ"
apply (intro iffD2[OF EXISTS.IH[of "Inr P # I" "Inr P # I'"]])
apply (elim enc_extend_interp)
apply (auto split: sum.splits split_if_asm) []
apply (blast dest!: wf_interp_for_formula_EXISTS[OF _ `w ≠ []`, of I, unfolded length(1)])
apply (blast dest!: wf_interp_for_formula_EXISTS[OF _ `w' ≠ []`, of I', unfolded length(2)[symmetric]])
apply assumption
done
ultimately show "satisfies (w, I) (EXISTS φ)" using length by (auto intro!: exI[of _ P])
qed
qed auto

lemma langMSO_Or:
assumes "wf_formula n (Or φ1 φ2)"
shows "langMSO n (Or φ1 φ2) ⊆
(langMSO n φ1 ∪ langMSO n φ2) ∩ {enc (w, I) | w I. length I = n ∧ wf_interp_for_formula (w, I) (Or φ1 φ2)}"

(is "_ ⊆ (?L1 ∪ ?L2) ∩ ?ENC")
proof (intro equalityI subsetI)
fix x assume "x ∈ langMSO n (Or φ1 φ2)"
then obtain w I where
*: "x = enc (w, I)" "wf_interp_for_formula (w, I) (Or φ1 φ2)" "length I = n" "length w > 0" and
"satisfies (w, I) φ1 ∨ satisfies (w, I) φ2" unfolding langMSO_def by auto
thus "x ∈ (?L1 ∪ ?L2) ∩ ?ENC"
proof (elim disjE)
assume "satisfies (w, I) φ1"
with * have "x ∈ ?L1" using assms unfolding langMSO_def by (fastforce)
with * show ?thesis by auto
next
assume "satisfies (w, I) φ2"
with * have "x ∈?L2" using assms unfolding langMSO_def by (fastforce)
with * show ?thesis by auto
qed
qed

lemma langMSO_And:
assumes "wf_formula n (And φ1 φ2)"
shows "langMSO n (And φ1 φ2) ⊆
langMSO n φ1 ∩ langMSO n φ2 ∩ {enc (w, I) | w I. length I = n ∧ wf_interp_for_formula (w, I) (And φ1 φ2)}"

(is "_ ⊆ ?L1 ∩ ?L2 ∩ ?ENC")
using assms unfolding langMSO_def by (fastforce)

subsection {* From MSO to Regular expressions *}

fun rexp_of :: "nat => 'a formula => ('a × bool list) rexp" where
"rexp_of n (Q a m) =
Inter (TIMES [rexp.Not Zero, arbitrary_except n [(m, True)] [a], rexp.Not Zero])
(ENC n (Q a m))"

| "rexp_of n (Less m1 m2) = (if m1 = m2 then Zero else
Inter (TIMES [rexp.Not Zero, arbitrary_except n [(m1, True)] Σ,
rexp.Not Zero, arbitrary_except n [(m2, True)] Σ,
rexp.Not Zero]) (ENC n (Less m1 m2)))"

| "rexp_of n (In m M) =
Inter (TIMES [rexp.Not Zero, arbitrary_except n [(min m M, True), (max m M, True)] Σ, rexp.Not Zero])
(ENC n (In m M))"

| "rexp_of n (Not φ) = Inter (rexp.Not (rexp_of n φ)) (ENC n (Not φ))"
| "rexp_of n (Or φ1 φ2) = Inter (Plus (rexp_of n φ1) (rexp_of n φ2)) (ENC n (Or φ1 φ2))"
| "rexp_of n (And φ1 φ2) = INTERSECT [rexp_of n φ1, rexp_of n φ2, ENC n (And φ1 φ2)]"
| "rexp_of n (Exists φ) = Pr (rexp_of (n + 1) φ)"
| "rexp_of n (EXISTS φ) = Pr (rexp_of (n + 1) φ)"

fun rexp_of_alt :: "nat => 'a formula => ('a × bool list) rexp" where
"rexp_of_alt n (Q a m) =
TIMES [rexp.Not Zero, arbitrary_except n [(m, True)] [a], rexp.Not Zero]"

| "rexp_of_alt n (Less m1 m2) = (if m1 = m2 then Zero else
TIMES [rexp.Not Zero, arbitrary_except n [(m1, True)] Σ,
rexp.Not Zero, arbitrary_except n [(m2, True)] Σ,
rexp.Not Zero])"

| "rexp_of_alt n (In m M) =
TIMES [rexp.Not Zero, arbitrary_except n [(min m M, True), (max m M, True)] Σ, rexp.Not Zero]"

| "rexp_of_alt n (Not φ) = rexp.Not (rexp_of_alt n φ)"
| "rexp_of_alt n (Or φ1 φ2) = Plus (rexp_of_alt n φ1) (rexp_of_alt n φ2)"
| "rexp_of_alt n (And φ1 φ2) = Inter (rexp_of_alt n φ1) (rexp_of_alt n φ2)"
| "rexp_of_alt n (Exists φ) = Pr (Inter (rexp_of_alt (n + 1) φ) (ENC (n + 1) φ))"
| "rexp_of_alt n (EXISTS φ) = Pr (Inter (rexp_of_alt (n + 1) φ) (ENC (n + 1) φ))"

definition "rexp_of' n φ = Inter (rexp_of_alt n φ) (ENC n φ)"
(*
abbreviation except where
"except m1 m2 ≡ if m1 < m2 then [(m1, True), (m2, False)] else [(m2, False), (m1, True)]"

fun rexp_of'' :: "nat => 'a formula => ('a × bool list) rexp" where
"rexp_of'' n (Q a m) =
TIMES [Star (arbitrary_except n [(m, False)] Σ),
arbitrary_except n [(m, True)] [a],
Star (arbitrary_except n [(m, False)] Σ)]"
| "rexp_of'' n (Less m1 m2) = (if m1 = m2 then Zero else
TIMES [Star (arbitrary_except n [(m1, False), (m2, False)] Σ),
arbitrary_except n (except m1 m2) Σ,
Star (arbitrary_except n [(m1, False), (m2, False)] Σ),
arbitrary_except n (except m2 m1) Σ,
Star (arbitrary_except n [(m1, False), (m2, False)] Σ)])"
| "rexp_of'' n (In m M) =
TIMES [Star (arbitrary_except n [(m, False)] Σ),
arbitrary_except n [(min m M, True), (max m M, True)] Σ,
Star (arbitrary_except n [(m, False)] Σ)]"
| "rexp_of'' n (Not φ) = Inter (rexp.Not (rexp_of'' n φ)) (ENC n (Not φ))"
| "rexp_of'' n (Or φ1 φ2) = Inter (Plus (rexp_of'' n φ1) (rexp_of'' n φ2)) (ENC n (Or φ1 φ2))"
| "rexp_of'' n (And φ1 φ2) = INTERSECT [rexp_of'' n φ1, rexp_of'' n φ2, ENC n (And φ1 φ2)]"
| "rexp_of'' n (Exists φ) = Pr (rexp_of'' (n + 1) φ)"
| "rexp_of'' n (EXISTS φ) = Pr (rexp_of'' (n + 1) φ)"
*)

lemma length_dec_interp[simp]: "length (dec_interp n FO x) = n"
unfolding dec_interp_def by auto

theorem langMSO_rexp_of: "wf_formula n φ ==> langMSO n φ = lang n (rexp_of n φ) - {[]}"
(is "_ ==> _ = ?L n φ")
proof (induct φ arbitrary: n)
case (Q a m)
show ?case
proof (intro equalityI subsetI)
fix x assume "x ∈ langMSO n (Q a m)"
then obtain w I where
*: "x = enc (w, I)" "wf_interp_for_formula (w, I) (Q a m)" "satisfies (w, I) (Q a m)"
"length I = n"
unfolding langMSO_def by blast
with Q(1) obtain p where p: "p < length w" "I ! m = Inl p" "w ! p = a"
by (auto simp: all_set_conv_all_nth split: sum.splits)
with *(1) have "x = take p (enc (w, I)) @ [enc_atom I p a] @ drop (p + 1) (enc (w, I))"
using id_take_nth_drop[of p "enc (w, I)"] by auto
moreover from *(4) Q(1) p(2)
have "[enc_atom I p a] ∈ lang n (arbitrary_except n [(m, True)] [a])"
by (intro enc_atom_lang_arbitrary_except_True) auto
moreover from *(2,4) have "take p (enc (w, I)) ∈ lang n (rexp.Not Zero)"
by (auto intro!: enc_atom_σ dest!: in_set_takeD)
moreover from *(2,4) have "drop (Suc p) (enc (w, I)) ∈ lang n (rexp.Not Zero)"
by (auto intro!: enc_atom_σ dest!: in_set_dropD)
ultimately show "x ∈ ?L n (Q a m)" using *(1,2,4)
unfolding rexp_of.simps lang.simps(5,8) rexp_of_list.simps Int_Diff lang_ENC[OF Q]
by (auto elim: ssubst simp del: o_apply append.simps lang.simps)
next
fix x assume x: "x ∈ ?L n (Q a m)"
with Q obtain w I p where m: "I ! m = Inl p" "m < length I" and
wI: "x = enc (w, I)" "length I = n" "wf_interp_for_formula (w, I) (Q a m)"
unfolding rexp_of.simps lang.simps lang_ENC[OF Q] Int_Diff by atomize_elim (auto split: sum.splits)
hence "wf_interp_for_formula (dec_word x, dec_interp n {m} x) (Q a m)" unfolding wI(1)
using enc_wf_interp[OF Q(1)[folded wI(2)]] by auto
moreover
from x obtain u1 u u2 where "x = u1 @ u @ u2" "u ∈ lang n (arbitrary_except n [(m, True)] [a])"
unfolding rexp_of.simps lang.simps rexp_of_list.simps using concE by fast
with Q(1) obtain v where v: "x = u1 @ [v] @ u2" "snd v ! m" "fst v = a"
using arbitrary_except[of u n m True "[a]"] by fastforce
hence u: "length u1 < length x" by auto
{ from v have "snd (x ! length u1) ! m" by auto
moreover
from m wI have "p < length x" "snd (x ! p) ! m"
by (fastforce intro: nth_mem split: sum.splits)+
moreover
from m wI have ex1: "∃!p. p < length x ∧ snd (x ! p) ! m" unfolding wI(1) by (intro enc_unique) auto
ultimately have "p = length u1" using u by auto
} note * = this
from v have "v = enc (w, I) ! length u1" unfolding wI(1) by simp
hence "a = w ! length u1" using nth_map[OF u, of fst] unfolding wI(1) v(3)[symmetric] by auto
with * m wI have "satisfies (dec_word x, dec_interp n {m} x) (Q a m)"
unfolding dec_word_enc[of w I, folded wI(1)]
by (auto simp del: enc.simps dest: dec_interp_not_Inr split: sum.splits)
(fastforce dest!: dec_interp_enc_Inl intro: nth_mem split: sum.splits)
moreover from wI have "wf_word n x" unfolding wf_word by (auto intro!: enc_atom_σ)
ultimately show "x ∈ langMSO n (Q a m)" unfolding langMSO_def using m wI(3)
by (auto simp del: enc.simps intro!: exI[of _ "dec_word x"] exI[of _ "dec_interp n {m} x"]
intro: sym[OF enc_dec[OF _ ballI[OF impI[OF enc_unique[of w I, folded wI(1)]]]]])
qed
next
case (Less m m')
show ?case
proof (cases "m = m'")
case False
thus ?thesis
proof (intro equalityI subsetI)
fix x assume "x ∈ langMSO n (Less m m')"
then obtain w I where
*: "x = enc (w, I)" "wf_interp_for_formula (w, I) (Less m m')" "satisfies (w, I) (Less m m')"
"length I = n"
unfolding langMSO_def by blast
with Less(1) obtain p q where pq: "p < length w" "I ! m = Inl p" "q < length w" "I ! m' = Inl q" "p < q"
by (auto simp: all_set_conv_all_nth split: sum.splits)
with *(1) have "x = take p (enc (w, I)) @ [enc_atom I p (w ! p)] @ drop (p + 1) (enc (w, I))"
using id_take_nth_drop[of p "enc (w, I)"] by auto
also have "drop (p + 1) (enc (w, I)) = take (q - p - 1) (drop (p + 1) (enc (w, I))) @
[enc_atom I q (w ! q)] @ drop (q - p) (drop (p + 1) (enc (w, I)))"
(is "_ = ?LHS")
using id_take_nth_drop[of "q - p - 1" "drop (p + 1) (enc (w, I))"] pq by auto
finally have "x = take p (enc (w, I)) @ [enc_atom I p (w ! p)] @ ?LHS" .
moreover from *(2,4) Less(1) pq(1,2)
have "[enc_atom I p (w ! p)] ∈ lang n (arbitrary_except n [(m, True)] Σ)"
by (intro enc_atom_lang_arbitrary_except_True) auto
moreover from *(2,4) Less(1) pq(3,4)
have "[enc_atom I q (w ! q)] ∈ lang n (arbitrary_except n [(m', True)] Σ)"
by (intro enc_atom_lang_arbitrary_except_True) auto
moreover from *(2,4) have "take p (enc (w, I)) ∈ lang n (rexp.Not Zero)"
by (auto intro!: enc_atom_σ dest!: in_set_takeD)
moreover from *(2,4) have "take (q - p - 1) (drop (Suc p) (enc (w, I))) ∈ lang n (rexp.Not Zero)"
by (auto intro!: enc_atom_σ dest!: in_set_dropD in_set_takeD)
moreover from *(2,4) have "drop (q - p) (drop (Suc p) (enc (w, I))) ∈ lang n (rexp.Not Zero)"
by (auto intro!: enc_atom_σ dest!: in_set_dropD)
ultimately show "x ∈ ?L n (Less m m')" using *(1,2,4)
unfolding rexp_of.simps lang.simps(5,8) rexp_of_list.simps Int_Diff lang_ENC[OF Less] if_not_P[OF False]
by (auto elim: ssubst simp del: o_apply append.simps lang.simps)
next
fix x assume x: "x ∈ ?L n (Less m m')"
with Less obtain w I where
wI: "x = enc (w, I)" "length I = n" "wf_interp_for_formula (w, I) (Less m m')"
unfolding rexp_of.simps lang.simps lang_ENC[OF Less] Int_Diff if_not_P[OF False]
by (fastforce split: sum.splits)
with Less obtain p p' where m: "I ! m = Inl p" "m < length I" "I ! m' = Inl p'" "m' < length I"
by (auto split: sum.splits)
with wI have "wf_interp_for_formula (dec_word x, dec_interp n {m, m'} x) (Less m m')" unfolding wI(1)
using enc_wf_interp[OF Less(1)[folded wI(2)]] by auto
moreover
from x obtain u1 u u2 u' u3 where "x = u1 @ u @ u2 @ u' @ u3"
"u ∈ lang n (arbitrary_except n [(m, True)] Σ)"
"u' ∈ lang n (arbitrary_except n [(m', True)] Σ)"
unfolding rexp_of.simps lang.simps rexp_of_list.simps if_not_P[OF False] using concE by fast
with Less(1) obtain v v' where v: "x = u1 @ [v] @ u2 @ [v'] @ u3" "snd v ! m" "snd v' ! m'"
using arbitrary_except[of u n m True Σ] arbitrary_except[of u' n m' True Σ] by fastforce
hence u: "length u1 < length x" and u': "Suc (length u1 + length u2) < length x" (is "?u' < _") by auto
{ from v have "snd (x ! length u1) ! m" by auto
moreover
from m wI have "p < length x" "snd (x ! p) ! m"
by (fastforce intro: nth_mem split: sum.splits)+
moreover
from m wI have ex1: "∃!p. p < length x ∧ snd (x ! p) ! m" unfolding wI(1) by (intro enc_unique) auto
ultimately have "p = length u1" using u by auto
}
{ from v have "snd (x ! ?u') ! m'" by (auto simp: nth_append)
moreover
from m wI have "p' < length x" "snd (x ! p') ! m'"
by (fastforce intro: nth_mem split: sum.splits)+
moreover
from m wI have ex1: "∃!p. p < length x ∧ snd (x ! p) ! m'" unfolding wI(1) by (intro enc_unique) auto
ultimately have "p' = ?u'" using u' by auto
} note * = this `p = length u1`
with * m wI have "satisfies (dec_word x, dec_interp n {m, m'} x) (Less m m')"
unfolding dec_word_enc[of w I, folded wI(1)]
by (auto simp del: enc.simps dest: dec_interp_not_Inr split: sum.splits)
(fastforce dest!: dec_interp_enc_Inl intro: nth_mem split: sum.splits)
moreover from wI have "wf_word n x" unfolding wf_word by (auto intro!: enc_atom_σ)
ultimately show "x ∈ langMSO n (Less m m')" unfolding langMSO_def using m wI(3)
by (auto simp del: enc.simps intro!: exI[of _ "dec_word x"] exI[of _ "dec_interp n {m, m'} x"]
intro: sym[OF enc_dec[OF _ ballI[OF impI[OF enc_unique[of w I, folded wI(1)]]]]])
qed
qed (simp add: langMSO_def del: o_apply)
next
case (In m M)
show ?case
proof (intro equalityI subsetI)
fix x assume "x ∈ langMSO n (In m M)"
then obtain w I where
*: "x = enc (w, I)" "wf_interp_for_formula (w, I) (In m M)" "satisfies (w, I) (In m M)"
"length I = n"
unfolding langMSO_def by blast
with In(1) obtain p P where p: "p < length w" "I ! m = Inl p" "I ! M = Inr P" "p ∈ P"
by (auto simp: all_set_conv_all_nth split: sum.splits)
with *(1) have "x = take p (enc (w, I)) @ [enc_atom I p (w ! p)] @ drop (p + 1) (enc (w, I))"
using id_take_nth_drop[of p "enc (w, I)"] by auto
moreover
have "[enc_atom I p (w ! p)] ∈ lang n (arbitrary_except n [(min m M, True), (max m M, True)] Σ)"
proof (cases "m < M")
case True with *(2,4) In(1) p show ?thesis
by (intro enc_atom_lang_arbitrary_except_True2) (auto simp: min_absorb1 max_absorb2)
next
case False with *(2,4) In(1) p show ?thesis
by (intro enc_atom_lang_arbitrary_except_True2) (auto simp: min_absorb2 max_absorb1)
qed
moreover from *(2,4) have "take p (enc (w, I)) ∈ lang n (rexp.Not Zero)"
by (auto intro!: enc_atom_σ dest!: in_set_takeD)
moreover from *(2,4) have "drop (Suc p) (enc (w, I)) ∈ lang n (rexp.Not Zero)"
by (auto intro!: enc_atom_σ dest!: in_set_dropD)
ultimately show "x ∈ ?L n (In m M)" using *(1,2,4)
unfolding rexp_of.simps lang.simps(5,8) rexp_of_list.simps Int_Diff lang_ENC[OF In]
by (auto elim: ssubst simp del: o_apply append.simps lang.simps)
next
fix x assume x: "x ∈ ?L n (In m M)"
with In obtain w I where wI: "x = enc (w, I)" "length I = n" "wf_interp_for_formula (w, I) (In m M)"
unfolding rexp_of.simps lang.simps lang_ENC[OF In] Int_Diff by (fastforce split: sum.splits)
with In obtain p P where m: "I ! m = Inl p" "m < length I" "I ! M = Inr P" "M < length I" by (auto split: sum.splits)
with wI have "wf_interp_for_formula (dec_word x, dec_interp n {m} x) (In m M)" unfolding wI(1)
using enc_wf_interp[OF In(1)[folded wI(2)]] by auto
moreover
from x obtain u1 u u2 where "x = u1 @ u @ u2"
"u ∈ lang n (arbitrary_except n [(min m M, True), (max m M, True)] Σ)"
unfolding rexp_of.simps lang.simps rexp_of_list.simps using concE by fast
with In(1) obtain v where v: "x = u1 @ [v] @ u2" and "snd v ! min m M" "snd v ! max m M"
using arbitrary_except2[of u n "min m M" True "max m M" True Σ] by fastforce
hence v': "snd v ! m" "snd v ! M"
by (induct "m < M") (auto simp: min_absorb1 min_absorb2 max_absorb1 max_absorb2)
from v have u: "length u1 < length x" by auto
{ from v v' have "snd (x ! length u1) ! m" by auto
moreover
from m wI have "p < length x" "snd (x ! p) ! m"
by (fastforce intro: nth_mem split: sum.splits)+
moreover
from m wI have ex1: "∃!p. p < length x ∧ snd (x ! p) ! m" unfolding wI(1) by (intro enc_unique) auto
ultimately have "p = length u1" using u by auto
} note * = this
from v v' have "v = enc (w, I) ! length u1" unfolding wI(1) by simp
with v'(2) m(3,4) u wI(1) have "length u1 ∈ P" by auto
with * m wI have "satisfies (dec_word x, dec_interp n {m} x) (In m M)"
unfolding dec_word_enc[of w I, folded wI(1)]
by (auto simp del: enc.simps dest: dec_interp_not_Inr dec_interp_not_Inl split: sum.splits)
(auto simp del: enc.simps dest!: dec_interp_enc_Inl dec_interp_enc_Inr dest: nth_mem split: sum.splits)
moreover from wI have "wf_word n x" unfolding wf_word by (auto intro!: enc_atom_σ)
ultimately show "x ∈ langMSO n (In m M)" unfolding langMSO_def using m wI(3)
by (auto simp del: enc.simps intro!: exI[of _ "dec_word x"] exI[of _ "dec_interp n {m} x"]
intro: sym[OF enc_dec[OF _ ballI[OF impI[OF enc_unique[of w I, folded wI(1)]]]]])
qed
next
case (Or φ1 φ2)
from Or(3) have IH1: "langMSO n φ1 = lang n (rexp_of n φ1) - {[]}"
by (intro Or(1)) auto
from Or(3) have IH2: "langMSO n φ2 = lang n (rexp_of n φ2) - {[]}"
by (intro Or(2)) auto
show ?case
proof (intro equalityI subsetI)
fix x assume "x ∈ langMSO n (Or φ1 φ2)" thus "x ∈ lang n (rexp_of n (Or φ1 φ2)) - {[]}"
using langMSO_Or[OF Or(3)] unfolding lang_ENC[OF Or(3)] rexp_of.simps lang.simps IH1 IH2 Int_Diff by auto
next
fix x assume "x ∈ lang n (rexp_of n (Or φ1 φ2)) - {[]}"
then obtain w I where or: "x ∈ langMSO n φ1 ∨ x ∈ langMSO n φ2" and wI: "x = enc (w, I)" "length I = n"
"wf_interp_for_formula (w, I) (Or φ1 φ2)"
unfolding lang_ENC[OF Or(3)] rexp_of.simps lang.simps IH1 IH2 Int_Diff by auto
have "satisfies (w, I) φ1 ∨ satisfies (w, I) φ2"
proof (intro mp[OF disj_mono[OF impI impI] or])
assume "x ∈ langMSO n φ1"
with wI(2,3) Or(3) show "satisfies (w, I) φ1"
unfolding langMSO_def wI(1) wf_interp_for_formula_Or[OF Or(3) wI(2)]
by (auto simp del: enc.simps dest!: iffD2[OF enc_welldef[of _ _ _ _ φ1]])
next
assume "x ∈ langMSO n φ2"
with wI(2,3) Or(3) show "satisfies (w, I) φ2"
unfolding langMSO_def wI(1) wf_interp_for_formula_Or[OF Or(3) wI(2)]
by (auto simp del: enc.simps dest!: iffD2[OF enc_welldef[of _ _ _ _ φ2]])
qed
with wI show "x ∈ langMSO n (Or φ1 φ2)" unfolding langMSO_def by auto
qed
next
case (And φ1 φ2)
from And(3) have IH1: "langMSO n φ1 = lang n (rexp_of n φ1) - {[]}"
by (intro And(1)) auto
from And(3) have IH2: "langMSO n φ2 = lang n (rexp_of n φ2) - {[]}"
by (intro And(2)) auto
show ?case
proof (intro equalityI subsetI)
fix x assume "x ∈ langMSO n (And φ1 φ2)" thus "x ∈ lang n (rexp_of n (And φ1 φ2)) - {[]}"
using langMSO_And[OF And(3)]
unfolding lang_ENC[OF And(3)] rexp_of.simps rexp_of_list.simps lang.simps IH1 IH2 Int_Diff by auto
next
fix x assume "x ∈ lang n (rexp_of n (And φ1 φ2)) - {[]}"
then obtain w I where "and": "x ∈ langMSO n φ1 ∧ x ∈ langMSO n φ2" and wI: "x = enc (w, I)" "length I = n"
"wf_interp_for_formula (w, I) (And φ1 φ2)"
unfolding lang_ENC[OF And(3)] rexp_of.simps rexp_of_list.simps lang.simps IH1 IH2 Int_Diff by auto
have "satisfies (w, I) φ1 ∧ satisfies (w, I) φ2"
proof (intro mp[OF conj_mono[OF impI impI] "and"])
assume "x ∈ langMSO n φ1"
with wI(2,3) And(3) show "satisfies (w, I) φ1"
unfolding langMSO_def wI(1) wf_interp_for_formula_And[OF And(3) wI(2)]
by (auto simp del: enc.simps dest!: iffD2[OF enc_welldef[of _ _ _ _ φ1]])
next
assume "x ∈ langMSO n φ2"
with wI(2,3) And(3) show "satisfies (w, I) φ2"
unfolding langMSO_def wI(1) wf_interp_for_formula_And[OF And(3) wI(2)]
by (auto simp del: enc.simps dest!: iffD2[OF enc_welldef[of _ _ _ _ φ2]])
qed
with wI show "x ∈ langMSO n (And φ1 φ2)" unfolding langMSO_def by auto
qed
next
case (Not φ)
hence IH: "?L n φ = langMSO n φ" by simp
show ?case
proof (intro equalityI subsetI)
fix x assume "x ∈ langMSO n (Not φ)"
then obtain w I where
*: "x = enc (w, I)" "wf_interp_for_formula (w, I) φ" "length I = n" "length w > 0"
and unsat: "¬ (satisfies (w, I) φ)"
unfolding langMSO_def by auto
{ assume "x ∈ ?L n φ"
with IH have "satisfies (w, I) φ" using enc_welldef[of _ _ w I φ, OF _ _ _ *(2)] Not(2)
unfolding *(1,3) langMSO_def by auto
}
with unsat have "x ∉ ?L n φ" by blast
with * show "x ∈ ?L n (Not φ)" unfolding rexp_of.simps lang.simps using lang_ENC[OF Not(2)]
by (auto simp: comp_def intro!: enc_atom_σ)
next
fix x assume "x ∈ ?L n (Not φ)"
with IH have "x ∈ lang n (ENC n (Not φ)) - {[]}" and x: "x ∉ langMSO n φ" by (auto simp del: o_apply)
then obtain w I where *: "x = enc (w, I)" "wf_interp_for_formula (w, I) (Not φ)" "length I = n"
unfolding lang_ENC[OF Not(2)] by blast
{ assume "¬ satisfies (w, I) (Not φ)"
with * have "x ∈ langMSO n φ" unfolding langMSO_def by auto
}
with x * show "x ∈ langMSO n (Not φ)" unfolding langMSO_def by blast
qed
next
case (Exists φ)
show ?case
proof (intro equalityI subsetI)
fix x assume "x ∈ langMSO n (Exists φ)"
then obtain w I p where
*: "x = enc (w, I)" "wf_interp_for_formula (w, I) (Exists φ)"
"length I = n" "length w > 0" "p ∈ {0 .. length w - 1}" "satisfies (w, Inl p # I) φ"
unfolding langMSO_def by auto
with Exists(2) have "enc (w, Inl p # I) ∈ ?L (Suc n) φ"
by (intro subsetD[OF equalityD1[OF Exists(1)], of "Suc n" "enc (w, Inl p # I)"])
(auto simp: langMSO_def nth_Cons' ord_less_eq_trans[OF le_imp_less_Suc Suc_pred[OF *(4)]]
split: split_if_asm sum.splits intro!: exI[of _ w] exI[of _ "Inl p # I"])
with *(1) show "x ∈ ?L n (Exists φ)"
by (auto simp: map_index intro!: image_eqI[of _ "map π"] simp del: o_apply) (auto simp: π_def)
next
fix x assume "x ∈ ?L n (Exists φ)"
then obtain x' where x: "x = map π x'" and "x' ∈ ?L (Suc n) φ" by (auto simp del: o_apply)
with Exists(2) have "x' ∈ langMSO (Suc n) φ"
by (intro subsetD[OF equalityD2[OF Exists(1)], of "Suc n" x'])
(auto split: split_if_asm sum.splits)
then obtain w I' where
*: "x' = enc (w, I')" "wf_interp_for_formula (w, I') φ" "length I' = Suc n" "satisfies (w, I') φ"
unfolding langMSO_def by auto
moreover then obtain I0 I where "I' = I0 # I" by (cases I') auto
moreover with Exists(2) *(2) obtain p where "I0 = Inl p" "p < length w"
by (auto simp: nth_Cons' split: sum.splits split_if_asm)
ultimately have "x = enc (w, I)" "wf_interp_for_formula (w, I) (Exists φ)" "length I = n"
"length w > 0" "satisfies (w, I) (Exists φ)"using Exists(2) unfolding x
by (auto simp: map_tl nth_Cons' split: split_if_asm simp del: o_apply) (auto simp: π_def)
thus "x ∈ langMSO n (Exists φ)" unfolding langMSO_def by (auto intro!: exI[of _ w] exI[of _ I])
qed
next
case (EXISTS φ)
show ?case
proof (intro equalityI subsetI)
fix x assume "x ∈ langMSO n (EXISTS φ)"
then obtain w I P where
*: "x = enc (w, I)" "wf_interp_for_formula (w, I) (EXISTS φ)"
"length I = n" "length w > 0" "P ⊆ {0 .. length w - 1}" "satisfies (w, Inr P # I) φ"
unfolding langMSO_def by auto
from *(4,5) have "∀p ∈ P. p < length w" by (cases w) auto
with *(2-4,6) EXISTS(2) have "enc (w, Inr P # I) ∈ ?L (Suc n) φ"
by (intro subsetD[OF equalityD1[OF EXISTS(1)], of "Suc n" "enc (w, Inr P # I)"])
(auto simp: langMSO_def nth_Cons' split: split_if_asm sum.splits
intro!: exI[of _ w] exI[of _ "Inr P # I"])
with *(1) show "x ∈ ?L n (EXISTS φ)"
by (auto simp: map_index intro!: image_eqI[of _ "map π"] simp del: o_apply) (auto simp: π_def)
next
fix x assume "x ∈ ?L n (EXISTS φ)"
then obtain x' where x: "x = map π x'" and x': "length x' > 0" and "x' ∈ ?L (Suc n) φ" by (auto simp del: o_apply)
with EXISTS(2) have "x' ∈ langMSO (Suc n) φ"
by (intro subsetD[OF equalityD2[OF EXISTS(1)], of "Suc n" x'])
(auto split: split_if_asm sum.splits)
then obtain w I' where
*: "x' = enc (w, I')" "wf_interp_for_formula (w, I') φ" "length I' = Suc n" "satisfies (w, I') φ"
unfolding langMSO_def by auto
moreover then obtain I0 I where "I' = I0 # I" by (cases I') auto
moreover with EXISTS(2) *(2) obtain P where "I0 = Inr P"
by (auto simp: nth_Cons' split: sum.splits split_if_asm)
moreover have "length w ≥ 1" using x' *(1) by (cases w) auto
ultimately have "x = enc (w, I)" "wf_interp_for_formula (w, I) (EXISTS φ)" "length I = n"
"length w > 0" "satisfies (w, I) (EXISTS φ)" using EXISTS(2) unfolding x
by (auto simp add: map_tl nth_Cons' split: split_if_asm
intro!: exI[of _ P] simp del: o_apply) (auto simp: π_def)
thus "x ∈ langMSO n (EXISTS φ)" unfolding langMSO_def by (auto intro!: exI[of _ w] exI[of _ I])
qed
qed

lemma wf_rexp_of: "wf_formula n φ ==> wf n (rexp_of n φ)"
by (induct φ arbitrary: n) (auto simp: wf_rexp_ENC intro: wf_rexp_arbitrary_except split: sum.splits split_if_asm)

lemma wf_rexp_of': "wf_formula n φ ==> wf n (rexp_of' n φ)"
unfolding rexp_of'_def
by (induct φ arbitrary: n) (auto simp: wf_rexp_ENC intro: wf_rexp_arbitrary_except split: sum.splits split_if_asm)


lemma ENC_Not: "ENC n (Not φ) = ENC n φ"
unfolding ENC_def by auto

lemma ENC_And:
"wf_formula n (And φ ψ) ==> lang n (ENC n (And φ ψ)) - {[]} ⊆ lang n (ENC n φ) ∩ lang n (ENC n ψ) - {[]}"
proof
fix x assume wf: "wf_formula n (And φ ψ)" and x: "x ∈ lang n (ENC n (And φ ψ)) - {[]}"
hence wf1: "wf_formula n φ" and wf2: "wf_formula n ψ" by auto
from x obtain w I where wI: "x = enc (w, I)" "wf_interp_for_formula (w, I) (And φ ψ)" "length I = n"
using lang_ENC[OF wf] by auto
hence "wf_interp_for_formula (w, I) φ" "wf_interp_for_formula (w, I) ψ"
using wf_interp_for_formula_And[OF wf] by auto
hence "x ∈ (lang n (ENC n φ) - {[]}) ∩ (lang n (ENC n ψ) - {[]})"
unfolding lang_ENC[OF wf1] lang_ENC[OF wf2] using wI by auto
thus "x ∈ lang n (ENC n φ) ∩ lang n (ENC n ψ) - {[]}" by blast
qed

lemma ENC_Or:
"wf_formula n (Or φ ψ) ==> lang n (ENC n (Or φ ψ)) - {[]} ⊆ lang n (ENC n φ) ∩ lang n (ENC n ψ) - {[]}"
proof
fix x assume wf: "wf_formula n (Or φ ψ)" and x: "x ∈ lang n (ENC n (Or φ ψ)) - {[]}"
hence wf1: "wf_formula n φ" and wf2: "wf_formula n ψ" by auto
from x obtain w I where wI: "x = enc (w, I)" "wf_interp_for_formula (w, I) (Or φ ψ)" "length I = n"
using lang_ENC[OF wf] by auto
hence "wf_interp_for_formula (w, I) φ" "wf_interp_for_formula (w, I) ψ"
using wf_interp_for_formula_Or[OF wf] by auto
hence "x ∈ (lang n (ENC n φ) - {[]}) ∩ (lang n (ENC n ψ) - {[]})"
unfolding lang_ENC[OF wf1] lang_ENC[OF wf2] using wI by auto
thus "x ∈ lang n (ENC n φ) ∩ lang n (ENC n ψ) - {[]}" by blast
qed

lemma project_enc: "map π (enc (w, x # I)) = enc (w, I)"
unfolding π_def by auto

lemma ENC_Exists:
"wf_formula n (Exists φ) ==> lang n (ENC n (Exists φ)) - {[]} = map π ` lang (Suc n) (ENC (Suc n) φ) - {[]}"
proof (intro equalityI subsetI)
fix x assume wf: "wf_formula n (Exists φ)" and x: "x ∈ lang n (ENC n (Exists φ)) - {[]}"
hence wf1: "wf_formula (Suc n) φ" by auto
from x obtain w I where wI: "x = enc (w, I)" "wf_interp_for_formula (w, I) (Exists φ)" "length I = n"
using lang_ENC[OF wf] by auto
with x have "w ≠ []" by (cases w) auto
from wI(2) obtain p where "p < length w" "wf_interp_for_formula (w, Inl p # I) φ"
using wf_interp_for_formula_Exists[OF wf[folded wI(3)] `w ≠ []`] by auto
with wI(3) have "x ∈ map π ` (lang (Suc n) (ENC (Suc n) φ) - {[]})"
unfolding wI(1) lang_ENC[OF wf1] project_enc[symmetric, of w I "Inl p"]
by (intro imageI CollectI exI[of _ w] exI[of _ "Inl p # I"]) auto
thus "x ∈ map π ` lang (Suc n) (ENC (Suc n) φ) - {[]}" by blast
next
fix x assume wf: "wf_formula n (Exists φ)" and "x ∈ map π ` lang (Suc n) (ENC (Suc n) φ) - {[]}"
hence wf1: "wf_formula (Suc n) φ" and "0 ∈ FOV φ" and x: "x ∈ map π ` (lang (Suc n) (ENC (Suc n) φ) - {[]})" by auto
from x obtain w I where wI: "x = map π (enc (w, I))" "wf_interp_for_formula (w, I) φ" "length I = Suc n"
using lang_ENC[OF wf1] by auto
with `0 ∈ FOV φ` obtain p I' where I: "I = Inl p # I'" by (cases I) (fastforce split: sum.splits)+
with wI have wtlI: "x = enc (w, I')" "length I' = n" unfolding π_def by auto
with x have "w ≠ []" by (cases w) auto
have "wf_interp_for_formula (w, I') (Exists φ)"
using wf_interp_for_formula_Exists[OF wf[folded wtlI(2)] `w ≠ []`]
wf_interp_for_formula_any_Inl[OF wI(2)[unfolded I]] ..
with wtlI show "x ∈ lang n (ENC n (Exists φ)) - {[]}" unfolding lang_ENC[OF wf] by blast
qed

lemma ENC_EXISTS:
"wf_formula n (EXISTS φ) ==> lang n (ENC n (EXISTS φ)) - {[]} = map π ` lang (Suc n) (ENC (Suc n) φ) - {[]}"
proof (intro equalityI subsetI)
fix x assume wf: "wf_formula n (EXISTS φ)" and x: "x ∈ lang n (ENC n (EXISTS φ)) - {[]}"
hence wf1: "wf_formula (Suc n) φ" by auto
from x obtain w I where wI: "x = enc (w, I)" "wf_interp_for_formula (w, I) (EXISTS φ)" "length I = n"
using lang_ENC[OF wf] by auto
with x have "w ≠ []" by (cases w) auto
from wI(2) obtain P where "∀p ∈ P. p < length w" "wf_interp_for_formula (w, Inr P # I) φ"
using wf_interp_for_formula_EXISTS[OF wf[folded wI(3)] `w ≠ []`] by auto
with wI(3) have "x ∈ map π ` (lang (Suc n) (ENC (Suc n) φ) - {[]})"
unfolding wI(1) lang_ENC[OF wf1] project_enc[symmetric, of w I "Inr P"]
by (intro imageI CollectI exI[of _ w] exI[of _ "Inr P # I"]) auto
thus "x ∈ map π ` lang (Suc n) (ENC (Suc n) φ) - {[]}" by blast
next
fix x assume wf: "wf_formula n (EXISTS φ)" and "x ∈ map π ` lang (Suc n) (ENC (Suc n) φ) - {[]}"
hence wf1: "wf_formula (Suc n) φ" and "0 ∈ SOV φ" and x: "x ∈ map π ` (lang (Suc n) (ENC (Suc n) φ) - {[]})" by auto
from x obtain w I where wI: "x = map π (enc (w, I))" "wf_interp_for_formula (w, I) φ" "length I = Suc n"
using lang_ENC[OF wf1] by auto
with `0 ∈ SOV φ` obtain P I' where I: "I = Inr P # I'" by (cases I) (fastforce split: sum.splits)+
with wI have wtlI: "x = enc (w, I')" "length I' = n" unfolding π_def by auto
with x have "w ≠ []" by (cases w) auto
have "wf_interp_for_formula (w, I') (EXISTS φ)"
using wf_interp_for_formula_EXISTS[OF wf[folded wtlI(2)] `w ≠ []`]
wf_interp_for_formula_any_Inr[OF wI(2)[unfolded I]] ..
with wtlI show "x ∈ lang n (ENC n (EXISTS φ)) - {[]}" unfolding lang_ENC[OF wf] by blast
qed

lemma langMSO_rexp_of_rexp_of':
"wf_formula n φ ==> lang n (rexp_of n φ) - {[]} = lang n (rexp_of' n φ) - {[]}"
unfolding rexp_of'_def proof (induction φ arbitrary: n)
case (Not φ)
hence "wf_formula n φ" by simp
with Not.IH show ?case unfolding rexp_of.simps rexp_of_alt.simps lang.simps ENC_Not by blast
next
case (And φ1 φ2)
hence wf1: "wf_formula n φ1" and wf2: "wf_formula n φ2" by force+
from And.IH(1)[OF wf1] And.IH(2)[OF wf2] show ?case using ENC_And[OF And.prems]
unfolding rexp_of.simps rexp_of_alt.simps lang.simps rexp_of_list.simps by blast
next
case (Or φ1 φ2)
hence wf1: "wf_formula n φ1" and wf2: "wf_formula n φ2" by force+
from Or.IH(1)[OF wf1] Or.IH(2)[OF wf2] show ?case using ENC_Or[OF Or.prems]
unfolding rexp_of.simps rexp_of_alt.simps lang.simps by blast
next
case (Exists φ)
hence wf: "wf_formula (n + 1) φ" by auto
have *: "!!A. map π ` A - {[]} = map π ` (A - {[]})" by auto
show ?case using ENC_Exists[OF Exists.prems]
unfolding rexp_of.simps rexp_of_alt.simps lang.simps * Exists.IH[OF wf] by auto
next
case (EXISTS φ)
hence wf: "wf_formula (n + 1) φ" by auto
have *: "!!A. map π ` A - {[]} = map π ` (A - {[]})" by auto
show ?case using ENC_EXISTS[OF EXISTS.prems]
unfolding rexp_of.simps rexp_of_alt.simps lang.simps * EXISTS.IH[OF wf] by auto
qed auto

theorem langMSO_rexp_of': "wf_formula n φ ==> langMSO n φ = lang n (rexp_of' n φ) - {[]}"
unfolding langMSO_rexp_of_rexp_of'[symmetric] by (rule langMSO_rexp_of)

end

end