Theory WS1S

theory WS1S
imports Formulas Stream
(* Author: Dmitriy Traytel *)

header {* WS1S *}

(*<*)
theory WS1S
imports Formulas Regular_Operators "~~/src/HOL/BNF/Examples/Stream"
begin
(*>*)

subsection {* Encodings *}

definition "cutSame x s = stake (LEAST n. sdrop n s = same x) s"

abbreviation "poss I ≡ (\<Union>x∈set I. case x of Inl p => {p} | Inr P => P)"

lemma shift_snth: "(xs @- s) !! n = (if n < length xs then xs ! n else s !! (n - length xs))"
by auto

lemma stream_map_stream_map2[simp]:
"stream_map f (stream_map2 g s1 s2) = stream_map2 (λx y. f (g x y)) s1 s2"
unfolding stream_map2_szip stream.map_comp' o_def split_def ..

lemma stream_map2_alt:
"(stream_map2 f s1 s2 = s) = (∀n. f (s1 !! n) (s2 !! n) = s !! n)"
unfolding stream_map2_szip stream_map_alt by auto

lemma snth_stream_map2[simp]:
"stream_map2 f s1 s2 !! n = f (s1 !! n) (s2 !! n)"
by (induct n arbitrary: s1 s2) auto

lemma stake_stream_map2[simp]:
"stake n (stream_map2 f s1 s2) = map (split f) (zip (stake n s1) (stake n s2))"
by (induct n arbitrary: s1 s2) auto

lemma sdrop_stream_map2[simp]:
"sdrop n (stream_map2 f s1 s2) = stream_map2 f (sdrop n s1) (sdrop n s2)"
by (induct n arbitrary: s1 s2) auto

lemma stake_szip[simp]:
"stake n (szip s1 s2) = zip (stake n s1) (stake n s2)"
by (induct n arbitrary: s1 s2) auto

lemma sdrop_szip[simp]: "sdrop n (szip s1 s2) = szip (sdrop n s1) (sdrop n s2)"
by (induct n arbitrary: s1 s2) auto

lemma take_stake: "take n (stake m s) = stake (min n m) s"
proof (induct m arbitrary: s n)
case (Suc m) thus ?case by (cases n) auto
qed simp

lemma drop_stake: "drop n (stake m s) = stake (m - n) (sdrop n s)"
proof (induct m arbitrary: s n)
case (Suc m) thus ?case by (cases n) auto
qed simp

lemma stream_map_same[simp]: "stream_map f (same x) = same (f x)"
by (coinduct rule: stream.coinduct[of "λs1 s2. s1 = stream_map f (same x) ∧ s2 = same (f x)"]) auto

lemma (in wellorder) min_Least:
"[|∃n. P n; ∃n. Q n|] ==> min (Least P) (Least Q) = (LEAST n. P n ∨ Q n)"
proof (intro sym[OF Least_equality])
fix y assume "P y ∨ Q y"
thus "min (Least P) (Least Q) ≤ y"
proof (elim disjE)
assume "P y"
hence "Least P ≤ y" by (auto intro: LeastI2_wellorder)
thus "min (Least P) (Least Q) ≤ y" unfolding min_def by auto
next
assume "Q y"
hence "Least Q ≤ y" by (auto intro: LeastI2_wellorder)
thus "min (Least P) (Least Q) ≤ y" unfolding min_def by auto
qed
qed (metis LeastI_ex min_def)

lemma sdrop_snth: "sdrop n s !! m = s !! (n + m)"
by (induct n arbitrary: m s) auto

context formula
begin

definition "any ≡ hd Σ"

lemma any_Σ[simp]: "any ∈ set Σ"
unfolding any_def by (auto simp: nonempty intro: someI[of _ "hd Σ"])

lemma enc_atom_any_σ[simp]: "length I = n ==> enc_atom I m any ∈ set (σ Σ n)"
by (auto simp: σ_def image_iff set_n_lists)

fun stream_enc :: "'a interp => ('a × bool list) stream" where
"stream_enc (w, I) = stream_map2 (enc_atom I) nats (w @- same any)"

lemma tl_stream_enc[simp]: "stream_map π (stream_enc (w, x # I)) = stream_enc (w, I)"
by (auto simp: comp_def π_def)

lemma enc_atom_max: "[|∀x∈set I. case x of Inl p => p ≤ n | Inr P => ∀p∈P. p ≤ n; n ≤ n'|] ==>
enc_atom I (Suc n') a = (a, replicate (length I) False)"

by (induct I) (auto split: sum.splits)

lemma ex_Loop_stream_enc:
assumes "∀x ∈ set I. case x of Inr P => finite P | _ => True"
shows "∃n. sdrop n (stream_enc (w, I)) = same (any, replicate (length I) False)"
proof -
from assms have "∃n > length w. ∀x∈set I. case x of Inl p => p ≤ n | Inr P => ∀p∈P. p ≤ n"
proof (induct I)
case (Cons x I)
then obtain n where IH: "length w < n"
"∀x∈set I. case x of Inl p => p ≤ n | Inr P => ∀p∈P. p ≤ n" by auto
thus ?case
proof (cases x)
case (Inl p)
with IH show ?thesis
by (intro exI[of _ "max p n"]) (fastforce split: sum.splits)
next
case (Inr P)
with IH Cons(2) show ?thesis
by (intro exI[of _ "max (Max P) n"]) (fastforce dest: Max_ge split: sum.splits)
qed
qed auto
then guess n by (elim exE conjE)
hence "sdrop (Suc n) (stream_enc (w, I)) = same (any, replicate (length I) False)"
(is "?s1 n = ?s2")
by (coinduct rule: stream.coinduct[of "λs1 s2. ∃n'≥ n. s1 = ?s1 n' ∧ s2 = ?s2"])
(auto simp: enc_atom_max dest: le_SucI)
thus ?thesis by blast
qed

lemma length_snth_enc[simp]: "length (snd (stream_enc (w, I) !! n)) = length I"
by auto

lemma stream_set_same: "[|y ∈ stream_set s; s = same x|] ==> y = x"
by (induct rule: stream_set_induct1) auto

lemma same_alt: "s = same x <-> stream_set s = {x}"
using stream_set_same[of _ "same x" x]
apply auto
apply (metis shd_stream_set)
apply (coinduct rule: stream.coinduct[of "λs1 s2. s2 = same x ∧ stream_set s1 = {x}"])
apply auto
apply (metis shd_stream_set singleton_iff)
apply (metis stl_stream_set singleton_iff)
by (metis (full_types) empty_iff insert_iff shd_stream_set stl_stream_set)

lemma sdrop_sameE: "[|sdrop n (w @- same y) = same y; p < length w; ¬ p < n|] ==> w ! p = y"
unfolding not_less same_alt
apply (induct p arbitrary: w n)
apply simp
apply (metis hd_conv_nth shd_stream_set shift_simps(1) singletonE stream_set_shift)
apply (case_tac w)
apply simp_all
apply (case_tac n)
apply simp_all
by (metis equals0D nth_mem singletonE subset_singletonD)

lemma less_length_cutSame:
"[|(w @- same y) !! p = a|] ==> a = y ∨ (p < length (cutSame y (w @- same y)) ∧ w ! p = a)"
unfolding cutSame_def length_stake
by (rule LeastI2_ex[OF exI[of _ "length w"]])
(auto simp: sdrop_shift shift_snth split: split_if_asm elim: sdrop_sameE)

lemma less_length_cutSame_Inl:
"[|(∀x ∈ set I. case x of Inr P => finite P | _ => True); r < length I; I ! r = Inl p|] ==>
p < length (cutSame (any, replicate (length I) False) (stream_enc (w, I)))"

unfolding cutSame_def length_stake
by (erule LeastI2_ex[OF ex_Loop_stream_enc ccontr])
(auto simp: stream_map2_alt dest!: add_diff_inverse,
metis (lifting, full_types) nth_map nth_replicate sum.simps(5))

lemma less_length_cutSame_Inr:
"[|(∀x ∈ set I. case x of Inr P => finite P | _ => True); r < length I; I ! r = Inr P|] ==>
∀p ∈ P. p < length (cutSame (any, replicate (length I) False) (stream_enc (w, I)))"

unfolding cutSame_def length_stake
by (rule ballI, erule LeastI2_ex[OF ex_Loop_stream_enc ccontr])
(auto simp: stream_map2_alt dest!: add_diff_inverse,
metis nth_map nth_replicate sum.simps(6))

fun enc :: "'a interp => ('a × bool list) list set" where
"enc (w, I) = {x. ∃n. x = (cutSame (any, replicate (length I) False) (stream_enc (w, I)) @
replicate n (any, replicate (length I) False))}"


lemma cutSame_all[simp]: "cutSame x (same x) = []"
unfolding cutSame_def by (auto intro: Least_equality)

lemma cutSame_stop[simp]:
assumes "x ≠ y"
shows "cutSame x (xs @- Stream y (same x)) = xs @ [y]" (is "cutSame x ?s = _")
proof -
have "(LEAST n. sdrop n ?s = same x) = Suc (length xs)"
proof (rule Least_equality)
show "sdrop (Suc (length xs)) ?s = same x"
by (metis sdrop_shift sdrop_simps(2) stream.sels(2))
next
fix m assume *: "sdrop m ?s = same x"
{ assume "m < Suc (length xs)"
hence "m ≤ length xs" by simp
then obtain ys where "sdrop m ?s = ys @- Stream y (same x)"
by atomize_elim (induct m arbitrary: xs, auto)
with * obtain "ys @- Stream y (same x) = same x" by simp
hence "Stream y (same x) = same x" by (metis sdrop_same sdrop_shift)
with assms have False by (metis same_simps(1) stream.sels(1))
}
thus "Suc (length xs) ≤ m" by (blast intro: leI)
qed
thus ?thesis unfolding cutSame_def
by (metis length_append_singleton shift.simps shift_append stake_shift)
qed

