header {* Deciding Equivalence of WS1S Formulas *}
theory WS1S_Equivalence_Checking
imports Smart_Constructors_Normalization WS1S_Normalization "~~/src/HOL/Library/Option_ord"
begin
type_synonym 'a T = "'a × bool list"
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_lderiv where [code del]: "norm_lderiv ≡ λΣ. embed.lderiv (ε Σ)"
definition norm_rderiv where [code del]: "norm_rderiv ≡ λΣ. embed.rderiv (ε Σ)"
definition norm_rderiv_and_add where [code del]: "norm_rderiv_and_add ≡ λΣ. embed.rderiv_and_add (ε Σ)"
definition quot where [code del]: "quot ≡ λΣ. embed.samequot_exec (ε Σ)"
interpretation embed "set o (σ (Σ :: 'a :: linorder list))" π "ε Σ"
where "embed.lderiv (ε Σ) = norm_lderiv Σ"
and "embed.rderiv (ε Σ) = norm_rderiv Σ"
and "embed.rderiv_and_add (ε Σ) = norm_rderiv_and_add Σ"
and "embed.samequot_exec (ε Σ) = quot Σ"
by (unfold_locales) (auto simp: norm_lderiv_def norm_rderiv_def norm_rderiv_and_add_def quot_def σ_def π_def ε_def)
definition norm_step' where [code del]:
"norm_step' ≡ λΣ. equivalence_checker.step' (σ Σ) (ε Σ) (Smart_Constructors_Normalization.norm :: 'a::linorder T rexp => 'a T rexp)"
definition norm_closure' where [code del]:
"norm_closure' ≡ λΣ. equivalence_checker.closure' (σ Σ) (ε Σ) (Smart_Constructors_Normalization.norm :: 'a::linorder T rexp => 'a T rexp)"
definition norm_check_eqv' where [code del]:
"norm_check_eqv' ≡ λΣ. equivalence_checker.check_eqv' (σ Σ) (ε Σ) (Smart_Constructors_Normalization.norm :: 'a::linorder T rexp => 'a T rexp)"
definition norm_step where [code del]:
"norm_step ≡ λΣ. equivalence_checker.step (σ Σ) (ε Σ) (Smart_Constructors_Normalization.norm :: 'a::linorder T rexp => 'a T rexp)"
definition norm_closure where [code del]:
"norm_closure ≡ λΣ. equivalence_checker.closure (σ Σ) (ε Σ) (Smart_Constructors_Normalization.norm :: 'a::linorder T rexp => 'a T rexp)"
definition norm_check_eqv where [code del]:
"norm_check_eqv ≡ λΣ. equivalence_checker.check_eqv (σ Σ) (ε Σ) (Smart_Constructors_Normalization.norm :: 'a::linorder T rexp => 'a T rexp)"
definition norm_check_eqv_counterexample where [code del]:
"norm_check_eqv_counterexample ≡ λΣ. equivalence_checker.check_eqv_counterexample (σ Σ) (ε Σ) (Smart_Constructors_Normalization.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: equivalence_checker "σ Σ" π "ε Σ" Smart_Constructors_Normalization.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])
abbreviation "ext Σ ≡ None # map Some (Σ :: 'a :: linorder list)"
definition any where [code del]:
"any ≡ λΣ. formula.any (ext Σ)"
definition pre_wf_formula where [code del]:
"pre_wf_formula ≡ λΣ. formula.pre_wf_formula (ext Σ)"
definition wf_formula where [code del]:
"wf_formula ≡ λΣ. formula.wf_formula (ext Σ)"
definition valid_ENC where [code del]: "valid_ENC ≡ λΣ. formula.valid_ENC (ext Σ)"
definition ENC where [code del]: "ENC ≡ λΣ. formula.ENC (ext Σ)"
definition rexp_of where [code del]: "rexp_of ≡ λΣ. formula.rexp_of (ext Σ)"
definition rexp_of_alt where [code del]: "rexp_of_alt ≡ λΣ. formula.rexp_of_alt (ext Σ)"
definition rexp_of' where [code del]: "rexp_of' ≡ λΣ. formula.rexp_of' (ext Σ)"
lemmas formula_defs = pre_wf_formula_def wf_formula_def any_def
rexp_of_def rexp_of'_def rexp_of_alt_def ENC_def valid_ENC_def FOV_def SOV_def
interpretation Φ: formula "ext (Σ :: 'a :: linorder list)"
where "alphabet.wf (set o σ Σ) = wf_rexp Σ"
and "embed.rderiv (ε Σ) = norm_rderiv Σ"
and "embed.rderiv_and_add (ε Σ) = norm_rderiv_and_add Σ"
and "embed.samequot_exec (ε Σ) = quot Σ"
and "Φ.any = any Σ"
and "Φ.pre_wf_formula = pre_wf_formula Σ"
and "Φ.wf_formula = wf_formula Σ"
and "Φ.rexp_of = rexp_of Σ"
and "Φ.rexp_of_alt = rexp_of_alt Σ"
and "Φ.rexp_of' = rexp_of' Σ"
and "Φ.valid_ENC = valid_ENC Σ"
and "Φ.ENC = ENC Σ"
by unfold_locales
(auto simp: σ_def π_def wf_rexp_def norm_rderiv_def norm_rderiv_and_add_def quot_def formula_defs)
definition check_eqv where
"check_eqv Σ n φ ψ <-> wf_formula Σ n (FOr φ ψ) ∧
norm_check_eqv' (ext Σ) n (rexp_of Σ n (norm φ)) (rexp_of Σ n (norm ψ))"
definition check_eqv_counterexample where
"check_eqv_counterexample Σ n φ ψ =
norm_check_eqv_counterexample (ext Σ) n (rexp_of Σ n (norm φ)) (rexp_of Σ n (norm ψ))"
definition check_eqv' where
"check_eqv' Σ n φ ψ <-> wf_formula Σ n (FOr φ ψ) ∧
norm_check_eqv' (ext Σ) n (rexp_of' Σ n (norm φ)) (rexp_of' Σ n (norm ψ))"
lemmas lang⇣W⇣S⇣1⇣S_rexp_of_norm = trans[OF sym[OF Φ.lang⇣W⇣S⇣1⇣S_norm] Φ.lang⇣W⇣S⇣1⇣S_rexp_of]
lemma soundness: "check_eqv Σ n φ ψ ==> Φ.lang⇣W⇣S⇣1⇣S Σ n φ = Φ.lang⇣W⇣S⇣1⇣S Σ n ψ"
by (rule box_equals[OF norm.soundness'
sym[OF trans[OF lang⇣W⇣S⇣1⇣S_rexp_of_norm]] sym[OF trans[OF lang⇣W⇣S⇣1⇣S_rexp_of_norm]]])
(auto simp: check_eqv_def split: sum.splits option.splits)
lemmas lang⇣W⇣S⇣1⇣S_rexp_of'_norm = trans[OF sym[OF Φ.lang⇣W⇣S⇣1⇣S_norm] Φ.lang⇣W⇣S⇣1⇣S_rexp_of']
lemma soundness': "check_eqv' Σ n φ ψ ==> Φ.lang⇣W⇣S⇣1⇣S Σ n φ = Φ.lang⇣W⇣S⇣1⇣S Σ n ψ"
by (rule box_equals[OF norm.soundness'
sym[OF trans[OF lang⇣W⇣S⇣1⇣S_rexp_of'_norm]] sym[OF trans[OF lang⇣W⇣S⇣1⇣S_rexp_of'_norm]]])
(auto simp: check_eqv'_def split: sum.splits option.splits)
lemma completeness:
assumes "Φ.lang⇣W⇣S⇣1⇣S Σ n φ = Φ.lang⇣W⇣S⇣1⇣S Σ n ψ" "wf_formula Σ n (FOr φ ψ)"
shows "check_eqv Σ n φ ψ"
using assms(2) unfolding check_eqv_def
by (intro conjI[OF assms(2) norm.completeness',
OF box_equals[OF assms(1) lang⇣W⇣S⇣1⇣S_rexp_of_norm lang⇣W⇣S⇣1⇣S_rexp_of_norm]])
(auto split: sum.splits option.splits intro!: Φ.wf_rexp_of)
lemma completeness':
assumes "Φ.lang⇣W⇣S⇣1⇣S Σ n φ = Φ.lang⇣W⇣S⇣1⇣S Σ n ψ" "wf_formula Σ n (FOr φ ψ)"
shows "check_eqv' Σ n φ ψ"
using assms(2) unfolding check_eqv'_def
by (intro conjI[OF assms(2) norm.completeness',
OF box_equals[OF assms(1) lang⇣W⇣S⇣1⇣S_rexp_of'_norm lang⇣W⇣S⇣1⇣S_rexp_of'_norm]])
(auto split: sum.splits option.splits intro!: Φ.wf_rexp_of')
end