src/HOL/IOA/meta_theory/IOA.ML
changeset 2056 93c093620c28
parent 2018 bcd69cc47cf0
child 2513 d708d8cdc8e8
equal deleted inserted replaced
2055:cc274e47f607 2056:93c093620c28
    52   by (asm_simp_tac (!simpset addsimps exec_rws) 1);
    52   by (asm_simp_tac (!simpset addsimps exec_rws) 1);
    53 qed "reachable_0";
    53 qed "reachable_0";
    54 
    54 
    55 goalw IOA.thy (reachable_def::exec_rws)
    55 goalw IOA.thy (reachable_def::exec_rws)
    56 "!!A. [| reachable A s; (s,a,t) : trans_of(A) |] ==> reachable A t";
    56 "!!A. [| reachable A s; (s,a,t) : trans_of(A) |] ==> reachable A t";
    57   by(Asm_full_simp_tac 1);
    57   by (Asm_full_simp_tac 1);
    58   by(safe_tac (!claset));
    58   by (safe_tac (!claset));
    59   by(res_inst_tac [("x","(%i.if i<n then fst ex i                    \
    59   by (res_inst_tac [("x","(%i.if i<n then fst ex i                    \
    60 \                            else (if i=n then Some a else None),    \
    60 \                            else (if i=n then Some a else None),    \
    61 \                         %i.if i<Suc n then snd ex i else t)")] bexI 1);
    61 \                         %i.if i<Suc n then snd ex i else t)")] bexI 1);
    62   by(res_inst_tac [("x","Suc(n)")] exI 1);
    62   by (res_inst_tac [("x","Suc(n)")] exI 1);
    63   by(Simp_tac 1);
    63   by (Simp_tac 1);
    64   by(Asm_simp_tac 1);
    64   by (Asm_simp_tac 1);
    65   by(REPEAT(rtac allI 1));
    65   by (REPEAT(rtac allI 1));
    66   by(res_inst_tac [("m","na"),("n","n")] (make_elim less_linear) 1);
    66   by (res_inst_tac [("m","na"),("n","n")] (make_elim less_linear) 1);
    67   be disjE 1;
    67   by (etac disjE 1);
    68   by(asm_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
    68   by (asm_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
    69   be disjE 1;
    69   by (etac disjE 1);
    70   by(Asm_simp_tac 1);
    70   by (Asm_simp_tac 1);
    71   by(Fast_tac 1);
    71   by (forward_tac [less_not_sym] 1);
    72   by(forward_tac [less_not_sym] 1);
    72   by (asm_simp_tac (!simpset addsimps [less_not_refl2,less_Suc_eq]) 1);
    73   by(asm_simp_tac (!simpset addsimps [less_not_refl2,less_Suc_eq]) 1);
       
    74 qed "reachable_n";
    73 qed "reachable_n";
    75 
    74 
    76 val [p1,p2] = goalw IOA.thy [invariant_def]
    75 val [p1,p2] = goalw IOA.thy [invariant_def]
    77   "[| !!s. s:starts_of(A) ==> P(s);                                          \
    76   "[| !!s. s:starts_of(A) ==> P(s);                                          \
    78 \     !!s t a. [|reachable A s; P(s)|] ==> (s,a,t): trans_of(A) --> P(t) |] \
    77 \     !!s t a. [|reachable A s; P(s)|] ==> (s,a,t): trans_of(A) --> P(t) |] \
    98   by (fast_tac (!claset addSIs [invariantI] addSDs [p1,p2]) 1);
    97   by (fast_tac (!claset addSIs [invariantI] addSDs [p1,p2]) 1);
    99 qed "invariantI1";
    98 qed "invariantI1";
   100 
    99 
   101 val [p1,p2] = goalw IOA.thy [invariant_def]
   100 val [p1,p2] = goalw IOA.thy [invariant_def]
   102 "[| invariant A P; reachable A s |] ==> P(s)";
   101 "[| invariant A P; reachable A s |] ==> P(s)";
   103    br(p2 RS (p1 RS spec RS mp))1;
   102    by (rtac (p2 RS (p1 RS spec RS mp)) 1);
   104 qed "invariantE";
   103 qed "invariantE";
   105 
   104 
   106 goal IOA.thy 
   105 goal IOA.thy 
   107 "actions(asig_comp a b) = actions(a) Un actions(b)";
   106 "actions(asig_comp a b) = actions(a) Un actions(b)";
   108   by(simp_tac (!simpset addsimps
   107   by (simp_tac (!simpset addsimps
   109                ([actions_def,asig_comp_def]@asig_projections)) 1);
   108                ([actions_def,asig_comp_def]@asig_projections)) 1);
   110   by(Fast_tac 1);
   109   by (Fast_tac 1);
   111 qed "actions_asig_comp";
   110 qed "actions_asig_comp";
   112 
   111 
   113 goal IOA.thy
   112 goal IOA.thy
   114 "starts_of(A || B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}";
   113 "starts_of(A || B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}";
   115   by(simp_tac (!simpset addsimps (par_def::ioa_projections)) 1);
   114   by (simp_tac (!simpset addsimps (par_def::ioa_projections)) 1);
   116 qed "starts_of_par";
   115 qed "starts_of_par";
   117 
   116 
   118 (* Every state in an execution is reachable *)
   117 (* Every state in an execution is reachable *)
   119 goalw IOA.thy [reachable_def] 
   118 goalw IOA.thy [reachable_def] 
   120 "!!A. ex:executions(A) ==> !n. reachable A (snd ex n)";
   119 "!!A. ex:executions(A) ==> !n. reachable A (snd ex n)";
   134 \     (fst(snd(snd(s))),a,fst(snd(snd(t)))):trans_of(C)                      \
   133 \     (fst(snd(snd(s))),a,fst(snd(snd(t)))):trans_of(C)                      \
   135 \   else fst(snd(snd(t)))=fst(snd(snd(s)))) &                                \
   134 \   else fst(snd(snd(t)))=fst(snd(snd(s)))) &                                \
   136 \  (if a:actions(asig_of(D)) then                                            \
   135 \  (if a:actions(asig_of(D)) then                                            \
   137 \     (snd(snd(snd(s))),a,snd(snd(snd(t)))):trans_of(D)                      \
   136 \     (snd(snd(snd(s))),a,snd(snd(snd(t)))):trans_of(D)                      \
   138 \   else snd(snd(snd(t)))=snd(snd(snd(s)))))";
   137 \   else snd(snd(snd(t)))=snd(snd(snd(s)))))";
   139   by(simp_tac (!simpset addsimps ([par_def,actions_asig_comp,Pair_fst_snd_eq]@
   138   by (simp_tac (!simpset addsimps ([par_def,actions_asig_comp,Pair_fst_snd_eq]@
   140                             ioa_projections)
   139                             ioa_projections)
   141                   setloop (split_tac [expand_if])) 1);
   140                   setloop (split_tac [expand_if])) 1);
   142 qed "trans_of_par4";
   141 qed "trans_of_par4";
   143 
   142 
   144 goal IOA.thy "starts_of(restrict ioa acts) = starts_of(ioa) &     \
   143 goal IOA.thy "starts_of(restrict ioa acts) = starts_of(ioa) &     \
   145 \             trans_of(restrict ioa acts) = trans_of(ioa) &       \
   144 \             trans_of(restrict ioa acts) = trans_of(ioa) &       \
   146 \             reachable (restrict ioa acts) s = reachable ioa s";
   145 \             reachable (restrict ioa acts) s = reachable ioa s";
   147 by(simp_tac (!simpset addsimps ([is_execution_fragment_def,executions_def,
   146 by (simp_tac (!simpset addsimps ([is_execution_fragment_def,executions_def,
   148                            reachable_def,restrict_def]@ioa_projections)) 1);
   147                            reachable_def,restrict_def]@ioa_projections)) 1);
   149 qed "cancel_restrict";
   148 qed "cancel_restrict";
   150 
   149 
   151 goal IOA.thy "asig_of(A || B) = asig_comp (asig_of A) (asig_of B)";
   150 goal IOA.thy "asig_of(A || B) = asig_comp (asig_of A) (asig_of B)";
   152   by(simp_tac (!simpset addsimps (par_def::ioa_projections)) 1);
   151   by (simp_tac (!simpset addsimps (par_def::ioa_projections)) 1);
   153 qed "asig_of_par";
   152 qed "asig_of_par";
   154 
   153 
   155 
   154 
   156 goal IOA.thy "externals(asig_of(A1||A2)) =    \
   155 goal IOA.thy "externals(asig_of(A1||A2)) =    \
   157 \  (externals(asig_of(A1)) Un externals(asig_of(A2)))";
   156 \  (externals(asig_of(A1)) Un externals(asig_of(A2)))";