author | paulson |
Tue, 16 Jul 1996 15:49:46 +0200 | |
changeset 1868 | 836950047d85 |
parent 1844 | 791e79473966 |
child 1894 | c2c8279d40f0 |
permissions | -rw-r--r-- |
966 | 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 |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
966
diff
changeset
|
16 |
"asig_of((x,y,z)) = x & starts_of((x,y,z)) = y & trans_of((x,y,z)) = z"; |
1266 | 17 |
by (simp_tac (!simpset addsimps ioa_projections) 1); |
966 | 18 |
qed "ioa_triple_proj"; |
19 |
||
20 |
goalw IOA.thy [ioa_def,state_trans_def,actions_def, is_asig_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
966
diff
changeset
|
21 |
"!!A. [| IOA(A); (s1,a,s2):trans_of(A) |] ==> a:actions(asig_of(A))"; |
966 | 22 |
by (REPEAT(etac conjE 1)); |
23 |
by (EVERY1[etac allE, etac impE, atac]); |
|
1266 | 24 |
by (Asm_full_simp_tac 1); |
966 | 25 |
qed "trans_in_actions"; |
26 |
||
27 |
||
28 |
goal IOA.thy "filter_oseq p (filter_oseq p s) = filter_oseq p s"; |
|
1266 | 29 |
by (simp_tac (!simpset addsimps [filter_oseq_def]) 1); |
966 | 30 |
by (rtac ext 1); |
31 |
by (Option.option.induct_tac "s(i)" 1); |
|
1266 | 32 |
by (Simp_tac 1); |
33 |
by (simp_tac (!simpset setloop (split_tac [expand_if])) 1); |
|
966 | 34 |
qed "filter_oseq_idemp"; |
35 |
||
1052 | 36 |
goalw IOA.thy [mk_trace_def,filter_oseq_def] |
37 |
"(mk_trace A s n = None) = \ |
|
966 | 38 |
\ (s(n)=None | (? a. s(n)=Some(a) & a ~: externals(asig_of(A)))) \ |
39 |
\ & \ |
|
1052 | 40 |
\ (mk_trace A s n = Some(a)) = \ |
966 | 41 |
\ (s(n)=Some(a) & a : externals(asig_of(A)))"; |
42 |
by (Option.option.induct_tac "s(n)" 1); |
|
1266 | 43 |
by (ALLGOALS (simp_tac (!simpset setloop (split_tac [expand_if])))); |
966 | 44 |
by (fast_tac HOL_cs 1); |
1052 | 45 |
qed "mk_trace_thm"; |
966 | 46 |
|
47 |
goalw IOA.thy [reachable_def] "!!A. s:starts_of(A) ==> reachable A s"; |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
966
diff
changeset
|
48 |
by (res_inst_tac [("x","(%i.None,%i.s)")] bexI 1); |
1266 | 49 |
by (Simp_tac 1); |
50 |
by (asm_simp_tac (!simpset addsimps exec_rws) 1); |
|
966 | 51 |
qed "reachable_0"; |
52 |
||
53 |
goalw IOA.thy (reachable_def::exec_rws) |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
966
diff
changeset
|
54 |
"!!A. [| reachable A s; (s,a,t) : trans_of(A) |] ==> reachable A t"; |
1266 | 55 |
by(Asm_full_simp_tac 1); |
966 | 56 |
by(safe_tac set_cs); |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
966
diff
changeset
|
57 |
by(res_inst_tac [("x","(%i.if i<n then fst ex i \ |
966 | 58 |
\ else (if i=n then Some a else None), \ |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
966
diff
changeset
|
59 |
\ %i.if i<Suc n then snd ex i else t)")] bexI 1); |
966 | 60 |
by(res_inst_tac [("x","Suc(n)")] exI 1); |
1266 | 61 |
by(Simp_tac 1); |
1673
d22110ddd0af
repaired critical proofs depending on the order inside non-confluent SimpSets
oheimb
parents:
1266
diff
changeset
|
62 |
by(asm_simp_tac (!simpset (*delsimps [less_Suc_eq]*)) 1); |
966 | 63 |
by(REPEAT(rtac allI 1)); |
64 |
by(res_inst_tac [("m","na"),("n","n")] (make_elim less_linear) 1); |
|
65 |
be disjE 1; |
|
1673
d22110ddd0af
repaired critical proofs depending on the order inside non-confluent SimpSets
oheimb
parents:
1266
diff
changeset
|
66 |
by(asm_simp_tac (!simpset addsimps [less_Suc_eq]) 1); |
966 | 67 |
be disjE 1; |
1266 | 68 |
by(Asm_simp_tac 1); |
966 | 69 |
by(fast_tac HOL_cs 1); |
70 |
by(forward_tac [less_not_sym] 1); |
|
1673
d22110ddd0af
repaired critical proofs depending on the order inside non-confluent SimpSets
oheimb
parents:
1266
diff
changeset
|
71 |
by(asm_simp_tac (!simpset addsimps [less_not_refl2,less_Suc_eq]) 1); |
966 | 72 |
qed "reachable_n"; |
73 |
||
74 |
val [p1,p2] = goalw IOA.thy [invariant_def] |
|
75 |
"[| !!s. s:starts_of(A) ==> P(s); \ |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
966
diff
changeset
|
76 |
\ !!s t a. [|reachable A s; P(s)|] ==> (s,a,t): trans_of(A) --> P(t) |] \ |
966 | 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); |
|
1266 | 85 |
by (Asm_simp_tac 1); |
966 | 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); \ |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
966
diff
changeset
|
94 |
\ !!s t a. reachable A s ==> P(s) --> (s,a,t):trans_of(A) --> P(t) \ |
966 | 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)"; |
|
1266 | 106 |
by(simp_tac (!simpset addsimps |
966 | 107 |
([actions_def,asig_comp_def]@asig_projections)) 1); |
1844 | 108 |
by(Fast_tac 1); |
966 | 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)}"; |
|
1266 | 113 |
by(simp_tac (!simpset addsimps (par_def::ioa_projections)) 1); |
966 | 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 |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
966
diff
changeset
|
124 |
"(s,a,t) : trans_of(A || B || C || D) = \ |
966 | 125 |
\ ((a:actions(asig_of(A)) | a:actions(asig_of(B)) | a:actions(asig_of(C)) | \ |
126 |
\ a:actions(asig_of(D))) & \ |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
966
diff
changeset
|
127 |
\ (if a:actions(asig_of(A)) then (fst(s),a,fst(t)):trans_of(A) \ |
966 | 128 |
\ else fst t=fst s) & \ |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
966
diff
changeset
|
129 |
\ (if a:actions(asig_of(B)) then (fst(snd(s)),a,fst(snd(t))):trans_of(B) \ |
966 | 130 |
\ else fst(snd(t))=fst(snd(s))) & \ |
131 |
\ (if a:actions(asig_of(C)) then \ |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
966
diff
changeset
|
132 |
\ (fst(snd(snd(s))),a,fst(snd(snd(t)))):trans_of(C) \ |
966 | 133 |
\ else fst(snd(snd(t)))=fst(snd(snd(s)))) & \ |
134 |
\ (if a:actions(asig_of(D)) then \ |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
966
diff
changeset
|
135 |
\ (snd(snd(snd(s))),a,snd(snd(snd(t)))):trans_of(D) \ |
966 | 136 |
\ else snd(snd(snd(t)))=snd(snd(snd(s)))))"; |
1266 | 137 |
by(simp_tac (!simpset addsimps ([par_def,actions_asig_comp,Pair_fst_snd_eq]@ |
966 | 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"; |
|
1266 | 145 |
by(simp_tac (!simpset addsimps ([is_execution_fragment_def,executions_def, |
966 | 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)"; |
|
1266 | 150 |
by(simp_tac (!simpset addsimps (par_def::ioa_projections)) 1); |
966 | 151 |
qed "asig_of_par"; |
1052 | 152 |
|
153 |
||
154 |
goal IOA.thy "externals(asig_of(A1||A2)) = \ |
|
155 |
\ (externals(asig_of(A1)) Un externals(asig_of(A2)))"; |
|
1266 | 156 |
by (asm_full_simp_tac (!simpset addsimps [externals_def,asig_of_par,asig_comp_def,asig_inputs_def,asig_outputs_def,Un_def,set_diff_def]) 1); |
1052 | 157 |
by (rtac set_ext 1); |
158 |
by (fast_tac set_cs 1); |
|
159 |
qed"externals_of_par"; |
|
160 |
||
161 |
goalw IOA.thy [externals_def,actions_def,compat_ioas_def,compat_asigs_def] |
|
162 |
"!! a. [| compat_ioas A1 A2; a:externals(asig_of(A1))|] ==> a~:internals(asig_of(A2))"; |
|
1266 | 163 |
by (Asm_full_simp_tac 1); |
1052 | 164 |
by (best_tac (set_cs addEs [equalityCE]) 1); |
165 |
qed"ext1_is_not_int2"; |
|
166 |
||
167 |
goalw IOA.thy [externals_def,actions_def,compat_ioas_def,compat_asigs_def] |
|
168 |
"!! a. [| compat_ioas A2 A1 ; a:externals(asig_of(A1))|] ==> a~:internals(asig_of(A2))"; |
|
1266 | 169 |
by (Asm_full_simp_tac 1); |
1052 | 170 |
by (best_tac (set_cs addEs [equalityCE]) 1); |
171 |
qed"ext2_is_not_int1"; |
|
172 |
||
173 |
val ext1_ext2_is_not_act2 = ext1_is_not_int2 RS int_and_ext_is_act; |
|
174 |
val ext1_ext2_is_not_act1 = ext2_is_not_int1 RS int_and_ext_is_act; |
|
175 |