Initial import from AFP.

This commit is contained in:
Achim D. Brucker 2020-12-19 08:46:46 +00:00
commit a7c68b6777
20 changed files with 5806 additions and 0 deletions

View File

@ -0,0 +1,649 @@
section \<open>Arithmetic Expressions\<close>
text\<open>
This theory defines a language of arithmetic expressions over variables and literal values. Here,
values are limited to integers and strings. Variables may be either inputs or registers. We also
limit ourselves to a simple arithmetic of addition, subtraction, and multiplication as a proof of
concept.
\<close>
theory AExp
imports Value_Lexorder VName FinFun.FinFun "HOL-Library.Option_ord"
begin
declare One_nat_def [simp del]
unbundle finfun_syntax
type_synonym registers = "nat \<Rightarrow>f value option"
type_synonym 'a datastate = "'a \<Rightarrow> value option"
text_raw\<open>\snip{aexptype}{1}{2}{%\<close>
datatype 'a aexp = L "value" | V 'a | Plus "'a aexp" "'a aexp" | Minus "'a aexp" "'a aexp" | Times "'a aexp" "'a aexp"
text_raw\<open>}%endsnip\<close>
fun is_lit :: "'a aexp \<Rightarrow> bool" where
"is_lit (L _) = True" |
"is_lit _ = False"
lemma aexp_induct_separate_V_cases [case_names L I R Plus Minus Times]:
"(\<And>x. P (L x)) \<Longrightarrow>
(\<And>x. P (V (I x))) \<Longrightarrow>
(\<And>x. P (V (R x))) \<Longrightarrow>
(\<And>x1a x2a. P x1a \<Longrightarrow> P x2a \<Longrightarrow> P (Plus x1a x2a)) \<Longrightarrow>
(\<And>x1a x2a. P x1a \<Longrightarrow> P x2a \<Longrightarrow> P (Minus x1a x2a)) \<Longrightarrow>
(\<And>x1a x2a. P x1a \<Longrightarrow> P x2a \<Longrightarrow> P (Times x1a x2a)) \<Longrightarrow>
P a"
by (metis aexp.induct vname.exhaust)
fun aval :: "'a aexp \<Rightarrow> 'a datastate \<Rightarrow> value option" where
"aval (L x) s = Some x" |
"aval (V x) s = s x" |
"aval (Plus a1 a2) s = value_plus (aval a1 s)(aval a2 s)" |
"aval (Minus a1 a2) s = value_minus (aval a1 s) (aval a2 s)" |
"aval (Times a1 a2) s = value_times (aval a1 s) (aval a2 s)"
lemma aval_plus_symmetry: "aval (Plus x y) s = aval (Plus y x) s"
by (simp add: value_plus_symmetry)
text \<open>A little syntax magic to write larger states compactly:\<close>
definition null_state ("<>") where
"null_state \<equiv> (K$ bot)"
no_notation finfun_update ("_'(_ $:= _')" [1000, 0, 0] 1000)
nonterminal fupdbinds and fupdbind
syntax
"_fupdbind" :: "'a \<Rightarrow> 'a \<Rightarrow> fupdbind" ("(2_ $:=/ _)")
"" :: "fupdbind \<Rightarrow> fupdbinds" ("_")
"_fupdbinds":: "fupdbind \<Rightarrow> fupdbinds \<Rightarrow> fupdbinds" ("_,/ _")
"_fUpdate" :: "'a \<Rightarrow> fupdbinds \<Rightarrow> 'a" ("_/'((_)')" [1000, 0] 900)
"_State" :: "fupdbinds => 'a" ("<_>")
translations
"_fUpdate f (_fupdbinds b bs)" \<rightleftharpoons> "_fUpdate (_fUpdate f b) bs"
"f(x$:=y)" \<rightleftharpoons> "CONST finfun_update f x y"
"_State ms" == "_fUpdate <> ms"
"_State (_updbinds b bs)" <= "_fUpdate (_State b) bs"
lemma empty_None: "<> = (K$ None)"
by (simp add: null_state_def bot_option_def)
lemma apply_empty_None [simp]: "<> $ x2 = None"
by (simp add: null_state_def bot_option_def)
definition input2state :: "value list \<Rightarrow> registers" where
"input2state n = fold (\<lambda>(k, v) f. f(k $:= Some v)) (enumerate 0 n) (K$ None)"
primrec input2state_prim :: "value list \<Rightarrow> nat \<Rightarrow> registers" where
"input2state_prim [] _ = (K$ None)" |
"input2state_prim (v#t) k = (input2state_prim t (k+1))(k $:= Some v)"
lemma input2state_append:
"input2state (i @ [a]) = (input2state i)(length i $:= Some a)"
apply (simp add: eq_finfun_All_ext finfun_All_def finfun_All_except_def)
apply clarify
by (simp add: input2state_def enumerate_eq_zip)
lemma input2state_out_of_bounds:
"i \<ge> length ia \<Longrightarrow> input2state ia $ i = None"
proof(induct ia rule: rev_induct)
case Nil
then show ?case
by (simp add: input2state_def)
next
case (snoc a as)
then show ?case
by (simp add: input2state_def enumerate_eq_zip)
qed
lemma input2state_within_bounds:
"input2state i $ x = Some a \<Longrightarrow> x < length i"
by (metis input2state_out_of_bounds not_le_imp_less option.distinct(1))
lemma input2state_empty: "input2state [] $ x1 = None"
by (simp add: input2state_out_of_bounds)
lemma input2state_nth:
"i < length ia \<Longrightarrow> input2state ia $ i = Some (ia ! i)"
proof(induct ia rule: rev_induct)
case Nil
then show ?case
by simp
next
case (snoc a ia)
then show ?case
apply (simp add: input2state_def enumerate_eq_zip)
by (simp add: finfun_upd_apply nth_append)
qed
lemma input2state_some:
"i < length ia \<Longrightarrow>
ia ! i = x \<Longrightarrow>
input2state ia $ i = Some x"
by (simp add: input2state_nth)
lemma input2state_take: "x1 < A \<Longrightarrow>
A \<le> length i \<Longrightarrow>
x = vname.I x1 \<Longrightarrow>
input2state i $ x1 = input2state (take A i) $ x1"
proof(induct i)
case Nil
then show ?case
by simp
next
case (Cons a i)
then show ?case
by (simp add: input2state_nth)
qed
lemma input2state_not_None:
"(input2state i $ x \<noteq> None) \<Longrightarrow> (x < length i)"
using input2state_within_bounds by blast
lemma input2state_Some:
"(\<exists>v. input2state i $ x = Some v) = (x < length i)"
apply standard
using input2state_within_bounds apply blast
by (simp add: input2state_nth)
lemma input2state_cons: "x1 > 0 \<Longrightarrow>
x1 < length ia \<Longrightarrow>
input2state (a # ia) $ x1 = input2state ia $ (x1-1)"
by (simp add: input2state_nth)
lemma input2state_cons_shift:
"input2state i $ x1 = Some a \<Longrightarrow> input2state (b # i) $ (Suc x1) = Some a"
proof(induct i rule: rev_induct)
case Nil
then show ?case
by (simp add: input2state_def)
next
case (snoc x xs)
then show ?case
using input2state_within_bounds[of xs x1 a]
using input2state_cons[of "Suc x1" "xs @ [x]" b]
apply simp
apply (case_tac "x1 < length xs")
apply simp
by (metis finfun_upd_apply input2state_append input2state_nth length_Cons length_append_singleton lessI list.sel(3) nth_tl)
qed
lemma input2state_exists: "\<exists>i. input2state i $ x1 = Some a"
proof(induct x1)
case 0
then show ?case
apply (rule_tac x="[a]" in exI)
by (simp add: input2state_def)
next
case (Suc x1)
then show ?case
apply clarify
apply (rule_tac x="a#i" in exI)
by (simp add: input2state_cons_shift)
qed
primrec repeat :: "nat \<Rightarrow> 'a \<Rightarrow> 'a list" where
"repeat 0 _ = []" |
"repeat (Suc m) a = a#(repeat m a)"
lemma length_repeat: "length (repeat n a) = n"
proof(induct n)
case 0
then show ?case
by simp
next
case (Suc a)
then show ?case
by simp
qed
lemma length_append_repeat: "length (i@(repeat a y)) \<ge> length i"
by simp
lemma length_input2state_repeat:
"input2state i $ x = Some a \<Longrightarrow> y < length (i @ repeat y a)"
by (metis append.simps(1) append_eq_append_conv input2state_within_bounds length_append length_repeat list.size(3) neqE not_add_less2 zero_order(3))
lemma input2state_double_exists:
"\<exists>i. input2state i $ x = Some a \<and> input2state i $ y = Some a"
apply (insert input2state_exists[of x a])
apply clarify
apply (case_tac "x \<ge> y")
apply (rule_tac x="list_update i y a" in exI)
apply (metis (no_types, lifting) input2state_within_bounds input2state_nth input2state_out_of_bounds le_trans length_list_update not_le_imp_less nth_list_update_eq nth_list_update_neq)
apply (rule_tac x="list_update (i@(repeat y a)) y a" in exI)
apply (simp add: not_le)
by (metis length_input2state_repeat input2state_nth input2state_out_of_bounds le_trans length_append_repeat length_list_update not_le_imp_less nth_append nth_list_update_eq nth_list_update_neq option.distinct(1))
lemma input2state_double_exists_2:
"x \<noteq> y \<Longrightarrow> \<exists>i. input2state i $ x = Some a \<and> input2state i $ y = Some a'"
apply (insert input2state_exists[of x a])
apply clarify
apply (case_tac "x \<ge> y")
apply (rule_tac x="list_update i y a'" in exI)
apply (metis (no_types, lifting) input2state_within_bounds input2state_nth input2state_out_of_bounds le_trans length_list_update not_le_imp_less nth_list_update_eq nth_list_update_neq)
apply (rule_tac x="list_update (i@(repeat y a')) y a'" in exI)
apply (simp add: not_le)
apply standard
apply (metis input2state_nth input2state_within_bounds le_trans length_append_repeat length_list_update linorder_not_le nth_append nth_list_update_neq order_refl)
by (metis input2state_nth length_append length_input2state_repeat length_list_update length_repeat nth_list_update_eq)
definition join_ir :: "value list \<Rightarrow> registers \<Rightarrow> vname datastate" where
"join_ir i r \<equiv> (\<lambda>x. case x of
R n \<Rightarrow> r $ n |
I n \<Rightarrow> (input2state i) $ n
)"
lemmas datastate = join_ir_def input2state_def
lemma join_ir_empty [simp]: "join_ir [] <> = (\<lambda>x. None)"
apply (rule ext)
apply (simp add: join_ir_def)
apply (case_tac x)
apply (simp add: input2state_def)
by (simp add: empty_None)
lemma join_ir_R [simp]: "(join_ir i r) (R n) = r $ n"
by (simp add: join_ir_def)
lemma join_ir_double_exists:
"\<exists>i r. join_ir i r v = Some a \<and> join_ir i r v' = Some a"
proof(cases v)
case (I x1)
then show ?thesis
apply (simp add: join_ir_def)
apply (cases v')
apply (simp add: input2state_double_exists input2state_exists)
using input2state_exists by auto
next
case (R x2)
then show ?thesis
apply (simp add: join_ir_def)
apply (cases v')
using input2state_exists apply force
using input2state_double_exists by fastforce
qed
lemma join_ir_double_exists_2:
"v \<noteq> v' \<Longrightarrow> \<exists>i r. join_ir i r v = Some a \<and> join_ir i r v' = Some a'"
proof(cases v)
case (I x1)
assume "v \<noteq> v'"
then show ?thesis using I input2state_exists
by (cases v', auto simp add: join_ir_def input2state_double_exists_2)
next
case (R x2)
assume "v \<noteq> v'"
then show ?thesis
apply (simp add: join_ir_def)
apply (cases v')
apply simp
using R input2state_exists apply auto[1]
apply (simp add: R)
apply (rule_tac x="<x2 $:= Some a,x2a $:= Some a'>" in exI)
by simp
qed
lemma exists_join_ir_ext: "\<exists>i r. join_ir i r v = s v"
apply (simp add: join_ir_def)
apply (case_tac "s v")
apply (cases v)
apply (rule_tac x="[]" in exI)
apply (simp add: input2state_out_of_bounds)
apply simp
apply (rule_tac x="<>" in exI)
apply simp
apply simp
apply (cases v)
apply simp
apply (simp add: input2state_exists)
apply simp
apply (rule_tac x="<x2 $:= Some a>" in exI)
apply simp
done
lemma join_ir_nth [simp]:
"i < length is \<Longrightarrow> join_ir is r (I i) = Some (is ! i)"
by (simp add: join_ir_def input2state_nth)
fun aexp_constrains :: "'a aexp \<Rightarrow> 'a aexp \<Rightarrow> bool" where
"aexp_constrains (L l) a = (L l = a)" |
"aexp_constrains (V v) v' = (V v = v')" |
"aexp_constrains (Plus a1 a2) v = ((Plus a1 a2) = v \<or> (Plus a1 a2) = v \<or> (aexp_constrains a1 v \<or> aexp_constrains a2 v))" |
"aexp_constrains (Minus a1 a2) v = ((Minus a1 a2) = v \<or> (aexp_constrains a1 v \<or> aexp_constrains a2 v))" |
"aexp_constrains (Times a1 a2) v = ((Times a1 a2) = v \<or> (aexp_constrains a1 v \<or> aexp_constrains a2 v))"
fun aexp_same_structure :: "'a aexp \<Rightarrow> 'a aexp \<Rightarrow> bool" where
"aexp_same_structure (L v) (L v') = True" |
"aexp_same_structure (V v) (V v') = True" |
"aexp_same_structure (Plus a1 a2) (Plus a1' a2') = (aexp_same_structure a1 a1' \<and> aexp_same_structure a2 a2')" |
"aexp_same_structure (Minus a1 a2) (Minus a1' a2') = (aexp_same_structure a1 a1' \<and> aexp_same_structure a2 a2')" |
"aexp_same_structure _ _ = False"
fun enumerate_aexp_inputs :: "vname aexp \<Rightarrow> nat set" where
"enumerate_aexp_inputs (L _) = {}" |
"enumerate_aexp_inputs (V (I n)) = {n}" |
"enumerate_aexp_inputs (V (R n)) = {}" |
"enumerate_aexp_inputs (Plus v va) = enumerate_aexp_inputs v \<union> enumerate_aexp_inputs va" |
"enumerate_aexp_inputs (Minus v va) = enumerate_aexp_inputs v \<union> enumerate_aexp_inputs va" |
"enumerate_aexp_inputs (Times v va) = enumerate_aexp_inputs v \<union> enumerate_aexp_inputs va"
lemma enumerate_aexp_inputs_list: "\<exists>l. enumerate_aexp_inputs a = set l"
proof(induct a)
case (L x)
then show ?case
by simp
next
case (V x)
then show ?case
apply (cases x)
apply (metis empty_set enumerate_aexp_inputs.simps(2) list.simps(15))
by simp
next
case (Plus a1 a2)
then show ?case
by (metis enumerate_aexp_inputs.simps(4) set_append)
next
case (Minus a1 a2)
then show ?case
by (metis enumerate_aexp_inputs.simps(5) set_append)
next
case (Times a1 a2)
then show ?case
by (metis enumerate_aexp_inputs.simps(6) set_append)
qed
fun enumerate_regs :: "vname aexp \<Rightarrow> nat set" where
"enumerate_regs (L _) = {}" |
"enumerate_regs (V (R n)) = {n}" |
"enumerate_regs (V (I _)) = {}" |
"enumerate_regs (Plus v va) = enumerate_regs v \<union> enumerate_regs va" |
"enumerate_regs (Minus v va) = enumerate_regs v \<union> enumerate_regs va" |
"enumerate_regs (Times v va) = enumerate_regs v \<union> enumerate_regs va"
lemma finite_enumerate_regs: "finite (enumerate_regs a)"
by(induct a rule: aexp_induct_separate_V_cases, auto)
lemma no_variables_aval: "enumerate_aexp_inputs a = {} \<Longrightarrow>
enumerate_regs a = {} \<Longrightarrow>
aval a s = aval a s'"
by (induct a rule: aexp_induct_separate_V_cases, auto)
lemma enumerate_aexp_inputs_not_empty:
"(enumerate_aexp_inputs a \<noteq> {}) = (\<exists>b c. enumerate_aexp_inputs a = set (b#c))"
using enumerate_aexp_inputs_list by fastforce
lemma aval_ir_take: "A \<le> length i \<Longrightarrow>
enumerate_regs a = {} \<Longrightarrow>
enumerate_aexp_inputs a \<noteq> {} \<Longrightarrow>
Max (enumerate_aexp_inputs a) < A \<Longrightarrow>
aval a (join_ir (take A i) r) = aval a (join_ir i ra)"
proof(induct a)
case (L x)
then show ?case
by simp
next
case (V x)
then show ?case
apply (cases x)
apply (simp add: join_ir_def input2state_nth)
by simp
next
case (Plus a1 a2)
then show ?case
apply (simp only: enumerate_aexp_inputs_not_empty[of "Plus a1 a2"])
apply (erule exE)+
apply (simp only: neq_Nil_conv List.linorder_class.Max.set_eq_fold)
apply (case_tac "fold max c b \<le> length i")
apply simp
apply (metis List.finite_set Max.union Plus.prems(4) enumerate_aexp_inputs.simps(4) enumerate_aexp_inputs_not_empty max_less_iff_conj no_variables_aval sup_bot.left_neutral sup_bot.right_neutral)
by simp
next
case (Minus a1 a2)
then show ?case
apply (simp only: enumerate_aexp_inputs_not_empty[of "Minus a1 a2"])
apply (erule exE)+
apply (simp only: neq_Nil_conv List.linorder_class.Max.set_eq_fold)
apply (case_tac "fold max c b \<le> length i")
apply simp
apply (metis List.finite_set Max.union Minus.prems(4) enumerate_aexp_inputs.simps(5) enumerate_aexp_inputs_not_empty max_less_iff_conj no_variables_aval sup_bot.left_neutral sup_bot.right_neutral)
by simp
next
case (Times a1 a2)
then show ?case
apply (simp only: enumerate_aexp_inputs_not_empty[of "Times a1 a2"])
apply (erule exE)+
apply (simp only: neq_Nil_conv List.linorder_class.Max.set_eq_fold)
apply (case_tac "fold max c b \<le> length i")
apply simp
apply (metis List.finite_set Max.union Times.prems(4) enumerate_aexp_inputs.simps(6) enumerate_aexp_inputs_not_empty max_less_iff_conj no_variables_aval sup_bot.left_neutral sup_bot.right_neutral)
by simp
qed
definition max_input :: "vname aexp \<Rightarrow> nat option" where
"max_input g = (let inputs = (enumerate_aexp_inputs g) in if inputs = {} then None else Some (Max inputs))"
definition max_reg :: "vname aexp \<Rightarrow> nat option" where
"max_reg g = (let regs = (enumerate_regs g) in if regs = {} then None else Some (Max regs))"
lemma max_reg_V_I: "max_reg (V (I n)) = None"
by (simp add: max_reg_def)
lemma max_reg_V_R: "max_reg (V (R n)) = Some n"
by (simp add: max_reg_def)
lemmas max_reg_V = max_reg_V_I max_reg_V_R
lemma max_reg_Plus: "max_reg (Plus a1 a2) = max (max_reg a1) (max_reg a2)"
apply (simp add: max_reg_def Let_def max_absorb2)
by (metis Max.union bot_option_def finite_enumerate_regs max_bot2 sup_Some sup_max)
lemma max_reg_Minus: "max_reg (Minus a1 a2) = max (max_reg a1) (max_reg a2)"
apply (simp add: max_reg_def Let_def max_absorb2)
by (metis Max.union bot_option_def finite_enumerate_regs max_bot2 sup_Some sup_max)
lemma max_reg_Times: "max_reg (Times a1 a2) = max (max_reg a1) (max_reg a2)"
apply (simp add: max_reg_def Let_def max_absorb2)
by (metis Max.union bot_option_def finite_enumerate_regs max_bot2 sup_Some sup_max)
lemma no_reg_aval_swap_regs:
"max_reg a = None \<Longrightarrow> aval a (join_ir i r) = aval a (join_ir i r')"
proof(induct a)
case (V x)
then show ?case
apply (cases x)
apply (simp add: join_ir_def)
by (simp add: join_ir_def max_reg_def)
next
case (Plus a1 a2)
then show ?case
by (metis (no_types, lifting) aval.simps(3) max.absorb2 max.cobounded2 max_reg_Plus sup_None_2 sup_max)
next
case (Minus a1 a2)
then show ?case
by (metis (no_types, lifting) aval.simps(4) max.cobounded2 max_def_raw max_reg_Minus sup_None_2 sup_max)
next
case (Times a1 a2)
then show ?case
proof -
have "bot = max_reg a2"
by (metis (no_types) Times.prems bot_option_def max.left_commute max_bot2 max_def_raw max_reg_Times)
then show ?thesis
by (metis Times.hyps(1) Times.hyps(2) Times.prems aval.simps(5) bot_option_def max_bot2 max_reg_Times)
qed
qed auto
lemma aval_reg_some_superset:
"\<forall>a. (r $ a \<noteq> None) \<longrightarrow> r $ a = r' $ a \<Longrightarrow>
aval a (join_ir i r) = Some v \<Longrightarrow>
aval a (join_ir i r') = Some v"
proof(induct a arbitrary: v rule: aexp_induct_separate_V_cases)
case (I x)
then show ?case
by (simp add: join_ir_def)
next
case (Plus x1a x2a)
then show ?case
apply simp
by (metis maybe_arith_int_not_None option.simps(3) value_plus_def)
next
case (Minus x1a x2a)
then show ?case
apply simp
by (metis maybe_arith_int_not_None option.simps(3) value_minus_def)
next
case (Times x1a x2a)
then show ?case
apply simp
by (metis maybe_arith_int_not_None option.simps(3) value_times_def)
qed auto
lemma aval_reg_none_superset:
"\<forall>a. (r $ a \<noteq> None) \<longrightarrow> r $ a = r' $ a \<Longrightarrow>
aval a (join_ir i r') = None \<Longrightarrow>
aval a (join_ir i r) = None"
proof(induct a)
case (V x)
then show ?case
apply (cases x)
apply (simp add: join_ir_def)
by auto
next
case (Plus a1 a2)
then show ?case
apply simp
by (metis (no_types, lifting) maybe_arith_int_None Plus.prems(1) aval_reg_some_superset value_plus_def)
next
case (Minus a1 a2)
then show ?case
apply simp
by (metis (no_types, lifting) maybe_arith_int_None Minus.prems(1) aval_reg_some_superset value_minus_def)
next
case (Times a1 a2)
then show ?case
apply simp
by (metis (no_types, lifting) maybe_arith_int_None Times.prems(1) aval_reg_some_superset value_times_def)
qed auto
lemma enumerate_regs_empty_reg_unconstrained:
"enumerate_regs a = {} \<Longrightarrow> \<forall>r. \<not> aexp_constrains a (V (R r))"
by (induct a rule: aexp_induct_separate_V_cases, auto)
lemma enumerate_aexp_inputs_empty_input_unconstrained:
"enumerate_aexp_inputs a = {} \<Longrightarrow> \<forall>r. \<not> aexp_constrains a (V (I r))"
by (induct a rule: aexp_induct_separate_V_cases, auto)
lemma input_unconstrained_aval_input_swap:
"\<forall>i. \<not> aexp_constrains a (V (I i)) \<Longrightarrow>
aval a (join_ir i r) = aval a (join_ir i' r)"
using join_ir_def
by (induct a rule: aexp_induct_separate_V_cases, auto)
lemma input_unconstrained_aval_register_swap:
"\<forall>i. \<not> aexp_constrains a (V (R i)) \<Longrightarrow>
aval a (join_ir i r) = aval a (join_ir i r')"
using join_ir_def
by (induct a rule: aexp_induct_separate_V_cases, auto)
lemma unconstrained_variable_swap_aval:
"\<forall>i. \<not> aexp_constrains a (V (I i)) \<Longrightarrow>
\<forall>r. \<not> aexp_constrains a (V (R r)) \<Longrightarrow>
aval a s = aval a s'"
by (induct a rule: aexp_induct_separate_V_cases, auto)
lemma max_input_I: "max_input (V (vname.I i)) = Some i"
by (simp add: max_input_def)
lemma max_input_Plus:
"max_input (Plus a1 a2) = max (max_input a1) (max_input a2)"
apply (simp add: max_input_def Let_def max.commute max_absorb2)
by (metis List.finite_set Max.union enumerate_aexp_inputs_list sup_Some sup_max)
lemma max_input_Minus:
"max_input (Minus a1 a2) = max (max_input a1) (max_input a2)"
apply (simp add: max_input_def Let_def max.commute max_absorb2)
by (metis List.finite_set Max.union enumerate_aexp_inputs_list sup_Some sup_max)
lemma max_input_Times:
"max_input (Times a1 a2) = max (max_input a1) (max_input a2)"
apply (simp add: max_input_def Let_def max.commute max_absorb2)
by (metis List.finite_set Max.union enumerate_aexp_inputs_list sup_Some sup_max)
lemma aval_take:
"max_input x < Some a \<Longrightarrow>
aval x (join_ir i r) = aval x (join_ir (take a i) r)"
proof(induct x rule: aexp_induct_separate_V_cases)
case (I x)
then show ?case
by (metis aval.simps(2) input2state_take join_ir_def le_cases less_option_Some max_input_I take_all vname.simps(5))
next
case (R x)
then show ?case
by (simp add: join_ir_def)
next
case (Plus x1a x2a)
then show ?case
by (simp add: max_input_Plus)
next
case (Minus x1a x2a)
then show ?case
by (simp add: max_input_Minus)
next
case (Times x1a x2a)
then show ?case
by (simp add: max_input_Times)
qed auto
lemma aval_no_reg_swap_regs: "max_input x < Some a \<Longrightarrow>
max_reg x = None \<Longrightarrow>
aval x (join_ir i ra) = aval x (join_ir (take a i) r)"
proof(induct x)
case (V x)
then show ?case
apply (cases x)
apply (metis aval_take enumerate_regs.simps(3) enumerate_regs_empty_reg_unconstrained input_unconstrained_aval_register_swap)
by (simp add: max_reg_def)
next
case (Plus x1 x2)
then show ?case
by (metis aval_take no_reg_aval_swap_regs)
next
case (Minus x1 x2)
then show ?case
by (metis aval_take no_reg_aval_swap_regs)
next
case (Times x1 x2)
then show ?case
by (metis aval_take no_reg_aval_swap_regs)
qed auto
fun enumerate_aexp_strings :: "'a aexp \<Rightarrow> String.literal set" where
"enumerate_aexp_strings (L (Str s)) = {s}" |
"enumerate_aexp_strings (L (Num s)) = {}" |
"enumerate_aexp_strings (V _) = {}" |
"enumerate_aexp_strings (Plus a1 a2) = enumerate_aexp_strings a1 \<union> enumerate_aexp_strings a2" |
"enumerate_aexp_strings (Minus a1 a2) = enumerate_aexp_strings a1 \<union> enumerate_aexp_strings a2" |
"enumerate_aexp_strings (Times a1 a2) = enumerate_aexp_strings a1 \<union> enumerate_aexp_strings a2"
fun enumerate_aexp_ints :: "'a aexp \<Rightarrow> int set" where
"enumerate_aexp_ints (L (Str s)) = {}" |
"enumerate_aexp_ints (L (Num s)) = {s}" |
"enumerate_aexp_ints (V _) = {}" |
"enumerate_aexp_ints (Plus a1 a2) = enumerate_aexp_ints a1 \<union> enumerate_aexp_ints a2" |
"enumerate_aexp_ints (Minus a1 a2) = enumerate_aexp_ints a1 \<union> enumerate_aexp_ints a2" |
"enumerate_aexp_ints (Times a1 a2) = enumerate_aexp_ints a1 \<union> enumerate_aexp_ints a2"
definition enumerate_vars :: "vname aexp \<Rightarrow> vname set" where
"enumerate_vars a = (image I (enumerate_aexp_inputs a)) \<union> (image R (enumerate_regs a))"
fun rename_regs :: "(nat \<Rightarrow> nat) \<Rightarrow> vname aexp \<Rightarrow> vname aexp" where
"rename_regs _ (L l) = (L l)" |
"rename_regs f (V (R r)) = (V (R (f r)))" |
"rename_regs _ (V v) = (V v)" |
"rename_regs f (Plus a b) = Plus (rename_regs f a) (rename_regs f b)" |
"rename_regs f (Minus a b) = Minus (rename_regs f a) (rename_regs f b)" |
"rename_regs f (Times a b) = Times (rename_regs f a) (rename_regs f b)"
definition eq_upto_rename :: "vname aexp \<Rightarrow> vname aexp \<Rightarrow> bool" where
"eq_upto_rename a1 a2 = (\<exists>f. bij f \<and> rename_regs f a1 = a2)"
end

View File

@ -0,0 +1,171 @@
subsection\<open>AExp Lexorder\<close>
text\<open>This theory defines a lexicographical ordering on arithmetic expressions such that we can build
orderings for guards and, subsequently, transitions. We make use of the previously established
orderings on variable names and values.\<close>
theory AExp_Lexorder
imports AExp Value_Lexorder
begin
text_raw\<open>\snip{height}{1}{2}{%\<close>
fun height :: "'a aexp \<Rightarrow> nat" where
"height (L l2) = 1" |
"height (V v2) = 1" |
"height (Plus e1 e2) = 1 + max (height e1) (height e2)" |
"height (Minus e1 e2) = 1 + max (height e1) (height e2)" |
"height (Times e1 e2) = 1 + max (height e1) (height e2)"
text_raw\<open>}%endsnip\<close>
instantiation aexp :: (linorder) linorder begin
fun less_aexp_aux :: "'a aexp \<Rightarrow> 'a aexp \<Rightarrow> bool" where
"less_aexp_aux (L l1) (L l2) = (l1 < l2)" |
"less_aexp_aux (L l1) _ = True" |
"less_aexp_aux (V v1) (L l1) = False" |
"less_aexp_aux (V v1) (V v2) = (v1 < v2)" |
"less_aexp_aux (V v1) _ = True" |
"less_aexp_aux (Plus e1 e2) (L l2) = False" |
"less_aexp_aux (Plus e1 e2) (V v2) = False" |
"less_aexp_aux (Plus e1 e2) (Plus e1' e2') = ((less_aexp_aux e1 e1') \<or> ((e1 = e1') \<and> (less_aexp_aux e2 e2')))"|
"less_aexp_aux (Plus e1 e2) _ = True" |
"less_aexp_aux (Minus e1 e2) (Minus e1' e2') = ((less_aexp_aux e1 e1') \<or> ((e1 = e1') \<and> (less_aexp_aux e2 e2')))" |
"less_aexp_aux (Minus e1 e2) (Times e1' e2') = True" |
"less_aexp_aux (Minus e1 e2) _ = False" |
"less_aexp_aux (Times e1 e2) (Times e1' e2') = ((less_aexp_aux e1 e1') \<or> ((e1 = e1') \<and> (less_aexp_aux e2 e2')))" |
"less_aexp_aux (Times e1 e2) _ = False"
definition less_aexp :: "'a aexp \<Rightarrow> 'a aexp \<Rightarrow> bool" where
"less_aexp a1 a2 = (
let
h1 = height a1;
h2 = height a2
in
if h1 = h2 then
less_aexp_aux a1 a2
else
h1 < h2
)"
definition less_eq_aexp :: "'a aexp \<Rightarrow> 'a aexp \<Rightarrow> bool"
where "less_eq_aexp e1 e2 \<equiv> (e1 < e2) \<or> (e1 = e2)"
declare less_aexp_def [simp]
lemma less_aexp_aux_antisym: "less_aexp_aux x y = (\<not>(less_aexp_aux y x) \<and> (x \<noteq> y))"
by (induct x y rule: less_aexp_aux.induct, auto)
lemma less_aexp_antisym: "(x::'a aexp) < y = (\<not>(y < x) \<and> (x \<noteq> y))"
apply (simp add: Let_def)
apply standard
using less_aexp_aux_antisym apply blast
apply (simp add: not_less)
apply clarify
by (induct x, auto)
lemma less_aexp_aux_trans: "less_aexp_aux x y \<Longrightarrow> less_aexp_aux y z \<Longrightarrow> less_aexp_aux x z"
proof (induct x y arbitrary: z rule: less_aexp_aux.induct)
case (1 l1 l2)
then show ?case by (cases z, auto)
next
case ("2_1" l1 v)
then show ?case by (cases z, auto)
next
case ("2_2" l1 v va)
then show ?case by (cases z, auto)
next
case ("2_3" l1 v va)
then show ?case by (cases z, auto)
next
case ("2_4" l1 v va)
then show ?case by (cases z, auto)
next
case (3 v1 l1)
then show ?case by (cases z, auto)
next
case (4 v1 v2)
then show ?case by (cases z, auto)
next
case ("5_1" v1 v va)
then show ?case by (cases z, auto)
next
case ("5_2" v1 v va)
then show ?case by (cases z, auto)
next
case ("5_3" v1 v va)
then show ?case by (cases z, auto)
next
case (6 e1 e2 l2)
then show ?case by (cases z, auto)
next
case (7 e1 e2 v2)
then show ?case by (cases z, auto)
next
case (8 e1 e2 e1' e2')
then show ?case by (cases z, auto)
next
case ("9_1" e1 e2 v va)
then show ?case by (cases z, auto)
next
case ("9_2" e1 e2 v va)
then show ?case by (cases z, auto)
next
case (10 e1 e2 e1' e2')
then show ?case by (cases z, auto)
next
case (11 e1 e2 e1' e2')
then show ?case by (cases z, auto)
next
case ("12_1" e1 e2 v)
then show ?case by (cases z, auto)
next
case ("12_2" e1 e2 v)
then show ?case by (cases z, auto)
next
case ("12_3" e1 e2 v va)
then show ?case by (cases z, auto)
next
case (13 e1 e2 e1' e2')
then show ?case by (cases z, auto)
next
case ("14_1" e1 e2 v)
then show ?case by (cases z, auto)
next
case ("14_2" e1 e2 v)
then show ?case by (cases z, auto)
next
case ("14_3" e1 e2 v va)
then show ?case by (cases z, auto)
next
case ("14_4" e1 e2 v va)
then show ?case by (cases z, auto)
qed
lemma less_aexp_trans: "(x::'a aexp) < y \<Longrightarrow> y < z \<Longrightarrow> x < z"
apply (simp add: Let_def)
apply standard
apply (metis AExp_Lexorder.less_aexp_aux_trans dual_order.asym)
by presburger
instance proof
fix x y z :: "'a aexp"
show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)"
by (metis less_aexp_antisym less_eq_aexp_def)
show "(x \<le> x)"
by (simp add: less_eq_aexp_def)
show "x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z"
by (metis less_aexp_trans less_eq_aexp_def)
show "x \<le> y \<Longrightarrow> y \<le> x \<Longrightarrow> x = y"
unfolding less_eq_aexp_def using less_aexp_antisym by blast
show "x \<le> y \<or> y \<le> x"
unfolding less_eq_aexp_def using less_aexp_antisym by blast
qed
end
lemma smaller_height: "height a1 < height a2 \<Longrightarrow> a1 < a2"
by simp
end

File diff suppressed because it is too large Load Diff

View File

@ -0,0 +1,281 @@
section\<open>LTL for EFSMs\<close>
text\<open>This theory builds off the \texttt{Linear\_Temporal\_Logic\_on\_Streams} theory from the HOL
library and defines functions to ease the expression of LTL properties over EFSMs. Since the LTL
operators effectively act over traces of models we must find a way to express models as streams.\<close>
theory EFSM_LTL
imports "Extended_Finite_State_Machines.EFSM" "HOL-Library.Linear_Temporal_Logic_on_Streams"
begin
text_raw\<open>\snip{statedef}{1}{2}{%\<close>
record state =
statename :: "nat option"
datastate :: registers
action :: action
"output" :: outputs
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{whitebox}{1}{2}{%\<close>
type_synonym whitebox_trace = "state stream"
text_raw\<open>}%endsnip\<close>
type_synonym property = "whitebox_trace \<Rightarrow> bool"
abbreviation label :: "state \<Rightarrow> String.literal" where
"label s \<equiv> fst (action s)"
abbreviation inputs :: "state \<Rightarrow> value list" where
"inputs s \<equiv> snd (action s)"
text_raw\<open>\snip{ltlStep}{1}{2}{%\<close>
fun ltl_step :: "transition_matrix \<Rightarrow> cfstate option \<Rightarrow> registers \<Rightarrow> action \<Rightarrow> (nat option \<times> outputs \<times> registers)" where
"ltl_step _ None r _ = (None, [], r)" |
"ltl_step e (Some s) r (l, i) = (let possibilities = possible_steps e s r l i in
if possibilities = {||} then (None, [], r)
else
let (s', t) = Eps (\<lambda>x. x |\<in>| possibilities) in
(Some s', (evaluate_outputs t i r), (evaluate_updates t i r))
)"
text_raw\<open>}%endsnip\<close>
lemma ltl_step_singleton:
"\<exists>t. possible_steps e n r (fst v) (snd v) = {|(aa, t)|} \<and> evaluate_outputs t (snd v) r = b \<and> evaluate_updates t (snd v) r = c\<Longrightarrow>
ltl_step e (Some n) r v = (Some aa, b, c)"
apply (cases v)
by auto
lemma ltl_step_none: "possible_steps e s r a b = {||} \<Longrightarrow> ltl_step e (Some s) r (a, b) = (None, [], r)"
by simp
lemma ltl_step_none_2: "possible_steps e s r (fst ie) (snd ie) = {||} \<Longrightarrow> ltl_step e (Some s) r ie = (None, [], r)"
by (metis ltl_step_none prod.exhaust_sel)
lemma ltl_step_alt: "ltl_step e (Some s) r t = (
let possibilities = possible_steps e s r (fst t) (snd t) in
if possibilities = {||} then
(None, [], r)
else
let (s', t') = Eps (\<lambda>x. x |\<in>| possibilities) in
(Some s', (apply_outputs (Outputs t') (join_ir (snd t) r)), (apply_updates (Updates t') (join_ir (snd t) r) r))
)"
by (case_tac t, simp add: Let_def)
lemma ltl_step_some:
assumes "possible_steps e s r l i = {|(s', t)|}"
and "evaluate_outputs t i r = p"
and "evaluate_updates t i r = r'"
shows "ltl_step e (Some s) r (l, i) = (Some s', p, r')"
by (simp add: assms)
lemma ltl_step_cases:
assumes invalid: "P (None, [], r)"
and valid: "\<forall>(s', t) |\<in>| (possible_steps e s r l i). P (Some s', (evaluate_outputs t i r), (evaluate_updates t i r))"
shows "P (ltl_step e (Some s) r (l, i))"
apply simp
apply (case_tac "possible_steps e s r l i")
apply (simp add: invalid)
apply simp
subgoal for x S'
apply (case_tac "SOME xa. xa = x \<or> xa |\<in>| S'")
apply simp
apply (insert assms(2))
apply (simp add: fBall_def Ball_def fmember_def)
by (metis (mono_tags, lifting) fst_conv prod.case_eq_if snd_conv someI_ex)
done
text\<open>The \texttt{make\_full\_observation} function behaves similarly to \texttt{observe\_execution}
from the \texttt{EFSM} theory. The main difference in behaviour is what is recorded. While the
observe execution function simply observes an execution of the EFSM to produce the corresponding
output for each action, the intention here is to record every detail of execution, including the
values of internal variables.
Thinking of each action as a step forward in time, there are five components which characterise
a given point in the execution of an EFSM. At each point, the model has a current control state and
data state. Each action has a label and some input parameters, and its execution may produce
some observableoutput. It is therefore sufficient to provide a stream of 5-tuples containing the
current control state, data state, the label and inputs of the action, and computed output. The
make full observation function can then be defined as in Figure 9.1, with an additional
function watch defined on top of this which starts the make full observation off in the
initial control state with the empty data state.
Careful inspection of the definition reveals another way that \texttt{make\_full\_observation}
differs from \texttt{observe\_execution}. Rather than taking a cfstate, it takes a cfstate option.
The reason for this is that we need to make our EFSM models complete. That is, we need them to be
able to respond to every action from every state like a DFA. If a model does not recognise a given
action in a given state, we cannot simply stop processing because we are working with necessarily
infinite traces. Since these traces are generated by observing action sequences, the make full
observation function must keep processing whether there is a viable transition or not.
To support this, the make full observation adds an implicit ``sink state'' to every EFSM it
processes by lifting control flow state indices from \texttt{nat} to \texttt{nat option} such that
state $n$ is seen as state \texttt{Some} $n$. The control flow state \texttt{None} represents a sink
state. If a model is unable to recognise a particular action from its current state, it moves into
the \texttt{None} state. From here, the behaviour is constant for the rest of the time --- the
control flow state remains None; the data state does not change, and no output is produced.\<close>
text_raw\<open>\snip{makeFullObservation}{1}{2}{%\<close>
primcorec make_full_observation :: "transition_matrix \<Rightarrow> cfstate option \<Rightarrow> registers \<Rightarrow> outputs \<Rightarrow> action stream \<Rightarrow> whitebox_trace" where
"make_full_observation e s d p i = (
let (s', o', d') = ltl_step e s d (shd i) in
\<lparr>statename = s, datastate = d, action=(shd i), output = p\<rparr>##(make_full_observation e s' d' o' (stl i))
)"
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{watch}{1}{2}{%\<close>
abbreviation watch :: "transition_matrix \<Rightarrow> action stream \<Rightarrow> whitebox_trace" where
"watch e i \<equiv> (make_full_observation e (Some 0) <> [] i)"
text_raw\<open>}%endsnip\<close>
subsection\<open>Expressing Properties\<close>
text\<open>In order to simplify the expression and understanding of properties, this theory defines a
number of named functions which can be used to express certain properties of EFSMs.\<close>
subsubsection\<open>State Equality\<close>
text\<open>The \textsc{state\_eq} takes a cfstate option representing a control flow state index and
returns true if this is the control flow state at the head of the full observation.\<close>
abbreviation state_eq :: "cfstate option \<Rightarrow> whitebox_trace \<Rightarrow> bool" where
"state_eq v s \<equiv> statename (shd s) = v"
lemma state_eq_holds: "state_eq s = holds (\<lambda>x. statename x = s)"
apply (rule ext)
by (simp add: holds_def)
lemma state_eq_None_not_Some: "state_eq None s \<Longrightarrow> \<not> state_eq (Some n) s"
by simp
subsubsection\<open>Label Equality\<close>
text\<open>The \textsc{label\_eq} function takes a string and returns true if this is equal to the label
at the head of the full observation.\<close>
abbreviation "label_eq v s \<equiv> fst (action (shd s)) = (String.implode v)"
lemma watch_label: "label_eq l (watch e t) = (fst (shd t) = String.implode l)"
by (simp add: )
subsubsection\<open>Input Equality\<close>
text\<open>The \textsc{input\_eq} function takes a value list and returns true if this is equal to the
input at the head of the full observation.\<close>
abbreviation "input_eq v s \<equiv> inputs (shd s) = v"
subsubsection\<open>Action Equality\<close>
text\<open>The \textsc{action\_eq} function takes a (label, value list) pair and returns true if this is
equal to the action at the head of the full observation. This effectively combines
\texttt{label\_eq} and \texttt{input\_eq} into one function.\<close>
abbreviation "action_eq e \<equiv> label_eq (fst e) aand input_eq (snd e)"
subsubsection\<open>Output Equality\<close>
text\<open>The \textsc{output\_eq} function takes a takes a value option list and returns true if this is
equal to the output at the head of the full observation.\<close>
abbreviation "output_eq v s \<equiv> output (shd s) = v"
text_raw\<open>\snip{ltlVName}{1}{2}{%\<close>
datatype ltl_vname = Ip nat | Op nat | Rg nat
text_raw\<open>}%endsnip\<close>
subsubsection\<open>Checking Arbitrary Expressions\<close>
text\<open>The \textsc{check\_exp} function takes a guard expression and returns true if the guard
expression evaluates to true in the given state.\<close>
type_synonym ltl_gexp = "ltl_vname gexp"
definition join_iro :: "value list \<Rightarrow> registers \<Rightarrow> outputs \<Rightarrow> ltl_vname datastate" where
"join_iro i r p = (\<lambda>x. case x of
Rg n \<Rightarrow> r $ n |
Ip n \<Rightarrow> Some (i ! n) |
Op n \<Rightarrow> p ! n
)"
lemma join_iro_R [simp]: "join_iro i r p (Rg n) = r $ n"
by (simp add: join_iro_def)
abbreviation "check_exp g s \<equiv> (gval g (join_iro (snd (action (shd s))) (datastate (shd s)) (output (shd s))) = trilean.true)"
lemma alw_ev: "alw f = not (ev (\<lambda>s. \<not>f s))"
by simp
lemma alw_state_eq_smap:
"alw (state_eq s) ss = alw (\<lambda>ss. shd ss = s) (smap statename ss)"
apply standard
apply (simp add: alw_iff_sdrop )
by (simp add: alw_mono alw_smap )
subsection\<open>Sink State\<close>
text\<open>Once the sink state is entered, it cannot be left and there are no outputs or updates
henceforth.\<close>
lemma shd_state_is_none: "(state_eq None) (make_full_observation e None r p t)"
by (simp add: )
lemma unfold_observe_none: "make_full_observation e None d p t = (\<lparr>statename = None, datastate = d, action=(shd t), output = p\<rparr>##(make_full_observation e None d [] (stl t)))"
by (simp add: stream.expand)
lemma once_none_always_none_aux:
assumes "\<exists> p r i. j = (make_full_observation e None r p) i"
shows "alw (state_eq None) j"
using assms apply coinduct
apply simp
by fastforce
lemma once_none_always_none: "alw (state_eq None) (make_full_observation e None r p t)"
using once_none_always_none_aux by blast
lemma once_none_nxt_always_none: "alw (nxt (state_eq None)) (make_full_observation e None r p t)"
using once_none_always_none
by (simp add: alw_iff_sdrop del: sdrop.simps)
lemma snth_sconst: "(\<forall>i. s !! i = h) = (s = sconst h)"
by (metis funpow_code_def id_funpow sdrop_simps(1) sdrop_siterate siterate.simps(1) smap_alt smap_sconst snth.simps(1) stream.map_id)
lemma alw_sconst: "(alw (\<lambda>xs. shd xs = h) t) = (t = sconst h)"
by (simp add: snth_sconst[symmetric] alw_iff_sdrop)
lemma smap_statename_None: "smap statename (make_full_observation e None r p i) = sconst None"
by (meson EFSM_LTL.alw_sconst alw_state_eq_smap once_none_always_none)
lemma alw_not_some: "alw (\<lambda>xs. statename (shd xs) \<noteq> Some s) (make_full_observation e None r p t)"
by (metis (mono_tags, lifting) alw_mono once_none_always_none option.distinct(1) )
lemma state_none: "((state_eq None) impl nxt (state_eq None)) (make_full_observation e s r p t)"
by (simp add: )
lemma state_none_2:
"(state_eq None) (make_full_observation e s r p t) \<Longrightarrow>
(state_eq None) (make_full_observation e s r p (stl t))"
by (simp add: )
lemma no_output_none_aux:
assumes "\<exists> p r i. j = (make_full_observation e None r []) i"
shows "alw (output_eq []) j"
using assms apply coinduct
apply simp
by fastforce
lemma no_output_none: "nxt (alw (output_eq [])) (make_full_observation e None r p t)"
using no_output_none_aux by auto
lemma nxt_alw: "nxt (alw P) s \<Longrightarrow> alw (nxt P) s"
by (simp add: alw_iff_sdrop)
lemma no_output_none_nxt: "alw (nxt (output_eq [])) (make_full_observation e None r p t)"
using nxt_alw no_output_none by blast
lemma no_output_none_if_empty: "alw (output_eq []) (make_full_observation e None r [] t)"
by (metis (mono_tags, lifting) alw_nxt make_full_observation.simps(1) no_output_none state.select_convs(4))
lemma no_updates_none_aux:
assumes "\<exists> p i. j = (make_full_observation e None r p) i"
shows "alw (\<lambda>x. datastate (shd x) = r) j"
using assms apply coinduct
by fastforce
lemma no_updates_none: "alw (\<lambda>x. datastate (shd x) = r) (make_full_observation e None r p t)"
using no_updates_none_aux by blast
lemma action_components: "(label_eq l aand input_eq i) s = (action (shd s) = (String.implode l, i))"
by (metis fst_conv prod.collapse snd_conv)
end

View File

@ -0,0 +1,341 @@
section\<open>FSet Utilities\<close>
text\<open>This theory provides various additional lemmas, definitions, and syntax over the fset data type.\<close>
theory FSet_Utils
imports "HOL-Library.FSet"
begin
notation (latex output)
"FSet.fempty" ("\<emptyset>") and
"FSet.fmember" ("\<in>")
syntax (ASCII)
"_fBall" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3ALL (_/:_)./ _)" [0, 0, 10] 10)
"_fBex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3EX (_/:_)./ _)" [0, 0, 10] 10)
"_fBex1" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3EX! (_/:_)./ _)" [0, 0, 10] 10)
syntax (input)
"_fBall" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3! (_/:_)./ _)" [0, 0, 10] 10)
"_fBex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3? (_/:_)./ _)" [0, 0, 10] 10)
"_fBex1" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3?! (_/:_)./ _)" [0, 0, 10] 10)
syntax
"_fBall" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3\<forall>(_/|\<in>|_)./ _)" [0, 0, 10] 10)
"_fBex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3\<exists>(_/|\<in>|_)./ _)" [0, 0, 10] 10)
"_fBnex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3\<nexists>(_/|\<in>|_)./ _)" [0, 0, 10] 10)
"_fBex1" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3\<exists>!(_/|\<in>|_)./ _)" [0, 0, 10] 10)
translations
"\<forall>x|\<in>|A. P" \<rightleftharpoons> "CONST fBall A (\<lambda>x. P)"
"\<exists>x|\<in>|A. P" \<rightleftharpoons> "CONST fBex A (\<lambda>x. P)"
"\<nexists>x|\<in>|A. P" \<rightleftharpoons> "CONST fBall A (\<lambda>x. \<not>P)"
"\<exists>!x|\<in>|A. P" \<rightharpoonup> "\<exists>!x. x |\<in>| A \<and> P"
lemma fset_of_list_remdups [simp]: "fset_of_list (remdups l) = fset_of_list l"
apply (induct l)
apply simp
by (simp add: finsert_absorb fset_of_list_elem)
definition "fSum \<equiv> fsum (\<lambda>x. x)"
lemma fset_both_sides: "(Abs_fset s = f) = (fset (Abs_fset s) = fset f)"
by (simp add: fset_inject)
lemma Abs_ffilter: "(ffilter f s = s') = ({e \<in> (fset s). f e} = (fset s'))"
by (simp add: ffilter_def fset_both_sides Abs_fset_inverse Set.filter_def)
lemma size_ffilter_card: "size (ffilter f s) = card ({e \<in> (fset s). f e})"
by (simp add: ffilter_def fset_both_sides Abs_fset_inverse Set.filter_def)
lemma ffilter_empty [simp]: "ffilter f {||} = {||}"
by auto
lemma ffilter_finsert:
"ffilter f (finsert a s) = (if f a then finsert a (ffilter f s) else (ffilter f s))"
apply simp
apply standard
apply (simp add: ffilter_def fset_both_sides Abs_fset_inverse)
apply auto[1]
apply (simp add: ffilter_def fset_both_sides Abs_fset_inverse)
by auto
lemma fset_equiv: "(f1 = f2) = (fset f1 = fset f2)"
by (simp add: fset_inject)
lemma finsert_equiv: "(finsert e f = f') = (insert e (fset f) = (fset f'))"
by (simp add: finsert_def fset_both_sides Abs_fset_inverse)
lemma filter_elements:
"x |\<in>| Abs_fset (Set.filter f (fset s)) = (x \<in> (Set.filter f (fset s)))"
by (metis ffilter.rep_eq fset_inverse notin_fset)
lemma sorted_list_of_fempty [simp]: "sorted_list_of_fset {||} = []"
by (simp add: sorted_list_of_fset_def)
lemma fmember_implies_member: "e |\<in>| f \<Longrightarrow> e \<in> fset f"
by (simp add: fmember_def)
lemma fold_union_ffUnion: "fold (|\<union>|) l {||} = ffUnion (fset_of_list l)"
by(induct l rule: rev_induct, auto)
lemma filter_filter:
"ffilter P (ffilter Q xs) = ffilter (\<lambda>x. Q x \<and> P x) xs"
by auto
lemma fsubset_strict:
"x2 |\<subset>| x1 \<Longrightarrow> \<exists>e. e |\<in>| x1 \<and> e |\<notin>| x2"
by auto
lemma fsubset:
"x2 |\<subset>| x1 \<Longrightarrow> \<nexists>e. e |\<in>| x2 \<and> e |\<notin>| x1"
by auto
lemma size_fsubset_elem:
assumes "\<exists>e. e |\<in>| x1 \<and> e |\<notin>| x2"
and "\<nexists>e. e |\<in>| x2 \<and> e |\<notin>| x1"
shows "size x2 < size x1"
using assms
apply (simp add: fmember_def)
by (metis card_seteq finite_fset linorder_not_le subsetI)
lemma size_fsubset: "x2 |\<subset>| x1 \<Longrightarrow> size x2 < size x1"
by (metis fsubset fsubset_strict size_fsubset_elem)
definition fremove :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
where [code_abbrev]: "fremove x A = A - {|x|}"
lemma arg_cong_ffilter:
"\<forall>e |\<in>| f. p e = p' e \<Longrightarrow> ffilter p f = ffilter p' f"
by auto
lemma ffilter_singleton: "f e \<Longrightarrow> ffilter f {|e|} = {|e|}"
apply (simp add: ffilter_def fset_both_sides Abs_fset_inverse)
by auto
lemma fset_eq_alt: "(x = y) = (x |\<subseteq>| y \<and> size x = size y)"
by (metis exists_least_iff le_less size_fsubset)
lemma ffold_empty [simp]: "ffold f b {||} = b"
by (simp add: ffold_def)
lemma sorted_list_of_fset_sort:
"sorted_list_of_fset (fset_of_list l) = sort (remdups l)"
by (simp add: fset_of_list.rep_eq sorted_list_of_fset.rep_eq sorted_list_of_set_sort_remdups)
lemma fMin_Min: "fMin (fset_of_list l) = Min (set l)"
by (simp add: fMin.F.rep_eq fset_of_list.rep_eq)
lemma sorted_hd_Min:
"sorted l \<Longrightarrow>
l \<noteq> [] \<Longrightarrow>
hd l = Min (set l)"
by (metis List.finite_set Min_eqI eq_iff hd_Cons_tl insertE list.set_sel(1) list.simps(15) sorted.simps(2))
lemma hd_sort_Min: "l \<noteq> [] \<Longrightarrow> hd (sort l) = Min (set l)"
by (metis sorted_hd_Min set_empty set_sort sorted_sort)
lemma hd_sort_remdups: "hd (sort (remdups l)) = hd (sort l)"
by (metis hd_sort_Min remdups_eq_nil_iff set_remdups)
lemma exists_fset_of_list: "\<exists>l. f = fset_of_list l"
using exists_fset_of_list by fastforce
lemma hd_sorted_list_of_fset:
"s \<noteq> {||} \<Longrightarrow> hd (sorted_list_of_fset s) = (fMin s)"
apply (insert exists_fset_of_list[of s])
apply (erule exE)
apply simp
apply (simp add: sorted_list_of_fset_sort fMin_Min hd_sort_remdups)
by (metis fset_of_list_simps(1) hd_sort_Min)
lemma fminus_filter_singleton:
"fset_of_list l |-| {|x|} = fset_of_list (filter (\<lambda>e. e \<noteq> x) l)"
by auto
lemma card_minus_fMin:
"s \<noteq> {||} \<Longrightarrow> card (fset s - {fMin s}) < card (fset s)"
by (metis Min_in bot_fset.rep_eq card_Diff1_less fMin.F.rep_eq finite_fset fset_equiv)
(* Provides a deterministic way to fold fsets similar to List.fold that works with the code generator *)
function ffold_ord :: "(('a::linorder) \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b \<Rightarrow> 'b" where
"ffold_ord f s b = (
if s = {||} then
b
else
let
h = fMin s;
t = s - {|h|}
in
ffold_ord f t (f h b)
)"
by auto
termination
apply (relation "measures [\<lambda>(a, s, ab). size s]")
apply simp
by (simp add: card_minus_fMin)
lemma sorted_list_of_fset_Cons:
"\<exists>h t. (sorted_list_of_fset (finsert s ss)) = h#t"
apply (simp add: sorted_list_of_fset_def)
by (cases "insort s (sorted_list_of_set (fset ss - {s}))", auto)
lemma list_eq_hd_tl:
"l \<noteq> [] \<Longrightarrow>
hd l = h \<Longrightarrow>
tl l = t \<Longrightarrow>
l = (h#t)"
by auto
lemma fset_of_list_sort: "fset_of_list l = fset_of_list (sort l)"
by (simp add: fset_of_list.abs_eq)
lemma exists_sorted_distinct_fset_of_list:
"\<exists>l. sorted l \<and> distinct l \<and> f = fset_of_list l"
by (metis distinct_sorted_list_of_set sorted_list_of_fset.rep_eq sorted_list_of_fset_simps(2) sorted_sorted_list_of_set)
lemma fset_of_list_empty [simp]: "(fset_of_list l = {||}) = (l = [])"
by (metis fset_of_list.rep_eq fset_of_list_simps(1) set_empty)
lemma ffold_ord_cons: assumes sorted: "sorted (h#t)"
and distinct: "distinct (h#t)"
shows "ffold_ord f (fset_of_list (h#t)) b = ffold_ord f (fset_of_list t) (f h b)"
proof-
have h_is_min: "h = fMin (fset_of_list (h#t))"
by (metis sorted fMin_Min list.sel(1) list.simps(3) sorted_hd_Min)
have remove_min: "fset_of_list t = (fset_of_list (h#t)) - {|h|}"
using distinct fset_of_list_elem by force
show ?thesis
apply (simp only: ffold_ord.simps[of f "fset_of_list (h#t)"])
by (metis h_is_min remove_min fset_of_list_empty list.distinct(1))
qed
lemma sorted_distinct_ffold_ord: assumes "sorted l"
and "distinct l"
shows "ffold_ord f (fset_of_list l) b = fold f l b"
using assms
apply (induct l arbitrary: b)
apply simp
by (metis distinct.simps(2) ffold_ord_cons fold_simps(2) sorted.simps(2))
lemma ffold_ord_fold_sorted: "ffold_ord f s b = fold f (sorted_list_of_fset s) b"
by (metis exists_sorted_distinct_fset_of_list sorted_distinct_ffold_ord distinct_remdups_id sorted_list_of_fset_sort sorted_sort_id)
context includes fset.lifting begin
lift_definition fprod :: "'a fset \<Rightarrow> 'b fset \<Rightarrow> ('a \<times> 'b) fset " (infixr "|\<times>|" 80) is "\<lambda>a b. fset a \<times> fset b"
by simp
lift_definition fis_singleton :: "'a fset \<Rightarrow> bool" is "\<lambda>A. is_singleton (fset A)".
end
lemma fprod_empty_l: "{||} |\<times>| a = {||}"
using bot_fset_def fprod.abs_eq by force
lemma fprod_empty_r: "a |\<times>| {||} = {||}"
by (simp add: fprod_def bot_fset_def Abs_fset_inverse)
lemmas fprod_empty = fprod_empty_l fprod_empty_r
lemma fprod_finsert: "(finsert a as) |\<times>| (finsert b bs) =
finsert (a, b) (fimage (\<lambda>b. (a, b)) bs |\<union>| fimage (\<lambda>a. (a, b)) as |\<union>| (as |\<times>| bs))"
apply (simp add: fprod_def fset_both_sides Abs_fset_inverse)
by auto
lemma fprod_member:
"x |\<in>| xs \<Longrightarrow>
y |\<in>| ys \<Longrightarrow>
(x, y) |\<in>| xs |\<times>| ys"
by (simp add: fmember_def fprod_def Abs_fset_inverse)
lemma fprod_subseteq:
"x |\<subseteq>| x' \<and> y |\<subseteq>| y' \<Longrightarrow> x |\<times>| y |\<subseteq>| x' |\<times>| y'"
apply (simp add: fprod_def less_eq_fset_def Abs_fset_inverse)
by auto
lemma fimage_fprod:
"(a, b) |\<in>| A |\<times>| B \<Longrightarrow> f a b |\<in>| (\<lambda>(x, y). f x y) |`| (A |\<times>| B)"
by force
lemma fprod_singletons: "{|a|} |\<times>| {|b|} = {|(a, b)|}"
apply (simp add: fprod_def)
by (metis fset_inverse fset_simps(1) fset_simps(2))
lemma fprod_equiv:
"(fset (f |\<times>| f') = s) = (((fset f) \<times> (fset f')) = s)"
by (simp add: fprod_def Abs_fset_inverse)
lemma fis_singleton_alt: "fis_singleton f = (\<exists>e. f = {|e|})"
by (metis fis_singleton.rep_eq fset_inverse fset_simps(1) fset_simps(2) is_singleton_def)
lemma singleton_singleton [simp]: "fis_singleton {|a|}"
by (simp add: fis_singleton_def)
lemma not_singleton_empty [simp]: "\<not> fis_singleton {||}"
apply (simp add: fis_singleton_def)
by (simp add: is_singleton_altdef)
lemma fis_singleton_fthe_elem:
"fis_singleton A \<longleftrightarrow> A = {|fthe_elem A|}"
by (metis fis_singleton_alt fthe_felem_eq)
lemma fBall_ffilter:
"\<forall>x |\<in>| X. f x \<Longrightarrow> ffilter f X = X"
by auto
lemma fBall_ffilter2:
"X = Y \<Longrightarrow>
\<forall>x |\<in>| X. f x \<Longrightarrow>
ffilter f X = Y"
by auto
lemma size_fset_of_list: "size (fset_of_list l) = length (remdups l)"
apply (induct l)
apply simp
by (simp add: fset_of_list.rep_eq insert_absorb)
lemma size_fsingleton: "(size f = 1) = (\<exists>e. f = {|e|})"
apply (insert exists_fset_of_list[of f])
apply clarify
apply (simp only: size_fset_of_list)
apply (simp add: fset_of_list_def fset_both_sides Abs_fset_inverse)
by (metis List.card_set One_nat_def card.insert card_1_singletonE card_empty empty_iff finite.intros(1))
lemma ffilter_mono: "(ffilter X xs = f) \<Longrightarrow> \<forall>x |\<in>| xs. X x = Y x \<Longrightarrow> (ffilter Y xs = f)"
by auto
lemma size_fimage: "size (fimage f s) \<le> size s"
apply (induct s)
apply simp
by (simp add: card_insert_if)
lemma size_ffilter: "size (ffilter P f) \<le> size f"
apply (induct f)
apply simp
apply (simp only: ffilter_finsert)
apply (case_tac "P x")
apply (simp add: fmember.rep_eq)
by (simp add: card_insert_if)
lemma fimage_size_le: "\<And>f s. size s \<le> n \<Longrightarrow> size (fimage f s) \<le> n"
using le_trans size_fimage by blast
lemma ffilter_size_le: "\<And>f s. size s \<le> n \<Longrightarrow> size (ffilter f s) \<le> n"
using dual_order.trans size_ffilter by blast
lemma set_membership_eq: "A = B \<longleftrightarrow> (\<lambda>x. Set.member x A) = (\<lambda>x. Set.member x B)"
apply standard
apply simp
by (meson equalityI subsetI)
lemmas ffilter_eq_iff = Abs_ffilter set_membership_eq fun_eq_iff
lemma size_le_1: "size f \<le> 1 = (f = {||} \<or> (\<exists>e. f = {|e|}))"
apply standard
apply (metis bot.not_eq_extremum gr_implies_not0 le_neq_implies_less less_one size_fsingleton size_fsubset)
by auto
lemma size_gt_1: "1 < size f \<Longrightarrow> \<exists>e1 e2 f'. e1 \<noteq> e2 \<and> f = finsert e1 (finsert e2 f')"
apply (induct f)
apply simp
apply (rule_tac x=x in exI)
by (metis finsertCI leD not_le_imp_less size_le_1)
end

View File

@ -0,0 +1,837 @@
subsection \<open>Guards Expressions\<close>
text\<open>
This theory defines the guard language of EFSMs which can be translated directly to and from
contexts. Boolean values true and false respectively represent the guards which are always and never
satisfied. Guards may test for (in)equivalence of two arithmetic expressions or be connected using
\textsc{nor} logic into compound expressions. The use of \textsc{nor} logic reduces the number of
subgoals when inducting over guard expressions.
We also define syntax hacks for the relations less than, less than or equal to, greater than or
equal to, and not equal to as well as the expression of logical conjunction, disjunction, and
negation in terms of nor logic.\<close>
theory GExp
imports AExp Trilean
begin
text_raw\<open>\snip{gexptype}{1}{2}{%\<close>
datatype 'a gexp = Bc bool | Eq "'a aexp" "'a aexp" | Gt "'a aexp" "'a aexp" | In 'a "value list" | Nor "'a gexp" "'a gexp"
text_raw\<open>}%endsnip\<close>
fun gval :: "'a gexp \<Rightarrow> 'a datastate \<Rightarrow> trilean" where
"gval (Bc True) _ = true" |
"gval (Bc False) _ = false" |
"gval (Gt a1 a2) s = value_gt (aval a1 s) (aval a2 s)" |
"gval (Eq a1 a2) s = value_eq (aval a1 s) (aval a2 s)" |
"gval (In v l) s = (case s v of None \<Rightarrow> invalid | Some vv \<Rightarrow> if vv \<in> set l then true else false)" |
"gval (Nor a1 a2) s = \<not>? ((gval a1 s) \<or>? (gval a2 s))"
text_raw\<open>\snip{connectives}{1}{2}{%\<close>
definition gNot :: "'a gexp \<Rightarrow> 'a gexp" where
"gNot g \<equiv> Nor g g"
definition gOr :: "'a gexp \<Rightarrow> 'a gexp \<Rightarrow> 'a gexp" (*infix "\<or>" 60*) where
"gOr v va \<equiv> Nor (Nor v va) (Nor v va)"
definition gAnd :: "'a gexp \<Rightarrow> 'a gexp \<Rightarrow> 'a gexp" (*infix "\<and>" 60*) where
"gAnd v va \<equiv> Nor (Nor v v) (Nor va va)"
definition gImplies :: "'a gexp \<Rightarrow> 'a gexp \<Rightarrow> 'a gexp" where
"gImplies p q \<equiv> gOr (gNot p) q"
definition Lt :: "'a aexp \<Rightarrow> 'a aexp \<Rightarrow> 'a gexp" (*infix "<" 60*) where
"Lt a b \<equiv> Gt b a"
definition Le :: "'a aexp \<Rightarrow> 'a aexp \<Rightarrow> 'a gexp" (*infix "\<le>" 60*) where
"Le v va \<equiv> gNot (Gt v va)"
definition Ge :: "'a aexp \<Rightarrow> 'a aexp \<Rightarrow> 'a gexp" (*infix "\<ge>" 60*) where
"Ge v va \<equiv> gNot (Lt v va)"
definition Ne :: "'a aexp \<Rightarrow> 'a aexp \<Rightarrow> 'a gexp" (*infix "\<noteq>" 60*) where
"Ne v va \<equiv> gNot (Eq v va)"
text_raw\<open>}%endsnip\<close>
lemma gval_Lt [simp]:
"gval (Lt a1 a2) s = value_gt (aval a2 s) (aval a1 s)"
by (simp add: Lt_def)
lemma gval_Le [simp]:
"gval (Le a1 a2) s = \<not>? (value_gt (aval a1 s) (aval a2 s))"
by (simp add: Le_def value_gt_def gNot_def maybe_or_idempotent)
lemma gval_Ge [simp]:
"gval (Ge a1 a2) s = \<not>? (value_gt (aval a2 s) (aval a1 s))"
by (simp add: Ge_def value_gt_def gNot_def maybe_or_idempotent)
lemma gval_Ne [simp]:
"gval (Ne a1 a2) s = \<not>? (value_eq (aval a1 s) (aval a2 s))"
by (simp add: Ne_def value_gt_def gNot_def maybe_or_idempotent)
lemmas connectives = gAnd_def gOr_def gNot_def Lt_def Le_def Ge_def Ne_def
lemma gval_gOr [simp]: "gval (gOr x y) r = (gval x r) \<or>? (gval y r)"
by (simp add: maybe_double_negation maybe_or_idempotent gOr_def)
lemma gval_gNot [simp]: "gval (gNot x) s = \<not>? (gval x s)"
by (simp add: maybe_or_idempotent gNot_def)
lemma gval_gAnd [simp]:
"gval (gAnd g1 g2) s = (gval g1 s) \<and>? (gval g2 s)"
by (simp add: de_morgans_1 maybe_double_negation maybe_or_idempotent gAnd_def)
lemma gAnd_commute: "gval (gAnd a b) s = gval (gAnd b a) s"
by (simp add: times_trilean_commutative)
lemma gOr_commute: "gval (gOr a b) s = gval (gOr b a) s"
by (simp add: plus_trilean_commutative gOr_def)
lemma gval_gAnd_True:
"(gval (gAnd g1 g2) s = true) = ((gval g1 s = true) \<and> gval g2 s = true)"
by (simp add: maybe_and_true)
lemma nor_equiv: "gval (gNot (gOr a b)) s = gval (Nor a b) s"
by simp
definition satisfiable :: "vname gexp \<Rightarrow> bool" where
"satisfiable g \<equiv> (\<exists>i r. gval g (join_ir i r) = true)"
definition "satisfiable_list l = satisfiable (fold gAnd l (Bc True))"
lemma unsatisfiable_false: "\<not> satisfiable (Bc False)"
by (simp add: satisfiable_def)
lemma satisfiable_true: "satisfiable (Bc True)"
by (simp add: satisfiable_def)
definition valid :: "vname gexp \<Rightarrow> bool" where
"valid g \<equiv> (\<forall>s. gval g s = true)"
lemma valid_true: "valid (Bc True)"
by (simp add: valid_def)
fun gexp_constrains :: "'a gexp \<Rightarrow> 'a aexp \<Rightarrow> bool" where
"gexp_constrains (Bc _) _ = False" |
"gexp_constrains (Eq a1 a2) a = (aexp_constrains a1 a \<or> aexp_constrains a2 a)" |
"gexp_constrains (Gt a1 a2) a = (aexp_constrains a1 a \<or> aexp_constrains a2 a)" |
"gexp_constrains (Nor g1 g2) a = (gexp_constrains g1 a \<or> gexp_constrains g2 a)" |
"gexp_constrains (In v l) a = aexp_constrains (V v) a"
fun contains_bool :: "'a gexp \<Rightarrow> bool" where
"contains_bool (Bc _) = True" |
"contains_bool (Nor g1 g2) = (contains_bool g1 \<or> contains_bool g2)" |
"contains_bool _ = False"
fun gexp_same_structure :: "'a gexp \<Rightarrow> 'a gexp \<Rightarrow> bool" where
"gexp_same_structure (Bc b) (Bc b') = (b = b')" |
"gexp_same_structure (Eq a1 a2) (Eq a1' a2') = (aexp_same_structure a1 a1' \<and> aexp_same_structure a2 a2')" |
"gexp_same_structure (Gt a1 a2) (Gt a1' a2') = (aexp_same_structure a1 a1' \<and> aexp_same_structure a2 a2')" |
"gexp_same_structure (Nor g1 g2) (Nor g1' g2') = (gexp_same_structure g1 g1' \<and> gexp_same_structure g2 g2')" |
"gexp_same_structure (In v l) (In v' l') = (v = v' \<and> l = l')" |
"gexp_same_structure _ _ = False"
lemma gval_foldr_true:
"(gval (foldr gAnd G (Bc True)) s = true) = (\<forall>g \<in> set G. gval g s = true)"
proof(induct G)
case (Cons a G)
then show ?case
apply (simp only: foldr.simps comp_def gval_gAnd maybe_and_true)
by simp
qed auto
fun enumerate_gexp_inputs :: "vname gexp \<Rightarrow> nat set" where
"enumerate_gexp_inputs (Bc _) = {}" |
"enumerate_gexp_inputs (Eq v va) = enumerate_aexp_inputs v \<union> enumerate_aexp_inputs va" |
"enumerate_gexp_inputs (Gt v va) = enumerate_aexp_inputs v \<union> enumerate_aexp_inputs va" |
"enumerate_gexp_inputs (In v va) = enumerate_aexp_inputs (V v)" |
"enumerate_gexp_inputs (Nor v va) = enumerate_gexp_inputs v \<union> enumerate_gexp_inputs va"
lemma enumerate_gexp_inputs_list: "\<exists>l. enumerate_gexp_inputs g = set l"
proof(induct g)
case (Eq x1a x2)
then show ?case
by (metis enumerate_aexp_inputs_list enumerate_gexp_inputs.simps(2) set_append)
next
case (Gt x1a x2)
then show ?case
by (metis enumerate_aexp_inputs_list enumerate_gexp_inputs.simps(3) set_append)
next
case (In x1a x2)
then show ?case
by (simp add: enumerate_aexp_inputs_list)
next
case (Nor g1 g2)
then show ?case
by (metis enumerate_gexp_inputs.simps(5) set_append)
qed auto
definition max_input :: "vname gexp \<Rightarrow> nat option" where
"max_input g = (let inputs = (enumerate_gexp_inputs g) in if inputs = {} then None else Some (Max inputs))"
definition max_input_list :: "vname gexp list \<Rightarrow> nat option" where
"max_input_list g = fold max (map (\<lambda>g. max_input g) g) None"
lemma max_input_list_cons:
"max_input_list (a # G) = max (max_input a) (max_input_list G)"
apply (simp add: max_input_list_def)
apply (cases "max_input a")
apply (simp add: max_def_raw)
by (metis (no_types, lifting) List.finite_set Max.insert Max.set_eq_fold fold_simps(1) list.set(2) max.assoc set_empty)
fun enumerate_regs :: "vname gexp \<Rightarrow> nat set" where
"enumerate_regs (Bc _) = {}" |
"enumerate_regs (Eq v va) = AExp.enumerate_regs v \<union> AExp.enumerate_regs va" |
"enumerate_regs (Gt v va) = AExp.enumerate_regs v \<union> AExp.enumerate_regs va" |
"enumerate_regs (In v va) = AExp.enumerate_regs (V v)" |
"enumerate_regs (Nor v va) = enumerate_regs v \<union> enumerate_regs va"
lemma finite_enumerate_regs: "finite (enumerate_regs g)"
using AExp.finite_enumerate_regs by (induct g, auto)
definition max_reg :: "vname gexp \<Rightarrow> nat option" where
"max_reg g = (let regs = (enumerate_regs g) in if regs = {} then None else Some (Max regs))"
lemma max_reg_gNot: "max_reg (gNot x) = max_reg x"
by (simp add: max_reg_def gNot_def)
lemma max_reg_Eq: "max_reg (Eq a b) = max (AExp.max_reg a) (AExp.max_reg b)"
apply (simp add: max_reg_def AExp.max_reg_def Let_def max_absorb2)
by (metis AExp.finite_enumerate_regs Max.union bot_option_def max_bot2 sup_Some sup_max)
lemma max_reg_Gt: "max_reg (Gt a b) = max (AExp.max_reg a) (AExp.max_reg b)"
apply (simp add: max_reg_def AExp.max_reg_def Let_def max_absorb2)
by (metis AExp.finite_enumerate_regs Max.union bot_option_def max_bot2 sup_Some sup_max)
lemma max_reg_Nor: "max_reg (Nor a b) = max (max_reg a) (max_reg b)"
apply (simp add: max_reg_def AExp.max_reg_def Let_def max_absorb2)
by (metis GExp.finite_enumerate_regs Max.union bot_option_def max_bot2 sup_Some sup_max)
lemma gval_In_cons:
"gval (In v (a # as)) s = (gval (Eq (V v) (L a)) s \<or>? gval (In v as) s)"
by (cases "s v", auto)
lemma possible_to_be_in: "s \<noteq> [] \<Longrightarrow> satisfiable (In v s)"
proof-
assume "s \<noteq> []"
have aux: "\<exists>v' i r. join_ir i r v = Some v' \<and> v' \<in> set s \<Longrightarrow>
\<exists>i r. (case join_ir i r v of None \<Rightarrow> false | Some v \<Rightarrow> if v \<in> set s then true else false) = true"
by (metis (mono_tags, lifting) option.simps(5))
show ?thesis
apply (simp add: satisfiable_def gval_In_cons)
apply (cases s)
apply (simp add: \<open>s \<noteq> []\<close>)
apply (cases v)
apply (case_tac "\<exists>(i::value list). length i > x1 \<and> i ! x1 = a")
apply clarsimp
subgoal for _ _ i by (rule exI[of _ i], intro exI, simp)
apply (metis gt_ex length_list_update length_repeat nth_list_update_eq)
apply (rule_tac exI)
apply (case_tac "\<exists>r. r $ x2 = Some a")
apply clarsimp
subgoal for _ _ _ r by (rule exI[of _ r], simp)
by (metis join_ir_R join_ir_double_exists)
qed
definition max_reg_list :: "vname gexp list \<Rightarrow> nat option" where
"max_reg_list g = (fold max (map (\<lambda>g. max_reg g) g) None)"
lemma max_reg_list_cons:
"max_reg_list (a # G) = max (max_reg a) (max_reg_list G)"
apply (simp add: max_reg_list_def)
by (metis (no_types, lifting) List.finite_set Max.insert Max.set_eq_fold fold.simps(1) id_apply list.simps(15) max.assoc set_empty)
lemma max_reg_list_append_singleton:
"max_reg_list (as@[bs]) = max (max_reg_list as) (max_reg_list [bs])"
apply (simp add: max_reg_list_def)
by (metis max.commute sup_None_2 sup_max)
lemma max_reg_list_append:
"max_reg_list (as@bs) = max (max_reg_list as) (max_reg_list bs)"
proof(induct bs rule: rev_induct)
case Nil
then show ?case
by (metis append_Nil2 fold_simps(1) list.simps(8) max_reg_list_def sup_None_2 sup_max)
next
case (snoc x xs)
then show ?case
by (metis append_assoc max.assoc max_reg_list_append_singleton)
qed
definition apply_guards :: "vname gexp list \<Rightarrow> vname datastate \<Rightarrow> bool" where
"apply_guards G s = (\<forall>g \<in> set (map (\<lambda>g. gval g s) G). g = true)"
lemma apply_guards_singleton[simp]: "(apply_guards [g] s) = (gval g s = true)"
by (simp add: apply_guards_def)
lemma apply_guards_empty [simp]: "apply_guards [] s"
by (simp add: apply_guards_def)
lemma apply_guards_cons:
"apply_guards (a # G) c = (gval a c = true \<and> apply_guards G c)"
by (simp add: apply_guards_def)
lemma apply_guards_double_cons:
"apply_guards (y # x # G) s = (gval (gAnd y x) s = true \<and> apply_guards G s)"
using apply_guards_cons gval_gAnd_True by blast
lemma apply_guards_append:
"apply_guards (a@a') s = (apply_guards a s \<and> apply_guards a' s)"
using apply_guards_def by auto
lemma apply_guards_foldr:
"apply_guards G s = (gval (foldr gAnd G (Bc True)) s = true)"
proof(induct G)
case Nil
then show ?case
by (simp add: apply_guards_def)
next
case (Cons a G)
then show ?case
by (metis apply_guards_cons foldr.simps(2) gval_gAnd_True o_apply)
qed
lemma rev_apply_guards: "apply_guards (rev G) s = apply_guards G s"
by (simp add: apply_guards_def)
lemma apply_guards_fold:
"apply_guards G s = (gval (fold gAnd G (Bc True)) s = true)"
using rev_apply_guards[symmetric]
by (simp add: foldr_conv_fold apply_guards_foldr)
lemma fold_apply_guards:
"(gval (fold gAnd G (Bc True)) s = true) = apply_guards G s"
by (simp add: apply_guards_fold)
lemma foldr_apply_guards:
"(gval (foldr gAnd G (Bc True)) s = true) = apply_guards G s"
by (simp add: apply_guards_foldr)
lemma apply_guards_subset:
"set g' \<subseteq> set g \<Longrightarrow> apply_guards g c \<longrightarrow> apply_guards g' c"
proof(induct g)
case (Cons a g)
then show ?case
using apply_guards_def by auto
qed auto
lemma apply_guards_subset_append:
"set G \<subseteq> set G' \<Longrightarrow> apply_guards (G @ G') s = apply_guards (G') s"
using apply_guards_append apply_guards_subset by blast
lemma apply_guards_rearrange:
"x \<in> set G \<Longrightarrow> apply_guards G s = apply_guards (x#G) s"
using apply_guards_def by auto
lemma apply_guards_condense: "\<exists>g. apply_guards G s = (gval g s = true)"
using apply_guards_fold by blast
lemma apply_guards_false_condense: "\<exists>g. (\<not>apply_guards G s) = (gval g s = false)"
using foldr_apply_guards gval.simps(2) not_true by blast
lemma max_input_Bc: "max_input (Bc x) = None"
by (simp add: max_input_def)
lemma max_input_Eq:
"max_input (Eq a1 a2) = max (AExp.max_input a1) (AExp.max_input a2)"
apply (simp add: AExp.max_input_def max_input_def Let_def max_absorb2)
by (metis List.finite_set Max.union bot_option_def enumerate_aexp_inputs_not_empty max_bot2 sup_Some sup_max)
lemma max_input_Gt:
"max_input (Gt a1 a2) = max (AExp.max_input a1) (AExp.max_input a2)"
apply (simp add: AExp.max_input_def max_input_def Let_def max_absorb2)
by (metis List.finite_set Max.union bot_option_def enumerate_aexp_inputs_not_empty max_bot2 sup_Some sup_max)
lemma gexp_max_input_Nor:
"max_input (Nor g1 g2) = max (max_input g1) (max_input g2)"
apply (simp add: AExp.max_input_def max_input_def Let_def max_absorb2)
by (metis List.finite_set Max.union enumerate_gexp_inputs_list less_eq_option_Some_None max_def sup_Some sup_max)
lemma gexp_max_input_In: "max_input (In v l) = AExp.max_input (V v)"
by (simp add: AExp.max_input_def GExp.max_input_def)
lemma gval_foldr_gOr_invalid:
"(gval (fold gOr l g) s = invalid) = (\<exists>g' \<in> (set (g#l)). gval g' s = invalid)"
proof(induct l rule: rev_induct)
case (snoc x xs)
then show ?case
by (simp, metis gval_gOr maybe_or_invalid)
qed auto
lemma gval_foldr_gOr_true:
"(gval (fold gOr l g) s = true) = ((\<exists>g' \<in> (set (g#l)). gval g' s = true) \<and> (\<forall>g' \<in> (set (g#l)). gval g' s \<noteq> invalid))"
proof(induct l rule: rev_induct)
case (snoc x xs)
then show ?case
apply (simp add: maybe_or_true)
using gval_foldr_gOr_invalid by auto
qed auto
lemma gval_foldr_gOr_false:
"(gval (fold gOr l g) s = false) = (\<forall>g' \<in> (set (g#l)). gval g' s = false)"
proof(induct l rule: rev_induct)
case (snoc x xs)
then show ?case
by (auto simp add: maybe_or_false)
qed auto
lemma gval_fold_gOr_rev: "gval (fold gOr (rev l) g) s = gval (fold gOr l g) s"
apply (cases "gval (fold gOr l g) s")
apply (simp, simp add: gval_foldr_gOr_true)
apply (simp, simp add: gval_foldr_gOr_false)
by (simp, simp add: gval_foldr_gOr_invalid)
lemma gval_fold_gOr_foldr: "gval (fold gOr l g) s = gval (foldr gOr l g) s"
by (simp add: foldr_conv_fold gval_fold_gOr_rev)
lemma gval_fold_gOr:
"gval (fold gOr (a # l) g) s = (gval a s \<or>? gval (fold gOr l g) s)"
by (simp only: gval_fold_gOr_foldr foldr.simps comp_def gval_gOr)
lemma gval_In_fold:
"gval (In v l) s = (if s v = None then invalid else gval (fold gOr (map (\<lambda>x. Eq (V v) (L x)) l) (Bc False)) s)"
proof(induct l)
case Nil
then show ?case
apply simp
apply (cases "s v")
apply simp
by auto
next
case (Cons a l)
then show ?case
apply (simp only: gval_In_cons)
apply (cases "s v")
apply simp
by (simp add: gval_fold_gOr del: fold.simps)
qed
fun fold_In :: "'a \<Rightarrow> value list \<Rightarrow> 'a gexp" where
"fold_In _ [] = Bc False" |
"fold_In v (l#t) = gOr (Eq (V v) (L l)) (fold_In v t)"
lemma gval_fold_In: "l \<noteq> [] \<Longrightarrow> gval (In v l) s = gval (fold_In v l) s"
proof(induct l)
next
case (Cons a l)
then show ?case
apply (case_tac "s v")
apply simp
apply simp
apply safe
apply simp
apply (metis fold_In.simps(1) gval.simps(2) plus_trilean.simps(4) plus_trilean.simps(5))
apply fastforce
apply fastforce
by fastforce
qed auto
lemma fold_maybe_or_invalid_base: "fold (\<or>?) l invalid = invalid"
proof(induct l)
case (Cons a l)
then show ?case
by (metis fold_simps(2) maybe_or_valid)
qed auto
lemma fold_maybe_or_true_base_never_false:
"fold (\<or>?) l true \<noteq> false"
proof(induct l)
case (Cons a l)
then show ?case
by (metis fold_maybe_or_invalid_base fold_simps(2) maybe_not.cases maybe_or_valid plus_trilean.simps(4) plus_trilean.simps(6))
qed auto
lemma fold_true_fold_false_not_invalid:
"fold (\<or>?) l true = true \<Longrightarrow>
fold (\<or>?) (rev l) false \<noteq> invalid"
proof(induct l)
case (Cons a l)
then show ?case
apply simp
by (metis fold_maybe_or_invalid_base maybe_or_invalid maybe_or_true)
qed auto
lemma fold_true_invalid_fold_rev_false_invalid:
"fold (\<or>?) l true = invalid \<Longrightarrow>
fold (\<or>?) (rev l) false = invalid"
proof(induct l)
case (Cons a l)
then show ?case
apply simp
by (metis maybe_or_true maybe_or_valid)
qed auto
lemma fold_maybe_or_rev:
"fold (\<or>?) l b = fold (\<or>?) (rev l) b"
proof(induct l)
case (Cons a l)
then show ?case
proof(induction a b rule: plus_trilean.induct)
case (1 uu)
then show ?case
by (simp add: fold_maybe_or_invalid_base)
next
case "2_1"
then show ?case
by (simp add: fold_maybe_or_invalid_base)
next
case "2_2"
then show ?case
by (simp add: fold_maybe_or_invalid_base)
next
case "3_1"
then show ?case
apply simp
by (metis add.assoc fold_maybe_or_true_base_never_false maybe_not.cases maybe_or_idempotent maybe_or_true)
next
case "3_2"
then show ?case
apply simp
apply (case_tac "fold (\<or>?) l true")
apply (simp add: eq_commute[of true])
apply (case_tac "fold (\<or>?) (rev l) false")
apply simp
apply simp
apply (simp add: fold_true_fold_false_not_invalid)
apply (simp add: fold_maybe_or_true_base_never_false)
by (simp add: fold_true_invalid_fold_rev_false_invalid)
next
case 4
then show ?case
by (simp add: maybe_or_zero)
next
case 5
then show ?case
by (simp add: maybe_or_zero)
qed
qed auto
lemma fold_maybe_or_cons:
"fold (\<or>?) (a#l) b = a \<or>? (fold (\<or>?) l b)"
by (metis fold_maybe_or_rev foldr.simps(2) foldr_conv_fold o_apply)
lemma gval_fold_gOr_map:
"gval (fold gOr l (Bc False)) s = fold (\<or>?) (map (\<lambda>g. gval g s) l) (false)"
proof(induct l)
case (Cons a l)
then show ?case
by (metis fold_maybe_or_cons gval_fold_gOr list.simps(9))
qed auto
lemma gval_unfold_first:
"gval (fold gOr (map (\<lambda>x. Eq (V v) (L x)) ls) (Eq (V v) (L l))) s =
gval (fold gOr (map (\<lambda>x. Eq (V v) (L x)) (l#ls)) (Bc False)) s"
proof(induct ls)
case Nil
then show ?case
apply (cases "s v")
apply simp
by (simp add: gOr_def)
next
case (Cons a ls)
then show ?case
proof -
have "gval (fold gOr (map (\<lambda>va. Eq (V v) (L va)) ls) (gOr (Eq (V v) (L l)) (Bc False))) s = gval (fold gOr (map (\<lambda>va. Eq (V v) (L va)) (l # ls)) (Bc False)) s"
by simp
then have "gval (fold gOr (map (\<lambda>va. Eq (V v) (L va)) (a # ls)) (Eq (V v) (L l))) s = gval (fold gOr (Eq (V v) (L a) # map (\<lambda>va. Eq (V v) (L va)) ls) (gOr (Eq (V v) (L l)) (Bc False))) s"
by (metis (no_types) Cons.hyps gval_fold_gOr list.simps(9))
then show ?thesis
by force
qed
qed
lemma fold_Eq_true:
"\<forall>v. fold (\<or>?) (map (\<lambda>x. if v = x then true else false) vs) true = true"
by(induct vs, auto)
lemma x_in_set_fold_eq:
"x \<in> set ll \<Longrightarrow>
fold (\<or>?) (map (\<lambda>xa. if x = xa then true else false) ll) false = true"
proof(induct ll)
case (Cons a ll)
then show ?case
apply simp
apply standard
apply (simp add: fold_Eq_true)
by auto
qed auto
lemma x_not_in_set_fold_eq:
"s v \<notin> Some ` set ll \<Longrightarrow>
false = fold (\<or>?) (map (\<lambda>x. if s v = Some x then true else false) ll) false"
by(induct ll, auto)
lemma gval_take: "max_input g < Some a \<Longrightarrow>
gval g (join_ir i r) = gval g (join_ir (take a i) r)"
proof(induct g)
case (Bc x)
then show ?case
by (metis (full_types) gval.simps(1) gval.simps(2))
next
case (Eq x1a x2)
then show ?case
by (metis aval_take gval.simps(4) max_input_Eq max_less_iff_conj)
next
case (Gt x1a x2)
then show ?case
by (metis aval_take gval.simps(3) max_input_Gt max_less_iff_conj)
next
case (Nor g1 g2)
then show ?case
by (simp add: maybe_not_eq gexp_max_input_Nor)
next
case (In v l)
then show ?case
apply (simp add: gexp_max_input_In)
using aval_take by fastforce
qed
lemma gval_fold_gAnd_append_singleton:
"gval (fold gAnd (a @ [G]) (Bc True)) s = gval (fold gAnd a (Bc True)) s \<and>? gval G s"
apply simp
using times_trilean_commutative by blast
lemma gval_fold_rev_true:
"gval (fold gAnd (rev G) (Bc True)) s = true \<Longrightarrow>
gval (fold gAnd G (Bc True)) s = true"
by (metis foldr_conv_fold gval_foldr_true rev_rev_ident set_rev)
lemma gval_fold_not_invalid_all_valid_contra:
"\<exists>g \<in> set G. gval g s = invalid \<Longrightarrow>
gval (fold gAnd G (Bc True)) s = invalid"
proof(induct G rule: rev_induct)
case (snoc a G)
then show ?case
apply (simp only: gval_fold_gAnd_append_singleton)
apply simp
using maybe_and_valid by blast
qed auto
lemma gval_fold_not_invalid_all_valid:
"gval (fold gAnd G (Bc True)) s \<noteq> invalid \<Longrightarrow>
\<forall>g \<in> set G. gval g s \<noteq> invalid"
using gval_fold_not_invalid_all_valid_contra by blast
lemma all_gval_not_false:
"(\<forall>g \<in> set G. gval g s \<noteq> false) = (\<forall>g \<in> set G. gval g s = true) \<or> (\<exists>g \<in> set G. gval g s = invalid)"
using trilean.exhaust by auto
lemma must_have_one_false_contra:
"\<forall>g \<in> set G. gval g s \<noteq> false \<Longrightarrow>
gval (fold gAnd G (Bc True)) s \<noteq> false"
using all_gval_not_false[of G s]
apply simp
apply (case_tac "(\<forall>g\<in>set G. gval g s = true)")
apply (metis (full_types) foldr_conv_fold gval_fold_rev_true gval_foldr_true not_true)
by (simp add: gval_fold_not_invalid_all_valid_contra)
lemma must_have_one_false:
"gval (fold gAnd G (Bc True)) s = false \<Longrightarrow>
\<exists>g \<in> set G. gval g s = false"
using must_have_one_false_contra by blast
lemma all_valid_fold:
"\<forall>g \<in> set G. gval g s \<noteq> invalid \<Longrightarrow>
gval (fold gAnd G (Bc True)) s \<noteq> invalid"
apply (induct G rule: rev_induct)
apply simp
by (simp add: maybe_and_invalid)
lemma one_false_all_valid_false:
"\<exists>g\<in>set G. gval g s = false \<Longrightarrow>
\<forall>g\<in>set G. gval g s \<noteq> invalid \<Longrightarrow>
gval (fold gAnd G (Bc True)) s = false"
by (metis (full_types) all_valid_fold foldr_conv_fold gval_foldr_true not_true rev_rev_ident set_rev)
lemma gval_fold_rev_false:
"gval (fold gAnd (rev G) (Bc True)) s = false \<Longrightarrow>
gval (fold gAnd G (Bc True)) s = false"
using must_have_one_false[of "rev G" s]
gval_fold_not_invalid_all_valid[of "rev G" s]
by (simp add: one_false_all_valid_false)
lemma fold_invalid_means_one_invalid:
"gval (fold gAnd G (Bc True)) s = invalid \<Longrightarrow>
\<exists>g \<in> set G. gval g s = invalid"
using all_valid_fold by blast
lemma gval_fold_rev_invalid:
"gval (fold gAnd (rev G) (Bc True)) s = invalid \<Longrightarrow>
gval (fold gAnd G (Bc True)) s = invalid"
using fold_invalid_means_one_invalid[of "rev G" s]
by (simp add: gval_fold_not_invalid_all_valid_contra)
lemma gval_fold_rev_equiv_fold:
"gval (fold gAnd (rev G) (Bc True)) s = gval (fold gAnd G (Bc True)) s"
apply (cases "gval (fold gAnd (rev G) (Bc True)) s")
apply (simp add: gval_fold_rev_true)
apply (simp add: gval_fold_rev_false)
by (simp add: gval_fold_rev_invalid)
lemma gval_fold_equiv_fold_rev:
"gval (fold gAnd G (Bc True)) s = gval (fold gAnd (rev G) (Bc True)) s"
by (simp add: gval_fold_rev_equiv_fold)
lemma gval_fold_equiv_gval_foldr:
"gval (fold gAnd G (Bc True)) s = gval (foldr gAnd G (Bc True)) s"
proof -
have "gval (fold gAnd G (Bc True)) s = gval (fold gAnd (rev G) (Bc True)) s"
using gval_fold_equiv_fold_rev by force
then show ?thesis
by (simp add: foldr_conv_fold)
qed
lemma gval_foldr_equiv_gval_fold:
"gval (foldr gAnd G (Bc True)) s = gval (fold gAnd G (Bc True)) s"
by (simp add: gval_fold_equiv_gval_foldr)
lemma gval_fold_cons:
"gval (fold gAnd (g # gs) (Bc True)) s = gval g s \<and>? gval (fold gAnd gs (Bc True)) s"
apply (simp only: apply_guards_fold gval_fold_equiv_gval_foldr)
by (simp only: foldr.simps comp_def gval_gAnd)
lemma gval_fold_take: "max_input_list G < Some a \<Longrightarrow>
a \<le> length i \<Longrightarrow>
max_input_list G \<le> Some (length i) \<Longrightarrow>
gval (fold gAnd G (Bc True)) (join_ir i r) = gval (fold gAnd G (Bc True)) (join_ir (take a i) r)"
proof(induct G)
case (Cons g gs)
then show ?case
apply (simp only: gval_fold_cons)
apply (simp add: max_input_list_cons)
using gval_take[of g a i r]
by simp
qed auto
primrec padding :: "nat \<Rightarrow> 'a list" where
"padding 0 = []" |
"padding (Suc m) = (Eps (\<lambda>x. True))#(padding m)"
definition take_or_pad :: "'a list \<Rightarrow> nat \<Rightarrow> 'a list" where
"take_or_pad a n = (if length a \<ge> n then take n a else a@(padding (n-length a)))"
lemma length_padding: "length (padding n) = n"
by (induct n, auto)
lemma length_take_or_pad: "length (take_or_pad a n) = n"
proof(induct n)
case 0
then show ?case
by (simp add: take_or_pad_def)
next
case (Suc n)
then show ?case
apply (simp add: take_or_pad_def)
apply standard
apply auto[1]
by (simp add: length_padding)
qed
fun enumerate_gexp_strings :: "'a gexp \<Rightarrow> String.literal set" where
"enumerate_gexp_strings (Bc _) = {}" |
"enumerate_gexp_strings (Eq a1 a2) = enumerate_aexp_strings a1 \<union> enumerate_aexp_strings a2" |
"enumerate_gexp_strings (Gt a1 a2) = enumerate_aexp_strings a1 \<union> enumerate_aexp_strings a2" |
"enumerate_gexp_strings (In v l) = fold (\<lambda>x acc. case x of Num n \<Rightarrow> acc | Str s \<Rightarrow> insert s acc) l {}" |
"enumerate_gexp_strings (Nor g1 g2) = enumerate_gexp_strings g1 \<union> enumerate_gexp_strings g2"
fun enumerate_gexp_ints :: "'a gexp \<Rightarrow> int set" where
"enumerate_gexp_ints (Bc _) = {}" |
"enumerate_gexp_ints (Eq a1 a2) = enumerate_aexp_ints a1 \<union> enumerate_aexp_ints a2" |
"enumerate_gexp_ints (Gt a1 a2) = enumerate_aexp_ints a1 \<union> enumerate_aexp_ints a2" |
"enumerate_gexp_ints (In v l) = fold (\<lambda>x acc. case x of Str s \<Rightarrow> acc | Num n \<Rightarrow> insert n acc) l {}" |
"enumerate_gexp_ints (Nor g1 g2) = enumerate_gexp_ints g1 \<union> enumerate_gexp_ints g2"
definition restricted_once :: "'a \<Rightarrow> 'a gexp list \<Rightarrow> bool" where
"restricted_once v G = (length (filter (\<lambda>g. gexp_constrains g (V v)) G) = 1)"
definition not_restricted :: "'a \<Rightarrow> 'a gexp list \<Rightarrow> bool" where
"not_restricted v G = (length (filter (\<lambda>g. gexp_constrains g (V v)) G) = 0)"
lemma restricted_once_cons:
"restricted_once v (g#gs) = ((gexp_constrains g (V v) \<and> not_restricted v gs) \<or> ((\<not> gexp_constrains g (V v)) \<and> restricted_once v gs))"
by (simp add: restricted_once_def not_restricted_def)
lemma not_restricted_cons:
"not_restricted v (g#gs) = ((\<not> gexp_constrains g (V v)) \<and> not_restricted v gs)"
by (simp add: not_restricted_def)
definition enumerate_vars :: "vname gexp \<Rightarrow> vname list" where
"enumerate_vars g = sorted_list_of_set ((image R (enumerate_regs g)) \<union> (image I (enumerate_gexp_inputs g)))"
fun rename_regs :: "(nat \<Rightarrow> nat) \<Rightarrow> vname gexp \<Rightarrow> vname gexp" where
"rename_regs _ (Bc b) = Bc b" |
"rename_regs f (Eq a1 a2) = Eq (AExp.rename_regs f a1) (AExp.rename_regs f a2)" |
"rename_regs f (Gt a1 a2) = Gt (AExp.rename_regs f a1) (AExp.rename_regs f a2)" |
"rename_regs f (In (R r) vs) = In (R (f r)) vs" |
"rename_regs f (In v vs) = In v vs" |
"rename_regs f (Nor g1 g2) = Nor (rename_regs f g1) (rename_regs f g2)"
definition eq_upto_rename :: "vname gexp \<Rightarrow> vname gexp \<Rightarrow> bool" where
"eq_upto_rename g1 g2 = (\<exists>f. bij f \<and> rename_regs f g1 = g2)"
lemma gval_reg_some_superset:
"\<forall>a. (r $ a \<noteq> None) \<longrightarrow> r $ a = r' $ a \<Longrightarrow>
x \<noteq> invalid \<Longrightarrow>
gval a (join_ir i r) = x \<Longrightarrow>
gval a (join_ir i r') = x"
proof(induct a arbitrary: x)
case (Bc b)
then show ?case by (cases b, auto)
next
case (Eq x1a x2)
then show ?case
apply (cases x)
apply simp
using value_eq_true[of "aval x1a (join_ir i r)" "aval x2 (join_ir i r)"]
apply clarsimp
apply (simp add: aval_reg_some_superset)
apply simp
using value_eq_false[of "aval x1a (join_ir i r)" "aval x2 (join_ir i r)"]
apply clarsimp
apply (simp add: aval_reg_some_superset)
by simp
next
case (Gt x1a x2)
then show ?case
apply (cases x)
apply simp
using value_gt_true_Some[of "aval x1a (join_ir i r)" "aval x2 (join_ir i r)"]
apply clarsimp
apply (simp add: aval_reg_some_superset)
apply simp
using value_gt_false_Some[of "aval x1a (join_ir i r)" "aval x2 (join_ir i r)"]
apply clarsimp
apply (simp add: aval_reg_some_superset)
by simp
next
case (In x1a x2)
then show ?case
apply simp
apply (case_tac "join_ir i r x1a")
apply simp
apply (case_tac "join_ir i r' x1a")
apply simp
apply (metis aval_reg_some_superset In.prems(1) aval.simps(2) option.distinct(1))
apply simp
by (metis (full_types) aval_reg_some_superset In.prems(1) aval.simps(2) option.inject)
next
case (Nor a1 a2)
then show ?case
apply simp
apply (cases x)
apply (simp add: maybe_negate_true maybe_or_false)
apply (simp add: maybe_negate_false maybe_or_true)
apply presburger
by simp
qed
lemma apply_guards_reg_some_superset:
"\<forall>a. (r $ a \<noteq> None) \<longrightarrow> r $ a = r' $ a \<Longrightarrow>
apply_guards G (join_ir i r) \<Longrightarrow>
apply_guards G (join_ir i r')"
apply (induct G)
apply simp
apply (simp add: apply_guards_cons)
using gval_reg_some_superset
by simp
end

View File

@ -0,0 +1,253 @@
subsection\<open>GExp Lexorder\<close>
text\<open>This theory defines a lexicographical ordering on guard expressions such that we can build
orderings for transitions. We make use of the previously established orderings on arithmetic
expressions.\<close>
theory
GExp_Lexorder
imports
"GExp"
"AExp_Lexorder"
"HOL-Library.List_Lexorder"
begin
fun height :: "'a gexp \<Rightarrow> nat" where
"height (Bc _) = 1" |
"height (Eq a1 a2) = 1 + max (AExp_Lexorder.height a1) (AExp_Lexorder.height a2)" |
"height (Gt a1 a2) = 1 + max (AExp_Lexorder.height a1) (AExp_Lexorder.height a2)" |
"height (In v l) = 2 + size l" |
"height (Nor g1 g2) = 1 + max (height g1) (height g2)"
instantiation gexp :: (linorder) linorder begin
fun less_gexp_aux :: "'a gexp \<Rightarrow> 'a gexp \<Rightarrow> bool" where
"less_gexp_aux (Bc b1) (Bc b2) = (b1 < b2)" |
"less_gexp_aux (Bc b1) _ = True" |
"less_gexp_aux (Eq e1 e2) (Bc b2) = False" |
"less_gexp_aux (Eq e1 e2) (Eq e1' e2') = ((e1 < e1') \<or> ((e1 = e1') \<and> (e2 < e2')))" |
"less_gexp_aux (Eq e1 e2) _ = True" |
"less_gexp_aux (Gt e1 e2) (Bc b2) = False" |
"less_gexp_aux (Gt e1 e2) (Eq e1' e2') = False" |
"less_gexp_aux (Gt e1 e2) (Gt e1' e2') = ((e1 < e1') \<or> ((e1 = e1') \<and> (e2 < e2')))" |
"less_gexp_aux (Gt e1 e2) _ = True" |
"less_gexp_aux (In vb vc) (Nor v va) = True" |
"less_gexp_aux (In vb vc) (In v va) = (vb < v \<or> (vb = v \<and> vc < va))" |
"less_gexp_aux (In vb vc) _ = False" |
"less_gexp_aux (Nor g1 g2) (Nor g1' g2') = ((less_gexp_aux g1 g1') \<or> ((g1 = g1') \<and> (less_gexp_aux g2 g2')))" |
"less_gexp_aux (Nor g1 g2) _ = False"
definition less_gexp :: "'a gexp \<Rightarrow> 'a gexp \<Rightarrow> bool" where
"less_gexp a1 a2 = (
let
h1 = height a1;
h2 = height a2
in
if h1 = h2 then
less_gexp_aux a1 a2
else
h1 < h2
)"
declare less_gexp_def [simp]
definition less_eq_gexp :: "'a gexp \<Rightarrow> 'a gexp \<Rightarrow> bool" where
"less_eq_gexp e1 e2 \<equiv> (e1 < e2) \<or> (e1 = e2)"
lemma less_gexp_aux_antisym: "less_gexp_aux x y = (\<not>(less_gexp_aux y x) \<and> (x \<noteq> y))"
proof (induct x y rule: less_gexp_aux.induct)
case (1 b1 b2)
then show ?case by auto
next
case ("2_1" b1 v va)
then show ?case by auto
next
case ("2_2" b1 v va)
then show ?case by auto
next
case ("2_3" b1 v va)
then show ?case by auto
next
case ("2_4" b1 v va)
then show ?case by auto
next
case (3 e1 e2 b2)
then show ?case by auto
next
case (4 e1 e2 e1' e2')
then show ?case
by (metis less_gexp_aux.simps(7) less_imp_not_less less_linear)
next
case ("5_1" e1 e2 v va)
then show ?case by auto
next
case ("5_2" e1 e2 v va)
then show ?case by auto
next
case ("5_3" e1 e2 v va)
then show ?case by auto
next
case (6 e1 e2 b2)
then show ?case by auto
next
case (7 e1 e2 e1' e2')
then show ?case by auto
next
case (8 e1 e2 e1' e2')
then show ?case
by (metis less_gexp_aux.simps(13) less_imp_not_less less_linear)
next
case ("9_1" e1 e2 v va)
then show ?case by auto
next
case ("9_2" e1 e2 v va)
then show ?case by auto
next
case (10 vb vc v va)
then show ?case by auto
next
case (11 vb vc v va)
then show ?case by auto
next
case ("12_1" vb vc v)
then show ?case by auto
next
case ("12_2" vb vc v va)
then show ?case by auto
next
case ("12_3" vb vc v va)
then show ?case by auto
next
case (13 g1 g2 g1' g2')
then show ?case by auto
next
case ("14_1" g1 g2 v)
then show ?case by auto
next
case ("14_2" g1 g2 v va)
then show ?case by auto
next
case ("14_3" g1 g2 v va)
then show ?case by auto
next
case ("14_4" g1 g2 v va)
then show ?case by auto
qed
lemma less_gexp_antisym: "(x::'a gexp) < y = (\<not>(y < x) \<and> (x \<noteq> y))"
apply (simp add: Let_def)
apply standard
using less_gexp_aux_antisym apply blast
apply clarsimp
by (induct x, auto)
lemma less_gexp_aux_trans: "less_gexp_aux x y \<Longrightarrow> less_gexp_aux y z \<Longrightarrow> less_gexp_aux x z"
proof(induct x y arbitrary: z rule: less_gexp_aux.induct)
case (1 b1 b2)
then show ?case by (cases z, auto)
next
case ("2_1" b1 v va)
then show ?case by (cases z, auto)
next
case ("2_2" b1 v va)
then show ?case by (cases z, auto)
next
case ("2_3" b1 v va)
then show ?case by (cases z, auto)
next
case ("2_4" b1 v va)
then show ?case by (cases z, auto)
next
case (3 e1 e2 b2)
then show ?case by (cases z, auto)
next
case (4 e1 e2 e1' e2')
then show ?case
apply (cases z)
apply simp
apply (metis dual_order.strict_trans less_gexp_aux.simps(7))
by auto
next
case ("5_1" e1 e2 v va)
then show ?case by (cases z, auto)
next
case ("5_2" e1 e2 v va)
then show ?case by (cases z, auto)
next
case ("5_3" e1 e2 v va)
then show ?case by (cases z, auto)
next
case (6 e1 e2 b2)
then show ?case by (cases z, auto)
next
case (7 e1 e2 e1' e2')
then show ?case by (cases z, auto)
next
case (8 e1 e2 e1' e2')
then show ?case
apply (cases z)
apply simp
apply simp
apply (metis dual_order.strict_trans less_gexp_aux.simps(13))
by auto
next
case ("9_1" e1 e2 v va)
then show ?case by (cases z, auto)
next
case ("9_2" e1 e2 v va)
then show ?case by (cases z, auto)
next
case (10 vb vc v va)
then show ?case by (cases z, auto)
next
case (11 vb vc v va)
then show ?case by (cases z, auto)
next
case ("12_1" vb vc v)
then show ?case by (cases z, auto)
next
case ("12_2" vb vc v va)
then show ?case by (cases z, auto)
next
case ("12_3" vb vc v va)
then show ?case by (cases z, auto)
next
case (13 g1 g2 g1' g2')
then show ?case by (cases z, auto)
next
case ("14_1" g1 g2 v)
then show ?case by (cases z, auto)
next
case ("14_2" g1 g2 v va)
then show ?case by (cases z, auto)
next
case ("14_3" g1 g2 v va)
then show ?case by (cases z, auto)
next
case ("14_4" g1 g2 v va)
then show ?case by (cases z, auto)
qed
lemma less_gexp_trans: "(x::'a gexp) < y \<Longrightarrow> y < z \<Longrightarrow> x < z"
apply (simp add: Let_def)
by (metis (no_types, lifting) dual_order.strict_trans less_gexp_aux_trans less_imp_not_less)
instance proof
fix x y z :: "'a gexp"
show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)"
by (metis less_gexp_antisym less_eq_gexp_def)
show "(x \<le> x)"
by (simp add: less_eq_gexp_def)
show "x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z"
by (metis less_eq_gexp_def less_gexp_trans)
show "x \<le> y \<Longrightarrow> y \<le> x \<Longrightarrow> x = y"
unfolding less_eq_gexp_def using less_gexp_antisym by blast
show "x \<le> y \<or> y \<le> x"
unfolding less_eq_gexp_def using less_gexp_antisym by blast
qed
end
end

View File

@ -0,0 +1,2 @@
# efsm-isabelle
Implementation of EFSMs in Isabelle/HOL

View File

@ -0,0 +1,30 @@
chapter AFP
session Extended_Finite_State_Machines (AFP) = "HOL-Library" +
options [timeout = 600]
sessions
FinFun
directories
"examples"
theories
"Trilean"
"Value"
"VName"
"AExp"
"AExp_Lexorder"
"GExp"
"GExp_Lexorder"
"FSet_Utils"
"Transition"
"Transition_Lexorder"
"EFSM"
"EFSM_LTL"
"examples/Drinks_Machine"
"examples/Drinks_Machine_2"
"examples/Drinks_Machine_LTL"
document_files
"root.tex"
"root.bib"

View File

@ -0,0 +1,188 @@
chapter\<open>Models\<close>
text\<open>In this chapter, we present our formalisation of EFSMs from \cite{foster2018}. We first
define transitions, before defining EFSMs as finite sets of transitions between states. Finally,
we provide a framework of function definitions and key lemmas such that LTL properties over EFSMs
can be more easily specified and proven.\<close>
section\<open>Transitions\<close>
text\<open>Here we define the transitions which make up EFSMs. As per \cite{foster2018}, each transition
has a label and an arity and, optionally, guards, outputs, and updates. To implement this, we use
the record type such that each component of the transition can be accessed.\<close>
theory Transition
imports GExp
begin
type_synonym label = String.literal
type_synonym arity = nat
type_synonym guard = "vname gexp"
type_synonym inputs = "value list"
type_synonym outputs = "value option list"
type_synonym output_function = "vname aexp"
type_synonym update_function = "nat \<times> vname aexp"
text_raw\<open>\snip{transitiontype}{1}{2}{%\<close>
record transition =
Label :: String.literal
Arity :: nat
Guards :: "guard list"
Outputs :: "output_function list"
Updates :: "update_function list"
text_raw\<open>}%endsnip\<close>
definition same_structure :: "transition \<Rightarrow> transition \<Rightarrow> bool" where
"same_structure t1 t2 = (
Label t1 = Label t2 \<and>
Arity t1 = Arity t2 \<and>
length (Outputs t1) = length (Outputs t2)
)"
definition enumerate_inputs :: "transition \<Rightarrow> nat set" where
"enumerate_inputs t = (\<Union> (set (map enumerate_gexp_inputs (Guards t)))) \<union>
(\<Union> (set (map enumerate_aexp_inputs (Outputs t)))) \<union>
(\<Union> (set (map (\<lambda>(_, u). enumerate_aexp_inputs u) (Updates t))))"
definition max_input :: "transition \<Rightarrow> nat option" where
"max_input t = (if enumerate_inputs t = {} then None else Some (Max (enumerate_inputs t)))"
definition total_max_input :: "transition \<Rightarrow> nat" where
"total_max_input t = (case max_input t of None \<Rightarrow> 0 | Some a \<Rightarrow> a)"
definition enumerate_regs :: "transition \<Rightarrow> nat set" where
"enumerate_regs t = (\<Union> (set (map GExp.enumerate_regs (Guards t)))) \<union>
(\<Union> (set (map AExp.enumerate_regs (Outputs t)))) \<union>
(\<Union> (set (map (\<lambda>(_, u). AExp.enumerate_regs u) (Updates t)))) \<union>
(\<Union> (set (map (\<lambda>(r, _). AExp.enumerate_regs (V (R r))) (Updates t))))"
definition max_reg :: "transition \<Rightarrow> nat option" where
"max_reg t = (if enumerate_regs t = {} then None else Some (Max (enumerate_regs t)))"
definition total_max_reg :: "transition \<Rightarrow> nat" where
"total_max_reg t = (case max_reg t of None \<Rightarrow> 0 | Some a \<Rightarrow> a)"
definition enumerate_ints :: "transition \<Rightarrow> int set" where
"enumerate_ints t = (\<Union> (set (map enumerate_gexp_ints (Guards t)))) \<union>
(\<Union> (set (map enumerate_aexp_ints (Outputs t)))) \<union>
(\<Union> (set (map (\<lambda>(_, u). enumerate_aexp_ints u) (Updates t)))) \<union>
(\<Union> (set (map (\<lambda>(r, _). enumerate_aexp_ints (V (R r))) (Updates t))))"
definition valid_transition :: "transition \<Rightarrow> bool" where
"valid_transition t = (\<forall>i \<in> enumerate_inputs t. i < Arity t)"
definition can_take :: "nat \<Rightarrow> vname gexp list \<Rightarrow> inputs \<Rightarrow> registers \<Rightarrow> bool" where
"can_take a g i r = (length i = a \<and> apply_guards g (join_ir i r))"
lemma can_take_empty [simp]: "length i = a \<Longrightarrow> can_take a [] i c"
by (simp add: can_take_def)
lemma can_take_subset_append:
assumes "set (Guards t) \<subseteq> set (Guards t')"
shows "can_take a (Guards t @ Guards t') i c = can_take a (Guards t') i c"
using assms
by (simp add: apply_guards_subset_append can_take_def)
definition "can_take_transition t i r = can_take (Arity t) (Guards t) i r"
lemmas can_take = can_take_def can_take_transition_def
lemma can_take_transition_empty_guard:
"Guards t = [] \<Longrightarrow> \<exists>i. can_take_transition t i c"
by (simp add: can_take_transition_def can_take_def Ex_list_of_length)
lemma can_take_subset: "length i = Arity t \<Longrightarrow>
Arity t = Arity t' \<Longrightarrow>
set (Guards t') \<subseteq> set (Guards t) \<Longrightarrow>
can_take_transition t i r \<Longrightarrow>
can_take_transition t' i r"
by (simp add: can_take_transition_def can_take_def apply_guards_subset)
lemma valid_list_can_take:
"\<forall>g \<in> set (Guards t). valid g \<Longrightarrow> \<exists>i. can_take_transition t i c"
by (simp add: can_take_transition_def can_take_def apply_guards_def valid_def Ex_list_of_length)
lemma cant_take_if:
"\<exists>g \<in> set (Guards t). gval g (join_ir i r) \<noteq> true \<Longrightarrow>
\<not> can_take_transition t i r"
using apply_guards_cons apply_guards_rearrange can_take_def can_take_transition_def by blast
definition apply_outputs :: "'a aexp list \<Rightarrow> 'a datastate \<Rightarrow> value option list" where
"apply_outputs p s = map (\<lambda>p. aval p s) p"
abbreviation "evaluate_outputs t i r \<equiv> apply_outputs (Outputs t) (join_ir i r)"
lemma apply_outputs_nth:
"i < length p \<Longrightarrow> apply_outputs p s ! i = aval (p ! i) s"
by (simp add: apply_outputs_def)
lemmas apply_outputs = datastate apply_outputs_def value_plus_def value_minus_def value_times_def
lemma apply_outputs_empty [simp]: "apply_outputs [] s = []"
by (simp add: apply_outputs_def)
lemma apply_outputs_preserves_length: "length (apply_outputs p s) = length p"
by (simp add: apply_outputs_def)
lemma apply_outputs_literal: assumes "P ! r = L v"
and "r < length P"
shows "apply_outputs P s ! r = Some v"
by (simp add: assms apply_outputs_nth)
lemma apply_outputs_register: assumes "r < length P"
shows "apply_outputs (list_update P r (V (R p))) (join_ir i c) ! r = c $ p"
by (metis apply_outputs_nth assms aval.simps(2) join_ir_R length_list_update nth_list_update_eq)
lemma apply_outputs_unupdated: assumes "ia \<noteq> r"
and "ia < length P"
shows "apply_outputs P j ! ia = apply_outputs (list_update P r v)j ! ia"
by (metis apply_outputs_nth assms(1) assms(2) length_list_update nth_list_update_neq)
definition apply_updates :: "update_function list \<Rightarrow> vname datastate \<Rightarrow> registers \<Rightarrow> registers" where
"apply_updates u old = fold (\<lambda>h r. r(fst h $:= aval (snd h) old)) u"
abbreviation "evaluate_updates t i r \<equiv> apply_updates (Updates t) (join_ir i r) r"
lemma apply_updates_cons: "ra \<noteq> r \<Longrightarrow>
apply_updates u (join_ir ia c) c $ ra = apply_updates ((r, a) # u) (join_ir ia c) c $ ra"
proof(induct u rule: rev_induct)
case Nil
then show ?case
by (simp add: apply_updates_def)
next
case (snoc u us)
then show ?case
apply (cases u)
apply (simp add: apply_updates_def)
by (case_tac "ra = aa", auto)
qed
lemma update_twice:
"apply_updates [(r, a), (r, b)] s regs = regs (r $:= aval b s)"
by (simp add: apply_updates_def)
lemma r_not_updated_stays_the_same:
"r \<notin> fst ` set U \<Longrightarrow> apply_updates U c d $ r = d $ r"
using apply_updates_def
by (induct U rule: rev_induct, auto)
definition rename_regs :: "(nat \<Rightarrow> nat) \<Rightarrow> transition \<Rightarrow> transition" where
"rename_regs f t = t\<lparr>
Guards := map (GExp.rename_regs f) (Guards t),
Outputs := map (AExp.rename_regs f) (Outputs t),
Updates := map (\<lambda>(r, u). (f r, AExp.rename_regs f u)) (Updates t)
\<rparr>"
definition eq_upto_rename_strong :: "transition \<Rightarrow> transition \<Rightarrow> bool" where
"eq_upto_rename_strong t1 t2 = (\<exists>f. bij f \<and> rename_regs f t1 = t2)"
inductive eq_upto_rename :: "transition \<Rightarrow> transition \<Rightarrow> bool" where
"Label t1 = Label t2 \<Longrightarrow>
Arity t2 = Arity t2 \<Longrightarrow>
apply_guards (map (GExp.rename_regs f) (Guards t1)) = apply_guards (Guards t2) \<Longrightarrow>
apply_outputs (map (AExp.rename_regs f) (Outputs t1)) = apply_outputs (Outputs t2) \<Longrightarrow>
apply_updates (map (\<lambda>(r, u). (f r, AExp.rename_regs f u)) (Updates t1)) = apply_updates (Updates t2) \<Longrightarrow>
eq_upto_rename t1 t2"
end

View File

@ -0,0 +1,30 @@
subsection\<open>Transition Lexorder\<close>
text\<open>This theory defines a lexicographical ordering on transitions such that we can convert from
the set representation of EFSMs to a sorted list that we can recurse over.\<close>
theory Transition_Lexorder
imports "Transition"
GExp_Lexorder
"HOL-Library.Product_Lexorder"
begin
instantiation "transition_ext" :: (linorder) linorder begin
definition less_transition_ext :: "'a::linorder transition_scheme \<Rightarrow> 'a transition_scheme \<Rightarrow> bool" where
"less_transition_ext t1 t2 = ((Label t1, Arity t1, Guards t1, Outputs t1, Updates t1, more t1) < (Label t2, Arity t2, Guards t2, Outputs t2, Updates t2, more t2))"
definition less_eq_transition_ext :: "'a::linorder transition_scheme \<Rightarrow> 'a transition_scheme \<Rightarrow> bool" where
"less_eq_transition_ext t1 t2 = (t1 < t2 \<or> t1 = t2)"
instance
apply standard
unfolding less_eq_transition_ext_def less_transition_ext_def
apply auto[1]
apply simp
using less_trans apply blast
using less_imp_not_less apply blast
by (metis Pair_inject equality neqE)
end
end

View File

@ -0,0 +1,281 @@
chapter\<open>Preliminaries\<close>
text\<open>In this chapter, we introduce the preliminaries, including a three-valued logic, variables,
arithmetic expressions and guard expressions.\<close>
section\<open>Three-Valued Logic\<close>
text\<open>Because our EFSMs are dynamically typed, we cannot rely on conventional Boolean logic when
evaluating expressions. For example, we may end up in the situation where we need to evaluate
the guard $r_1 > 5$. This is fine if $r_1$ holds a numeric value, but if $r_1$ evaluates to a
string, this causes problems. We cannot simply evaluate to \emph{false} because then the negation
would evaluate to \emph{true.} Instead, we need a three-valued logic such that we can meaningfully
evaluate nonsensical guards.
The \texttt{trilean} datatype is used to implement three-valued Bochvar logic
\cite{bochvar1981}. Here we prove that the logic is an idempotent semiring, define a partial order,
and prove some other useful lemmas.\<close>
theory Trilean
imports Main
begin
datatype trilean = true | false | invalid
instantiation trilean :: semiring begin
fun times_trilean :: "trilean \<Rightarrow> trilean \<Rightarrow> trilean" where
"times_trilean _ invalid = invalid" |
"times_trilean invalid _ = invalid" |
"times_trilean true true = true" |
"times_trilean _ false = false" |
"times_trilean false _ = false"
fun plus_trilean :: "trilean \<Rightarrow> trilean \<Rightarrow> trilean" where
"plus_trilean invalid _ = invalid" |
"plus_trilean _ invalid = invalid" |
"plus_trilean true _ = true" |
"plus_trilean _ true = true" |
"plus_trilean false false = false"
abbreviation maybe_and :: "trilean \<Rightarrow> trilean \<Rightarrow> trilean" (infixl "\<and>?" 70) where
"maybe_and x y \<equiv> x * y"
abbreviation maybe_or :: "trilean \<Rightarrow> trilean \<Rightarrow> trilean" (infixl "\<or>?" 65) where
"maybe_or x y \<equiv> x + y"
lemma plus_trilean_assoc:
"a \<or>? b \<or>? c = a \<or>? (b \<or>? c)"
proof(induct a b arbitrary: c rule: plus_trilean.induct)
case (1 uu)
then show ?case
by simp
next
case "2_1"
then show ?case
by simp
next
case "2_2"
then show ?case
by simp
next
case "3_1"
then show ?case
by (metis plus_trilean.simps(2) plus_trilean.simps(4) trilean.exhaust)
next
case "3_2"
then show ?case
by (metis plus_trilean.simps(3) plus_trilean.simps(5) plus_trilean.simps(6) plus_trilean.simps(7) trilean.exhaust)
next
case 4
then show ?case
by (metis plus_trilean.simps(2) plus_trilean.simps(3) plus_trilean.simps(4) plus_trilean.simps(5) plus_trilean.simps(6) trilean.exhaust)
next
case 5
then show ?case
by (metis plus_trilean.simps(6) plus_trilean.simps(7) trilean.exhaust)
qed
lemma plus_trilean_commutative: "a \<or>? b = b \<or>? a"
proof(induct a b rule: plus_trilean.induct)
case (1 uu)
then show ?case
by (metis plus_trilean.simps(1) plus_trilean.simps(2) plus_trilean.simps(3) trilean.exhaust)
qed auto
lemma times_trilean_commutative: "a \<and>? b = b \<and>? a"
by (metis (mono_tags) times_trilean.simps trilean.distinct(5) trilean.exhaust)
lemma times_trilean_assoc:
"a \<and>? b \<and>? c = a \<and>? (b \<and>? c)"
proof(induct a b arbitrary: c rule: plus_trilean.induct)
case (1 uu)
then show ?case
by (metis (mono_tags, lifting) times_trilean.simps(1) times_trilean_commutative)
next
case "2_1"
then show ?case
by (metis (mono_tags, lifting) times_trilean.simps(1) times_trilean_commutative)
next
case "2_2"
then show ?case
by (metis (mono_tags, lifting) times_trilean.simps(1) times_trilean_commutative)
next
case "3_1"
then show ?case
by (metis times_trilean.simps(1) times_trilean.simps(4) times_trilean.simps(5) trilean.exhaust)
next
case "3_2"
then show ?case
by (metis times_trilean.simps(1) times_trilean.simps(5) times_trilean.simps(6) times_trilean.simps(7) trilean.exhaust)
next
case 4
then show ?case
by (metis times_trilean.simps(1) times_trilean.simps(4) times_trilean.simps(5) times_trilean.simps(7) trilean.exhaust)
next
case 5
then show ?case
by (metis (full_types) times_trilean.simps(1) times_trilean.simps(6) times_trilean.simps(7) trilean.exhaust)
qed
lemma trilean_distributivity_1:
"(a \<or>? b) \<and>? c = a \<and>? c \<or>? b \<and>? c"
proof(induct a b rule: times_trilean.induct)
case (1 uu)
then show ?case
by (metis (mono_tags, lifting) plus_trilean.simps(1) plus_trilean_commutative times_trilean.simps(1) times_trilean_commutative)
next
case "2_1"
then show ?case
by (metis (mono_tags, lifting) plus_trilean.simps(1) times_trilean.simps(1) times_trilean_commutative)
next
case "2_2"
then show ?case
by (metis (mono_tags, lifting) plus_trilean.simps(1) times_trilean.simps(1) times_trilean_commutative)
next
case 3
then show ?case
apply simp
by (metis (no_types, hide_lams) plus_trilean.simps(1) plus_trilean.simps(4) plus_trilean.simps(7) times_trilean.simps(1) times_trilean.simps(4) times_trilean.simps(5) trilean.exhaust)
next
case "4_1"
then show ?case
apply simp
by (metis (no_types, hide_lams) plus_trilean.simps(1) plus_trilean.simps(5) plus_trilean.simps(7) times_trilean.simps(1) times_trilean.simps(4) times_trilean.simps(5) times_trilean.simps(6) times_trilean.simps(7) trilean.exhaust)
next
case "4_2"
then show ?case
apply simp
by (metis (no_types, hide_lams) plus_trilean.simps(1) plus_trilean.simps(7) times_trilean.simps(1) times_trilean.simps(6) times_trilean.simps(7) trilean.exhaust)
next
case 5
then show ?case
apply simp
by (metis (no_types, hide_lams) plus_trilean.simps(1) plus_trilean.simps(6) plus_trilean.simps(7) times_trilean.simps(1) times_trilean.simps(4) times_trilean.simps(5) times_trilean.simps(6) times_trilean.simps(7) trilean.exhaust)
qed
instance
apply standard
apply (simp add: plus_trilean_assoc)
apply (simp add: plus_trilean_commutative)
apply (simp add: times_trilean_assoc)
apply (simp add: trilean_distributivity_1)
using times_trilean_commutative trilean_distributivity_1 by auto
end
lemma maybe_or_idempotent: "a \<or>? a = a"
by (cases a) auto
lemma maybe_and_idempotent: "a \<and>? a = a"
by (cases a) auto
instantiation trilean :: ord begin
definition less_eq_trilean :: "trilean \<Rightarrow> trilean \<Rightarrow> bool" where
"less_eq_trilean a b = (a + b = b)"
definition less_trilean :: "trilean \<Rightarrow> trilean \<Rightarrow> bool" where
"less_trilean a b = (a \<le> b \<and> a \<noteq> b)"
declare less_trilean_def less_eq_trilean_def [simp]
instance
by standard
end
instantiation trilean :: uminus begin
fun maybe_not :: "trilean \<Rightarrow> trilean" ("\<not>? _" [60] 60) where
"\<not>? true = false" |
"\<not>? false = true" |
"\<not>? invalid = invalid"
instance
by standard
end
lemma maybe_and_one: "true \<and>? x = x"
by (cases x, auto)
lemma maybe_or_zero: "false \<or>? x = x"
by (cases x, auto)
lemma maybe_double_negation: "\<not>? \<not>? x = x"
by (cases x, auto)
lemma maybe_negate_true: "(\<not>? x = true) = (x = false)"
by (cases x, auto)
lemma maybe_negate_false: "(\<not>? x = false) = (x = true)"
by (cases x, auto)
lemma maybe_and_true: "(x \<and>? y = true) = (x = true \<and> y = true)"
using times_trilean.elims by blast
lemma maybe_and_not_true:
"(x \<and>? y \<noteq> true) = (x \<noteq> true \<or> y \<noteq> true)"
by (simp add: maybe_and_true)
lemma negate_valid: "(\<not>? x \<noteq> invalid) = (x \<noteq> invalid)"
by (metis maybe_double_negation maybe_not.simps(3))
lemma maybe_and_valid:
"x \<and>? y \<noteq> invalid \<Longrightarrow> x \<noteq> invalid \<and> y \<noteq> invalid"
using times_trilean.elims by blast
lemma maybe_or_valid:
"x \<or>? y \<noteq> invalid \<Longrightarrow> x \<noteq> invalid \<and> y \<noteq> invalid"
using plus_trilean.elims by blast
lemma maybe_or_false:
"(x \<or>? y = false) = (x = false \<and> y = false)"
using plus_trilean.elims by blast
lemma maybe_or_true:
"(x \<or>? y = true) = ((x = true \<or> y = true) \<and> x \<noteq> invalid \<and> y \<noteq> invalid)"
using plus_trilean.elims by blast
lemma maybe_not_invalid: "(\<not>? x = invalid) = (x = invalid)"
by (metis maybe_double_negation maybe_not.simps(3))
lemma maybe_or_invalid:
"(x \<or>? y = invalid) = (x = invalid \<or> y = invalid)"
using plus_trilean.elims by blast
lemma maybe_and_invalid:
"(x \<and>? y = invalid) = (x = invalid \<or> y = invalid)"
using times_trilean.elims by blast
lemma maybe_and_false:
"(x \<and>? y = false) = ((x = false \<or> y = false) \<and> x \<noteq> invalid \<and> y \<noteq> invalid)"
using times_trilean.elims by blast
lemma invalid_maybe_and: "invalid \<and>? x = invalid"
using maybe_and_valid by blast
lemma maybe_not_eq: "(\<not>? x = \<not>? y) = (x = y)"
by (metis maybe_double_negation)
lemma de_morgans_1:
"\<not>? (a \<or>? b) = (\<not>?a) \<and>? (\<not>?b)"
by (metis (no_types, hide_lams) add.commute invalid_maybe_and maybe_and_idempotent maybe_and_one maybe_not.elims maybe_not.simps(1) maybe_not.simps(3) maybe_not_invalid maybe_or_zero plus_trilean.simps(1) plus_trilean.simps(4) times_trilean.simps(1) times_trilean_commutative trilean.exhaust trilean.simps(6))
lemma de_morgans_2:
"\<not>? (a \<and>? b) = (\<not>?a) \<or>? (\<not>?b)"
by (metis de_morgans_1 maybe_double_negation)
lemma not_true: "(x \<noteq> true) = (x = false \<or> x = invalid)"
by (metis (no_types, lifting) maybe_not.cases trilean.distinct(1) trilean.distinct(3))
lemma pull_negation: "(x = \<not>? y) = (\<not>? x = y)"
using maybe_double_negation by auto
lemma comp_fun_commute_maybe_or: "comp_fun_commute maybe_or"
apply standard
apply (simp add: comp_def)
apply (rule ext)
by (simp add: add.left_commute)
lemma comp_fun_commute_maybe_and: "comp_fun_commute maybe_and"
apply standard
apply (simp add: comp_def)
apply (rule ext)
by (metis add.left_commute de_morgans_2 maybe_not_eq)
end

View File

@ -0,0 +1,44 @@
section\<open>Variables\<close>
text\<open>Variables can either be inputs or registers. Here we define the \texttt{vname} datatype which
allows us to write expressions in terms of variables and case match during evaluation. We also make
the \texttt{vname} datatype a member of linorder such that we can establish a linear order on
arithmetic expressions, guards, and subsequently transitions.\<close>
theory VName
imports Main
begin
text_raw\<open>\snip{vnametype}{1}{2}{%\<close>
datatype vname = I nat | R nat
text_raw\<open>}%endsnip\<close>
instantiation vname :: linorder begin
text_raw\<open>\snip{vnameorder}{1}{2}{%\<close>
fun less_vname :: "vname \<Rightarrow> vname \<Rightarrow> bool" where
"(I n1) < (R n2) = True" |
"(R n1) < (I n2) = False" |
"(I n1) < (I n2) = (n1 < n2)" |
"(R n1) < (R n2) = (n1 < n2)"
text_raw\<open>}%endsnip\<close>
definition less_eq_vname :: "vname \<Rightarrow> vname \<Rightarrow> bool" where
"less_eq_vname v1 v2 = (v1 < v2 \<or> v1 = v2)"
declare less_eq_vname_def [simp]
instance
apply standard
apply (metis (full_types) dual_order.asym less_eq_vname_def less_vname.simps(2) less_vname.simps(3) less_vname.simps(4) vname.exhaust)
apply simp
subgoal for x y z
apply (induct x y rule: less_vname.induct)
apply (metis less_eq_vname_def less_vname.elims(2) less_vname.elims(3) vname.simps(4))
apply simp
apply (metis less_eq_vname_def less_trans less_vname.elims(3) less_vname.simps(3) vname.simps(4))
by (metis le_less_trans less_eq_vname_def less_imp_le_nat less_vname.elims(2) less_vname.simps(4) vname.simps(4))
apply (metis dual_order.asym less_eq_vname_def less_vname.elims(2) less_vname.simps(3) less_vname.simps(4))
subgoal for x y
by (induct x y rule: less_vname.induct, auto)
done
end
end

View File

@ -0,0 +1,99 @@
section\<open>Values\<close>
text\<open>Our EFSM implementation can currently handle integers and strings. Here we define a sum type
which combines these. We also define an arithmetic in terms of values such that EFSMs do not need
to be strongly typed.\<close>
theory Value
imports Trilean
begin
text_raw\<open>\snip{valuetype}{1}{2}{%\<close>
datatype "value" = Num int | Str String.literal
text_raw\<open>}%endsnip\<close>
fun is_Num :: "value \<Rightarrow> bool" where
"is_Num (Num _) = True" |
"is_Num (Str _) = False"
text_raw\<open>\snip{maybeIntArith}{1}{2}{%\<close>
fun maybe_arith_int :: "(int \<Rightarrow> int \<Rightarrow> int) \<Rightarrow> value option \<Rightarrow> value option \<Rightarrow> value option" where
"maybe_arith_int f (Some (Num x)) (Some (Num y)) = Some (Num (f x y))" |
"maybe_arith_int _ _ _ = None"
text_raw\<open>}%endsnip\<close>
lemma maybe_arith_int_not_None:
"maybe_arith_int f a b \<noteq> None = (\<exists>n n'. a = Some (Num n) \<and> b = Some (Num n'))"
using maybe_arith_int.elims maybe_arith_int.simps(1) by blast
lemma maybe_arith_int_Some:
"maybe_arith_int f a b = Some (Num x) = (\<exists>n n'. a = Some (Num n) \<and> b = Some (Num n') \<and> f n n' = x)"
using maybe_arith_int.elims maybe_arith_int.simps(1) by blast
lemma maybe_arith_int_None:
"(maybe_arith_int f a1 a2 = None) = (\<nexists>n n'. a1 = Some (Num n) \<and> a2 = Some (Num n'))"
using maybe_arith_int_not_None by blast
lemma maybe_arith_int_Not_Num:
"(\<forall>n. maybe_arith_int f a1 a2 \<noteq> Some (Num n)) = (maybe_arith_int f a1 a2 = None)"
by (metis maybe_arith_int.elims option.distinct(1))
lemma maybe_arith_int_never_string: "maybe_arith_int f a b \<noteq> Some (Str x)"
using maybe_arith_int.elims by blast
definition "value_plus = maybe_arith_int (+)"
lemma value_plus_never_string: "value_plus a b \<noteq> Some (Str x)"
by (simp add: value_plus_def maybe_arith_int_never_string)
lemma value_plus_symmetry: "value_plus x y = value_plus y x"
apply (induct x y rule: maybe_arith_int.induct)
by (simp_all add: value_plus_def)
definition "value_minus = maybe_arith_int (-)"
lemma value_minus_never_string: "value_minus a b \<noteq> Some (Str x)"
by (simp add: maybe_arith_int_never_string value_minus_def)
definition "value_times = maybe_arith_int (*)"
lemma value_times_never_string: "value_times a b \<noteq> Some (Str x)"
by (simp add: maybe_arith_int_never_string value_times_def)
fun MaybeBoolInt :: "(int \<Rightarrow> int \<Rightarrow> bool) \<Rightarrow> value option \<Rightarrow> value option \<Rightarrow> trilean" where
"MaybeBoolInt f (Some (Num a)) (Some (Num b)) = (if f a b then true else false)" |
"MaybeBoolInt _ _ _ = invalid"
lemma MaybeBoolInt_not_num_1:
"\<forall>n. r \<noteq> Some (Num n) \<Longrightarrow> MaybeBoolInt f n r = invalid"
using MaybeBoolInt.elims by blast
definition value_gt :: "value option \<Rightarrow> value option \<Rightarrow> trilean" where
"value_gt a b \<equiv> MaybeBoolInt (>) a b"
fun value_eq :: "value option \<Rightarrow> value option \<Rightarrow> trilean" where
"value_eq None _ = invalid" |
"value_eq _ None = invalid" |
"value_eq (Some a) (Some b) = (if a = b then true else false)"
lemma value_eq_true: "(value_eq a b = true) = (\<exists>x y. a = Some x \<and> b = Some y \<and> x = y)"
by (cases a; cases b, auto)
lemma value_eq_false: "(value_eq a b = false) = (\<exists>x y. a = Some x \<and> b = Some y \<and> x \<noteq> y)"
by (cases a; cases b, auto)
lemma value_gt_true_Some: "value_gt a b = true \<Longrightarrow> (\<exists>x. a = Some x) \<and> (\<exists>y. b = Some y)"
by (cases a; cases b, auto simp: value_gt_def)
lemma value_gt_true: "(value_gt a b = true) = (\<exists>x y. a = Some (Num x) \<and> b = Some (Num y) \<and> x > y)"
apply (cases a)
using value_gt_true_Some apply blast
apply (cases b)
using value_gt_true_Some apply blast
subgoal for aa bb by (cases aa; cases bb, auto simp: value_gt_def)
done
lemma value_gt_false_Some: "value_gt a b = false \<Longrightarrow> (\<exists>x. a = Some x) \<and> (\<exists>y. b = Some y)"
by (cases a; cases b, auto simp: value_gt_def)
end

View File

@ -0,0 +1,41 @@
subsection\<open>Value Lexorder\<close>
text\<open>This theory defines a lexicographical ordering on values such that we can build orderings for
arithmetic expressions and guards. Here, numbers are defined as less than strings, else the natural
ordering on the respective datatypes is used.\<close>
theory Value_Lexorder
imports Value
begin
instantiation "value" :: linorder begin
text_raw\<open>\snip{valueorder}{1}{2}{%\<close>
fun less_value :: "value \<Rightarrow> value \<Rightarrow> bool" where
"(Num n) < (Str s) = True" |
"(Str s) < (Num n) = False" |
"(Str s1) < (Str s2) = (s1 < s2)" |
"(Num n1) < (Num n2) = (n1 < n2)"
text_raw\<open>}%endsnip\<close>
definition less_eq_value :: "value \<Rightarrow> value \<Rightarrow> bool" where
"less_eq_value v1 v2 = (v1 < v2 \<or> v1 = v2)"
declare less_eq_value_def [simp]
instance
apply standard
apply (simp, metis less_imp_not_less less_value.elims(2) less_value.simps(3) less_value.simps(4))
apply simp
subgoal for x y z
apply (induct x y rule: less_value.induct)
apply (metis less_eq_value_def less_value.elims(3) less_value.simps(2) value.simps(4))
apply simp
apply (metis dual_order.strict_trans less_eq_value_def less_value.elims(2) less_value.simps(3) value.distinct(1))
by (cases z, auto)
apply (metis less_eq_value_def less_imp_not_less less_value.elims(2) less_value.simps(3) less_value.simps(4))
subgoal for x y
by (induct x y rule: less_value.induct, auto)
done
end
end

View File

@ -0,0 +1,70 @@
@article{bochvar1981,
author = {Bochvar, D. A.},
doi = {10.1080/01445348108837023},
issn = {14645149},
journal = {History and Philosophy of Logic},
month = {jan},
number = {1-2},
pages = {87--112},
title = {{On a three-valued logical calculus and its application to the analysis of the paradoxes of the classical extended functional calculus}},
url = {http://www.tandfonline.com/doi/abs/10.1080/01445348108837023},
volume = {2},
year = {1981}
}
@InCollection{foster2018,
address = {Heidelberg},
author = {Michael Foster and Ramsay G. Taylor and Achim D. Brucker
and John Derrick},
booktitle = {ICFEM},
doi = {10.1007/978-3-030-02450-5},
editor = {Jin Song Dong and Jing Sun},
isbn = {978-3-030-02449-9},
keywords = {model inference, state machine models, EFSM},
language = {USenglish},
location = {Gold Coast, Australia},
number = {11232},
pages = {373--387},
pdf = {https://www.brucker.ch/bibliography/download/2018/foster.ea-efsm-2018.pdf},
publisher = {Springer-Verlag},
series = {Lecture Notes in Computer Science},
title = {Formalising Extended Finite State Machine Transition
Merging},
url = {https://www.brucker.ch/bibliography/abstract/foster.ea-efsm-2018},
year = {2018}
}
@InCollection{foster2019,
address = {Heidelberg},
author = {Michael Foster and Achim D. Brucker and Ramsay G. Taylor
and Siobh{\'a}n North and John Derrick},
booktitle = {Software Engineering and Formal Methods (SEFM)},
doi = {10.1007/978-3-030-30446-1_14},
editor = {Peter C. {\"O}lveczky and Gwen Sala{\"u}n},
isbn = {3-540-25109-X},
keywords = {EFSM Inference, Model Inference, Reverse Engineering},
language = {USenglish},
location = {Oslo},
number = {11724},
pages = {257--272},
pdf = {https://www.brucker.ch/bibliography/download/2019/foster.ea-incorporating-2019.pdf},
publisher = {Springer-Verlag},
series = {Lecture Notes in Computer Science},
title = {Incorporating Data into EFSM Inference},
url = {https://www.brucker.ch/bibliography/abstract/foster.ea-incorporating-2019},
year = {2019}
}
@inproceedings{lorenzoli2008,
address = {New York, New York, USA},
author = {Lorenzoli, Davide and Mariani, Leonardo and Pezz{\`{e}}, Mauro},
booktitle = {Proceedings of the 13th international conference on Software engineering - ICSE '08},
doi = {10.1145/1368088.1368157},
isbn = {9781605580791},
keywords = {dynamic analysis,gk-tail,model synthesis},
pages = {501},
publisher = {ACM Press},
title = {{Automatic generation of software behavioral models}},
url = {http://portal.acm.org/citation.cfm?doid=1368088.1368157},
year = {2008}
}

View File

@ -0,0 +1,148 @@
\documentclass[10pt,DIV16,a4paper,abstract=true,twoside=semi,openright]{scrreprt}
\usepackage[USenglish]{babel}
\usepackage[numbers, sort&compress]{natbib}
\usepackage{isabelle,isabellesym}
\usepackage{booktabs}
\usepackage{paralist}
\usepackage{graphicx}
\usepackage{amssymb}
\usepackage{xspace}
\usepackage{xcolor}
\usepackage{hyperref}
\usepackage{enumitem} % Mess about with itemize, enumerate, description styles
\newcommand\mydescriptionlabel[1]{\hspace{\leftmargini}\textbf{#1}}
\newenvironment{where}{%
\let\descriptionlabel\mydescriptionlabel
\description[itemsep=0em, font=\normalfont]
}{%
\enddescription
}
\pagestyle{headings}
\isabellestyle{default}
\setcounter{tocdepth}{1}
\newcommand{\ie}{i.\,e.\xspace}
\newcommand{\eg}{e.\,g.\xspace}
\newcommand{\thy}{\isabellecontext}
\renewcommand{\isamarkupsection}[1]{%
\begingroup%
\def\isacharunderscore{\textunderscore}%
\section{#1 (\thy)}%
\def\isacharunderscore{-}%
\expandafter\label{sec:\isabellecontext}%
\endgroup%
}
\newcommand{\orcidID}[1]{} % temp. hack
\newcommand{\repeatisanl}[1]
{\ifnum#1=0\else\isanewline\repeatisanl{\numexpr#1-1}\fi}
\newcommand{\snip}[4]{\repeatisanl#2#4\repeatisanl#3}
\newcommand{\DefineSnippet}[2]{%
\expandafter\newcommand\csname snippet--#1\endcsname{%
\begin{quote}
\begin{isabelle}
#2
\end{isabelle}
\end{quote}}}
\newcommand{\Snippet}[1]{%
\ifcsname snippet--#1\endcsname{\csname snippet--#1\endcsname}%
\else+++++++ERROR: Snippet ``#1 not defined''+++++++ \fi}
\title{A Formal Model of Extended Finite State Machines}%
\author{%
\begin{minipage}{.8\textwidth}
\centering
Michael~Foster\footnotemark[1]\orcidID{0000-0001-8233-9873}%
\qquad\qquad%
Achim~D.~Brucker\footnotemark[2]\orcidID{0000-0002-6355-1200}%
\\%
Ramsay~G.~Taylor\footnotemark[1]\orcidID{0000-0002-4036-7590}%
\qquad\qquad%
John~Derrick\footnotemark[1]\orcidID{0000-0002-6631-8914}%
\end{minipage}
}
\publishers{%
\footnotemark[1]~Department of Computer Science, The University of Sheffield, Sheffield, UK\texorpdfstring{\\}{, }%
\texttt{\{%
\href{mailto:jmafoster1@sheffield.ac.uk}{jmafoster1},%
\href{mailto:r.g.taylor@sheffield.ac.uk}{r.g.taylor},%
\href{mailto:j.derrick@sheffield.ac.uk}{j.derrick}%
\}@sheffield.ac.uk}\\[2em]%
\footnotemark[2]~%
Department of Computer Science, University of Exeter, Exeter, UK\texorpdfstring{\\}{, }%
\href{mailto:a.brucker@exeter.ac.uk}{\texttt{a.brucker@exeter.ac.uk}}%
}
\begin{document}
\maketitle
\begin{abstract}
In this AFP entry, we provide a formalisation of extended finite state machines (EFSMs) where models are represented as finite sets of transitions between states. EFSMs execute traces to produce observable outputs. We also define various simulation and equality metrics for EFSMs in terms of traces and prove their strengths in relation to each other. Another key contribution is a framework of function definitions such that LTL properties can be phrased over EFSMs. Finally, we provide a simple example case study in the form of a drinks machine.
\begin{quote}
\bigskip
\noindent{\textbf{Keywords:} Extended Finite State Machines, Automata, Linear Temporal Logic}
\end{quote}
\end{abstract}
\tableofcontents
\cleardoublepage
\chapter{Introduction}
This AFP entry formalises extended finite state machines (EFSMs) as defined in \cite{foster2018}. Here, models maintain both a \emph{control flow state} and a \emph{data state}, which takes the form of a set of \emph{registers} to which values may be assigned. Transitions may take additional input parameters, and may impose guard conditions on the values of both inputs and registers. Additionally, transitions may produce observable outputs and update the data state by evaluating arithmetic functions over inputs and registers.
As defined in \cite{foster2018}, an EFSM is a tuple, $(S, s_0, T)$ where
\begin{where}
\item [$S$] is a finite non-empty set of states.
\item [$s_0 \in S$]is the initial state.
\item [$T$] is the transition matrix $T:(S \times S) \to \mathcal{P}(L \times \mathbb{N} \times G \times F \times U)$ with rows representing origin states and columns representing destination states.
\end{where}
In $T$
\begin{where}
\item [$L$] is a finite set of transition labels
\item [$\mathbb{N}$] gives the transition \emph{arity} (the number of input parameters), which may be zero.
\item [$G$] is a finite set of Boolean guard functions $G:(I \times R) \to \mathbb{B}$.
\item [$F$] is a finite set of \emph{output functions} $F:(I \times R) \to O$.
\item [$U$] is a finite set of \emph{update functions} $U:(I \times R) \to R$.
\end{where}
In $G$, $F$, and $U$
\begin{where}
\item [$I$] is a list $[i_0, i_1, \ldots, i_{m-1}]$ of values representing the inputs of a transition, which is empty if the arity is zero.
\item [$R$] is a mapping from variables $[r_0, r_1, \ldots]$, representing each register of the machine, to their values.
\item [$O$] is a list $[o_0, o_1, \ldots, o_{n-1}]$ of values, which may be empty, representing the outputs of a transition.
\end{where}
EFSM transitions have five components: label, arity, guards, outputs, and updates. Transition labels are strings, and the arities natural numbers. Guards have a defined type of \emph{guard expression} (\texttt{gexp}) and the outputs and updates are defined using \emph{arithmetic expressions} (\texttt{aexp}). Outputs are simply a list of expressions to be evaluated. Updates are a list of pairs with the first element being the index of the register to be updated, and the second element being an arithmetic expression to be evaluated.
\begin{figure}
\centering
\includegraphics[height=\textheight]{session_graph}
\caption{The Dependency Graph of the Isabelle Theories.\label{fig:session-graph}}
\end{figure}
The rest of this document is automatically generated from the
formalization in Isabelle/HOL, i.e., all content is checked by
Isabelle. Overall, the structure of this document follows the
theory dependencies (see \autoref{fig:session-graph}):
\nocite{foster.ea:efsm:2018}
\clearpage
% \chapter{Theories}
\input{session}
{\small
\bibliographystyle{abbrvnat}
\bibliography{root}
}
\end{document}
\endinput
%%% Local Variables:
%%% mode: latex
%%% TeX-master: t
%%% End:

View File

@ -0,0 +1,282 @@
chapter\<open>Examples\<close>
text\<open>In this chapter, we provide some examples of EFSMs and proofs over them. We first present a
formalisation of a simple drinks machine. Next, we prove observational equivalence of an alternative
model. Finally, we prove some temporal properties of the first example.\<close>
section\<open>Drinks Machine\<close>
text\<open>This theory formalises a simple drinks machine. The \emph{select} operation takes one
argument - the desired beverage. The \emph{coin} operation also takes one parameter representing
the value of the coin. The \emph{vend} operation has two flavours - one which dispenses the drink if
the customer has inserted enough money, and one which dispenses nothing if the user has not inserted
sufficient funds.
We first define a datatype \emph{statemane} which corresponds to $S$ in the formal definition.
Note that, while statename has four elements, the drinks machine presented here only requires three
states. The fourth element is included here so that the \emph{statename} datatype may be used in
the next example.\<close>
theory Drinks_Machine
imports "Extended_Finite_State_Machines.EFSM"
begin
text_raw\<open>\snip{selectdef}{1}{2}{%\<close>
definition select :: "transition" where
"select \<equiv> \<lparr>
Label = STR ''select'',
Arity = 1,
Guards = [],
Outputs = [],
Updates = [
(1, V (I 0)),
(2, L (Num 0))
]
\<rparr>"
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{coindef}{1}{2}{%\<close>
definition coin :: "transition" where
"coin \<equiv> \<lparr>
Label = STR ''coin'',
Arity = 1,
Guards = [],
Outputs = [Plus (V (R 2)) (V (I 0))],
Updates = [
(1, V (R 1)),
(2, Plus (V (R 2)) (V (I 0)))
]
\<rparr>"
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{venddef}{1}{2}{%\<close>
definition vend:: "transition" where
"vend\<equiv> \<lparr>
Label = STR ''vend'',
Arity = 0,
Guards = [(Ge (V (R 2)) (L (Num 100)))],
Outputs = [(V (R 1))],
Updates = [(1, V (R 1)), (2, V (R 2))]
\<rparr>"
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{vendfaildef}{1}{2}{%\<close>
definition vend_fail :: "transition" where
"vend_fail \<equiv> \<lparr>
Label = STR ''vend'',
Arity = 0,
Guards = [(Lt (V (R 2)) (L (Num 100)))],
Outputs = [],
Updates = [(1, V (R 1)), (2, V (R 2))]
\<rparr>"
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{drinksdef}{1}{2}{%\<close>
definition drinks :: "transition_matrix" where
"drinks \<equiv> {|
((0,1), select),
((1,1), coin),
((1,1), vend_fail),
((1,2), vend)
|}"
text_raw\<open>}%endsnip\<close>
lemmas transitions = select_def coin_def vend_def vend_fail_def
lemma apply_updates_vend: "apply_updates (Updates vend) (join_ir [] r) r = r"
by (simp add: vend_def apply_updates_def)
lemma drinks_states: "S drinks = {|0, 1, 2|}"
apply (simp add: S_def drinks_def)
by auto
lemma possible_steps_0:
"length i = 1 \<Longrightarrow>
possible_steps drinks 0 r (STR ''select'') i = {|(1, select)|}"
apply (simp add: possible_steps_singleton drinks_def)
apply safe
by (simp_all add: transitions apply_guards_def)
lemma first_step_select:
"(s', t) |\<in>| possible_steps drinks 0 r aa b \<Longrightarrow> s' = 1 \<and> t = select"
apply (simp add: possible_steps_def fimage_def ffilter_def fmember_def Abs_fset_inverse Set.filter_def drinks_def)
apply safe
by (simp_all add: transitions)
lemma drinks_vend_insufficient:
"r $ 2 = Some (Num x1) \<Longrightarrow>
x1 < 100 \<Longrightarrow>
possible_steps drinks 1 r (STR ''vend'') [] = {|(1, vend_fail)|}"
apply (simp add: possible_steps_singleton drinks_def)
apply safe
by (simp_all add: transitions apply_guards_def value_gt_def join_ir_def connectives)
lemma drinks_vend_invalid:
"\<nexists>n. r $ 2 = Some (Num n) \<Longrightarrow>
possible_steps drinks 1 r (STR ''vend'') [] = {||}"
apply (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
by (simp add: MaybeBoolInt_not_num_1 value_gt_def)
lemma possible_steps_1_coin:
"length i = 1 \<Longrightarrow> possible_steps drinks 1 r (STR ''coin'') i = {|(1, coin)|}"
apply (simp add: possible_steps_singleton drinks_def)
apply safe
by (simp_all add: transitions)
lemma possible_steps_2_vend:
"\<exists>n. r $ 2 = Some (Num n) \<and> n \<ge> 100 \<Longrightarrow>
possible_steps drinks 1 r (STR ''vend'') [] = {|(2, vend)|}"
apply (simp add: possible_steps_singleton drinks_def)
apply safe
by (simp_all add: transitions apply_guards_def value_gt_def join_ir_def connectives)
lemma recognises_from_2:
"recognises_execution drinks 1 <1 $:= d, 2 $:= Some (Num 100)> [(STR ''vend'', [])]"
apply (rule recognises_execution.step)
apply (rule_tac x="(2, vend)" in fBexI)
apply simp
by (simp add: possible_steps_2_vend)
lemma recognises_from_1a:
"recognises_execution drinks 1 <1 $:= d, 2 $:= Some (Num 50)> [(STR ''coin'', [Num 50]), (STR ''vend'', [])]"
apply (rule recognises_execution.step)
apply (rule_tac x="(1, coin)" in fBexI)
apply (simp add: apply_updates_def coin_def finfun_update_twist value_plus_def recognises_from_2)
by (simp add: possible_steps_1_coin)
lemma recognises_from_1: "recognises_execution drinks 1 <2 $:= Some (Num 0), 1 $:= Some d>
[(STR ''coin'', [Num 50]), (STR ''coin'', [Num 50]), (STR ''vend'', [])]"
apply (rule recognises_execution.step)
apply (rule_tac x="(1, coin)" in fBexI)
apply (simp add: apply_updates_def coin_def value_plus_def finfun_update_twist recognises_from_1a)
by (simp add: possible_steps_1_coin)
lemma purchase_coke:
"observe_execution drinks 0 <> [(STR ''select'', [Str ''coke'']), (STR ''coin'', [Num 50]), (STR ''coin'', [Num 50]), (STR ''vend'', [])] =
[[], [Some (Num 50)], [Some (Num 100)], [Some (Str ''coke'')]]"
by (simp add: possible_steps_0 possible_steps_1_coin possible_steps_2_vend transitions
apply_outputs_def apply_updates_def value_plus_def)
lemma rejects_input:
"l \<noteq> STR ''coin'' \<Longrightarrow>
l \<noteq> STR ''vend'' \<Longrightarrow>
\<not> recognises_execution drinks 1 d' [(l, i)]"
apply (rule no_possible_steps_rejects)
by (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
lemma rejects_recognises_prefix: "l \<noteq> STR ''coin'' \<Longrightarrow>
l \<noteq> STR ''vend'' \<Longrightarrow>
\<not> (recognises drinks [(STR ''select'', [Str ''coke'']), (l, i)])"
apply (rule trace_reject_later)
apply (simp add: possible_steps_0 select_def join_ir_def input2state_def)
using rejects_input by blast
lemma rejects_termination:
"observe_execution drinks 0 <> [(STR ''select'', [Str ''coke'']), (STR ''rejects'', [Num 50]), (STR ''coin'', [Num 50])] = [[]]"
apply (rule observe_execution_step)
apply (simp add: step_def possible_steps_0 select_def)
apply (rule observe_execution_no_possible_step)
by (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
(* Part of Example 2 in Foster et. al. *)
lemma r2_0_vend:
"can_take_transition vend i r \<Longrightarrow>
\<exists>n. r $ 2 = Some (Num n) \<and> n \<ge> 100" (* You can't take vendimmediately after taking select *)
apply (simp add: can_take_transition_def can_take_def vend_def apply_guards_def maybe_negate_true maybe_or_false connectives value_gt_def)
using MaybeBoolInt.elims by force
lemma drinks_vend_sufficient: "r $ 2 = Some (Num x1) \<Longrightarrow>
x1 \<ge> 100 \<Longrightarrow>
possible_steps drinks 1 r (STR ''vend'') [] = {|(2, vend)|}"
using possible_steps_2_vend by blast
lemma drinks_end: "possible_steps drinks 2 r a b = {||}"
apply (simp add: possible_steps_def drinks_def transitions)
by force
lemma drinks_vend_r2_String:
"r $ 2 = Some (value.Str x2) \<Longrightarrow>
possible_steps drinks 1 r (STR ''vend'') [] = {||}"
apply (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
by (simp add: value_gt_def)
lemma drinks_vend_r2_rejects:
"\<nexists>n. r $ 2 = Some (Num n) \<Longrightarrow> step drinks 1 r (STR ''vend'') [] = None"
apply (rule no_possible_steps_1)
apply (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
by (simp add: MaybeBoolInt_not_num_1 value_gt_def)
lemma drinks_0_rejects:
"\<not> (fst a = STR ''select'' \<and> length (snd a) = 1) \<Longrightarrow>
(possible_steps drinks 0 r (fst a) (snd a)) = {||}"
apply (simp add: drinks_def possible_steps_def transitions)
by force
lemma drinks_vend_empty: "(possible_steps drinks 0 <> (STR ''vend'') []) = {||}"
using drinks_0_rejects
by auto
lemma drinks_1_rejects:
"fst a = STR ''coin'' \<longrightarrow> length (snd a) \<noteq> 1 \<Longrightarrow>
a \<noteq> (STR ''vend'', []) \<Longrightarrow>
possible_steps drinks 1 r (fst a) (snd a) = {||}"
apply (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
by (metis prod.collapse)
lemma drinks_rejects_future: "\<not> recognises_execution drinks 2 d ((l, i)#t)"
apply (rule no_possible_steps_rejects)
by (simp add: possible_steps_empty drinks_def)
lemma drinks_1_rejects_trace:
assumes not_vend: "e \<noteq> (STR ''vend'', [])"
and not_coin: "\<nexists>i. e = (STR ''coin'', [i])"
shows "\<not> recognises_execution drinks 1 r (e # es)"
proof-
show ?thesis
apply (cases e, simp)
subgoal for a b
apply (rule no_possible_steps_rejects)
apply (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
apply (case_tac b)
using not_vend apply simp
using not_coin by auto
done
qed
lemma rejects_state_step: "s > 1 \<Longrightarrow> step drinks s r l i = None"
apply (rule no_possible_steps_1)
by (simp add: possible_steps_empty drinks_def)
lemma invalid_other_states:
"s > 1 \<Longrightarrow> \<not> recognises_execution drinks s r ((aa, b) # t)"
apply (rule no_possible_steps_rejects)
by (simp add: possible_steps_empty drinks_def)
lemma vend_ge_100:
"possible_steps drinks 1 r l i = {|(2, vend)|} \<Longrightarrow>
\<not>? value_gt (Some (Num 100)) (r $ 2) = trilean.true"
apply (insert possible_steps_apply_guards[of drinks 1 r l i 2 vend])
by (simp add: possible_steps_def apply_guards_def vend_def)
lemma drinks_no_possible_steps_1:
assumes not_coin: "\<not> (a = STR ''coin'' \<and> length b = 1)"
and not_vend: "\<not> (a = STR ''vend'' \<and> b = [])"
shows "possible_steps drinks 1 r a b = {||}"
using drinks_1_rejects not_coin not_vend by auto
lemma possible_steps_0_not_select: "a \<noteq> STR ''select'' \<Longrightarrow>
possible_steps drinks 0 <> a b = {||}"
apply (simp add: possible_steps_def ffilter_def fset_both_sides Abs_fset_inverse Set.filter_def drinks_def)
apply safe
by (simp_all add: select_def)
lemma possible_steps_select_wrong_arity: "a = STR ''select'' \<Longrightarrow>
length b \<noteq> 1 \<Longrightarrow>
possible_steps drinks 0 <> a b = {||}"
apply (simp add: possible_steps_def ffilter_def fset_both_sides Abs_fset_inverse Set.filter_def drinks_def)
apply safe
by (simp_all add: select_def)
lemma possible_steps_0_invalid:
"\<not> (l = STR ''select'' \<and> length i = 1) \<Longrightarrow>
possible_steps drinks 0 <> l i = {||}"
using possible_steps_0_not_select possible_steps_select_wrong_arity by fastforce
end

View File

@ -0,0 +1,222 @@
section\<open>An Observationally Equivalent Model\<close>
text\<open>This theory defines a second formalisation of the drinks machine example which produces
identical output to the first model. This property is called \emph{observational equivalence} and is
discussed in more detail in \cite{foster2018}.\<close>
theory Drinks_Machine_2
imports Drinks_Machine
begin
(* Only imports Drinks_Machine_LTL so I can easily open all three at once for testing *)
definition vend_nothing :: "transition" where
"vend_nothing \<equiv> \<lparr>
Label = (STR ''vend''),
Arity = 0,
Guards = [],
Outputs = [],
Updates = [(1, V (R 1)), (2, V (R 2))]
\<rparr>"
lemmas transitions = Drinks_Machine.transitions vend_nothing_def
definition drinks2 :: transition_matrix where
"drinks2 = {|
((0,1), select),
((1,1), vend_nothing),
((1,2), coin),
((2,2), coin),
((2,2), vend_fail),
((2,3), vend)
|}"
lemma possible_steps_0:
"length i = 1 \<Longrightarrow>
possible_steps drinks2 0 r ((STR ''select'')) i = {|(1, select)|}"
apply (simp add: possible_steps_def drinks2_def transitions)
by force
lemma possible_steps_1:
"length i = 1 \<Longrightarrow>
possible_steps drinks2 1 r ((STR ''coin'')) i = {|(2, coin)|}"
apply (simp add: possible_steps_def drinks2_def transitions)
by force
lemma possible_steps_2_coin:
"length i = 1 \<Longrightarrow>
possible_steps drinks2 2 r ((STR ''coin'')) i = {|(2, coin)|}"
apply (simp add: possible_steps_def drinks2_def transitions)
by force
lemma possible_steps_2_vend:
"r $ 2 = Some (Num n) \<Longrightarrow>
n \<ge> 100 \<Longrightarrow>
possible_steps drinks2 2 r ((STR ''vend'')) [] = {|(3, vend)|}"
apply (simp add: possible_steps_singleton drinks2_def)
apply safe
by (simp_all add: transitions apply_guards_def value_gt_def join_ir_def connectives)
lemma recognises_first_select:
"recognises_execution drinks 0 r ((aa, b) # as) \<Longrightarrow> aa = STR ''select'' \<and> length b = 1"
using recognises_must_be_possible_step[of drinks 0 r "(aa, b)" as]
apply simp
apply clarify
by (metis first_step_select recognises_possible_steps_not_empty drinks_0_rejects fst_conv snd_conv)
lemma drinks2_vend_insufficient:
"possible_steps drinks2 1 r ((STR ''vend'')) [] = {|(1, vend_nothing)|}"
apply (simp add: possible_steps_def drinks2_def transitions)
by force
lemma drinks2_vend_insufficient2:
"r $ 2 = Some (Num x1) \<Longrightarrow>
x1 < 100 \<Longrightarrow>
possible_steps drinks2 2 r ((STR ''vend'')) [] = {|(2, vend_fail)|}"
apply (simp add: possible_steps_singleton drinks2_def)
apply safe
by (simp_all add: transitions apply_guards_def value_gt_def join_ir_def connectives)
lemma drinks2_vend_sufficient: "r $ 2 = Some (Num x1) \<Longrightarrow>
\<not> x1 < 100 \<Longrightarrow>
possible_steps drinks2 2 r ((STR ''vend'')) [] = {|(3, vend)|}"
apply (simp add: possible_steps_singleton drinks2_def)
apply safe
by (simp_all add: transitions apply_guards_def value_gt_def join_ir_def connectives)
lemma recognises_1_2: "recognises_execution drinks 1 r t \<longrightarrow> recognises_execution drinks2 2 r t"
proof(induct t arbitrary: r)
case (Cons a as)
then show ?case
apply (cases a)
apply (simp add: recognises_step_equiv)
apply (case_tac "a=(STR ''vend'', [])")
apply (case_tac "\<exists>n. r$2 = Some (Num n)")
apply clarify
apply (case_tac "n < 100")
apply (simp add: drinks_vend_insufficient drinks2_vend_insufficient2)
apply (simp add: drinks_vend_sufficient drinks2_vend_sufficient)
apply (metis recognises_prim recognises_prim.elims(3) drinks_rejects_future)
using drinks_vend_invalid apply blast
apply (case_tac "\<exists>i. a=(STR ''coin'', [i])")
apply clarify
apply (simp add: possible_steps_1_coin possible_steps_2_coin)
by (metis recognises_execution.simps drinks_1_rejects_trace)
qed auto
lemma drinks_reject_0_2:
"\<nexists>i. a = (STR ''select'', [i]) \<Longrightarrow>
possible_steps drinks 0 r (fst a) (snd a) = {||}"
apply (rule drinks_0_rejects)
by (cases a, case_tac "snd a", auto)
lemma purchase_coke:
"observe_execution drinks2 0 <> [((STR ''select''), [Str ''coke'']), ((STR ''coin''), [Num 50]), ((STR ''coin''), [Num 50]), ((STR ''vend''), [])] =
[[], [Some (Num 50)], [Some (Num 100)], [Some (Str ''coke'')]]"
by (simp add: possible_steps_0 select_def apply_updates_def
possible_steps_1 coin_def apply_outputs finfun_update_twist
possible_steps_2_coin possible_steps_2_vend vend_def)
lemma drinks2_0_invalid:
"\<not> (aa = (STR ''select'') \<and> length (b) = 1) \<Longrightarrow>
(possible_steps drinks2 0 <> aa b) = {||}"
apply (simp add: drinks2_def possible_steps_def transitions)
by force
lemma drinks2_vend_r2_none:
"r $ 2 = None \<Longrightarrow> possible_steps drinks2 2 r ((STR ''vend'')) [] = {||}"
apply (simp add: possible_steps_empty drinks2_def can_take_transition_def can_take_def transitions)
by (simp add: value_gt_def)
lemma drinks2_end: "possible_steps drinks2 3 r a b = {||}"
apply (simp add: possible_steps_def drinks2_def transitions)
by force
lemma drinks2_vend_r2_String: "r $ 2 = Some (value.Str x2) \<Longrightarrow>
possible_steps drinks2 2 r ((STR ''vend'')) [] = {||}"
apply (simp add: possible_steps_empty drinks2_def)
apply safe
by (simp_all add: transitions can_take_transition_def can_take_def value_gt_def)
lemma drinks2_2_invalid:
"fst a = (STR ''coin'') \<longrightarrow> length (snd a) \<noteq> 1 \<Longrightarrow>
a \<noteq> ((STR ''vend''), []) \<Longrightarrow>
possible_steps drinks2 2 r (fst a) (snd a) = {||}"
apply (simp add: possible_steps_empty drinks2_def transitions can_take_transition_def can_take_def)
by (metis prod.collapse)
lemma drinks2_1_invalid:
"\<not>(a = (STR ''coin'') \<and> length b = 1) \<Longrightarrow>
\<not>(a = (STR ''vend'') \<and> b = []) \<Longrightarrow>
possible_steps drinks2 1 r a b = {||}"
apply (simp add: possible_steps_empty drinks2_def)
apply safe
by (simp_all add: transitions can_take_transition_def can_take_def value_gt_def)
lemma drinks2_vend_invalid:
"\<nexists>n. r $ 2 = Some (Num n) \<Longrightarrow>
possible_steps drinks2 2 r (STR ''vend'') [] = {||}"
apply (simp add: possible_steps_empty drinks2_def)
apply safe
by (simp_all add: transitions can_take_transition_def can_take_def value_gt_def MaybeBoolInt_not_num_1)
lemma equiv_1_2: "executionally_equivalent drinks 1 r drinks2 2 r x"
proof(induct x arbitrary: r)
case (Cons a t)
then show ?case
apply (cases a, clarsimp)
apply (simp add: executionally_equivalent_step)
apply (case_tac "fst a = STR ''coin'' \<and> length (snd a) = 1")
apply (simp add: Drinks_Machine.possible_steps_1_coin possible_steps_2_coin)
apply (case_tac "a = (STR ''vend'', [])")
defer using drinks2_2_invalid drinks_no_possible_steps_1 apply auto[1]
apply (case_tac "\<exists>n. r $ 2 = Some (Num n)")
defer using drinks_vend_invalid drinks2_vend_invalid apply simp
apply clarify
subgoal for aa b n
apply (case_tac "n < 100")
apply (simp add: Drinks_Machine.drinks_vend_insufficient drinks2_vend_insufficient2)
apply (simp add: Drinks_Machine.drinks_vend_sufficient drinks2_vend_sufficient)
apply (induct t)
apply simp
subgoal for aa t apply (case_tac aa, clarsimp)
by (simp add: executionally_equivalent_step drinks_end drinks2_end)
done
done
qed auto
lemma equiv_1_1: "r$2 = Some (Num 0) \<Longrightarrow> executionally_equivalent drinks 1 r drinks2 1 r x"
proof(induct x)
case (Cons a t)
then show ?case
apply (cases a, clarsimp)
subgoal for aa b
apply (simp add: executionally_equivalent_step)
apply (case_tac "aa = STR ''coin'' \<and> length b = 1")
apply (simp add: possible_steps_1_coin possible_steps_1 equiv_1_2)
apply (case_tac "a = (STR ''vend'', [])")
apply clarsimp
apply (simp add: drinks_vend_insufficient drinks2_vend_insufficient)
apply (simp add: vend_fail_def vend_nothing_def apply_updates_def)
apply (metis finfun_upd_triv)
by (simp add: drinks2_1_invalid drinks_no_possible_steps_1)
done
qed auto
(* Corresponds to Example 3 in Foster et. al. *)
lemma executional_equivalence: "executionally_equivalent drinks 0 <> drinks2 0 <> t"
proof(induct t)
case (Cons a t)
then show ?case
apply (cases a, clarify)
subgoal for aa b
apply (simp add: executionally_equivalent_step)
apply (case_tac "aa = STR ''select'' \<and> length b = 1")
apply (simp add: Drinks_Machine.possible_steps_0 possible_steps_0)
apply (simp add: apply_updates_def select_def equiv_1_1)
by (simp add: drinks2_0_invalid possible_steps_0_invalid)
done
qed auto
lemma observational_equivalence: "trace_equivalent drinks drinks2"
by (simp add: executional_equivalence executionally_equivalent_trace_equivalent)
end

View File

@ -0,0 +1,342 @@
section\<open>Temporal Properties\<close>
text\<open>This theory presents some examples of temporal properties over the simple drinks machine.\<close>
theory Drinks_Machine_LTL
imports "Drinks_Machine" "Extended_Finite_State_Machines.EFSM_LTL"
begin
declare One_nat_def [simp del]
lemma P_ltl_step_0:
assumes invalid: "P (None, [], <>)"
assumes select: "l = STR ''select'' \<longrightarrow> P (Some 1, [], <1 $:= Some (hd i), 2 $:= Some (Num 0)>)"
shows "P (ltl_step drinks (Some 0) <> (l, i))"
proof-
have length_i: "\<exists>d. (l, i) = (STR ''select'', [d]) \<Longrightarrow> length i = 1"
by (induct i, auto)
have length_i_2: "\<forall>d. i \<noteq> [d] \<Longrightarrow> length i \<noteq> 1"
by (induct i, auto)
show ?thesis
apply (case_tac "\<exists>d. (l, i) = (STR ''select'', [d])")
apply (simp add: possible_steps_0 length_i select_def apply_updates_def)
using select apply auto[1]
by (simp add: possible_steps_0_invalid length_i_2 invalid)
qed
lemma P_ltl_step_1:
assumes invalid: "P (None, [], r)"
assumes coin: "l = STR ''coin'' \<longrightarrow> P (Some 1, [value_plus (r $ 2) (Some (hd i))], r(2 $:= value_plus (r $ 2) (Some (i ! 0))))"
assumes vend_fail: "value_gt (Some (Num 100)) (r $ 2) = trilean.true \<longrightarrow> P (Some 1, [],r)"
assumes vend: "\<not>? value_gt (Some (Num 100)) (r $ 2) = trilean.true \<longrightarrow> P (Some 2, [r$1], r)"
shows "P (ltl_step drinks (Some 1) r (l, i))"
proof-
have length_i: "\<And>s. \<exists>d. (l, i) = (s, [d]) \<Longrightarrow> length i = 1"
by (induct i, auto)
have length_i_2: "\<forall>d. i \<noteq> [d] \<Longrightarrow> length i \<noteq> 1"
by (induct i, auto)
show ?thesis
apply (case_tac "\<exists>d. (l, i) = (STR ''coin'', [d])")
apply (simp add: possible_steps_1_coin length_i coin_def apply_outputs_def apply_updates_def)
using coin apply auto[1]
apply (case_tac "(l, i) = (STR ''vend'', [])")
apply (case_tac "\<exists>n. r $ 2 = Some (Num n)")
apply clarsimp
subgoal for n
apply (case_tac "n \<ge> 100")
apply (simp add: drinks_vend_sufficient vend_def apply_updates_def apply_outputs_def)
apply (metis finfun_upd_triv possible_steps_2_vend vend vend_ge_100)
apply (simp add: drinks_vend_insufficient vend_fail_def apply_updates_def apply_outputs_def)
apply (metis MaybeBoolInt.simps(1) finfun_upd_triv not_less value_gt_def vend_fail)
done
apply (simp add: drinks_vend_invalid invalid)
by (simp add: drinks_no_possible_steps_1 length_i_2 invalid)
qed
lemma LTL_r2_not_always_gt_100: "not (alw (check_exp (Gt (V (Rg 2)) (L (Num 100))))) (watch drinks i)"
using value_gt_def by auto
lemma drinks_step_2_none: "ltl_step drinks (Some 2) r e = (None, [], r)"
by (simp add: drinks_end ltl_step_none_2)
lemma one_before_two_2:
"alw (\<lambda>x. statename (shd (stl x)) = Some 2 \<longrightarrow> statename (shd x) = Some 1) (make_full_observation drinks (Some 2) r [r $ 1] x2a)"
proof(coinduction)
case alw
then show ?case
apply (simp add: drinks_step_2_none)
by (metis (mono_tags, lifting) alw_mono nxt.simps once_none_nxt_always_none option.distinct(1))
qed
lemma one_before_two_aux:
assumes "\<exists> p r i. j = nxt (make_full_observation drinks (Some 1) r p) i"
shows "alw (\<lambda>x. nxt (state_eq (Some 2)) x \<longrightarrow> state_eq (Some 1) x) j"
using assms apply(coinduct)
apply simp
apply clarify
apply standard
apply simp
apply simp
subgoal for r i
apply (case_tac "shd (stl i)")
apply (simp del: ltl_step.simps)
apply (rule P_ltl_step_1)
apply (rule disjI2)
apply (rule alw_mono[of "nxt (state_eq None)"])
apply (simp add: once_none_nxt_always_none)
apply simp
apply auto[1]
apply auto[1]
apply simp
by (simp add: one_before_two_2)
done
lemma LTL_nxt_2_means_vend:
"alw (nxt (state_eq (Some 2)) impl (state_eq (Some 1))) (watch drinks i)"
proof(coinduction)
case alw
then show ?case
apply (case_tac "shd i")
apply (simp del: ltl_step.simps)
apply (rule P_ltl_step_0)
apply simp
apply (rule disjI2)
apply (rule alw_mono[of "nxt (state_eq None)"])
apply (simp add: once_none_nxt_always_none)
using one_before_two_aux by auto
qed
lemma costsMoney_aux:
assumes "\<exists>p r i. j = (nxt (make_full_observation drinks (Some 1) r p) i)"
shows "alw (\<lambda>xs. nxt (state_eq (Some 2)) xs \<longrightarrow> check_exp (Ge (V (Rg 2)) (L (Num 100))) xs) j"
using assms apply coinduct
apply clarsimp
subgoal for r i
apply (case_tac "shd (stl i)")
apply (simp del: ltl_step.simps)
apply (rule P_ltl_step_1)
apply simp
apply (rule disjI2)
apply (rule alw_mono[of "nxt (state_eq None)"])
apply (simp add: once_none_nxt_always_none)
apply simp
apply auto[1]
apply auto[1]
apply simp
apply standard
apply (rule disjI2)
apply (rule alw_mono[of "nxt (state_eq None)"])
apply (metis (no_types, lifting) drinks_step_2_none fst_conv make_full_observation.sel(2) nxt.simps nxt_alw once_none_always_none_aux)
by simp
done
(* costsMoney: THEOREM drinks |- G(X(cfstate=State_2) => gval(value_ge(r_2, Some(NUM(100))))); *)
lemma LTL_costsMoney:
"(alw (nxt (state_eq (Some 2)) impl (check_exp (Ge (V (Rg 2)) (L (Num 100)))))) (watch drinks i)"
proof(coinduction)
case alw
then show ?case
apply (cases "shd i")
subgoal for l ip
apply (case_tac "l = STR ''select'' \<and> length ip = 1")
defer
apply (simp add: possible_steps_0_invalid)
apply (rule disjI2)
apply (rule alw_mono[of "nxt (state_eq None)"])
apply (simp add: once_none_nxt_always_none)
apply (simp add: )
apply (simp add: possible_steps_0 select_def)
apply (rule disjI2)
apply (simp only: nxt.simps[symmetric])
using costsMoney_aux by auto
done
qed
lemma LTL_costsMoney_aux:
"(alw (not (check_exp (Ge (V (Rg 2)) (L (Num 100)))) impl (not (nxt (state_eq (Some 2)))))) (watch drinks i)"
by (metis (no_types, lifting) LTL_costsMoney alw_mono)
lemma implode_select: "String.implode ''select'' = STR ''select''"
by (metis Literal.rep_eq String.implode_explode_eq zero_literal.rep_eq)
lemma implode_coin: "String.implode ''coin'' = STR ''coin''"
by (metis Literal.rep_eq String.implode_explode_eq zero_literal.rep_eq)
lemma implode_vend: "String.implode ''vend'' = STR ''vend''"
by (metis Literal.rep_eq String.implode_explode_eq zero_literal.rep_eq)
lemmas implode_labels = implode_select implode_coin implode_vend
lemma LTL_neverReachS2:"(((((action_eq (''select'', [Str ''coke''])))
aand
(nxt ((action_eq (''coin'', [Num 100])))))
aand
(nxt (nxt((label_eq ''vend'' aand (input_eq []))))))
impl
(nxt (nxt (nxt (state_eq (Some 2))))))
(watch drinks i)"
apply (simp add: implode_labels)
apply (cases i)
apply clarify
apply simp
apply (simp add: possible_steps_0 select_def)
apply (case_tac "shd x2", clarify)
apply (simp add: possible_steps_1_coin coin_def value_plus_def finfun_update_twist apply_updates_def)
apply (case_tac "shd (stl x2)", clarify)
by (simp add: drinks_vend_sufficient )
lemma ltl_step_not_select:
"\<nexists>i. e = (STR ''select'', [i]) \<Longrightarrow>
ltl_step drinks (Some 0) r e = (None, [], r)"
apply (cases e, clarify)
subgoal for a b
apply (rule ltl_step_none)
apply (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def select_def)
by (cases e, case_tac b, auto)
done
lemma ltl_step_select:
"ltl_step drinks (Some 0) <> (STR ''select'', [i]) = (Some 1, [], <1 $:= Some i, 2 $:= Some (Num 0)>)"
apply (rule ltl_step_some[of _ _ _ _ _ _ select])
apply (simp add: possible_steps_0)
apply (simp add: select_def)
by (simp add: select_def finfun_update_twist apply_updates_def)
lemma ltl_step_not_coin_or_vend:
"\<nexists>i. e = (STR ''coin'', [i]) \<Longrightarrow>
e \<noteq> (STR ''vend'', []) \<Longrightarrow>
ltl_step drinks (Some 1) r e = (None, [], r)"
apply (cases e)
subgoal for a b
apply (simp del: ltl_step.simps)
apply (rule ltl_step_none)
apply (simp add: possible_steps_empty drinks_def can_take_transition_def can_take_def transitions)
by (case_tac e, case_tac b, auto)
done
lemma ltl_step_coin:
"\<exists>p r'. ltl_step drinks (Some 1) r (STR ''coin'', [i]) = (Some 1, p, r')"
by (simp add: possible_steps_1_coin)
lemma alw_tl:
"alw \<phi> (make_full_observation e (Some 0) <> [] xs) \<Longrightarrow>
alw \<phi>
(make_full_observation e (fst (ltl_step e (Some 0) <> (shd xs))) (snd (snd (ltl_step e (Some 0) <> (shd xs))))
(fst (snd (ltl_step e (Some 0) <> (shd xs)))) (stl xs))"
by auto
lemma stop_at_none:
"alw (\<lambda>xs. output (shd (stl xs)) = [Some (EFSM.Str drink)] \<longrightarrow> check_exp (Ge (V (Rg 2)) (L (Num 100))) xs)
(make_full_observation drinks None r p t)"
apply (rule alw_mono[of "nxt (output_eq [])"])
apply (simp add: no_output_none_nxt)
by simp
lemma drink_costs_money_aux:
assumes "\<exists>p r t. j = make_full_observation drinks (Some 1) r p t"
shows "alw (\<lambda>xs. output (shd (stl xs)) = [Some (EFSM.Str drink)] \<longrightarrow> check_exp (Ge (V (Rg 2)) (L (Num 100))) xs) j"
using assms apply coinduct
apply clarsimp
apply (case_tac "shd t")
apply (simp del: ltl_step.simps)
apply (rule P_ltl_step_1)
apply simp
apply (rule disjI2)
apply (rule alw_mono[of "nxt (output_eq [])"])
apply (simp add: no_output_none_nxt)
apply simp
apply (simp add: Str_def value_plus_never_string)
apply auto[1]
apply auto[1]
apply simp
apply standard
apply (rule disjI2)
apply (rule alw_mono[of "nxt (output_eq [])"])
apply (simp add: drinks_step_2_none no_output_none_if_empty nxt_alw)
by simp
lemma LTL_drinks_cost_money:
"alw (nxt (output_eq [Some (Str drink)]) impl (check_exp (Ge (V (Rg 2)) (L (Num 100))))) (watch drinks t)"
proof(coinduction)
case alw
then show ?case
apply (case_tac "shd t")
apply (simp del: ltl_step.simps)
apply (rule P_ltl_step_0)
apply simp
apply (rule disjI2)
apply (rule alw_mono[of "nxt (output_eq [])"])
apply (simp add: no_output_none_nxt)
apply simp
apply simp
using drink_costs_money_aux
apply simp
by blast
qed
lemma steps_1_invalid:
"\<nexists>i. (a, b) = (STR ''coin'', [i]) \<Longrightarrow>
\<nexists>i. (a, b) = (STR ''vend'', []) \<Longrightarrow>
possible_steps drinks 1 r a b = {||}"
apply (simp add: possible_steps_empty drinks_def transitions can_take_transition_def can_take_def)
by (induct b, auto)
lemma output_vend_aux:
assumes "\<exists>p r t. j = make_full_observation drinks (Some 1) r p t"
shows "alw (\<lambda>xs. label_eq ''vend'' xs \<and> output (shd (stl xs)) = [Some d] \<longrightarrow> check_exp (Ge (V (Rg 2)) (L (Num 100))) xs) j"
using assms apply coinduct
apply clarsimp
subgoal for r t
apply (case_tac "shd t")
apply (simp add: implode_vend del: ltl_step.simps)
apply (rule P_ltl_step_1)
apply simp
apply (rule disjI2)
apply (rule alw_mono[of "nxt (output_eq [])"])
apply (simp add: no_output_none_nxt)
apply simp
apply auto[1]
apply auto[1]
apply simp
apply standard
apply (rule disjI2)
apply (rule alw_mono[of "nxt (output_eq [])"])
apply (simp add: drinks_step_2_none no_output_none_if_empty nxt_alw)
by simp
done
text_raw\<open>\snip{outputVend}{1}{2}{%\<close>
lemma LTL_output_vend:
"alw (((label_eq ''vend'') aand (nxt (output_eq [Some d]))) impl
(check_exp (Ge (V (Rg 2)) (L (Num 100))))) (watch drinks t)"
text_raw\<open>}%endsnip\<close>
proof(coinduction)
case alw
then show ?case
apply (simp add: implode_vend)
apply (case_tac "shd t")
apply (simp del: ltl_step.simps)
apply (rule P_ltl_step_0)
apply simp
apply (rule disjI2)
apply (rule alw_mono[of "nxt (output_eq [])"])
apply (simp add: no_output_none_nxt)
apply simp
apply simp
subgoal for a b
using output_vend_aux[of "(make_full_observation drinks (Some 1)
<1 $:= Some (hd b), 2 $:= Some (Num 0)> [] (stl t))" d]
using implode_vend by auto
done
qed
text_raw\<open>\snip{outputVendUnfolded}{1}{2}{%\<close>
lemma LTL_output_vend_unfolded:
"alw (\<lambda>xs. (label (shd xs) = STR ''vend'' \<and>
nxt (\<lambda>s. output (shd s) = [Some d]) xs) \<longrightarrow>
\<not>? value_gt (Some (Num 100)) (datastate (shd xs) $ 2) = trilean.true)
(watch drinks t)"
text_raw\<open>}%endsnip\<close>
apply (insert LTL_output_vend[of d t])
by (simp add: implode_vend)
end