lemma cutSame_shift_same: "∃n. w = cutSame x (w @- same x) @ replicate n x"
proof (induct w rule: rev_induct)
case (snoc a w)
then obtain n where "w = cutSame x (w @- same x) @ replicate n x" by blast
thus ?case
by (cases "a = x")
(auto simp: same_unfold[symmetric] replicate_append_same[symmetric] intro!: exI[of _ "Suc n"])
qed simp

lemma set_cutSame: "set (cutSame x (w @- same x)) ⊆ set w"
proof (induct w rule: rev_induct)
case (snoc a w)
thus ?case by (cases "a = x") (auto simp: same_unfold[symmetric])
qed simp

lemma stream_enc_cutSame:
assumes "(∀x ∈ set I. case x of Inr P => finite P | _ => True)"
shows "stream_enc (w, I) = cutSame (any, replicate (length I) False) (stream_enc (w, I)) @-
same (any, replicate (length I) False)"

unfolding cutSame_def
by (rule trans[OF sym[OF stake_sdrop] arg_cong2[of _ _ _ _ "op @-", OF refl]])
(rule LeastI_ex[OF ex_Loop_stream_enc[OF assms]])

lemma map_fst_zip_min[simp]: "map fst (zip xs ys) ≡ take (min (length xs) (length ys)) xs"
proof (induct ys arbitrary: xs)
case Cons thus ?case by (case_tac xs) auto
qed simp

lemma map_snd_zip_min[simp]: "map snd (zip xs ys) ≡ take (min (length xs) (length ys)) ys"
proof (induct ys arbitrary: xs)
case Cons thus ?case by (case_tac xs) auto
qed simp

lemma stream_enc_enc:
assumes "(∀x ∈ set I. case x of Inr P => finite P | _ => True)" and v: "v ∈ enc (w, I)"
shows "stream_enc (w, I) = v @- same (any, replicate (length I) False)"
(is "?s = ?v @- same ?F")
proof -
from assms(1) obtain n where "sdrop n (stream_enc (w, I)) = same ?F" by (metis ex_Loop_stream_enc)
moreover from v obtain m where "?v = cutSame ?F ?s @ replicate m ?F" by auto
ultimately show "?s = v @- same ?F"
by (auto simp del: stream_enc.simps intro: stream_enc_cutSame[OF assms(1)])
qed

lemma stream_enc_enc_some:
assumes "(∀x ∈ set I. case x of Inr P => finite P | _ => True)"
shows "stream_enc (w, I) = (SOME v. v ∈ enc (w, I)) @- same (any, replicate (length I) False)"
by (rule stream_enc_enc[OF assms], rule someI_ex) auto

lemma enc_unique_length: "v ∈ enc (w, I) ==> ∀v'. length v' = length v ∧ v' ∈ enc (w, I) --> v = v'"
by auto

lemma sdrop_same: "sdrop n s = same x ==> n ≤ m ==> s !! m = x"
by (metis le_iff_add sdrop_snth snth_same)

lemma fin_cutSame_tl:
assumes "∃n. sdrop n s = same x"
shows "fin_cutSame (π x) (map π (cutSame x s)) = cutSame (π x) (stream_map π s)"
proof -
def min "LEAST n. sdrop n s = same x"
from assms have min: "sdrop min s = same x" "!!m. sdrop m s = same x ==> min ≤ m"
unfolding min_def by (auto intro: LeastI Least_le)
have Ex: "∃n. drop n (map π (stake min s)) = replicate (length (map π (stake min s)) - n) (π x)"
by (auto intro: exI[of _ "length (map π (stake min s))"])
have "fin_cutSame (π x) (map π (cutSame x s)) =
map π (stake (LEAST n.
map π (stake (min - n) (sdrop n s)) = replicate (min - n) (π x) ∨ sdrop n s = same x) s)"

unfolding fin_cutSame_def cutSame_def take_map take_stake min_Least[OF Ex assms, folded min_def]
min_def[symmetric] by (auto simp: drop_map drop_stake)
also have "(λn. map π (stake (min - n) (sdrop n s)) = replicate (min - n) (π x) ∨ sdrop n s = same x) =
(λn. stream_map π (sdrop n s) = same (π x))"

proof (rule ext, unfold stream_map_alt snth_same, safe)
fix n m
assume "map π (stake (min - n) (sdrop n s)) = replicate (min - n) (π x)"
hence "∀y∈set (stake (min - n) (sdrop n s)). π y = π x"
by (intro iffD1[OF map_eq_conv]) (metis length_stake map_replicate_const)
hence "∀i<min - n. π (sdrop n s !! i) = π x"
unfolding all_set_conv_all_nth by (auto simp: sdrop_snth)
thus "π (sdrop n s !! m) = π x"
proof (cases "m < min - n")
case False
hence "min ≤ n + m" by linarith
hence "sdrop n s !! m = x" unfolding sdrop_snth by (rule sdrop_same[OF min(1)])
thus ?thesis by simp
qed auto
next
fix n
assume "∀m. π (sdrop n s !! m) = π x"
thus "map π (stake (min - n) (sdrop n s)) = replicate (min - n) (π x)"
unfolding stake_stream_map[symmetric] by (metis snth_same stake_same stream_map_alt)
qed auto
finally show ?thesis unfolding cutSame_def sdrop_stream_map stake_stream_map .
qed

lemma tl_enc[simp]:
assumes "∀x ∈ set (x # I). case x of Inr P => finite P | _ => True"
shows "SAMEQUOT (any, replicate (length I) False) (map π ` enc (w, x # I)) = enc (w, I)"
unfolding SAMEQUOT_def
by (fastforce simp: assms π_def
fin_cutSame_tl[OF ex_Loop_stream_enc[OF assms], unfolded π_def, simplified, symmetric])

lemma encD:
"[|v ∈ enc (w, I); (∀x ∈ set I. case x of Inr P => finite P | _ => True)|] ==>
v = map (split (enc_atom I)) (zip [0 ..< length v] (stake (length v) (w @- same any)))"

by (erule box_equals[OF sym[OF arg_cong[of _ _ "stake (length v)" ,OF stream_enc_enc]]])
(auto simp: stake_shift sdrop_shift stake_add[symmetric] simp del: stake_add)

lemma enc_Inl: "[|x ∈ enc (w, I); (∀x ∈ set I. case x of Inr P => finite P | _ => True);
m < length I; I ! m = Inl p|] ==> p < length x ∧ snd (x ! p) ! m"

by (auto dest!: less_length_cutSame_Inl[of _ _ _ w] simp: nth_append cutSame_def)

lemma enc_Inr: assumes "x ∈ enc (w, I)" "∀x ∈ set I. case x of Inr P => finite P | _ => True"
"M < length I" "I ! M = Inr P"
shows "p ∈ P <-> p < length x ∧ snd (x ! p) ! M"
proof
assume "p ∈ P" with assms show "p < length x ∧ snd (x ! p) ! M"
by (auto dest!: less_length_cutSame_Inr[of _ _ _ w] simp: nth_append cutSame_def)
next
assume "p < length x ∧ snd (x ! p) ! M"
thus "p ∈ P" using assms by (subst (asm) (2) encD[OF assms(1,2)]) auto
qed

lemma enc_length:
assumes "enc (w, I) = enc (w', I')"
shows "length I = length I'"
proof -
let ?cL = "λw I. cutSame (any, replicate (length I) False) (stream_enc (w, I))"
let ?w = "λw I m. ?cL w I @ replicate (m - length (?cL w I)) (any, replicate (length I) False)"
let ?max = "max (length (?cL w I)) (length (?cL w' I')) + 1"
from assms have "?w w I ?max ∈ enc (w, I)" "?w w' I' ?max ∈ enc (w', I')" by auto
hence "?w w I ?max = ?w w' I' ?max" using enc_unique_length assms by (simp del: enc.simps)
moreover have "last (?w w I ?max) = (any, replicate (length I) False)"
"last (?w w' I' ?max) = (any, replicate (length I') False)" by auto
ultimately show "length I = length I'" by auto
qed

lemma enc_stream_enc:
"[|(∀x ∈ set I. case x of Inr P => finite P | _ => True);
(∀x ∈ set I'. case x of Inr P => finite P | _ => True);
enc (w, I) = enc (w', I')|] ==> stream_enc (w, I) = stream_enc (w', I')"

by (rule box_equals[OF _ sym[OF stream_enc_enc_some] sym[OF stream_enc_enc_some]])
(auto dest: enc_length simp del: enc.simps)

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


fun satisfies :: "'a interp => 'a formula => bool" (infix "\<Turnstile>" 50) where
"(w, I) \<Turnstile> FQ a m = ((case I ! m of Inl p => if p < length w then w ! p else any) = a)"
| "(w, I) \<Turnstile> FLess m1 m2 = ((case I ! m1 of Inl p => p) < (case I ! m2 of Inl p => p))"
| "(w, I) \<Turnstile> FIn m M = ((case I ! m of Inl p => p) ∈ (case I ! M of Inr P => P))"
| "(w, I) \<Turnstile> FNot φ = (¬ (w, I) \<Turnstile> φ)"
| "(w, I) \<Turnstile> (FOr φ1 φ2) = ((w, I) \<Turnstile> φ1 ∨ (w, I) \<Turnstile> φ2)"
| "(w, I) \<Turnstile> (FAnd φ1 φ2) = ((w, I) \<Turnstile> φ1 ∧ (w, I) \<Turnstile> φ2)"
| "(w, I) \<Turnstile> (FExists φ) = (∃p. (w, Inl p # I) \<Turnstile> φ)"
| "(w, I) \<Turnstile> (FEXISTS φ) = (∃P. finite P ∧ (w, Inr P # I) \<Turnstile> φ)"

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

lemma encD_ex: "[|x ∈ enc (w, I); (∀x ∈ set I. case x of Inr P => finite P | _ => True)|] ==>
∃n. x = map (split (enc_atom I)) (zip [0 ..< n] (stake n (w @- same any)))"

by (auto dest!: encD simp del: enc.simps)

lemma enc_set_σ: "[|x ∈ enc (w, I); (∀x ∈ set I. case x of Inr P => finite P | _ => True);
length I = n; a ∈ set x; set w ⊆ set Σ|] ==> a ∈ set (σ Σ n)"

