168
|
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 |
|
156
|
9 |
open IOA Asig;
|
|
10 |
|
168
|
11 |
val ioa_projections = [asig_of_def, starts_of_def, trans_of_def];
|
|
12 |
|
156
|
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";
|
168
|
17 |
by (simp_tac (SS addsimps ioa_projections) 1);
|
171
|
18 |
qed "ioa_triple_proj";
|
156
|
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);
|
171
|
25 |
qed "trans_in_actions";
|
156
|
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);
|
171
|
34 |
qed "filter_oseq_idemp";
|
156
|
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);
|
171
|
45 |
qed "mk_behaviour_thm";
|
156
|
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)>::('action,'a)execution")] bexI 1);
|
|
49 |
by (simp_tac SS 1);
|
|
50 |
by (asm_simp_tac (SS addsimps exec_rws) 1);
|
171
|
51 |
qed "reachable_0";
|
156
|
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,fst(ex,i),if(i=n,Some(a),None)), \
|
|
58 |
\ %i.if(i<Suc(n),snd(ex,i),t)>")] bexI 1);
|
|
59 |
by(res_inst_tac [("x","Suc(n)")] exI 1);
|
|
60 |
by(simp_tac SS 1);
|
|
61 |
by(asm_simp_tac (SS delsimps [less_Suc_eq]) 1);
|
|
62 |
by(REPEAT(rtac allI 1));
|
|
63 |
by(res_inst_tac [("m","na"),("n","n")] (make_elim less_linear) 1);
|
|
64 |
be disjE 1;
|
|
65 |
by(asm_simp_tac SS 1);
|
|
66 |
be disjE 1;
|
|
67 |
by(asm_simp_tac SS 1);
|
|
68 |
by(fast_tac HOL_cs 1);
|
|
69 |
by(forward_tac [less_not_sym] 1);
|
|
70 |
by(asm_simp_tac (SS addsimps [less_not_refl2]) 1);
|
171
|
71 |
qed "reachable_n";
|
156
|
72 |
|
|
73 |
val [p1,p2] = goalw IOA.thy [invariant_def]
|
|
74 |
"[| !!s. s:starts_of(A) ==> P(s); \
|
|
75 |
\ !!s t a. [|reachable(A,s); P(s)|] ==> <s,a,t>: trans_of(A) --> P(t) |] \
|
|
76 |
\ ==> invariant(A,P)";
|
|
77 |
by (rewrite_goals_tac(reachable_def::Let_def::exec_rws));
|
|
78 |
by (safe_tac set_cs);
|
|
79 |
by (res_inst_tac [("Q","reachable(A,snd(ex,n))")] conjunct1 1);
|
|
80 |
by (nat_ind_tac "n" 1);
|
|
81 |
by (fast_tac (set_cs addIs [p1,reachable_0]) 1);
|
|
82 |
by (eres_inst_tac[("x","n1")]allE 1);
|
|
83 |
by (eres_inst_tac[("P","%x.!a::'action.?Q(x,a)"),
|
|
84 |
("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])));
|
171
|
90 |
qed "invariantI";
|
156
|
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);
|
171
|
97 |
qed "invariantI1";
|
156
|
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;
|
171
|
102 |
qed "invariantE";
|
156
|
103 |
|
|
104 |
goal IOA.thy
|
|
105 |
"actions(asig_comp(a,b)) = actions(a) Un actions(b)";
|
214
|
106 |
by(simp_tac (prod_ss addsimps
|
168
|
107 |
([actions_def,asig_comp_def]@asig_projections)) 1);
|
156
|
108 |
by(fast_tac eq_cs 1);
|
171
|
109 |
qed "actions_asig_comp";
|
156
|
110 |
|
|
111 |
goal IOA.thy
|
|
112 |
"starts_of(A || B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}";
|
168
|
113 |
by(simp_tac (SS addsimps (par_def::ioa_projections)) 1);
|
171
|
114 |
qed "starts_of_par";
|
156
|
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);
|
171
|
120 |
qed "states_of_exec_reachable";
|
156
|
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)), <fst(s),a,fst(t)>:trans_of(A), fst(t)=fst(s)) & \
|
|
128 |
\ if(a:actions(asig_of(B)), <fst(snd(s)),a,fst(snd(t))>:trans_of(B), \
|
|
129 |
\ fst(snd(t))=fst(snd(s))) & \
|
|
130 |
\ if(a:actions(asig_of(C)), \
|
|
131 |
\ <fst(snd(snd(s))),a,fst(snd(snd(t)))>:trans_of(C), \
|
|
132 |
\ fst(snd(snd(t)))=fst(snd(snd(s)))) & \
|
|
133 |
\ if(a:actions(asig_of(D)), \
|
|
134 |
\ <snd(snd(snd(s))),a,snd(snd(snd(t)))>:trans_of(D), \
|
|
135 |
\ snd(snd(snd(t)))=snd(snd(snd(s)))))";
|
168
|
136 |
by(simp_tac (SS addsimps ([par_def,actions_asig_comp,Pair_fst_snd_eq]@
|
|
137 |
ioa_projections)
|
156
|
138 |
setloop (split_tac [expand_if])) 1);
|
171
|
139 |
qed "trans_of_par4";
|
156
|
140 |
|
|
141 |
goal IOA.thy "starts_of(restrict(ioa,acts)) = starts_of(ioa) & \
|
|
142 |
\ trans_of(restrict(ioa,acts)) = trans_of(ioa) & \
|
|
143 |
\ reachable(restrict(ioa,acts),s) = reachable(ioa,s)";
|
168
|
144 |
by(simp_tac (SS addsimps ([is_execution_fragment_def,executions_def,
|
|
145 |
reachable_def,restrict_def]@ioa_projections)) 1);
|
171
|
146 |
qed "cancel_restrict";
|
156
|
147 |
|
|
148 |
goal IOA.thy "asig_of(A || B) = asig_comp(asig_of(A),asig_of(B))";
|
168
|
149 |
by(simp_tac (SS addsimps (par_def::ioa_projections)) 1);
|
171
|
150 |
qed "asig_of_par";
|