Theory Formulas

theory Formulas
imports Regular_Operators List_More RBT_Set
(* Author: Dmitriy Traytel *)

header {* Monadic Second-Order Logic Formulas *}

(*<*)
theory Formulas
imports Regular_Operators List_More "~~/src/HOL/Library/RBT_Set"
begin

hide_const (open) RBT.map RBT.fold
(*>*)

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

subsection {* Syntax and Semantics of MSO *}

datatype 'a formula =
FQ 'a nat
| FLess nat nat
| FIn nat nat
| FNot "'a formula"
| FOr "'a formula" "'a formula"
| FAnd "'a formula" "'a formula"
| FExists "'a formula"
| FEXISTS "'a formula"

primrec FOV :: "'a formula => nat set" where
"FOV (FQ a m) = {m}"
| "FOV (FLess m1 m2) = {m1, m2}"
| "FOV (FIn m M) = {m}"
| "FOV (FNot φ) = FOV φ"
| "FOV (FOr φ1 φ2) = FOV φ1 ∪ FOV φ2"
| "FOV (FAnd φ1 φ2) = FOV φ1 ∪ FOV φ2"
| "FOV (FExists φ) = (λx. x - 1) ` (FOV φ - {0})"
| "FOV (FEXISTS φ) = (λx. x - 1) ` FOV φ"

primrec SOV :: "'a formula => nat set" where
"SOV (FQ a m) = {}"
| "SOV (FLess m1 m2) = {}"
| "SOV (FIn m M) = {M}"
| "SOV (FNot φ) = SOV φ"
| "SOV (FOr φ1 φ2) = SOV φ1 ∪ SOV φ2"
| "SOV (FAnd φ1 φ2) = SOV φ1 ∪ SOV φ2"
| "SOV (FExists φ) = (λx. x - 1) ` SOV φ"
| "SOV (FEXISTS φ) = (λx. x - 1) ` (SOV φ - {0})"

definition "σ = (λΣ n. concat (map (λbs. map (λa. (a, bs)) Σ) (List.n_lists n [True, False])))"
definition "π = (λ(a, bs). (a, tl bs))"
definition "ε = (λΣ (a::'a, bs). if a ∈ set Σ then [(a, True # bs), (a, False # bs)] else [])"

locale formula = embed "set o (σ Σ)" π "ε Σ" for Σ :: "'a :: linorder list" +
assumes nonempty: "Σ ≠ []"
begin

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


(* Normal form: quantified variables are used in the body *)
primrec pre_wf_formula :: "nat => 'a formula => bool" where
"pre_wf_formula n (FQ a m) = (a ∈ set Σ ∧ m < n)"
| "pre_wf_formula n (FLess m1 m2) = (m1 < n ∧ m2 < n)"
| "pre_wf_formula n (FIn m M) = (m < n ∧ M < n)"
| "pre_wf_formula n (FNot φ) = pre_wf_formula n φ"
| "pre_wf_formula n (FOr φ1 φ2) = (pre_wf_formula n φ1 ∧ pre_wf_formula n φ2)"
| "pre_wf_formula n (FAnd φ1 φ2) = (pre_wf_formula n φ1 ∧ pre_wf_formula n φ2)"
| "pre_wf_formula n (FExists φ) = (pre_wf_formula (n + 1) φ ∧ 0 ∈ FOV φ ∧ 0 ∉ SOV φ)"
| "pre_wf_formula n (FEXISTS φ) = (pre_wf_formula (n + 1) φ ∧ 0 ∉ FOV φ ∧ 0 ∈ SOV φ)"

abbreviation "closed ≡ pre_wf_formula 0"

definition [simp]: "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)

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 wf_rexp_arbitrary_except:
"[|length pbs ≤ n; set xs ⊆ set Σ|] ==> wf n (arbitrary_except n pbs xs)"
by (auto simp: arbitrary_except_def) (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

end

end