apply (auto dest!: encD_ex simp: in_set_zip simp del: enc.simps)
apply (case_tac "na < length w")
apply (auto intro!: enc_atom_σ)
done

definition "positions_in_row s i =
Option.these (stream_set (stream_map2 (λp (_, bs). if nth bs i then Some p else None) nats s))"


lemma positions_in_row: "positions_in_row s i = {p. snd (s !! p) ! i}"
unfolding positions_in_row_def these_def stream_map2_szip stream.set_natural' stream_set_range
by (auto split: split_if_asm intro!: image_eqI[of _ the] split: prod.splits)

lemma positions_in_row_unique: "∃!p. snd (s !! p) ! i ==>
the_elem (positions_in_row s i) = (THE p. snd (s !! p) ! i)"

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

lemma positions_in_row_nth: "∃!p. snd (s !! p) ! i ==>
snd (s !! the_elem (positions_in_row s i)) ! i"

unfolding positions_in_row_unique by (rule the1I2) auto

definition "dec_word s = cutSame any (stream_map fst s)"

lemma dec_word_stream_enc: "dec_word (stream_enc (w, I)) = cutSame any (w @- same any)"
unfolding dec_word_def by (auto intro!: arg_cong[of _ _ "cutSame any"] simp: stream_map2_alt)

definition "stream_dec n FO s = map (λi.
if i ∈ FO
then Inl (the_elem (positions_in_row s i))
else Inr (positions_in_row s i)) [0..<n]"


lemma stream_dec_Inl: "[|i ∈ FO; i < n|] ==> ∃p. stream_dec n FO s ! i = Inl p"
unfolding stream_dec_def using nth_map[of n "[0..<n]"] by auto

lemma stream_dec_not_Inr: "[|stream_dec n FO s ! i = Inr P; i ∈ FO; i < n|] ==> False"
unfolding stream_dec_def using nth_map[of n "[0..<n]"] by auto

lemma stream_dec_Inr: "[|i ∉ FO; i < n|] ==> ∃P. stream_dec n FO s ! i = Inr P"
unfolding stream_dec_def using nth_map[of n "[0..<n]"] by auto

lemma stream_dec_not_Inl: "[|stream_dec n FO s ! i = Inl p; i ∉ FO; i < n|] ==> False"
unfolding stream_dec_def using nth_map[of n "[0..<n]"] by auto

lemma Inr_dec_finite: "[|∀i<n. finite {p. snd (s !! p) ! i}; Inr P ∈ set (stream_dec n FO s)|] ==>
finite P"

unfolding stream_dec_def by (auto simp: positions_in_row)

lemma enc_atom_dec:
"[|∀p. length (snd (s !! p)) = n; ∀i ∈ FO. i < n --> (∃!p. snd (s !! p) ! i); a = fst (s !! p)|] ==>
enc_atom (stream_dec n FO s) p a = s !! p"

unfolding stream_dec_def
by (rule sym, subst surjective_pairing[of "s !! p"])
(auto intro!: nth_equalityI simp: positions_in_row simp del: pair_collapse split: split_if_asm,
(metis positions_in_row positions_in_row_nth)+)

lemma length_stream_dec[simp]: "length (stream_dec n FO x) = n"
unfolding stream_dec_def by auto

lemma stream_enc_dec:
"[|∃n. sdrop n (stream_map fst s) = same any;
stream_all (λx. length (snd x) = n) s; ∀i ∈ FO. (∃!p. snd (s !! p) ! i)|] ==>
stream_enc (dec_word s, stream_dec n FO s) = s"

unfolding dec_word_def snth_fromN
by (drule LeastI_ex)
(auto intro!: enc_atom_dec simp: stream_map2_alt cutSame_def
simp del: stake_stream_map sdrop_stream_map
intro!: trans[OF arg_cong2[of _ _ _ _ "op !!"] snth_stream_map]
trans[OF arg_cong2[of _ _ _ _ "op @-"] stake_sdrop])

lemma stream_enc_unique:
"i < length I ==> ∃p. I ! i = Inl p ==> ∃!p. snd (stream_enc (w, I) !! p) ! i"
by auto

lemma stream_dec_enc_Inl:
"[|stream_dec n FO (stream_enc (w, I)) ! i = Inl p'; I ! i = Inl p; i ∈ FO; i < n; length I = n|] ==>
p = p'"

unfolding stream_dec_def
by (auto intro!: trans[OF _ sym[OF positions_in_row_unique[OF stream_enc_unique]]]
simp del: stream_enc.simps) simp

lemma stream_dec_enc_Inr:
"[|stream_dec n FO (stream_enc (w, I)) ! i = Inr P'; I ! i = Inr P; i ∉ FO; i < n; length I = n|] ==>
P = P'"

unfolding stream_dec_def positions_in_row by auto

lemma Collect_snth: "{p. P (Stream x s !! p)} ⊆ {0} ∪ Suc ` {p. P (s !! p)}"
unfolding image_def by (auto simp: gr0_conv_Suc)

lemma finite_True_in_row: "∀i < n. finite {p. snd ((w @- same (any, replicate n False)) !! p) ! i}"
by (induct w) (auto intro: finite_subset[OF Collect_snth])

lemma lang_ENC:
assumes "wf_formula n φ"
shows "lang n (ENC n φ) = \<Union>{enc (w, I) | w I . length I = n ∧ wf_interp_for_formula (w, I) φ}"
(is "?L = ?R")
proof (intro equalityI subsetI)
fix x assume L: "x ∈ ?L"
hence *: "set x ⊆ set (σ Σ n)" using wf_lang_wf_word[OF wf_rexp_ENC] by (auto simp: wf_word)
let ?s = "x @- same (any, replicate n False)"
have "list_all (λbs. length (snd bs) = n) x"
using bspec[OF wf_lang_wf_word[OF wf_rexp_ENC], OF `x ∈ ?L`]
by (auto simp: list_all_iff wf_word) (auto simp: σ_def set_n_lists)
hence "stream_all (λx. length (snd x) = n) (x @- same (any, replicate n False))"
by (auto simp only: stream_all_shift stream_all_same length_replicate snd_conv)
moreover
{ fix m assume "m ∈ FOV φ"
with assms have "m < n" by (auto simp: max_idx_vars)
with L `m ∈ FOV φ` obtain u z v where uzv: "x = u @ z @ v"
"u ∈ star (lang n (arbitrary_except n [(m, False)] Σ))"
"z ∈ lang n (arbitrary_except n [(m, True)] Σ)"
"v ∈ star (lang n (arbitrary_except n [(m, False)] Σ))" unfolding ENC_def
by (auto simp: wf_rexp_valid_ENC finite_FOV dest!: iffD1[OF lang_flatten_INTERSECT, rotated -1])
(fastforce simp: valid_ENC_def)
with `m < n` have "∃!p. snd (x ! p) ! m ∧ p < length x"
proof (intro ex1I[of _ "length u"])
fix p assume "m < n" "snd (x ! p) ! m ∧ p < length x"
with star_arbitrary_except[OF uzv(2)] arbitrary_except[OF uzv(3)] star_arbitrary_except[OF uzv(4)]
show "p = length u" by (cases rule: nat_less_cases) (auto simp: nth_append uzv(1))
qed (auto dest!: arbitrary_except)
then obtain p where p: "p < length x" "snd (x ! p) ! m"
"!!q. snd (x ! q) ! m ∧ q < length x --> q = p" by auto
hence "∃!p. snd (?s !! p) ! m"
proof (intro ex1I[of _ p])
fix q from p `m < n` show "snd (?s !! q) ! m ==> q = p" by (cases "q < length x") auto
qed auto
}
moreover have "sdrop (length x) (stream_map fst (x @- same (any, replicate n False))) = same any"
unfolding sdrop_stream_map by (subst sdrop_shift[OF refl refl]) simp
ultimately have enc_dec: "stream_enc (dec_word ?s, stream_dec n (FOV φ) ?s) =
x @- same (any, replicate n False)"
by (intro stream_enc_dec) auto
def I "stream_dec n (FOV φ) ?s"
with assms have "wf_interp_for_formula (dec_word ?s, I) φ" unfolding I_def dec_word_def
by (auto dest: stream_dec_not_Inr stream_dec_not_Inl simp :σ_def max_idx_vars
dest!: set_mp[OF set_cutSame[of any "map fst x"]] set_mp[OF *] split: sum.splits)
(auto simp: stream_dec_def positions_in_row finite_True_in_row)
moreover have "length I = n" unfolding I_def by simp
moreover have "x ∈ enc (dec_word ?s, I)" unfolding I_def
by (simp add: enc_dec cutSame_shift_same del: stream_enc.simps)
ultimately show "x ∈ ?R" by blast
next
fix x assume "x ∈ ?R"
then obtain w I where I: "x ∈ enc (w, I)" "wf_interp_for_formula (w, I) φ" "length I = n" by blast
{ fix i from I(2) have "(w @- same any) !! i ∈ set Σ" by (cases "i < length w") auto } note * = this
from I have "x @- same (any, replicate (length I) False) = stream_enc (w, I)" (is "x @- ?F = ?s")
by (intro stream_enc_enc[symmetric]) auto
with * `length I = n` have "∀x ∈ set x. length (snd x) = n ∧ fst x ∈ set Σ"
by (auto dest!: shift_snth_less[of _ _ ?F, symmetric] simp: in_set_conv_nth)
thus "x ∈ ?L"
proof (cases "FOV φ = {}")
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 assume "r ∈ FOV φ"
with I(2) obtain p where p: "I ! r = Inl p" by (cases "I ! r") auto
from `r ∈ FOV φ` assms I(2,3) have r: "r < length I" by (auto dest!: max_idx_vars)
from p I(1,2) r have "p < length x"
using less_length_cutSame_Inl[of I r p w] by auto
with p I r * have "[x ! p] ∈ lang n (arbitrary_except n [(r, True)] Σ)"
by (subst encD[of x]) (auto intro!: enc_atom_lang_arbitrary_except_True)
moreover
from p I r * have "take p x ∈ star (lang n (arbitrary_except n [(r, False)] Σ))"
by (subst encD[of x]) (auto simp: in_set_conv_nth intro!: Ball_starI enc_atom_lang_arbitrary_except_False)
moreover
from p I r * have "drop (Suc p) x ∈ star (lang n (arbitrary_except n [(r, False)] Σ))"
by (subst encD[of x]) (auto simp: in_set_conv_nth simp del: snth.simps intro!: Ball_starI enc_atom_lang_arbitrary_except_False)
ultimately have "take p x @ [x ! p] @ drop (p + 1) x ∈ lang n (valid_ENC n r)"
using `0 < n` unfolding valid_ENC_def by (auto simp del: append.simps)
hence "x ∈ lang n (valid_ENC n r)" using id_take_nth_drop[OF `p < length x`] by auto
}
with False lang_flatten_INTERSECT[OF finite nonempty wf_rexp] show ?thesis by (auto simp: ENC_def)
qed (simp add: ENC_def, auto simp: σ_def set_n_lists image_iff)
qed

