oheimb@4530: (* Title: HOL/IOA/IOA.thy mueller@3078: ID: $Id$ mueller@3078: Author: Tobias Nipkow & Konrad Slind mueller@3078: Copyright 1994 TU Muenchen mueller@3078: *) mueller@3078: wenzelm@17288: header {* The I/O automata of Lynch and Tuttle *} wenzelm@17288: wenzelm@17288: theory IOA wenzelm@17288: imports Asig wenzelm@17288: begin mueller@3078: mueller@3078: types mueller@3078: 'a seq = "nat => 'a" mueller@3078: 'a oseq = "nat => 'a option" mueller@3078: ('a,'b)execution = "'a oseq * 'b seq" mueller@3078: ('a,'s)transition = "('s * 'a * 's)" mueller@3078: ('a,'s)ioa = "'a signature * 's set * ('a,'s)transition set" mueller@3078: mueller@3078: consts mueller@3078: mueller@3078: (* IO automata *) mueller@3078: state_trans::"['action signature, ('action,'state)transition set] => bool" mueller@3078: asig_of ::"('action,'state)ioa => 'action signature" mueller@3078: starts_of ::"('action,'state)ioa => 'state set" mueller@3078: trans_of ::"('action,'state)ioa => ('action,'state)transition set" mueller@3078: IOA ::"('action,'state)ioa => bool" mueller@3078: mueller@3078: (* Executions, schedules, and traces *) mueller@3078: wenzelm@17288: is_execution_fragment ::"[('action,'state)ioa, ('action,'state)execution] => bool" mueller@3078: has_execution ::"[('action,'state)ioa, ('action,'state)execution] => bool" mueller@3078: executions :: "('action,'state)ioa => ('action,'state)execution set" mueller@3078: mk_trace :: "[('action,'state)ioa, 'action oseq] => 'action oseq" mueller@3078: reachable :: "[('action,'state)ioa, 'state] => bool" mueller@3078: invariant :: "[('action,'state)ioa, 'state=>bool] => bool" mueller@3078: has_trace :: "[('action,'state)ioa, 'action oseq] => bool" mueller@3078: traces :: "('action,'state)ioa => 'action oseq set" mueller@3078: NF :: "'a oseq => 'a oseq" mueller@3078: mueller@3078: (* Composition of action signatures and automata *) mueller@3078: compatible_asigs ::"('a => 'action signature) => bool" mueller@3078: asig_composition ::"('a => 'action signature) => 'action signature" mueller@3078: compatible_ioas ::"('a => ('action,'state)ioa) => bool" mueller@3078: ioa_composition ::"('a => ('action, 'state)ioa) =>('action,'a => 'state)ioa" mueller@3078: mueller@3078: (* binary composition of action signatures and automata *) mueller@3078: compat_asigs ::"['action signature, 'action signature] => bool" mueller@3078: asig_comp ::"['action signature, 'action signature] => 'action signature" mueller@3078: compat_ioas ::"[('action,'s)ioa, ('action,'t)ioa] => bool" wenzelm@17288: par ::"[('a,'s)ioa, ('a,'t)ioa] => ('a,'s*'t)ioa" (infixr "||" 10) mueller@3078: mueller@3078: (* Filtering and hiding *) mueller@3078: filter_oseq :: "('a => bool) => 'a oseq => 'a oseq" mueller@3078: mueller@3078: restrict_asig :: "['a signature, 'a set] => 'a signature" mueller@3078: restrict :: "[('a,'s)ioa, 'a set] => ('a,'s)ioa" mueller@3078: mueller@3078: (* Notions of correctness *) mueller@3078: ioa_implements :: "[('action,'state1)ioa, ('action,'state2)ioa] => bool" mueller@3078: mueller@3078: (* Instantiation of abstract IOA by concrete actions *) mueller@3078: rename:: "('a, 'b)ioa => ('c => 'a option) => ('c,'b)ioa" mueller@3078: mueller@3078: defs mueller@3078: wenzelm@17288: state_trans_def: wenzelm@17288: "state_trans asig R == wenzelm@17288: (!triple. triple:R --> fst(snd(triple)):actions(asig)) & mueller@3078: (!a. (a:inputs(asig)) --> (!s1. ? s2. (s1,a,s2):R))" mueller@3078: mueller@3078: wenzelm@17288: asig_of_def: "asig_of == fst" wenzelm@17288: starts_of_def: "starts_of == (fst o snd)" wenzelm@17288: trans_of_def: "trans_of == (snd o snd)" mueller@3078: wenzelm@17288: ioa_def: wenzelm@17288: "IOA(ioa) == (is_asig(asig_of(ioa)) & wenzelm@17288: (~ starts_of(ioa) = {}) & mueller@3078: state_trans (asig_of ioa) (trans_of ioa))" mueller@3078: mueller@3078: mueller@3078: (* An execution fragment is modelled with a pair of sequences: mueller@3078: * the first is the action options, the second the state sequence. mueller@3078: * Finite executions have None actions from some point on. mueller@3078: *******) wenzelm@17288: is_execution_fragment_def: wenzelm@17288: "is_execution_fragment A ex == wenzelm@17288: let act = fst(ex); state = snd(ex) wenzelm@17288: in !n a. (act(n)=None --> state(Suc(n)) = state(n)) & mueller@3078: (act(n)=Some(a) --> (state(n),a,state(Suc(n))):trans_of(A))" mueller@3078: mueller@3078: wenzelm@17288: executions_def: wenzelm@17288: "executions(ioa) == {e. snd e 0:starts_of(ioa) & mueller@3078: is_execution_fragment ioa e}" mueller@3078: wenzelm@17288: wenzelm@17288: reachable_def: mueller@3078: "reachable ioa s == (? ex:executions(ioa). ? n. (snd ex n) = s)" mueller@3078: mueller@3078: wenzelm@17288: invariant_def: "invariant A P == (!s. reachable A s --> P(s))" mueller@3078: mueller@3078: mueller@3078: (* Restrict the trace to those members of the set s *) wenzelm@17288: filter_oseq_def: wenzelm@17288: "filter_oseq p s == wenzelm@17288: (%i. case s(i) wenzelm@17288: of None => None mueller@3078: | Some(x) => if p x then Some x else None)" mueller@3078: mueller@3078: wenzelm@17288: mk_trace_def: wenzelm@3842: "mk_trace(ioa) == filter_oseq(%a. a:externals(asig_of(ioa)))" mueller@3078: mueller@3078: mueller@3078: (* Does an ioa have an execution with the given trace *) wenzelm@17288: has_trace_def: wenzelm@17288: "has_trace ioa b == mueller@3078: (? ex:executions(ioa). b = mk_trace ioa (fst ex))" mueller@3078: wenzelm@17288: normal_form_def: wenzelm@17288: "NF(tr) == @nf. ? f. mono(f) & (!i. nf(i)=tr(f(i))) & wenzelm@17288: (!j. j ~: range(f) --> nf(j)= None) & mueller@3078: (!i. nf(i)=None --> (nf (Suc i)) = None) " wenzelm@17288: mueller@3078: (* All the traces of an ioa *) mueller@3078: wenzelm@17288: traces_def: mueller@3078: "traces(ioa) == {trace. ? tr. trace=NF(tr) & has_trace ioa tr}" mueller@3078: mueller@3078: (* wenzelm@17288: traces_def: mueller@3078: "traces(ioa) == {tr. has_trace ioa tr}" mueller@3078: *) wenzelm@17288: wenzelm@17288: compat_asigs_def: wenzelm@17288: "compat_asigs a1 a2 == wenzelm@17288: (((outputs(a1) Int outputs(a2)) = {}) & wenzelm@17288: ((internals(a1) Int actions(a2)) = {}) & mueller@3078: ((internals(a2) Int actions(a1)) = {}))" mueller@3078: mueller@3078: wenzelm@17288: compat_ioas_def: mueller@3078: "compat_ioas ioa1 ioa2 == compat_asigs (asig_of(ioa1)) (asig_of(ioa2))" mueller@3078: mueller@3078: wenzelm@17288: asig_comp_def: wenzelm@17288: "asig_comp a1 a2 == wenzelm@17288: (((inputs(a1) Un inputs(a2)) - (outputs(a1) Un outputs(a2)), wenzelm@17288: (outputs(a1) Un outputs(a2)), mueller@3078: (internals(a1) Un internals(a2))))" mueller@3078: mueller@3078: wenzelm@17288: par_def: wenzelm@17288: "(ioa1 || ioa2) == wenzelm@17288: (asig_comp (asig_of ioa1) (asig_of ioa2), wenzelm@17288: {pr. fst(pr):starts_of(ioa1) & snd(pr):starts_of(ioa2)}, wenzelm@17288: {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr)) wenzelm@17288: in (a:actions(asig_of(ioa1)) | a:actions(asig_of(ioa2))) & wenzelm@17288: (if a:actions(asig_of(ioa1)) then wenzelm@17288: (fst(s),a,fst(t)):trans_of(ioa1) wenzelm@17288: else fst(t) = fst(s)) wenzelm@17288: & wenzelm@17288: (if a:actions(asig_of(ioa2)) then wenzelm@17288: (snd(s),a,snd(t)):trans_of(ioa2) mueller@3078: else snd(t) = snd(s))})" mueller@3078: mueller@3078: wenzelm@17288: restrict_asig_def: wenzelm@17288: "restrict_asig asig actns == wenzelm@17288: (inputs(asig) Int actns, outputs(asig) Int actns, mueller@3078: internals(asig) Un (externals(asig) - actns))" mueller@3078: mueller@3078: wenzelm@17288: restrict_def: wenzelm@17288: "restrict ioa actns == mueller@3078: (restrict_asig (asig_of ioa) actns, starts_of(ioa), trans_of(ioa))" mueller@3078: mueller@3078: wenzelm@17288: ioa_implements_def: wenzelm@17288: "ioa_implements ioa1 ioa2 == wenzelm@17288: ((inputs(asig_of(ioa1)) = inputs(asig_of(ioa2))) & wenzelm@17288: (outputs(asig_of(ioa1)) = outputs(asig_of(ioa2))) & mueller@3078: traces(ioa1) <= traces(ioa2))" wenzelm@17288: wenzelm@17288: rename_def: wenzelm@17288: "rename ioa ren == wenzelm@17288: (({b. ? x. Some(x)= ren(b) & x : inputs(asig_of(ioa))}, wenzelm@17288: {b. ? x. Some(x)= ren(b) & x : outputs(asig_of(ioa))}, wenzelm@17288: {b. ? x. Some(x)= ren(b) & x : internals(asig_of(ioa))}), wenzelm@17288: starts_of(ioa) , wenzelm@17288: {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr)) wenzelm@17288: in mueller@3078: ? x. Some(x) = ren(a) & (s,x,t):trans_of(ioa)})" mueller@3078: wenzelm@19801: wenzelm@19801: declare Let_def [simp] wenzelm@19801: wenzelm@19801: lemmas ioa_projections = asig_of_def starts_of_def trans_of_def wenzelm@19801: and exec_rws = executions_def is_execution_fragment_def wenzelm@19801: wenzelm@19801: lemma ioa_triple_proj: wenzelm@19801: "asig_of(x,y,z) = x & starts_of(x,y,z) = y & trans_of(x,y,z) = z" wenzelm@19801: apply (simp add: ioa_projections) wenzelm@19801: done wenzelm@19801: wenzelm@19801: lemma trans_in_actions: wenzelm@19801: "[| IOA(A); (s1,a,s2):trans_of(A) |] ==> a:actions(asig_of(A))" wenzelm@19801: apply (unfold ioa_def state_trans_def actions_def is_asig_def) wenzelm@19801: apply (erule conjE)+ wenzelm@19801: apply (erule allE, erule impE, assumption) wenzelm@19801: apply simp wenzelm@19801: done wenzelm@19801: wenzelm@19801: wenzelm@19801: lemma filter_oseq_idemp: "filter_oseq p (filter_oseq p s) = filter_oseq p s" wenzelm@19801: apply (simp add: filter_oseq_def) wenzelm@19801: apply (rule ext) wenzelm@19801: apply (case_tac "s i") wenzelm@19801: apply simp_all wenzelm@19801: done wenzelm@19801: wenzelm@19801: lemma mk_trace_thm: wenzelm@19801: "(mk_trace A s n = None) = wenzelm@19801: (s(n)=None | (? a. s(n)=Some(a) & a ~: externals(asig_of(A)))) wenzelm@19801: & wenzelm@19801: (mk_trace A s n = Some(a)) = wenzelm@19801: (s(n)=Some(a) & a : externals(asig_of(A)))" wenzelm@19801: apply (unfold mk_trace_def filter_oseq_def) wenzelm@19801: apply (case_tac "s n") wenzelm@19801: apply auto wenzelm@19801: done wenzelm@19801: wenzelm@19801: lemma reachable_0: "s:starts_of(A) ==> reachable A s" wenzelm@19801: apply (unfold reachable_def) wenzelm@19801: apply (rule_tac x = "(%i. None, %i. s)" in bexI) wenzelm@19801: apply simp wenzelm@19801: apply (simp add: exec_rws) wenzelm@19801: done wenzelm@19801: wenzelm@19801: lemma reachable_n: wenzelm@19801: "!!A. [| reachable A s; (s,a,t) : trans_of(A) |] ==> reachable A t" wenzelm@19801: apply (unfold reachable_def exec_rws) wenzelm@19801: apply (simp del: bex_simps) wenzelm@19801: apply (simp (no_asm_simp) only: split_tupled_all) wenzelm@19801: apply safe wenzelm@19801: apply (rename_tac ex1 ex2 n) wenzelm@19801: apply (rule_tac x = "(%i. if i P(s)" wenzelm@19801: and p2: "!!s t a. [|reachable A s; P(s)|] ==> (s,a,t): trans_of(A) --> P(t)" wenzelm@19801: shows "invariant A P" wenzelm@19801: apply (unfold invariant_def reachable_def Let_def exec_rws) wenzelm@19801: apply safe wenzelm@19801: apply (rename_tac ex1 ex2 n) wenzelm@19801: apply (rule_tac Q = "reachable A (ex2 n) " in conjunct1) wenzelm@19801: apply simp wenzelm@19801: apply (induct_tac n) wenzelm@19801: apply (fast intro: p1 reachable_0) wenzelm@19801: apply (erule_tac x = na in allE) wenzelm@19801: apply (case_tac "ex1 na", simp_all) wenzelm@19801: apply safe wenzelm@19801: apply (erule p2 [THEN mp]) wenzelm@19801: apply (fast dest: reachable_n)+ wenzelm@19801: done wenzelm@19801: wenzelm@19801: lemma invariantI1: wenzelm@19801: "[| !!s. s : starts_of(A) ==> P(s); wenzelm@19801: !!s t a. reachable A s ==> P(s) --> (s,a,t):trans_of(A) --> P(t) wenzelm@19801: |] ==> invariant A P" wenzelm@19801: apply (blast intro!: invariantI) wenzelm@19801: done wenzelm@19801: wenzelm@19801: lemma invariantE: wenzelm@19801: "[| invariant A P; reachable A s |] ==> P(s)" wenzelm@19801: apply (unfold invariant_def) wenzelm@19801: apply blast wenzelm@19801: done wenzelm@19801: wenzelm@19801: lemma actions_asig_comp: wenzelm@19801: "actions(asig_comp a b) = actions(a) Un actions(b)" wenzelm@19801: apply (auto simp add: actions_def asig_comp_def asig_projections) wenzelm@19801: done wenzelm@19801: wenzelm@19801: lemma starts_of_par: wenzelm@19801: "starts_of(A || B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}" wenzelm@19801: apply (simp add: par_def ioa_projections) wenzelm@19801: done wenzelm@19801: wenzelm@19801: (* Every state in an execution is reachable *) wenzelm@19801: lemma states_of_exec_reachable: wenzelm@19801: "ex:executions(A) ==> !n. reachable A (snd ex n)" wenzelm@19801: apply (unfold reachable_def) wenzelm@19801: apply fast wenzelm@19801: done wenzelm@19801: wenzelm@19801: wenzelm@19801: lemma trans_of_par4: wenzelm@19801: "(s,a,t) : trans_of(A || B || C || D) = wenzelm@19801: ((a:actions(asig_of(A)) | a:actions(asig_of(B)) | a:actions(asig_of(C)) | wenzelm@19801: a:actions(asig_of(D))) & wenzelm@19801: (if a:actions(asig_of(A)) then (fst(s),a,fst(t)):trans_of(A) wenzelm@19801: else fst t=fst s) & wenzelm@19801: (if a:actions(asig_of(B)) then (fst(snd(s)),a,fst(snd(t))):trans_of(B) wenzelm@19801: else fst(snd(t))=fst(snd(s))) & wenzelm@19801: (if a:actions(asig_of(C)) then wenzelm@19801: (fst(snd(snd(s))),a,fst(snd(snd(t)))):trans_of(C) wenzelm@19801: else fst(snd(snd(t)))=fst(snd(snd(s)))) & wenzelm@19801: (if a:actions(asig_of(D)) then wenzelm@19801: (snd(snd(snd(s))),a,snd(snd(snd(t)))):trans_of(D) wenzelm@19801: else snd(snd(snd(t)))=snd(snd(snd(s)))))" wenzelm@19801: (*SLOW*) wenzelm@19801: apply (simp (no_asm) add: par_def actions_asig_comp Pair_fst_snd_eq ioa_projections) wenzelm@19801: done wenzelm@19801: wenzelm@19801: lemma cancel_restrict: "starts_of(restrict ioa acts) = starts_of(ioa) & wenzelm@19801: trans_of(restrict ioa acts) = trans_of(ioa) & wenzelm@19801: reachable (restrict ioa acts) s = reachable ioa s" wenzelm@19801: apply (simp add: is_execution_fragment_def executions_def wenzelm@19801: reachable_def restrict_def ioa_projections) wenzelm@19801: done wenzelm@19801: wenzelm@19801: lemma asig_of_par: "asig_of(A || B) = asig_comp (asig_of A) (asig_of B)" wenzelm@19801: apply (simp add: par_def ioa_projections) wenzelm@19801: done wenzelm@19801: wenzelm@19801: wenzelm@19801: lemma externals_of_par: "externals(asig_of(A1||A2)) = wenzelm@19801: (externals(asig_of(A1)) Un externals(asig_of(A2)))" wenzelm@19801: apply (simp add: externals_def asig_of_par asig_comp_def berghofe@26806: asig_inputs_def asig_outputs_def Un_def set_diff_eq) wenzelm@19801: apply blast wenzelm@19801: done wenzelm@19801: wenzelm@19801: lemma ext1_is_not_int2: wenzelm@19801: "[| compat_ioas A1 A2; a:externals(asig_of(A1))|] ==> a~:internals(asig_of(A2))" wenzelm@19801: apply (unfold externals_def actions_def compat_ioas_def compat_asigs_def) wenzelm@19801: apply auto wenzelm@19801: done wenzelm@19801: wenzelm@19801: lemma ext2_is_not_int1: wenzelm@19801: "[| compat_ioas A2 A1 ; a:externals(asig_of(A1))|] ==> a~:internals(asig_of(A2))" wenzelm@19801: apply (unfold externals_def actions_def compat_ioas_def compat_asigs_def) wenzelm@19801: apply auto wenzelm@19801: done wenzelm@19801: wenzelm@19801: lemmas ext1_ext2_is_not_act2 = ext1_is_not_int2 [THEN int_and_ext_is_act] wenzelm@19801: and ext1_ext2_is_not_act1 = ext2_is_not_int1 [THEN int_and_ext_is_act] wenzelm@17288: wenzelm@17288: end