Theory Equivalence_Relations
section ‹Equivalence Relations for LTL formulas›
theory Equivalence_Relations
imports
  LTL
begin
subsection ‹Language Equivalence›
definition ltl_lang_equiv :: "'a ltln ⇒ 'a ltln ⇒ bool" (infix ‹∼⇩L› 75)
where
  "φ ∼⇩L ψ ≡ ∀w. w ⊨⇩n φ ⟷ w ⊨⇩n ψ"
lemma ltl_lang_equiv_equivp:
  "equivp (∼⇩L)"
  unfolding ltl_lang_equiv_def
  by (simp add: equivpI reflp_def symp_def transp_def)
lemma ltl_lang_equiv_and_true[simp]:
  "φ⇩1 and⇩n φ⇩2 ∼⇩L true⇩n ⟷ φ⇩1 ∼⇩L true⇩n ∧ φ⇩2 ∼⇩L true⇩n"
  unfolding ltl_lang_equiv_def by auto
lemma ltl_lang_equiv_and_false[intro, simp]:
  "φ⇩1 ∼⇩L false⇩n ⟹ φ⇩1 and⇩n φ⇩2 ∼⇩L false⇩n"
  "φ⇩2 ∼⇩L false⇩n ⟹ φ⇩1 and⇩n φ⇩2 ∼⇩L false⇩n"
  unfolding ltl_lang_equiv_def by auto
lemma ltl_lang_equiv_or_false[simp]:
  "φ⇩1 or⇩n φ⇩2 ∼⇩L false⇩n ⟷ φ⇩1 ∼⇩L false⇩n ∧ φ⇩2 ∼⇩L false⇩n"
  unfolding ltl_lang_equiv_def by auto
lemma ltl_lang_equiv_or_const[intro, simp]:
  "φ⇩1 ∼⇩L true⇩n ⟹ φ⇩1 or⇩n φ⇩2 ∼⇩L true⇩n"
  "φ⇩2 ∼⇩L true⇩n ⟹ φ⇩1 or⇩n φ⇩2 ∼⇩L true⇩n"
  unfolding ltl_lang_equiv_def by auto
subsection ‹Propositional Equivalence›
fun ltl_prop_entailment :: "'a ltln set ⇒ 'a ltln ⇒ bool" (infix ‹⊨⇩P› 80)
where
  "𝒜 ⊨⇩P true⇩n = True"
| "𝒜 ⊨⇩P false⇩n = False"
| "𝒜 ⊨⇩P φ and⇩n ψ = (𝒜 ⊨⇩P φ ∧ 𝒜 ⊨⇩P ψ)"
| "𝒜 ⊨⇩P φ or⇩n ψ = (𝒜 ⊨⇩P φ ∨ 𝒜 ⊨⇩P ψ)"
| "𝒜 ⊨⇩P φ = (φ ∈ 𝒜)"
lemma ltl_prop_entailment_monotonI[intro]:
  "S ⊨⇩P φ ⟹ S ⊆ S' ⟹ S' ⊨⇩P φ"
  by (induction φ) auto
lemma ltl_models_equiv_prop_entailment:
  "w ⊨⇩n φ ⟷ {ψ. w ⊨⇩n ψ} ⊨⇩P φ"
  by (induction φ) auto
definition ltl_prop_equiv :: "'a ltln ⇒ 'a ltln ⇒ bool" (infix ‹∼⇩P› 75)
where
  "φ ∼⇩P ψ ≡ ∀𝒜. 𝒜 ⊨⇩P φ ⟷ 𝒜 ⊨⇩P ψ"
definition ltl_prop_implies :: "'a ltln ⇒ 'a ltln ⇒ bool" (infix ‹⟶⇩P› 75)
where
  "φ ⟶⇩P ψ ≡ ∀𝒜. 𝒜 ⊨⇩P φ ⟶ 𝒜 ⊨⇩P ψ"
lemma ltl_prop_implies_equiv:
  "φ ∼⇩P ψ ⟷ (φ ⟶⇩P ψ ∧ ψ ⟶⇩P φ)"
  unfolding ltl_prop_equiv_def ltl_prop_implies_def by meson
lemma ltl_prop_equiv_equivp:
  "equivp (∼⇩P)"
  by (simp add: ltl_prop_equiv_def equivpI reflp_def symp_def transp_def)
lemma ltl_prop_equiv_trans[trans]:
  "φ ∼⇩P ψ ⟹ ψ ∼⇩P χ ⟹ φ ∼⇩P χ"
  by (simp add: ltl_prop_equiv_def)
lemma ltl_prop_equiv_true:
  "φ ∼⇩P true⇩n ⟷ {} ⊨⇩P φ"
  using bot.extremum ltl_prop_entailment.simps(1) ltl_prop_equiv_def by blast
lemma ltl_prop_equiv_false:
  "φ ∼⇩P false⇩n ⟷ ¬ UNIV ⊨⇩P φ"
  by (meson ltl_prop_entailment.simps(2) ltl_prop_entailment_monotonI ltl_prop_equiv_def top_greatest)
lemma ltl_prop_equiv_true_implies_true:
  "x ∼⇩P true⇩n ⟹ x ⟶⇩P y ⟹ y ∼⇩P true⇩n"
  by (simp add: ltl_prop_equiv_def ltl_prop_implies_def)
lemma ltl_prop_equiv_false_implied_by_false:
  "y ∼⇩P false⇩n ⟹ x ⟶⇩P y ⟹ x ∼⇩P false⇩n"
  by (simp add: ltl_prop_equiv_def ltl_prop_implies_def)
lemma ltl_prop_implication_implies_ltl_implication:
  "w ⊨⇩n φ ⟹ φ ⟶⇩P ψ ⟹ w ⊨⇩n ψ"
  using ltl_models_equiv_prop_entailment ltl_prop_implies_def by blast
lemma ltl_prop_equiv_implies_ltl_lang_equiv:
  "φ ∼⇩P ψ ⟹ φ ∼⇩L ψ"
  using ltl_lang_equiv_def ltl_prop_implication_implies_ltl_implication ltl_prop_implies_equiv by blast
lemma ltl_prop_equiv_lt_ltl_lang_equiv[simp]:
  "(∼⇩P) ≤ (∼⇩L)"
  using ltl_prop_equiv_implies_ltl_lang_equiv by blast
subsection ‹Constants Equivalence›