subsection {* Welldefinedness of enc wrt. Models *}

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

by (auto simp: nth_Cons' split: split_if_asm)

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

by (auto simp: nth_Cons' split: split_if_asm)

lemma wf_interp_for_formula_FEXISTS:
"[|wf_formula (length I) (FEXISTS φ)|]==>
wf_interp_for_formula (w, I) (FEXISTS φ) <-> (∀P. finite P --> wf_interp_for_formula (w, Inr P # I) φ)"

by (auto simp: nth_Cons' split: split_if_asm)

lemma wf_interp_for_formula_any_Inr: "wf_interp_for_formula (w, Inr P # I) φ ==>
∀P. finite P --> wf_interp_for_formula (w, Inr P # I) φ"

by (auto simp: nth_Cons' split: split_if_asm)

lemma wf_interp_for_formula_FOr:
"wf_interp_for_formula (w, I) (FOr φ1 φ2) =
(wf_interp_for_formula (w, I) φ1 ∧ wf_interp_for_formula (w, I) φ2)"

by auto

lemma wf_interp_for_formula_FAnd:
"wf_interp_for_formula (w, I) (FAnd φ1 φ2) =
(wf_interp_for_formula (w, I) φ1 ∧ wf_interp_for_formula (w, I) φ2)"

by auto

lemma enc_wf_interp:
"[|wf_formula (length I) φ; wf_interp_for_formula (w, I) φ; x ∈ enc (w, I)|] ==>
wf_interp_for_formula (dec_word (x @- same (any, replicate (length I) False)),
stream_dec (length I) (FOV φ) (x @- same (any, replicate (length I) False))) φ"

using
stream_dec_Inl[of _ "FOV φ" "length I" "stream_enc (w, I)", OF _ bspec[OF max_idx_vars]]
stream_dec_Inr[of _ "FOV φ" "length I" "stream_enc (w, I)", OF _ bspec[OF max_idx_vars]]
by (auto split: sum.splits intro: Inr_dec_finite[OF finite_True_in_row] simp: max_idx_vars dec_word_def
dest!: stream_dec_not_Inl stream_dec_not_Inr set_mp[OF set_cutSame] simp del: stream_enc.simps)
(auto simp: cutSame_def in_set_zip stream_map2_alt shift_snth)

lemma enc_atom_welldef: "∀x a. enc_atom I x a = enc_atom I' x a ==> m < length I ==>
(case (I ! m, I' ! m) of (Inl p, Inl q) => p = q | (Inr P, Inr Q) => P = Q | _ => True)"

proof (induct "length I" arbitrary: I I' m)
case (Suc n I I')
then obtain x xs x' xs' where *: "I = x # xs" "I' = x' # xs'"
by (fastforce simp: Suc_length_conv map_eq_Cons_conv)
with Suc show ?case
proof (cases m)
case 0 thus ?thesis using Suc(3) unfolding *
by (cases x x' rule: sum.exhaust[case_product sum.exhaust]) auto
qed auto
qed simp

lemma stream_enc_welldef: "[|stream_enc (w, I) = stream_enc (w', I'); wf_formula (length I) φ;
wf_interp_for_formula (w, I) φ; wf_interp_for_formula (w', I') φ|] ==>
(w, I) \<Turnstile> φ <-> (w', I') \<Turnstile> φ"

proof (induction φ arbitrary: w w' I I')
case (FQ a m) thus ?case using enc_atom_welldef[of I I' m]
by (simp split: sum.splits add: stream_map2_alt shift_snth) (metis snth_same)
next
case (FLess m1 m2) thus ?case using enc_atom_welldef[of I I' m1] enc_atom_welldef[of I I' m2]
by (auto split: sum.splits simp add: stream_map2_alt)
next
case (FIn m M) thus ?case using enc_atom_welldef[of I I' m] enc_atom_welldef[of I I' M]
by (auto split: sum.splits simp add: stream_map2_alt)
next
case (FOr φ1 φ2) show ?case unfolding satisfies.simps(5)
proof (intro disj_cong)
from FOr(3-6) show "(w, I) \<Turnstile> φ1 <-> (w', I') \<Turnstile> φ1"
by (intro FOr(1)) auto
next
from FOr(3-6) show "(w, I) \<Turnstile> φ2 <-> (w', I') \<Turnstile> φ2"
by (intro FOr(2)) auto
qed
next
case (FAnd φ1 φ2) show ?case unfolding satisfies.simps(6)
proof (intro conj_cong)
from FAnd(3-6) show "(w, I) \<Turnstile> φ1 <-> (w', I') \<Turnstile> φ1"
by (intro FAnd(1)) auto
next
from FAnd(3-6) show "(w, I) \<Turnstile> φ2 <-> (w', I') \<Turnstile> φ2"
by (intro FAnd(2)) auto
qed
next
case (FExists φ)
hence length: "length I' = length I" by (metis length_snth_enc)
show ?case
proof
assume "(w, I) \<Turnstile> FExists φ"
with FExists.prems(3) obtain p where "(w, Inl p # I) \<Turnstile> φ" by auto
moreover
with FExists.prems have "(w', Inl p # I') \<Turnstile> φ"
apply (intro iffD1[OF FExists.IH[of w "Inl p # I" w' "Inl p # I'"]])
apply (auto simp: stream_map2_alt split: sum.splits) []
apply (auto split: sum.splits split_if_asm) []
apply (blast dest!: wf_interp_for_formula_FExists)
apply (blast dest!: wf_interp_for_formula_FExists[of I', unfolded length])
apply assumption
done
ultimately show "(w', I') \<Turnstile> FExists φ" by auto
next
assume "(w', I') \<Turnstile> FExists φ"
with FExists.prems(1,2,4) obtain p where "(w', Inl p # I') \<Turnstile> φ" by auto
moreover
with FExists.prems have "(w, Inl p # I) \<Turnstile> φ"
apply (intro iffD2[OF FExists.IH[of w "Inl p # I" w' "Inl p # I'"]])
apply (auto simp: stream_map2_alt split: sum.splits) []
apply (auto split: sum.splits split_if_asm) []
apply (blast dest!: wf_interp_for_formula_FExists)
apply (blast dest!: wf_interp_for_formula_FExists[of I', unfolded length])
apply assumption
done
ultimately show "(w, I) \<Turnstile> FExists φ" by auto
qed
next
case (FEXISTS φ)
hence length: "length I' = length I" by (metis length_snth_enc)
show ?case
proof
assume "(w, I) \<Turnstile> FEXISTS φ"
with FEXISTS.prems(3) obtain P where "finite P" "(w, Inr P # I) \<Turnstile> φ" by auto
moreover
with FEXISTS.prems have "(w', Inr P # I') \<Turnstile> φ"
apply (intro iffD1[OF FEXISTS.IH[of w "Inr P # I" w' "Inr P # I'"]])
apply (auto simp: stream_map2_alt split: sum.splits) []
apply (auto split: sum.splits split_if_asm) []
apply (blast dest!: wf_interp_for_formula_FEXISTS)
apply (blast dest!: wf_interp_for_formula_FEXISTS[of I', unfolded length])
apply assumption
done
ultimately show "(w', I') \<Turnstile> FEXISTS φ" by auto
next
assume "(w', I') \<Turnstile> FEXISTS φ"
with FEXISTS.prems(1,2,4) obtain P where "finite P" "(w', Inr P # I') \<Turnstile> φ" by auto
moreover
with FEXISTS.prems have "(w, Inr P # I) \<Turnstile> φ"
apply (intro iffD2[OF FEXISTS.IH[of w "Inr P # I" w' "Inr P # I'"]])
apply (auto simp: stream_map2_alt split: sum.splits) []
apply (auto split: sum.splits split_if_asm) []
apply (blast dest!: wf_interp_for_formula_FEXISTS)
apply (blast dest!: wf_interp_for_formula_FEXISTS[of I', unfolded length])
apply assumption
done
ultimately show "(w, I) \<Turnstile> FEXISTS φ" by auto
qed
qed auto

lemma langWS1S_FOr:
assumes "wf_formula n (FOr φ1 φ2)"
shows "langWS1S n (FOr φ1 φ2) ⊆
(langWS1S n φ1 ∪ langWS1S n φ2) ∩ \<Union>{enc (w, I) | w I. length I = n ∧ wf_interp_for_formula (w, I) (FOr φ1 φ2)}"

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

lemma langWS1S_FAnd:
assumes "wf_formula n (FAnd φ1 φ2)"
shows "langWS1S n (FAnd φ1 φ2) ⊆
langWS1S n φ1 ∩ langWS1S n φ2 ∩ \<Union>{enc (w, I) | w I. length I = n ∧ wf_interp_for_formula (w, I) (FAnd φ1 φ2)}"

using assms unfolding langWS1S_def by (fastforce simp del: enc.simps)

subsection {* From WS1S to Regular expressions *}

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

| "rexp_of n (FLess 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 (FLess m1 m2)))"

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

| "rexp_of n (FNot φ) = Inter (rexp.Not (rexp_of n φ)) (ENC n (FNot φ))"
| "rexp_of n (FOr φ1 φ2) = Inter (Plus (rexp_of n φ1) (rexp_of n φ2)) (ENC n (FOr φ1 φ2))"
| "rexp_of n (FAnd φ1 φ2) = INTERSECT [rexp_of n φ1, rexp_of n φ2, ENC n (FAnd φ1 φ2)]"
| "rexp_of n (FExists φ) = samequot_exec (any, replicate n False) (Pr (rexp_of (n + 1) φ))"
| "rexp_of n (FEXISTS φ) = samequot_exec (any, replicate n False) (Pr (rexp_of (n + 1) φ))"

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

| "rexp_of_alt n (FLess 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 (FIn m M) =
TIMES [rexp.Not Zero, arbitrary_except n [(min m M, True), (max m M, True)] Σ, rexp.Not Zero]"

| "rexp_of_alt n (FNot φ) = rexp.Not (rexp_of_alt n φ)"
| "rexp_of_alt n (FOr φ1 φ2) = Plus (rexp_of_alt n φ1) (rexp_of_alt n φ2)"
| "rexp_of_alt n (FAnd φ1 φ2) = Inter (rexp_of_alt n φ1) (rexp_of_alt n φ2)"
| "rexp_of_alt n (FExists φ) = samequot_exec (any, replicate n False) (Pr (Inter (rexp_of_alt (n + 1) φ) (ENC (Suc n) φ)))"
| "rexp_of_alt n (FEXISTS φ) = samequot_exec (any, replicate n False) (Pr (Inter (rexp_of_alt (n + 1) φ) (ENC (Suc n) φ)))"

definition "rexp_of' n φ = Inter (rexp_of_alt n φ) (ENC n φ)"

lemma enc_eqI:
assumes "x ∈ enc (w, I)" "x ∈ enc (w', I')" "wf_interp_for_formula (w, I) φ" "wf_interp_for_formula (w', I') φ"
"length I = length I'"
shows "enc (w, I) = enc (w', I')"
proof -
from assms have "stream_enc (w, I) = stream_enc (w', I')"
by (intro box_equals[OF _ stream_enc_enc[symmetric] stream_enc_enc[symmetric]]) auto
thus ?thesis using assms(5) by auto
qed

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

by (intro stream_enc_welldef) (auto simp del: stream_enc.simps intro!: enc_stream_enc)

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

by (intro enc_eq_welldef[OF enc_eqI])

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 intro!: wf_samequot_exec,
auto simp: σ_def set_n_lists image_iff)

theorem langWS1S_rexp_of: "wf_formula n φ ==> langWS1S n φ = lang n (rexp_of n φ)"
(is "_ ==> _ = ?L n φ")
proof (induct φ arbitrary: n)
case (FQ a m)
show ?case
proof (intro equalityI subsetI)
fix x assume "x ∈ langWS1S n (FQ a m)"
then obtain w I where
*: "x ∈ enc (w, I)" "wf_interp_for_formula (w, I) (FQ a m)" "length I = n" "(w, I) \<Turnstile> FQ a m"
unfolding langWS1S_def by blast
hence x_alt: "x = map (split (enc_atom I)) (zip [0 ..< length x] (stake (length x) (w @- same any)))"
by (intro encD) auto
from FQ(1) *(2,4) obtain p where p: "I ! m = Inl p"
by (auto simp: all_set_conv_all_nth enc_def split: sum.splits)
with FQ(1) * have p_less: "p < length x"
by (auto simp del: stream_enc.simps intro: trans_less_add1[OF less_length_cutSame_Inl])
hence enc_atom: "x ! p = enc_atom I p ((w @- same any) !! p)" (is "_ = enc_atom _ _ ?p")
by (subst x_alt, simp)
with *(1) p_less(1) have "x = take p x @ [enc_atom I p ?p] @ drop (p + 1) x"
using id_take_nth_drop[of p x] by auto
moreover
from *(2,3,4) FQ(1) p have "[enc_atom I p ?p] ∈ lang n (arbitrary_except n [(m, True)] [a])"
by (intro enc_atom_lang_arbitrary_except_True) auto
moreover from *(2,3) have "take p x ∈ lang n (rexp.Not Zero)"
by (subst x_alt) (auto simp: in_set_zip shift_snth intro!: enc_atom_σ dest!: in_set_takeD)
moreover from *(2,3) have "drop (Suc p) x ∈ lang n (rexp.Not Zero)"
by (subst x_alt) (auto simp: in_set_zip shift_snth intro!: enc_atom_σ dest!: in_set_dropD)
ultimately show "x ∈ ?L n (FQ a m)" using *(1,2,3)
unfolding rexp_of.simps lang.simps(5,8) rexp_of_list.simps lang_ENC[OF FQ]
by (auto elim: ssubst simp del: o_apply append.simps lang.simps enc.simps)
next
fix x let ?x = "x @- same (any, replicate n False)"
assume x: "x ∈ ?L n (FQ a m)"
with FQ obtain w I where
I: "x ∈ enc (w, I)" "length I = n" "wf_interp_for_formula (w, I) (FQ a m)"
unfolding rexp_of.simps lang.simps lang_ENC[OF FQ] by fastforce
hence stream_enc: "stream_enc (w, I) = ?x" using stream_enc_enc by auto
from I FQ obtain p where m: "I ! m = Inl p" "m < length I" by (auto split: sum.splits)
with I have "wf_interp_for_formula (dec_word ?x, stream_dec n {m} ?x) (FQ a m)" unfolding I(1)
using enc_wf_interp[OF FQ(1)[folded I(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 FQ(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
from v have u: "length u1 < length x" by auto
{ from v have "snd (x ! length u1) ! m" by auto
moreover
from m I have "p < length x" "snd (x ! p) ! m" by (auto dest: enc_Inl simp del: enc.simps)
moreover
from m I have ex1: "∃!p. snd (stream_enc (w, I) !! p) ! m" by (intro stream_enc_unique) auto
ultimately have "p = length u1" unfolding stream_enc using u I(3) by auto
} note * = this
from v have "v = x ! length u1" by simp
with u have "?x !! length u1 = v" by (auto simp: shift_snth)
with * m I v(3) have "(dec_word ?x, stream_dec n {m} ?x) \<Turnstile> FQ a m"
using stream_enc_enc[OF _ I(1), symmetric] less_length_cutSame[of w "any" "length u1" a]
by (auto simp del: enc.simps stream_enc.simps simp: dec_word_stream_enc dest!:
stream_dec_enc_Inl stream_dec_not_Inr split: sum.splits)
(auto simp: stream_map2_alt cutSame_def)
moreover from m I(2)
have stream_enc_dec: "stream_enc (dec_word (stream_enc (w, I)), stream_dec n {m} (stream_enc (w, I))) = stream_enc (w, I)"
by (intro stream_enc_dec)
(auto simp: stream_map2_alt sdrop_snth shift_snth intro: stream_enc_unique,
auto simp: stream_map2_szip stream.set_natural')
moreover from I have "wf_word n x" unfolding wf_word by (auto elim: enc_set_σ simp del: enc.simps)
ultimately show "x ∈ langWS1S n (FQ a m)" unfolding langWS1S_def using m I(1,3)
by (auto simp del: enc.simps stream_enc.simps intro!: exI[of _ "enc (dec_word ?x, stream_dec n {m} ?x)"],
fastforce simp del: enc.simps stream_enc.simps,
auto simp del: stream_enc.simps simp: stream_enc[symmetric] I(2))
qed
next
case (FLess m m')
show ?case
proof (cases "m = m'")
case False
thus ?thesis
proof (intro equalityI subsetI)
fix x assume "x ∈ langWS1S n (FLess m m')"
then obtain w I where
*: "x ∈ enc (w, I)" "wf_interp_for_formula (w, I) (FLess m m')" "length I = n" "(w, I) \<Turnstile> FLess m m'"
unfolding langWS1S_def by blast
hence x_alt: "x = map (split (enc_atom I)) (zip [0 ..< length x] (stake (length x) (w @- same any)))"
by (intro encD) auto
from FLess(1) *(2,4) obtain p q where pq: "I ! m = Inl p" "I ! m' = Inl q" "p < q"
by (auto simp: all_set_conv_all_nth enc_def split: sum.splits)
with FLess(1) *(1,2,3) have pq_less: "p < length x" "q < length x"
by (auto simp del: stream_enc.simps intro!: trans_less_add1[OF less_length_cutSame_Inl])
hence enc_atom: "x ! p = enc_atom I p ((w @- same any) !! p)" (is "_ = enc_atom _ _ ?p")
"x ! q = enc_atom I q ((w @- same any) !! q)" (is "_ = enc_atom _ _ ?q") by (subst x_alt, simp)+
with *(1) pq_less(1) have "x = take p x @ [enc_atom I p ?p] @ drop (p + 1) x"
using id_take_nth_drop[of p x] by auto
also have "drop (p + 1) x = take (q - p - 1) (drop (p + 1) x) @
[enc_atom I q ?q] @ drop (q - p) (drop (p + 1) x)"
(is "_ = ?LHS")
using id_take_nth_drop[of "q - p - 1" "drop (p + 1) x"] pq pq_less(2) enc_atom(2) by auto
finally have "x = take p x @ [enc_atom I p ?p] @ ?LHS" .
moreover from *(2,3) FLess(1) pq(1)
have "[enc_atom I p ?p] ∈ lang n (arbitrary_except n [(m, True)] Σ)"
by (intro enc_atom_lang_arbitrary_except_True) (auto simp: shift_snth)
moreover from *(2,3) FLess(1) pq(2)
have "[enc_atom I q ?q] ∈ lang n (arbitrary_except n [(m', True)] Σ)"
by (intro enc_atom_lang_arbitrary_except_True) (auto simp: shift_snth)
moreover from *(2,3) have "take p x ∈ lang n (rexp.Not Zero)"
by (subst x_alt) (auto simp: in_set_zip shift_snth intro!: enc_atom_σ dest!: in_set_takeD)
moreover from *(2,3) have "take (q - p - 1) (drop (Suc p) x) ∈ lang n (rexp.Not Zero)"
by (subst x_alt) (auto simp: in_set_zip shift_snth intro!: enc_atom_σ dest!: in_set_dropD in_set_takeD)
moreover from *(2,3) have "drop (q - p) (drop (Suc p) x) ∈ lang n (rexp.Not Zero)"
by (subst x_alt) (auto simp: in_set_zip shift_snth intro!: enc_atom_σ dest!: in_set_dropD)
ultimately show "x ∈ ?L n (FLess m m')" using *(1,2,3)
unfolding rexp_of.simps lang.simps(5,8) rexp_of_list.simps Int_Diff lang_ENC[OF FLess] if_not_P[OF False]
by (auto elim: ssubst simp del: o_apply append.simps lang.simps enc.simps)
next
fix x let ?x = "x @- same (any, replicate n False)"
assume x: "x ∈ ?L n (FLess m m')"
with FLess obtain w I where
I: "x ∈ enc (w, I)" "length I = n" "wf_interp_for_formula (w, I) (FLess m m')"
unfolding rexp_of.simps lang.simps lang_ENC[OF FLess] if_not_P[OF False] by fastforce
hence stream_enc: "stream_enc (w, I) = x @- same (any, replicate n False)" using stream_enc_enc by auto
from I FLess 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 I have "wf_interp_for_formula (dec_word ?x, stream_dec n {m, m'} ?x) (FLess m m')" unfolding I(1)
using enc_wf_interp[OF FLess(1)[folded I(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 FLess(1) obtain v v' where v: "x = u1 @ [v] @ u2 @ [v'] @ u3"
"snd v ! m" "snd v' ! m'" "fst v ∈ set Σ" "fst v' ∈ set Σ"
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 I have "p < length x" "snd (x ! p) ! m" by (auto dest: enc_Inl simp del: enc.simps)
moreover
from m I have ex1: "∃!p. snd (stream_enc (w, I) !! p) ! m" by (intro stream_enc_unique) auto
ultimately have "p = length u1" unfolding stream_enc using u I(3) by auto
}
{ from v have "snd (x ! ?u') ! m'" by (auto simp: nth_append)
moreover
from m I have "p' < length x" "snd (x ! p') ! m'" by (auto dest: enc_Inl simp del: enc.simps)
moreover
from m I have ex1: "∃!p. snd (stream_enc (w, I) !! p) ! m'" unfolding I(1) by (intro stream_enc_unique) auto
ultimately have "p' = ?u'" unfolding stream_enc using u' I(3) by auto (metis shift_snth_less)
} note * = this `p = length u1`
with m I have "(dec_word ?x, stream_dec n {m, m'} ?x) \<Turnstile> FLess m m'"
using stream_enc_enc[OF _ I(1), symmetric]
by (auto dest: stream_dec_not_Inr stream_dec_enc_Inl split: sum.splits simp del: stream_enc.simps)
moreover from m I(2)
have stream_enc_dec: "stream_enc (dec_word (stream_enc (w, I)), stream_dec n {m, m'} (stream_enc (w, I))) = stream_enc (w, I)"
by (intro stream_enc_dec)
(auto simp: stream_map2_alt sdrop_snth shift_snth intro: stream_enc_unique,
auto simp: stream_map2_szip stream.set_natural')
moreover from I have "wf_word n x" unfolding wf_word by (auto elim: enc_set_σ simp del: enc.simps)
ultimately show "x ∈ langWS1S n (FLess m m')" unfolding langWS1S_def using m I(1,3)
by (auto simp del: enc.simps stream_enc.simps intro!: exI[of _ "enc (dec_word ?x, stream_dec n {m, m'} ?x)"],
fastforce simp del: enc.simps stream_enc.simps,
auto simp del: stream_enc.simps simp: stream_enc[symmetric] I(2))
qed
qed (simp add: langWS1S_def del: o_apply)
next
case (FIn m M)
show ?case
proof (intro equalityI subsetI)
fix x assume "x ∈ langWS1S n (FIn m M)"
then obtain w I where
*: "x ∈ enc (w, I)" "wf_interp_for_formula (w, I) (FIn m M)" "length I = n" "(w, I) \<Turnstile> FIn m M"
unfolding langWS1S_def by blast
hence x_alt: "x = map (split (enc_atom I)) (zip [0 ..< length x] (stake (length x) (w @- same any)))"
by (intro encD) auto
from FIn(1) *(2,4) obtain p P where p: "I ! m = Inl p" "I ! M = Inr P" "p ∈ P"
by (auto simp: all_set_conv_all_nth enc_def split: sum.splits)
with FIn(1) *(1,2,3) have p_less: "p < length x" "∀p ∈ P. p < length x"
by (auto simp del: stream_enc.simps intro: trans_less_add1[OF less_length_cutSame_Inl]
trans_less_add1[OF bspec[OF less_length_cutSame_Inr]])
hence enc_atom: "x ! p = enc_atom I p ((w @- same any) !! p)" (is "_ = enc_atom _ _ ?p")
"∀p ∈ P. x ! p = enc_atom I p ((w @- same any) !! p)" (is "Ball _ (λp. _ = enc_atom _ _ (?P p))")
by (subst x_alt, simp)+
with *(1) p_less(1) have "x = take p x @ [enc_atom I p ?p] @ drop (p + 1) x"
using id_take_nth_drop[of p x] by auto
moreover
have "[enc_atom I p ?p] ∈ lang n (arbitrary_except n [(min m M, True), (max m M, True)] Σ)"
proof (cases "m < M")
case True with *(2,3) FIn(1) p show ?thesis
by (intro enc_atom_lang_arbitrary_except_True2) (auto simp: min_absorb1 max_absorb2 shift_snth)
next
case False with *(2,3) FIn(1) p show ?thesis
by (intro enc_atom_lang_arbitrary_except_True2) (auto simp: min_absorb2 max_absorb1 shift_snth)
qed
moreover from *(2,3) have "take p x ∈ lang n (rexp.Not Zero)"
by (subst x_alt) (auto simp: in_set_zip shift_snth intro!: enc_atom_σ dest!: in_set_takeD)
moreover from *(2,3) have "drop (Suc p) x ∈ lang n (rexp.Not Zero)"
by (subst x_alt) (auto simp: in_set_zip shift_snth intro!: enc_atom_σ dest!: in_set_dropD)
ultimately show "x ∈ ?L n (FIn m M)" using *(1,2,3)
unfolding rexp_of.simps lang.simps(5,8) rexp_of_list.simps Int_Diff lang_ENC[OF FIn]
by (auto elim: ssubst simp del: o_apply append.simps lang.simps enc.simps)
next
fix x let ?x = "x @- same (any, replicate n False)"
assume x: "x ∈ ?L n (FIn m M)"
with FIn obtain w I where
I: "x ∈ enc (w, I)" "length I = n" "wf_interp_for_formula (w, I) (FIn m M)"
unfolding rexp_of.simps lang.simps lang_ENC[OF FIn] by fastforce
hence stream_enc: "stream_enc (w, I) = ?x" using stream_enc_enc by auto
from I FIn 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 I have "wf_interp_for_formula (dec_word ?x, stream_dec n {m} ?x) (FIn m M)" unfolding I(1)
using enc_wf_interp[OF FIn(1)[folded I(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 FIn(1) obtain v where v: "x = u1 @ [v] @ u2" and "snd v ! min m M" "snd v ! max m M" "fst v ∈ set Σ"
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 I have "p < length x" "snd (x ! p) ! m" by (auto dest: enc_Inl simp del: enc.simps)
moreover
from m I have ex1: "∃!p. snd (stream_enc (w, I) !! p) ! m" by (intro stream_enc_unique) auto
ultimately have "p = length u1" unfolding stream_enc using u I(3) by auto
} note * = this
from v v' have "v = x ! length u1" by simp
with v'(2) m(3,4) u I(1,3) have "length u1 ∈ P" by (auto dest!: enc_Inr simp del: enc.simps)
with * m I have "(dec_word ?x, stream_dec n {m} ?x) \<Turnstile> FIn m M"
using stream_enc_enc[OF _ I(1), symmetric]
by (auto simp del: stream_enc.simps dest: stream_dec_not_Inr stream_dec_not_Inl
stream_dec_enc_Inl stream_dec_enc_Inr split: sum.splits)
moreover from m I(2)
have stream_enc_dec: "stream_enc (dec_word (stream_enc (w, I)), stream_dec n {m} (stream_enc (w, I))) = stream_enc (w, I)"
by (intro stream_enc_dec)
(auto simp: stream_map2_alt sdrop_snth shift_snth intro: stream_enc_unique,
auto simp: stream_map2_szip stream.set_natural')
moreover from I have "wf_word n x" unfolding wf_word by (auto elim: enc_set_σ simp del: enc.simps)
ultimately show "x ∈ langWS1S n (FIn m M)" unfolding langWS1S_def using m I(1,3)
by (auto simp del: enc.simps stream_enc.simps intro!: exI[of _ "enc (dec_word ?x, stream_dec n {m} ?x)"],
fastforce simp del: enc.simps stream_enc.simps,
auto simp del: stream_enc.simps simp: stream_enc[symmetric] I(2))
qed
next
case (FOr φ1 φ2)
from FOr(3) have IH1: "langWS1S n φ1 = lang n (rexp_of n φ1)"
by (intro FOr(1)) auto
from FOr(3) have IH2: "langWS1S n φ2 = lang n (rexp_of n φ2)"
by (intro FOr(2)) auto
show ?case
proof (intro equalityI subsetI)
fix x assume "x ∈ langWS1S n (FOr φ1 φ2)" thus "x ∈ lang n (rexp_of n (FOr φ1 φ2))"
using langWS1S_FOr[OF FOr(3)] unfolding lang_ENC[OF FOr(3)] rexp_of.simps lang.simps IH1 IH2 by blast
next
fix x assume "x ∈ lang n (rexp_of n (FOr φ1 φ2))"
then obtain w I where or: "x ∈ langWS1S n φ1 ∨ x ∈ langWS1S n φ2" and I: "x ∈ enc (w, I)" "length I = n"
"wf_interp_for_formula (w, I) (FOr φ1 φ2)"
unfolding lang_ENC[OF FOr(3)] rexp_of.simps lang.simps IH1 IH2 Int_Diff by auto
have "(w, I) \<Turnstile> φ1 ∨ (w, I) \<Turnstile> φ2"
proof (intro mp[OF disj_mono[OF impI impI] or])
assume "x ∈ langWS1S n φ1"
with I FOr(3) show "(w, I) \<Turnstile> φ1"
unfolding langWS1S_def I(1) wf_interp_for_formula_FOr
by (auto dest!: enc_welldef[of x w I _ _ φ1] simp del: enc.simps)
next
assume "x ∈ langWS1S n φ2"
with I FOr(3) show "(w, I) \<Turnstile> φ2"
unfolding langWS1S_def I(1) wf_interp_for_formula_FOr
by (auto dest!: enc_welldef[of x w I _ _ φ2] simp del: enc.simps)
qed
with I show "x ∈ langWS1S n (FOr φ1 φ2)" unfolding langWS1S_def by auto
qed
next
case (FAnd φ1 φ2)
from FAnd(3) have IH1: "langWS1S n φ1 = lang n (rexp_of n φ1)"
by (intro FAnd(1)) auto
from FAnd(3) have IH2: "langWS1S n φ2 = lang n (rexp_of n φ2)"
by (intro FAnd(2)) auto
show ?case
proof (intro equalityI subsetI)
fix x assume "x ∈ langWS1S n (FAnd φ1 φ2)" thus "x ∈ lang n (rexp_of n (FAnd φ1 φ2))"
using langWS1S_FAnd[OF FAnd(3)]
unfolding lang_ENC[OF FAnd(3)] rexp_of.simps rexp_of_list.simps lang.simps IH1 IH2 Int_assoc
by blast
next
fix x assume "x ∈ lang n (rexp_of n (FAnd φ1 φ2))"
then obtain w I where "and": "x ∈ langWS1S n φ1 ∧ x ∈ langWS1S n φ2" and I: "x ∈ enc (w, I)" "length I = n"
"wf_interp_for_formula (w, I) (FAnd φ1 φ2)"
unfolding lang_ENC[OF FAnd(3)] rexp_of.simps rexp_of_list.simps lang.simps IH1 IH2 Int_Diff by auto
have "(w, I) \<Turnstile> φ1 ∧ (w, I) \<Turnstile> φ2"
proof (intro mp[OF conj_mono[OF impI impI] "and"])
assume "x ∈ langWS1S n φ1"
with I FAnd(3) show "(w, I) \<Turnstile> φ1"
unfolding langWS1S_def I(1) wf_interp_for_formula_FAnd
by (auto dest!: enc_welldef[of x w I _ _ φ1] simp del: enc.simps)
next
assume "x ∈ langWS1S n φ2"
with I FAnd(3) show "(w, I) \<Turnstile> φ2"
unfolding langWS1S_def I(1) wf_interp_for_formula_FAnd
by (auto dest!: enc_welldef[of x w I _ _ φ2] simp del: enc.simps)
qed
with I show "x ∈ langWS1S n (FAnd φ1 φ2)" unfolding langWS1S_def by auto
qed
next
case (FNot φ)
hence IH: "?L n φ = langWS1S n φ" by simp
show ?case
proof (intro equalityI subsetI)
fix x assume "x ∈ langWS1S n (FNot φ)"
then obtain w I where
*: "x ∈ enc (w, I)" "wf_interp_for_formula (w, I) φ" "length I = n" and unsat: "¬ (w, I) \<Turnstile> φ"
unfolding langWS1S_def by auto
{ assume "x ∈ ?L n φ"
hence "(w, I) \<Turnstile> φ" using enc_welldef[of x w I _ _ φ, OF *(1) _ _ _ *(2)] FNot(2)
unfolding *(3) langWS1S_def IH by auto
}
with unsat have "x ∉ ?L n φ" by blast
with * show "x ∈ ?L n (FNot φ)" unfolding rexp_of.simps lang.simps using lang_ENC[OF FNot(2)]
by (auto simp del: enc.simps simp: comp_def intro: enc_set_σ)
next
fix x assume "x ∈ ?L n (FNot φ)"
with IH have "x ∈ lang n (ENC n (FNot φ))" and x: "x ∉ langWS1S n φ" by (auto simp del: o_apply)
then obtain w I where *: "x ∈ enc (w, I)" "wf_interp_for_formula (w, I) (FNot φ)" "length I = n"
unfolding lang_ENC[OF FNot(2)] by blast
{ assume "¬ (w, I) \<Turnstile> FNot φ"
with * have "x ∈ langWS1S n φ" unfolding langWS1S_def by auto
}
with x * show "x ∈ langWS1S n (FNot φ)" unfolding langWS1S_def by blast
qed
next
case (FExists φ)
have σ: "(any, replicate n False) ∈ (set o σ Σ) n" by (auto simp: σ_def set_n_lists image_iff)
from FExists(2) have wf: "wf n (Pr (rexp_of (Suc n) φ))" by (fastforce intro: wf_rexp_of)
note lang_quot = lang_samequot_exec[OF wf σ]
show ?case
proof (intro equalityI subsetI)
fix x assume "x ∈ langWS1S n (FExists φ)"
then obtain w I p where
*: "x ∈ enc (w, I)" "wf_interp_for_formula (w, I) (FExists φ)" "length I = n" "(w, Inl p # I) \<Turnstile> φ"
unfolding langWS1S_def by auto
with FExists(2) have "enc (w, Inl p # I) ⊆ ?L (Suc n) φ"
by (subst FExists(1)[of "Suc n", symmetric])
(fastforce simp del: enc.simps simp: langWS1S_def nth_Cons' intro!: exI[of _ "enc (w, Inl p # I)"])+
thus "x ∈ ?L n (FExists φ)" using *(1,2,3)
by (auto simp: lang_quot simp del: o_apply enc.simps elim: set_mp[OF SAMEQUOT_mono[OF image_mono]])
next
fix x assume "x ∈ ?L n (FExists φ)"
then obtain x' m where "x' ∈ ?L (Suc n) φ" and
x: "x = fin_cutSame (any, replicate n False) (map π x') @ replicate m (any, replicate n False)"
by (auto simp: lang_quot SAMEQUOT_def simp del: o_apply enc.simps)
with FExists(2) have "x' ∈ langWS1S (Suc n) φ"
by (intro subsetD[OF equalityD2[OF FExists(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" "(w, I') \<Turnstile> φ"
unfolding langWS1S_def by blast
moreover then obtain I0 I where "I' = I0 # I" by (cases I') auto
moreover with FExists(2) *(2) obtain p where "I0 = Inl p"
by (auto simp: nth_Cons' split: sum.splits split_if_asm)
ultimately have "x ∈ enc (w, I)" "wf_interp_for_formula (w, I) (FExists φ)" "length I = n"
"(w, I) \<Turnstile> FExists φ" using FExists(2) fin_cutSame_tl[OF ex_Loop_stream_enc, of "Inl p # I" w]
unfolding x by (auto simp add: π_def nth_Cons' split: split_if_asm)
thus "x ∈ langWS1S n (FExists φ)" unfolding langWS1S_def by (auto intro!: exI[of _ I])
qed
next
case (FEXISTS φ)
have σ: "(any, replicate n False) ∈ (set o σ Σ) n" by (auto simp: σ_def set_n_lists image_iff)
from FEXISTS(2) have wf: "wf n (Pr (rexp_of (Suc n) φ))" by (fastforce intro: wf_rexp_of)
note lang_quot = lang_samequot_exec[OF wf σ]
show ?case
proof (intro equalityI subsetI)
fix x assume "x ∈ langWS1S n (FEXISTS φ)"
then obtain w I P where
*: "x ∈ enc (w, I)" "wf_interp_for_formula (w, I) (FEXISTS φ)" "length I = n" "finite P" "(w, Inr P # I) \<Turnstile> φ"
unfolding langWS1S_def by auto
with FEXISTS(2) have "enc (w, Inr P # I) ⊆ ?L (Suc n) φ"
by (subst FEXISTS(1)[of "Suc n", symmetric])
(fastforce simp del: enc.simps simp: langWS1S_def nth_Cons' intro!: exI[of _ "enc (w, Inr P # I)"])+
thus "x ∈ ?L n (FEXISTS φ)" using *(1,2,3,4)
by (auto simp: lang_quot simp del: o_apply enc.simps elim: set_mp[OF SAMEQUOT_mono[OF image_mono]])
next
fix x assume "x ∈ ?L n (FEXISTS φ)"
then obtain x' m where "x' ∈ ?L (Suc n) φ" and
x: "x = fin_cutSame (any, replicate n False) (map π x') @ replicate m (any, replicate n False)"
by (auto simp: lang_quot SAMEQUOT_def simp del: o_apply enc.simps)
with FEXISTS(2) have "x' ∈ langWS1S (Suc n) φ"
by (intro subsetD[OF equalityD2[OF FEXISTS(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" "(w, I') \<Turnstile> φ"
unfolding langWS1S_def by blast
moreover then obtain I0 I where "I' = I0 # I" by (cases I') auto
moreover with FEXISTS(2) *(2) obtain P where "I0 = Inr P" "finite P"
by (auto simp: nth_Cons' split: sum.splits split_if_asm)
ultimately have "x ∈ enc (w, I)" "wf_interp_for_formula (w, I) (FEXISTS φ)" "length I = n"
"(w, I) \<Turnstile> FEXISTS φ" using FEXISTS(2) fin_cutSame_tl[OF ex_Loop_stream_enc, of "Inr P # I"]
unfolding x by (auto simp: nth_Cons' π_def split: split_if_asm)
thus "x ∈ langWS1S n (FEXISTS φ)" unfolding langWS1S_def by (auto intro!: exI[of _ I])
qed
qed

lemma wf_rexp_of_alt: "wf_formula n φ ==> wf n (rexp_of_alt n φ)"
by (induct φ arbitrary: n)
(auto simp: wf_rexp_ENC intro: wf_rexp_arbitrary_except intro!: wf_samequot_exec,
auto simp: σ_def set_n_lists image_iff)

lemma wf_rexp_of': "wf_formula n φ ==> wf n (rexp_of' n φ)"
unfolding rexp_of'_def by (auto intro: wf_rexp_of_alt wf_rexp_ENC)

lemma ENC_FNot: "ENC n (FNot φ) = ENC n φ"
unfolding ENC_def by auto

lemma ENC_FAnd:
"wf_formula n (FAnd φ ψ) ==> lang n (ENC n (FAnd φ ψ)) ⊆ lang n (ENC n φ) ∩ lang n (ENC n ψ)"
proof
fix x assume wf: "wf_formula n (FAnd φ ψ)" and x: "x ∈ lang n (ENC n (FAnd φ ψ))"
hence wf1: "wf_formula n φ" and wf2: "wf_formula n ψ" by auto
from x obtain w I where I: "x ∈ enc (w, I)" "wf_interp_for_formula (w, I) (FAnd φ ψ)" "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_FAnd by auto
thus "x ∈ lang n (ENC n φ) ∩ lang n (ENC n ψ)"
unfolding lang_ENC[OF wf1] lang_ENC[OF wf2] using I by blast
qed

lemma ENC_FOr:
"wf_formula n (FOr φ ψ) ==> lang n (ENC n (FOr φ ψ)) ⊆ lang n (ENC n φ) ∩ lang n (ENC n ψ)"
proof
fix x assume wf: "wf_formula n (FOr φ ψ)" and x: "x ∈ lang n (ENC n (FOr φ ψ))"
hence wf1: "wf_formula n φ" and wf2: "wf_formula n ψ" by auto
from x obtain w I where I: "x ∈ enc (w, I)" "wf_interp_for_formula (w, I) (FOr φ ψ)" "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_FOr by auto
thus "x ∈ lang n (ENC n φ) ∩ lang n (ENC n ψ)"
unfolding lang_ENC[OF wf1] lang_ENC[OF wf2] using I by blast
qed

lemma ENC_FExists:
"wf_formula n (FExists φ) ==> lang n (ENC n (FExists φ)) =
SAMEQUOT (any, replicate n False) (map π ` lang (Suc n) (ENC (Suc n) φ))"
(is "_ ==> ?L = ?R")
proof (intro equalityI subsetI)
fix x assume wf: "wf_formula n (FExists φ)" and x: "x ∈ ?L"
hence wf1: "wf_formula (Suc n) φ" by auto
from x obtain w I where I: "x ∈ enc (w, I)" "wf_interp_for_formula (w, I) (FExists φ)" "length I = n"
using lang_ENC[OF wf] by auto
from I(2) obtain p where "wf_interp_for_formula (w, Inl p # I) φ"
using wf_interp_for_formula_FExists[OF wf[folded I(3)]] by blast
with I(3) show "x ∈ ?R"
unfolding lang_ENC[OF wf1] using I(1) tl_enc[of "Inl p" I, symmetric]
by (simp del: enc.simps)
(fastforce simp del: enc.simps elim!: set_rev_mp[OF _ SAMEQUOT_mono[OF image_mono]]
intro: exI[of _ "enc (w, Inl p # I)"])
next
fix x assume wf: "wf_formula n (FExists φ)" and x: "x ∈ ?R"
hence wf1: "wf_formula (Suc n) φ" and "0 ∈ FOV φ" by auto
from x obtain w I where I: "x ∈ SAMEQUOT (any, replicate n False) (map π ` enc (w, I))"
"wf_interp_for_formula (w, I) φ" "length I = Suc n"
using lang_ENC[OF wf1] unfolding SAMEQUOT_def by fast
with `0 ∈ FOV φ` obtain p I' where I': "I = Inl p # I'" by (cases I) (fastforce split: sum.splits)+
with I have wtlI: "x ∈ enc (w, I')" "length I' = n" using tl_enc[of "Inl p" I' w] by auto
have "wf_interp_for_formula (w, I') (FExists φ)"
using wf_interp_for_formula_FExists[OF wf[folded wtlI(2)]]
wf_interp_for_formula_any_Inl[OF I(2)[unfolded I']] ..
with wtlI show "x ∈ ?L" unfolding lang_ENC[OF wf] by blast
qed

lemma ENC_FEXISTS:
"wf_formula n (FEXISTS φ) ==> lang n (ENC n (FEXISTS φ)) =
SAMEQUOT (any, replicate n False) (map π ` lang (Suc n) (ENC (Suc n) φ))"
(is "_ ==> ?L = ?R")
proof (intro equalityI subsetI)
fix x assume wf: "wf_formula n (FEXISTS φ)" and x: "x ∈ ?L"
hence wf1: "wf_formula (Suc n) φ" by auto
from x obtain w I where I: "x ∈ enc (w, I)" "wf_interp_for_formula (w, I) (FEXISTS φ)" "length I = n"
using lang_ENC[OF wf] by auto
from I(2) obtain P where "wf_interp_for_formula (w, Inr P # I) φ"
using wf_interp_for_formula_FEXISTS[OF wf[folded I(3)]] by blast
with I(3) show "x ∈ ?R"
unfolding lang_ENC[OF wf1] using I(1) tl_enc[of "Inr P" I, symmetric]
by (simp del: enc.simps)
(fastforce simp del: enc.simps elim!: set_rev_mp[OF _ SAMEQUOT_mono[OF image_mono]]
intro: exI[of _ "enc (w, Inr P # I)"])
next
fix x assume wf: "wf_formula n (FEXISTS φ)" and x: "x ∈ ?R"
hence wf1: "wf_formula (Suc n) φ" and "0 ∈ SOV φ" by auto
from x obtain w I where I: "x ∈ SAMEQUOT (any, replicate n False) (map π ` enc (w, I))"
"wf_interp_for_formula (w, I) φ" "length I = Suc n"
using lang_ENC[OF wf1] unfolding SAMEQUOT_def by fast
with `0 ∈ SOV φ` obtain P I' where I': "I = Inr P # I'" by (cases I) (fastforce split: sum.splits)+
with I have wtlI: "x ∈ enc (w, I')" "length I' = n" using tl_enc[of "Inr P" I' w] by auto
have "wf_interp_for_formula (w, I') (FEXISTS φ)"
using wf_interp_for_formula_FEXISTS[OF wf[folded wtlI(2)]]
wf_interp_for_formula_any_Inr[OF I(2)[unfolded I']] ..
with wtlI show "x ∈ ?L" unfolding lang_ENC[OF wf] by blast
qed

lemma langWS1S_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 (FNot φ)
hence "wf_formula n φ" by simp
with FNot.IH show ?case unfolding rexp_of.simps rexp_of_alt.simps lang.simps ENC_FNot by blast
next
case (FAnd φ1 φ2)
hence wf1: "wf_formula n φ1" and wf2: "wf_formula n φ2" by force+
from FAnd.IH(1)[OF wf1] FAnd.IH(2)[OF wf2] show ?case using ENC_FAnd[OF FAnd.prems]
unfolding rexp_of.simps rexp_of_alt.simps lang.simps rexp_of_list.simps by blast
next
case (FOr φ1 φ2)
hence wf1: "wf_formula n φ1" and wf2: "wf_formula n φ2" by force+
from FOr.IH(1)[OF wf1] FOr.IH(2)[OF wf2] show ?case using ENC_FOr[OF FOr.prems]
unfolding rexp_of.simps rexp_of_alt.simps lang.simps by blast
next
case (FExists φ)
from FExists(2) have IH: "lang (n + 1) (rexp_of (n + 1) φ) =
lang (n + 1) (Inter (rexp_of_alt (n + 1) φ) (ENC (n + 1) φ))"
by (intro FExists.IH) auto
have σ: "(any, replicate n False) ∈ (set o σ Σ) n" by (auto simp: σ_def set_n_lists image_iff)
from FExists(2) have wf: "wf n (Pr (rexp.Inter (rexp_of_alt (n + 1) φ) (ENC (n + 1) φ)))"
"wf n (Pr (rexp_of (n + 1) φ))" by (fastforce intro!: wf_rexp_of wf_rexp_of_alt wf_rexp_ENC)+
note lang_quot = lang_samequot_exec[OF wf(1) σ] lang_samequot_exec[OF wf(2) σ]
show ?case unfolding rexp_of.simps rexp_of_alt.simps lang.simps IH lang_quot Suc_eq_plus1
ENC_FExists[OF FExists.prems, unfolded Suc_eq_plus1] by (auto simp add: SAMEQUOT_def)
next
case (FEXISTS φ)
from FEXISTS(2) have IH: "lang (n + 1) (rexp_of (n + 1) φ) =
lang (n + 1) (Inter (rexp_of_alt (n + 1) φ) (ENC (n + 1) φ))"
by (intro FEXISTS.IH) auto
have σ: "(any, replicate n False) ∈ (set o σ Σ) n" by (auto simp: σ_def set_n_lists image_iff)
from FEXISTS(2) have wf: "wf n (Pr (rexp.Inter (rexp_of_alt (n + 1) φ) (ENC (n + 1) φ)))"
"wf n (Pr (rexp_of (n + 1) φ))" by (fastforce intro: wf_rexp_of wf_rexp_of_alt wf_rexp_ENC)+
note lang_quot = lang_samequot_exec[OF wf(1) σ] lang_samequot_exec[OF wf(2) σ]
show ?case unfolding rexp_of.simps rexp_of_alt.simps lang.simps IH lang_quot Suc_eq_plus1
ENC_FEXISTS[OF FEXISTS.prems, unfolded Suc_eq_plus1] by (auto simp add: SAMEQUOT_def)
qed auto

theorem langWS1S_rexp_of': "wf_formula n φ ==> langWS1S n φ = lang n (rexp_of' n φ)"
unfolding langWS1S_rexp_of_rexp_of'[symmetric] by (rule langWS1S_rexp_of)

end

end