src/HOL/IOA/meta_theory/IOA.ML
changeset 966 3fd66f245ad7
child 972 e61b058d58d2
equal deleted inserted replaced
965:24eef3860714 966:3fd66f245ad7
       
     1 (*  Title:      HOL/IOA/meta_theory/IOA.ML
       
     2     ID:         $Id$
       
     3     Author:     Tobias Nipkow & Konrad Slind
       
     4     Copyright   1994  TU Muenchen
       
     5 
       
     6 The I/O automata of Lynch and Tuttle.
       
     7 *)
       
     8 
       
     9 open IOA Asig;
       
    10 
       
    11 val ioa_projections = [asig_of_def, starts_of_def, trans_of_def];
       
    12 
       
    13 val exec_rws = [executions_def,is_execution_fragment_def];
       
    14 
       
    15 goal IOA.thy
       
    16 "asig_of(<x,y,z>) = x & starts_of(<x,y,z>) = y & trans_of(<x,y,z>) = z";
       
    17   by (simp_tac (SS addsimps ioa_projections) 1);
       
    18   qed "ioa_triple_proj";
       
    19 
       
    20 goalw IOA.thy [ioa_def,state_trans_def,actions_def, is_asig_def]
       
    21   "!!A. [| IOA(A); <s1,a,s2>:trans_of(A) |] ==> a:actions(asig_of(A))";
       
    22   by (REPEAT(etac conjE 1));
       
    23   by (EVERY1[etac allE, etac impE, atac]);
       
    24   by (asm_full_simp_tac SS 1);
       
    25 qed "trans_in_actions";
       
    26 
       
    27 
       
    28 goal IOA.thy "filter_oseq p (filter_oseq p s) = filter_oseq p s";
       
    29   by (simp_tac (SS addsimps [filter_oseq_def]) 1);
       
    30   by (rtac ext 1);
       
    31   by (Option.option.induct_tac "s(i)" 1);
       
    32   by (simp_tac SS 1);
       
    33   by (simp_tac (SS setloop (split_tac [expand_if])) 1);
       
    34 qed "filter_oseq_idemp";
       
    35 
       
    36 goalw IOA.thy [mk_behaviour_def,filter_oseq_def]
       
    37 "(mk_behaviour A s n = None) =                                        \
       
    38 \  (s(n)=None | (? a. s(n)=Some(a) & a ~: externals(asig_of(A))))     \
       
    39 \  &                                                                  \
       
    40 \  (mk_behaviour A s n = Some(a)) =                                   \
       
    41 \   (s(n)=Some(a) & a : externals(asig_of(A)))";
       
    42   by (Option.option.induct_tac "s(n)" 1);
       
    43   by (ALLGOALS (simp_tac (SS setloop (split_tac [expand_if]))));
       
    44   by (fast_tac HOL_cs 1);
       
    45 qed "mk_behaviour_thm";
       
    46 
       
    47 goalw IOA.thy [reachable_def] "!!A. s:starts_of(A) ==> reachable A s";
       
    48   by (res_inst_tac [("x","<%i.None,%i.s>")] bexI 1);
       
    49   by (simp_tac SS 1);
       
    50   by (asm_simp_tac (SS addsimps exec_rws) 1);
       
    51 qed "reachable_0";
       
    52 
       
    53 goalw IOA.thy (reachable_def::exec_rws)
       
    54 "!!A. [| reachable A s; <s,a,t> : trans_of(A) |] ==> reachable A t";
       
    55   by(asm_full_simp_tac SS 1);
       
    56   by(safe_tac set_cs);
       
    57   by(res_inst_tac [("x","<%i.if i<n then fst ex i                    \
       
    58 \                            else (if i=n then Some a else None),    \
       
    59 \                         %i.if i<Suc n then snd ex i else t>")] bexI 1);
       
    60   by(res_inst_tac [("x","Suc(n)")] exI 1);
       
    61   by(simp_tac SS 1);
       
    62   by(asm_simp_tac (SS delsimps [less_Suc_eq]) 1);
       
    63   by(REPEAT(rtac allI 1));
       
    64   by(res_inst_tac [("m","na"),("n","n")] (make_elim less_linear) 1);
       
    65   be disjE 1;
       
    66   by(asm_simp_tac SS 1);
       
    67   be disjE 1;
       
    68   by(asm_simp_tac SS 1);
       
    69   by(fast_tac HOL_cs 1);
       
    70   by(forward_tac [less_not_sym] 1);
       
    71   by(asm_simp_tac (SS addsimps [less_not_refl2]) 1);
       
    72 qed "reachable_n";
       
    73 
       
    74 val [p1,p2] = goalw IOA.thy [invariant_def]
       
    75   "[| !!s. s:starts_of(A) ==> P(s);                                          \
       
    76 \     !!s t a. [|reachable A s; P(s)|] ==> <s,a,t>: trans_of(A) --> P(t) |] \
       
    77 \  ==> invariant A P";
       
    78   by (rewrite_goals_tac(reachable_def::Let_def::exec_rws));
       
    79   by (safe_tac set_cs);
       
    80   by (res_inst_tac [("Q","reachable A (snd ex n)")] conjunct1 1);
       
    81   by (nat_ind_tac "n" 1);
       
    82   by (fast_tac (set_cs addIs [p1,reachable_0]) 1);
       
    83   by (eres_inst_tac[("x","n1")]allE 1);
       
    84   by (eres_inst_tac[("P","%x.!a.?Q x a"), ("opt","fst ex n1")] optE 1);
       
    85   by (asm_simp_tac HOL_ss 1);
       
    86   by (safe_tac HOL_cs);
       
    87   by (etac (p2 RS mp) 1);
       
    88   by (ALLGOALS(fast_tac(set_cs addDs [hd Option.option.inject RS iffD1,
       
    89                                       reachable_n])));
       
    90 qed "invariantI";
       
    91 
       
    92 val [p1,p2] = goal IOA.thy
       
    93  "[| !!s. s : starts_of(A) ==> P(s); \
       
    94 \   !!s t a. reachable A s ==> P(s) --> <s,a,t>:trans_of(A) --> P(t) \
       
    95 \ |] ==> invariant A P";
       
    96   by (fast_tac (HOL_cs addSIs [invariantI] addSDs [p1,p2]) 1);
       
    97 qed "invariantI1";
       
    98 
       
    99 val [p1,p2] = goalw IOA.thy [invariant_def]
       
   100 "[| invariant A P; reachable A s |] ==> P(s)";
       
   101    br(p2 RS (p1 RS spec RS mp))1;
       
   102 qed "invariantE";
       
   103 
       
   104 goal IOA.thy 
       
   105 "actions(asig_comp a b) = actions(a) Un actions(b)";
       
   106   by(simp_tac (prod_ss addsimps
       
   107                ([actions_def,asig_comp_def]@asig_projections)) 1);
       
   108   by(fast_tac eq_cs 1);
       
   109 qed "actions_asig_comp";
       
   110 
       
   111 goal IOA.thy
       
   112 "starts_of(A || B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}";
       
   113   by(simp_tac (SS addsimps (par_def::ioa_projections)) 1);
       
   114 qed "starts_of_par";
       
   115 
       
   116 (* Every state in an execution is reachable *)
       
   117 goalw IOA.thy [reachable_def] 
       
   118 "!!A. ex:executions(A) ==> !n. reachable A (snd ex n)";
       
   119   by (fast_tac set_cs 1);
       
   120 qed "states_of_exec_reachable";
       
   121 
       
   122 
       
   123 goal IOA.thy 
       
   124 "<s,a,t> : trans_of(A || B || C || D) =                                      \
       
   125 \ ((a:actions(asig_of(A)) | a:actions(asig_of(B)) | a:actions(asig_of(C)) |  \
       
   126 \   a:actions(asig_of(D))) &                                                 \
       
   127 \  (if a:actions(asig_of(A)) then <fst(s),a,fst(t)>:trans_of(A)              \
       
   128 \   else fst t=fst s) &                                                      \
       
   129 \  (if a:actions(asig_of(B)) then <fst(snd(s)),a,fst(snd(t))>:trans_of(B)    \
       
   130 \   else fst(snd(t))=fst(snd(s))) &                                          \
       
   131 \  (if a:actions(asig_of(C)) then                                            \
       
   132 \     <fst(snd(snd(s))),a,fst(snd(snd(t)))>:trans_of(C)                      \
       
   133 \   else fst(snd(snd(t)))=fst(snd(snd(s)))) &                                \
       
   134 \  (if a:actions(asig_of(D)) then                                            \
       
   135 \     <snd(snd(snd(s))),a,snd(snd(snd(t)))>:trans_of(D)                      \
       
   136 \   else snd(snd(snd(t)))=snd(snd(snd(s)))))";
       
   137   by(simp_tac (SS addsimps ([par_def,actions_asig_comp,Pair_fst_snd_eq]@
       
   138                             ioa_projections)
       
   139                   setloop (split_tac [expand_if])) 1);
       
   140 qed "trans_of_par4";
       
   141 
       
   142 goal IOA.thy "starts_of(restrict ioa acts) = starts_of(ioa) &     \
       
   143 \             trans_of(restrict ioa acts) = trans_of(ioa) &       \
       
   144 \             reachable (restrict ioa acts) s = reachable ioa s";
       
   145 by(simp_tac (SS addsimps ([is_execution_fragment_def,executions_def,
       
   146                            reachable_def,restrict_def]@ioa_projections)) 1);
       
   147 qed "cancel_restrict";
       
   148 
       
   149 goal IOA.thy "asig_of(A || B) = asig_comp (asig_of A) (asig_of B)";
       
   150   by(simp_tac (SS addsimps (par_def::ioa_projections)) 1);
       
   151 qed "asig_of_par";