Theory CommCSL
section ‹CommCSL›
text ‹In this file, we define the assertion language and the rules of CommCSL.›
theory CommCSL
imports Lang StateModel
begin
definition no_guard :: "('i, 'a) heap ⇒ bool" where
"no_guard h ⟷ get_gs h = None ∧ (∀k. get_gu h k = None)"
typedef 'a precondition = "{ pre :: ('a ⇒ 'a ⇒ bool) |pre. ∀a b. pre a b ⟶ (pre b a ∧ pre a a) }"
using Sup2_E by auto
lemma charact_rep_prec:
assumes "Rep_precondition pre a b"
shows "Rep_precondition pre b a ∧ Rep_precondition pre a a"
using Rep_precondition assms by fastforce
typedef ('i, 'a) indexed_precondition = "{ pre :: ('i ⇒ 'a ⇒ 'a ⇒ bool) |pre. ∀a b k. pre k a b ⟶ (pre k b a ∧ pre k a a) }"
using Sup2_E by auto
lemma charact_rep_indexed_prec:
assumes "Rep_indexed_precondition pre k a b"
shows "Rep_indexed_precondition pre k b a ∧ Rep_indexed_precondition pre k a a"
by (metis (no_types, lifting) Abs_indexed_precondition_cases Rep_indexed_precondition_cases Rep_indexed_precondition_inverse assms mem_Collect_eq)
type_synonym 'a list_exp = "store ⇒ 'a list"
subsection ‹Assertion Language›