Up to index of Isabelle/HOL/MSO
theory MSO_Equivalence_Checking(* Author: Dmitriy Traytel *)
header {* Deciding Equivalence of MSO Formulas *}
(*<*)
theory MSO_Equivalence_Checking
imports MSO_Normalization Smart_Constructors_Normalization
begin
(*>*)
type_synonym 'a T = "'a × bool list"
abbreviation "ε ≡ λΣ (a::'a, bs). if a ∈ set Σ then [(a, True # bs), (a, False # bs)] else []"
abbreviation "\<LL> ≡ λΣ. project.lang (set o (σ Σ)) π"
definition wf_rexp where [code del]:
"wf_rexp Σ = alphabet.wf (set o σ Σ)"
interpretation project "set o σ Σ" π
where "alphabet.wf (set o σ Σ) = wf_rexp Σ"
by (unfold_locales) (auto simp: σ_def π_def wf_rexp_def)
definition norm_ederiv where [code del]:
"norm_ederiv ≡ λΣ. embed.ederiv (ε Σ)"
interpretation embed "set o (σ (Σ :: 'a :: linorder list))" π "ε Σ"
where "embed.ederiv (ε Σ) = norm_ederiv Σ"
by (unfold_locales) (auto simp: norm_ederiv_def σ_def π_def)
definition norm_step' where [code del]:
"norm_step' ≡ λΣ. (*efficient_*)equivalence_checker.step' (σ Σ) (ε Σ) (norm :: 'a::linorder T rexp => 'a T rexp)"
definition norm_closure' where [code del]:
"norm_closure' ≡ λΣ. (*efficient_*)equivalence_checker.closure' (σ Σ) (ε Σ) (norm :: 'a::linorder T rexp => 'a T rexp)"
definition norm_check_eqv' where [code del]:
"norm_check_eqv' ≡ λΣ. (*efficient_*)equivalence_checker.check_eqv' (σ Σ) (ε Σ) (norm :: 'a::linorder T rexp => 'a T rexp)"
definition norm_step where [code del]:
"norm_step ≡ λΣ. (*efficient_*)equivalence_checker.step (σ Σ) (ε Σ) (norm :: 'a::linorder T rexp => 'a T rexp)"
definition norm_closure where [code del]:
"norm_closure ≡ λΣ. (*efficient_*)equivalence_checker.closure (σ Σ) (ε Σ) (norm :: 'a::linorder T rexp => 'a T rexp)"
definition norm_check_eqv where [code del]:
"norm_check_eqv ≡ λΣ. (*efficient_*)equivalence_checker.check_eqv (σ Σ) (ε Σ) (norm :: 'a::linorder T rexp => 'a T rexp)"
definition norm_check_eqv_counterexample where [code del]:
"norm_check_eqv_counterexample ≡ λΣ. (*efficient_*)equivalence_checker.check_eqv_counterexample (σ Σ) (ε Σ) (norm :: 'a::linorder T rexp => 'a T rexp)"
lemmas norm_defs = wf_rexp_def
norm_check_eqv_def norm_closure_def norm_step_def norm_check_eqv_counterexample_def
norm_check_eqv'_def norm_closure'_def norm_step'_def
interpretation norm: (*efficient_*)equivalence_checker "σ Σ" π "ε Σ" norm "\<LL> Σ"
where "norm.check_eqv' = norm_check_eqv' Σ"
and "norm.check_eqv = norm_check_eqv Σ"
and "norm.check_eqv_counterexample = norm_check_eqv_counterexample Σ"
and "norm.closure' = norm_closure' Σ"
and "norm.closure = norm_closure Σ"
and "norm.step' = norm_step' Σ"
and "norm.step = norm_step Σ"
by unfold_locales (auto simp: norm_defs trans[OF lang_norm[OF iffD2[OF ACI_norm_wf]] ACI_norm_lang])
interpretation MSO where "alphabet.wf (set o σ Σ) = wf_rexp Σ"
by (unfold_locales) (auto simp: σ_def π_def wf_rexp_def)
definition check_eqv where
"check_eqv Σ n φ ψ <-> wf_formula Σ n (Or φ ψ) ∧
norm_check_eqv' Σ n (Plus (rexp_of Σ n (norm φ)) One) (Plus (rexp_of Σ n (norm ψ)) One)"
definition check_eqv_counterexample where
"check_eqv_counterexample Σ n φ ψ =
norm_check_eqv_counterexample Σ n (Plus (rexp_of Σ n (norm φ)) One) (Plus (rexp_of Σ n (norm ψ)) One)"
definition check_eqv' where
"check_eqv' Σ n φ ψ <-> wf_formula Σ n (Or φ ψ) ∧
norm_check_eqv' Σ n (Plus (rexp_of' Σ n (norm φ)) One) (Plus (rexp_of' Σ n (norm ψ)) One)"
(*
definition check_eqv'' where
"check_eqv'' Σ n φ ψ <-> wf_formula Σ n (Or φ ψ) ∧
norm_check_eqv Σ n (Plus (rexp_of'' Σ n (norm φ)) One) (Plus (rexp_of'' Σ n (norm ψ)) One)"
*)
lemma lang_Plus_Zero: "\<LL> Σ n (Plus r One) = \<LL> Σ n (Plus s One) <-> \<LL> Σ n r - {[]} = \<LL> Σ n s - {[]}"
by auto
lemmas lang⇣M⇣S⇣O_rexp_of_norm = trans[OF sym[OF lang⇣M⇣S⇣O_norm] lang⇣M⇣S⇣O_rexp_of]
lemma soundness: "check_eqv Σ n φ ψ ==> lang⇣M⇣S⇣O Σ n φ = lang⇣M⇣S⇣O Σ n ψ"
by (rule box_equals[OF iffD1[OF lang_Plus_Zero, OF norm.soundness']
sym[OF trans[OF lang⇣M⇣S⇣O_rexp_of_norm]] sym[OF trans[OF lang⇣M⇣S⇣O_rexp_of_norm]]])
(auto simp: check_eqv_def split: sum.splits option.splits)
lemmas lang⇣M⇣S⇣O_rexp_of'_norm = trans[OF sym[OF lang⇣M⇣S⇣O_norm] lang⇣M⇣S⇣O_rexp_of']
lemma soundness': "check_eqv' Σ n φ ψ ==> lang⇣M⇣S⇣O Σ n φ = lang⇣M⇣S⇣O Σ n ψ"
by (rule box_equals[OF iffD1[OF lang_Plus_Zero, OF norm.soundness']
sym[OF trans[OF lang⇣M⇣S⇣O_rexp_of'_norm]] sym[OF trans[OF lang⇣M⇣S⇣O_rexp_of'_norm]]])
(auto simp: check_eqv'_def split: sum.splits option.splits)
lemma completeness:
assumes "lang⇣M⇣S⇣O Σ n φ = lang⇣M⇣S⇣O Σ n ψ" "wf_formula Σ n (Or φ ψ)"
shows "check_eqv Σ n φ ψ"
using assms(2) unfolding check_eqv_def
by (intro conjI[OF assms(2) norm.completeness'[OF iffD2[OF lang_Plus_Zero]],
OF box_equals[OF assms(1) lang⇣M⇣S⇣O_rexp_of_norm lang⇣M⇣S⇣O_rexp_of_norm]])
(auto simp: wf_rexp_def[symmetric] split: sum.splits option.splits intro!: wf_rexp_of)
lemma completeness':
assumes "lang⇣M⇣S⇣O Σ n φ = lang⇣M⇣S⇣O Σ n ψ" "wf_formula Σ n (Or φ ψ)"
shows "check_eqv' Σ n φ ψ"
using assms(2) unfolding check_eqv'_def
by (intro conjI[OF assms(2) norm.completeness'[OF iffD2[OF lang_Plus_Zero]],
OF box_equals[OF assms(1) lang⇣M⇣S⇣O_rexp_of'_norm lang⇣M⇣S⇣O_rexp_of'_norm]])
(auto simp: wf_rexp_def[symmetric] split: sum.splits option.splits intro!: wf_rexp_of')
end