| author | blanchet | 
| Mon, 25 Oct 2010 09:29:43 +0200 | |
| changeset 40115 | e5ed638e49b0 | 
| parent 36452 | d37c6eed8117 | 
| permissions | -rw-r--r-- | 
| 3071 | 1 | (* Title: HOLCF/IOA/meta_theory/Automata.thy | 
| 12218 | 2 | Author: Olaf Müller, Konrad Slind, Tobias Nipkow | 
| 17233 | 3 | *) | 
| 3071 | 4 | |
| 17233 | 5 | header {* The I/O automata of Lynch and Tuttle in HOLCF *}
 | 
| 3071 | 6 | |
| 17233 | 7 | theory Automata | 
| 8 | imports Asig | |
| 9 | begin | |
| 10 | ||
| 36452 | 11 | default_sort type | 
| 17233 | 12 | |
| 3071 | 13 | types | 
| 17233 | 14 |   ('a, 's) transition = "'s * 'a * 's"
 | 
| 15 |   ('a, 's) ioa = "'a signature * 's set * ('a,'s)transition set * ('a set set) * ('a set set)"
 | |
| 3071 | 16 | |
| 17 | consts | |
| 17233 | 18 | |
| 3071 | 19 | (* IO automata *) | 
| 3521 | 20 | |
| 21 |   asig_of        ::"('a,'s)ioa => 'a signature"
 | |
| 22 |   starts_of      ::"('a,'s)ioa => 's set"
 | |
| 23 |   trans_of       ::"('a,'s)ioa => ('a,'s)transition set"
 | |
| 24 |   wfair_of       ::"('a,'s)ioa => ('a set) set"
 | |
| 25 |   sfair_of       ::"('a,'s)ioa => ('a set) set"
 | |
| 26 | ||
| 27 |   is_asig_of     ::"('a,'s)ioa => bool"
 | |
| 17233 | 28 |   is_starts_of   ::"('a,'s)ioa => bool"
 | 
| 29 |   is_trans_of    ::"('a,'s)ioa => bool"
 | |
| 30 |   input_enabled  ::"('a,'s)ioa => bool"
 | |
| 31 |   IOA            ::"('a,'s)ioa => bool"
 | |
| 3071 | 32 | |
| 4559 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 33 | (* constraints for fair IOA *) | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 34 | |
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 35 |   fairIOA        ::"('a,'s)ioa => bool"
 | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 36 |   input_resistant::"('a,'s)ioa => bool"
 | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 37 | |
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 38 | (* enabledness of actions and action sets *) | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 39 | |
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 40 |   enabled        ::"('a,'s)ioa => 'a => 's => bool"
 | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 41 |   Enabled    ::"('a,'s)ioa => 'a set => 's => bool"
 | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 42 | |
| 17233 | 43 | (* action set keeps enabled until probably disabled by itself *) | 
| 4559 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 44 | |
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 45 |   en_persistent  :: "('a,'s)ioa => 'a set => bool"
 | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 46 | |
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 47 | (* post_conditions for actions and action sets *) | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 48 | |
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 49 |   was_enabled        ::"('a,'s)ioa => 'a => 's => bool"
 | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 50 |   set_was_enabled    ::"('a,'s)ioa => 'a set => 's => bool"
 | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 51 | |
| 23778 | 52 | (* invariants *) | 
| 3071 | 53 |   invariant     :: "[('a,'s)ioa, 's=>bool] => bool"
 | 
| 54 | ||
| 55 | (* binary composition of action signatures and automata *) | |
| 56 | asig_comp ::"['a signature, 'a signature] => 'a signature" | |
| 3521 | 57 |   compatible   ::"[('a,'s)ioa, ('a,'t)ioa] => bool"
 | 
| 22808 | 58 |   par          ::"[('a,'s)ioa, ('a,'t)ioa] => ('a,'s*'t)ioa"  (infixr "||" 10)
 | 
| 3071 | 59 | |
| 3521 | 60 | (* hiding and restricting *) | 
| 61 | hide_asig :: "['a signature, 'a set] => 'a signature" | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35174diff
changeset | 62 |   hide          :: "[('a,'s)ioa, 'a set] => ('a,'s)ioa"
 | 
| 3071 | 63 | restrict_asig :: "['a signature, 'a set] => 'a signature" | 
| 64 |   restrict      :: "[('a,'s)ioa, 'a set] => ('a,'s)ioa"
 | |
| 65 | ||
| 66 | (* renaming *) | |
| 3521 | 67 |   rename_set    :: "'a set => ('c => 'a option) => 'c set"
 | 
| 68 |   rename        :: "('a, 'b)ioa => ('c => 'a option) => ('c,'b)ioa"
 | |
| 3071 | 69 | |
| 22808 | 70 | notation (xsymbols) | 
| 71 | par (infixr "\<parallel>" 10) | |
| 3071 | 72 | |
| 73 | ||
| 23778 | 74 | inductive | 
| 75 |   reachable :: "('a, 's) ioa => 's => bool"
 | |
| 76 |   for C :: "('a, 's) ioa"
 | |
| 77 | where | |
| 78 | reachable_0: "s : starts_of C ==> reachable C s" | |
| 79 | | reachable_n: "[| reachable C s; (s, a, t) : trans_of C |] ==> reachable C t" | |
| 3071 | 80 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23778diff
changeset | 81 | abbreviation | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23778diff
changeset | 82 |   trans_of_syn  ("_ -_--_-> _" [81,81,81,81] 100) where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23778diff
changeset | 83 | "s -a--A-> t == (s,a,t):trans_of A" | 
| 3071 | 84 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23778diff
changeset | 85 | notation (xsymbols) | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23778diff
changeset | 86 |   trans_of_syn  ("_ \<midarrow>_\<midarrow>_\<longrightarrow> _" [81,81,81,81] 100)
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23778diff
changeset | 87 | |
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23778diff
changeset | 88 | abbreviation "act A == actions (asig_of A)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23778diff
changeset | 89 | abbreviation "ext A == externals (asig_of A)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23778diff
changeset | 90 | abbreviation int where "int A == internals (asig_of A)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23778diff
changeset | 91 | abbreviation "inp A == inputs (asig_of A)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23778diff
changeset | 92 | abbreviation "out A == outputs (asig_of A)" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23778diff
changeset | 93 | abbreviation "local A == locals (asig_of A)" | 
| 3433 
2de17c994071
added deadlock freedom, polished definitions and proofs
 mueller parents: 
3275diff
changeset | 94 | |
| 3071 | 95 | defs | 
| 96 | ||
| 97 | (* --------------------------------- IOA ---------------------------------*) | |
| 98 | ||
| 17233 | 99 | asig_of_def: "asig_of == fst" | 
| 100 | starts_of_def: "starts_of == (fst o snd)" | |
| 101 | trans_of_def: "trans_of == (fst o snd o snd)" | |
| 102 | wfair_of_def: "wfair_of == (fst o snd o snd o snd)" | |
| 103 | sfair_of_def: "sfair_of == (snd o snd o snd o snd)" | |
| 104 | ||
| 105 | is_asig_of_def: | |
| 106 | "is_asig_of A == is_asig (asig_of A)" | |
| 107 | ||
| 108 | is_starts_of_def: | |
| 109 |   "is_starts_of A ==  (~ starts_of A = {})"
 | |
| 110 | ||
| 111 | is_trans_of_def: | |
| 112 | "is_trans_of A == | |
| 113 | (!triple. triple:(trans_of A) --> fst(snd(triple)):actions(asig_of A))" | |
| 114 | ||
| 115 | input_enabled_def: | |
| 116 | "input_enabled A == | |
| 117 | (!a. (a:inputs(asig_of A)) --> (!s1. ? s2. (s1,a,s2):(trans_of A)))" | |
| 3071 | 118 | |
| 119 | ||
| 17233 | 120 | ioa_def: | 
| 121 | "IOA A == (is_asig_of A & | |
| 122 | is_starts_of A & | |
| 3521 | 123 | is_trans_of A & | 
| 124 | input_enabled A)" | |
| 3071 | 125 | |
| 126 | ||
| 17233 | 127 | invariant_def: "invariant A P == (!s. reachable A s --> P(s))" | 
| 3071 | 128 | |
| 129 | ||
| 130 | (* ------------------------- parallel composition --------------------------*) | |
| 131 | ||
| 132 | ||
| 17233 | 133 | compatible_def: | 
| 134 | "compatible A B == | |
| 135 |   (((out A Int out B) = {}) &
 | |
| 136 |    ((int A Int act B) = {}) &
 | |
| 3433 
2de17c994071
added deadlock freedom, polished definitions and proofs
 mueller parents: 
3275diff
changeset | 137 |    ((int B Int act A) = {}))"
 | 
| 3071 | 138 | |
| 17233 | 139 | asig_comp_def: | 
| 140 | "asig_comp a1 a2 == | |
| 141 | (((inputs(a1) Un inputs(a2)) - (outputs(a1) Un outputs(a2)), | |
| 142 | (outputs(a1) Un outputs(a2)), | |
| 3071 | 143 | (internals(a1) Un internals(a2))))" | 
| 144 | ||
| 17233 | 145 | par_def: | 
| 146 | "(A || B) == | |
| 147 | (asig_comp (asig_of A) (asig_of B), | |
| 148 |        {pr. fst(pr):starts_of(A) & snd(pr):starts_of(B)},
 | |
| 149 |        {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr))
 | |
| 150 | in (a:act A | a:act B) & | |
| 151 | (if a:act A then | |
| 152 | (fst(s),a,fst(t)):trans_of(A) | |
| 153 | else fst(t) = fst(s)) | |
| 154 | & | |
| 155 | (if a:act B then | |
| 156 | (snd(s),a,snd(t)):trans_of(B) | |
| 3521 | 157 | else snd(t) = snd(s))}, | 
| 158 | wfair_of A Un wfair_of B, | |
| 159 | sfair_of A Un sfair_of B)" | |
| 160 | ||
| 3071 | 161 | |
| 162 | (* ------------------------ hiding -------------------------------------------- *) | |
| 163 | ||
| 17233 | 164 | restrict_asig_def: | 
| 165 | "restrict_asig asig actns == | |
| 166 | (inputs(asig) Int actns, | |
| 167 | outputs(asig) Int actns, | |
| 3071 | 168 | internals(asig) Un (externals(asig) - actns))" | 
| 169 | ||
| 17233 | 170 | (* Notice that for wfair_of and sfair_of nothing has to be changed, as | 
| 171 | changes from the outputs to the internals does not touch the locals as | |
| 3521 | 172 | a whole, which is of importance for fairness only *) | 
| 3071 | 173 | |
| 17233 | 174 | restrict_def: | 
| 175 | "restrict A actns == | |
| 176 | (restrict_asig (asig_of A) actns, | |
| 177 | starts_of A, | |
| 3521 | 178 | trans_of A, | 
| 179 | wfair_of A, | |
| 180 | sfair_of A)" | |
| 181 | ||
| 17233 | 182 | hide_asig_def: | 
| 183 | "hide_asig asig actns == | |
| 184 | (inputs(asig) - actns, | |
| 185 | outputs(asig) - actns, | |
| 3521 | 186 | internals(asig) Un actns)" | 
| 187 | ||
| 17233 | 188 | hide_def: | 
| 189 | "hide A actns == | |
| 190 | (hide_asig (asig_of A) actns, | |
| 191 | starts_of A, | |
| 3521 | 192 | trans_of A, | 
| 193 | wfair_of A, | |
| 194 | sfair_of A)" | |
| 3071 | 195 | |
| 196 | (* ------------------------- renaming ------------------------------------------- *) | |
| 17233 | 197 | |
| 198 | rename_set_def: | |
| 199 |   "rename_set A ren == {b. ? x. Some x = ren b & x : A}"
 | |
| 3521 | 200 | |
| 17233 | 201 | rename_def: | 
| 202 | "rename ioa ren == | |
| 203 | ((rename_set (inp ioa) ren, | |
| 204 | rename_set (out ioa) ren, | |
| 205 | rename_set (int ioa) ren), | |
| 206 | starts_of ioa, | |
| 207 |    {tr. let s = fst(tr); a = fst(snd(tr));  t = snd(snd(tr))
 | |
| 208 | in | |
| 3521 | 209 | ? x. Some(x) = ren(a) & (s,x,t):trans_of ioa}, | 
| 210 |    {rename_set s ren | s. s: wfair_of ioa},
 | |
| 211 |    {rename_set s ren | s. s: sfair_of ioa})"
 | |
| 3071 | 212 | |
| 4559 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 213 | (* ------------------------- fairness ----------------------------- *) | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 214 | |
| 17233 | 215 | fairIOA_def: | 
| 216 | "fairIOA A == (! S : wfair_of A. S<= local A) & | |
| 4559 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 217 | (! S : sfair_of A. S<= local A)" | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 218 | |
| 17233 | 219 | input_resistant_def: | 
| 220 | "input_resistant A == ! W : sfair_of A. ! s a t. | |
| 4559 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 221 | reachable A s & reachable A t & a:inp A & | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 222 | Enabled A W s & s -a--A-> t | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 223 | --> Enabled A W t" | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 224 | |
| 17233 | 225 | enabled_def: | 
| 4559 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 226 | "enabled A a s == ? t. s-a--A-> t" | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 227 | |
| 17233 | 228 | Enabled_def: | 
| 4559 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 229 | "Enabled A W s == ? w:W. enabled A w s" | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 230 | |
| 17233 | 231 | en_persistent_def: | 
| 232 | "en_persistent A W == ! s a t. Enabled A W s & | |
| 233 | a ~:W & | |
| 234 | s -a--A-> t | |
| 4559 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 235 | --> Enabled A W t" | 
| 17233 | 236 | was_enabled_def: | 
| 4559 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 237 | "was_enabled A a t == ? s. s-a--A-> t" | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 238 | |
| 17233 | 239 | set_was_enabled_def: | 
| 4559 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 240 | "set_was_enabled A W t == ? w:W. was_enabled A w t" | 
| 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3521diff
changeset | 241 | |
| 19741 | 242 | |
| 243 | declare split_paired_Ex [simp del] | |
| 244 | ||
| 245 | lemmas ioa_projections = asig_of_def starts_of_def trans_of_def wfair_of_def sfair_of_def | |
| 246 | ||
| 247 | ||
| 248 | subsection "asig_of, starts_of, trans_of" | |
| 249 | ||
| 250 | lemma ioa_triple_proj: | |
| 251 | "((asig_of (x,y,z,w,s)) = x) & | |
| 252 | ((starts_of (x,y,z,w,s)) = y) & | |
| 253 | ((trans_of (x,y,z,w,s)) = z) & | |
| 254 | ((wfair_of (x,y,z,w,s)) = w) & | |
| 255 | ((sfair_of (x,y,z,w,s)) = s)" | |
| 256 | apply (simp add: ioa_projections) | |
| 257 | done | |
| 258 | ||
| 259 | lemma trans_in_actions: | |
| 260 | "[| is_trans_of A; (s1,a,s2):trans_of(A) |] ==> a:act A" | |
| 261 | apply (unfold is_trans_of_def actions_def is_asig_def) | |
| 262 | apply (erule allE, erule impE, assumption) | |
| 263 | apply simp | |
| 264 | done | |
| 265 | ||
| 266 | lemma starts_of_par: | |
| 267 | "starts_of(A || B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}"
 | |
| 268 | apply (simp add: par_def ioa_projections) | |
| 269 | done | |
| 270 | ||
| 271 | lemma trans_of_par: | |
| 272 | "trans_of(A || B) = {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr))  
 | |
| 273 | in (a:act A | a:act B) & | |
| 274 | (if a:act A then | |
| 275 | (fst(s),a,fst(t)):trans_of(A) | |
| 276 | else fst(t) = fst(s)) | |
| 277 | & | |
| 278 | (if a:act B then | |
| 279 | (snd(s),a,snd(t)):trans_of(B) | |
| 280 | else snd(t) = snd(s))}" | |
| 281 | ||
| 282 | apply (simp add: par_def ioa_projections) | |
| 283 | done | |
| 284 | ||
| 285 | ||
| 286 | subsection "actions and par" | |
| 287 | ||
| 288 | lemma actions_asig_comp: | |
| 289 | "actions(asig_comp a b) = actions(a) Un actions(b)" | |
| 290 | apply (simp (no_asm) add: actions_def asig_comp_def asig_projections) | |
| 291 | apply blast | |
| 292 | done | |
| 293 | ||
| 294 | lemma asig_of_par: "asig_of(A || B) = asig_comp (asig_of A) (asig_of B)" | |
| 295 | apply (simp add: par_def ioa_projections) | |
| 296 | done | |
| 297 | ||
| 298 | ||
| 299 | lemma externals_of_par: "ext (A1||A2) = | |
| 300 | (ext A1) Un (ext A2)" | |
| 301 | apply (simp add: externals_def asig_of_par asig_comp_def | |
| 26806 | 302 | asig_inputs_def asig_outputs_def Un_def set_diff_eq) | 
| 19741 | 303 | apply blast | 
| 304 | done | |
| 305 | ||
| 306 | lemma actions_of_par: "act (A1||A2) = | |
| 307 | (act A1) Un (act A2)" | |
| 308 | apply (simp add: actions_def asig_of_par asig_comp_def | |
| 26806 | 309 | asig_inputs_def asig_outputs_def asig_internals_def Un_def set_diff_eq) | 
| 19741 | 310 | apply blast | 
| 311 | done | |
| 312 | ||
| 313 | lemma inputs_of_par: "inp (A1||A2) = | |
| 314 | ((inp A1) Un (inp A2)) - ((out A1) Un (out A2))" | |
| 315 | apply (simp add: actions_def asig_of_par asig_comp_def | |
| 26806 | 316 | asig_inputs_def asig_outputs_def Un_def set_diff_eq) | 
| 19741 | 317 | done | 
| 318 | ||
| 319 | lemma outputs_of_par: "out (A1||A2) = | |
| 320 | (out A1) Un (out A2)" | |
| 321 | apply (simp add: actions_def asig_of_par asig_comp_def | |
| 26806 | 322 | asig_outputs_def Un_def set_diff_eq) | 
| 19741 | 323 | done | 
| 324 | ||
| 325 | lemma internals_of_par: "int (A1||A2) = | |
| 326 | (int A1) Un (int A2)" | |
| 327 | apply (simp add: actions_def asig_of_par asig_comp_def | |
| 26806 | 328 | asig_inputs_def asig_outputs_def asig_internals_def Un_def set_diff_eq) | 
| 19741 | 329 | done | 
| 330 | ||
| 331 | ||
| 332 | subsection "actions and compatibility" | |
| 333 | ||
| 334 | lemma compat_commute: "compatible A B = compatible B A" | |
| 335 | apply (simp add: compatible_def Int_commute) | |
| 336 | apply auto | |
| 337 | done | |
| 338 | ||
| 339 | lemma ext1_is_not_int2: | |
| 340 | "[| compatible A1 A2; a:ext A1|] ==> a~:int A2" | |
| 341 | apply (unfold externals_def actions_def compatible_def) | |
| 342 | apply simp | |
| 343 | apply blast | |
| 344 | done | |
| 345 | ||
| 346 | (* just commuting the previous one: better commute compatible *) | |
| 347 | lemma ext2_is_not_int1: | |
| 348 | "[| compatible A2 A1 ; a:ext A1|] ==> a~:int A2" | |
| 349 | apply (unfold externals_def actions_def compatible_def) | |
| 350 | apply simp | |
| 351 | apply blast | |
| 352 | done | |
| 353 | ||
| 354 | lemmas ext1_ext2_is_not_act2 = ext1_is_not_int2 [THEN int_and_ext_is_act, standard] | |
| 355 | lemmas ext1_ext2_is_not_act1 = ext2_is_not_int1 [THEN int_and_ext_is_act, standard] | |
| 356 | ||
| 357 | lemma intA_is_not_extB: | |
| 358 | "[| compatible A B; x:int A |] ==> x~:ext B" | |
| 359 | apply (unfold externals_def actions_def compatible_def) | |
| 360 | apply simp | |
| 361 | apply blast | |
| 362 | done | |
| 363 | ||
| 364 | lemma intA_is_not_actB: | |
| 365 | "[| compatible A B; a:int A |] ==> a ~: act B" | |
| 366 | apply (unfold externals_def actions_def compatible_def is_asig_def asig_of_def) | |
| 367 | apply simp | |
| 368 | apply blast | |
| 369 | done | |
| 370 | ||
| 371 | (* the only one that needs disjointness of outputs and of internals and _all_ acts *) | |
| 372 | lemma outAactB_is_inpB: | |
| 373 | "[| compatible A B; a:out A ;a:act B|] ==> a : inp B" | |
| 374 | apply (unfold asig_outputs_def asig_internals_def actions_def asig_inputs_def | |
| 375 | compatible_def is_asig_def asig_of_def) | |
| 376 | apply simp | |
| 377 | apply blast | |
| 378 | done | |
| 379 | ||
| 380 | (* needed for propagation of input_enabledness from A,B to A||B *) | |
| 381 | lemma inpAAactB_is_inpBoroutB: | |
| 382 | "[| compatible A B; a:inp A ;a:act B|] ==> a : inp B | a: out B" | |
| 383 | apply (unfold asig_outputs_def asig_internals_def actions_def asig_inputs_def | |
| 384 | compatible_def is_asig_def asig_of_def) | |
| 385 | apply simp | |
| 386 | apply blast | |
| 387 | done | |
| 388 | ||
| 389 | ||
| 390 | subsection "input_enabledness and par" | |
| 391 | ||
| 392 | ||
| 393 | (* ugly case distinctions. Heart of proof: | |
| 394 | 1. inpAAactB_is_inpBoroutB ie. internals are really hidden. | |
| 395 | 2. inputs_of_par: outputs are no longer inputs of par. This is important here *) | |
| 396 | lemma input_enabled_par: | |
| 397 | "[| compatible A B; input_enabled A; input_enabled B|] | |
| 398 | ==> input_enabled (A||B)" | |
| 399 | apply (unfold input_enabled_def) | |
| 400 | apply (simp add: Let_def inputs_of_par trans_of_par) | |
| 32152 | 401 | apply (tactic "safe_tac (global_claset_of @{theory Fun})")
 | 
| 19741 | 402 | apply (simp add: inp_is_act) | 
| 403 | prefer 2 | |
| 404 | apply (simp add: inp_is_act) | |
| 405 | (* a: inp A *) | |
| 406 | apply (case_tac "a:act B") | |
| 407 | (* a:act B *) | |
| 408 | apply (erule_tac x = "a" in allE) | |
| 409 | apply simp | |
| 410 | apply (drule inpAAactB_is_inpBoroutB) | |
| 411 | apply assumption | |
| 412 | apply assumption | |
| 413 | apply (erule_tac x = "a" in allE) | |
| 414 | apply simp | |
| 415 | apply (erule_tac x = "aa" in allE) | |
| 416 | apply (erule_tac x = "b" in allE) | |
| 417 | apply (erule exE) | |
| 418 | apply (erule exE) | |
| 419 | apply (rule_tac x = " (s2,s2a) " in exI) | |
| 420 | apply (simp add: inp_is_act) | |
| 421 | (* a~: act B*) | |
| 422 | apply (simp add: inp_is_act) | |
| 423 | apply (erule_tac x = "a" in allE) | |
| 424 | apply simp | |
| 425 | apply (erule_tac x = "aa" in allE) | |
| 426 | apply (erule exE) | |
| 427 | apply (rule_tac x = " (s2,b) " in exI) | |
| 428 | apply simp | |
| 429 | ||
| 430 | (* a:inp B *) | |
| 431 | apply (case_tac "a:act A") | |
| 432 | (* a:act A *) | |
| 433 | apply (erule_tac x = "a" in allE) | |
| 434 | apply (erule_tac x = "a" in allE) | |
| 435 | apply (simp add: inp_is_act) | |
| 436 | apply (frule_tac A1 = "A" in compat_commute [THEN iffD1]) | |
| 437 | apply (drule inpAAactB_is_inpBoroutB) | |
| 438 | back | |
| 439 | apply assumption | |
| 440 | apply assumption | |
| 441 | apply simp | |
| 442 | apply (erule_tac x = "aa" in allE) | |
| 443 | apply (erule_tac x = "b" in allE) | |
| 444 | apply (erule exE) | |
| 445 | apply (erule exE) | |
| 446 | apply (rule_tac x = " (s2,s2a) " in exI) | |
| 447 | apply (simp add: inp_is_act) | |
| 448 | (* a~: act B*) | |
| 449 | apply (simp add: inp_is_act) | |
| 450 | apply (erule_tac x = "a" in allE) | |
| 451 | apply (erule_tac x = "a" in allE) | |
| 452 | apply simp | |
| 453 | apply (erule_tac x = "b" in allE) | |
| 454 | apply (erule exE) | |
| 455 | apply (rule_tac x = " (aa,s2) " in exI) | |
| 456 | apply simp | |
| 457 | done | |
| 458 | ||
| 459 | ||
| 460 | subsection "invariants" | |
| 461 | ||
| 462 | lemma invariantI: | |
| 463 | "[| !!s. s:starts_of(A) ==> P(s); | |
| 464 | !!s t a. [|reachable A s; P(s)|] ==> (s,a,t): trans_of(A) --> P(t) |] | |
| 465 | ==> invariant A P" | |
| 466 | apply (unfold invariant_def) | |
| 467 | apply (rule allI) | |
| 468 | apply (rule impI) | |
| 23778 | 469 | apply (rule_tac x = "s" in reachable.induct) | 
| 19741 | 470 | apply assumption | 
| 471 | apply blast | |
| 472 | apply blast | |
| 473 | done | |
| 474 | ||
| 475 | lemma invariantI1: | |
| 476 | "[| !!s. s : starts_of(A) ==> P(s); | |
| 477 | !!s t a. reachable A s ==> P(s) --> (s,a,t):trans_of(A) --> P(t) | |
| 478 | |] ==> invariant A P" | |
| 479 | apply (blast intro: invariantI) | |
| 480 | done | |
| 481 | ||
| 482 | lemma invariantE: "[| invariant A P; reachable A s |] ==> P(s)" | |
| 483 | apply (unfold invariant_def) | |
| 484 | apply blast | |
| 485 | done | |
| 486 | ||
| 487 | ||
| 488 | subsection "restrict" | |
| 489 | ||
| 490 | ||
| 491 | lemmas reachable_0 = reachable.reachable_0 | |
| 492 | and reachable_n = reachable.reachable_n | |
| 493 | ||
| 494 | lemma cancel_restrict_a: "starts_of(restrict ioa acts) = starts_of(ioa) & | |
| 495 | trans_of(restrict ioa acts) = trans_of(ioa)" | |
| 496 | apply (simp add: restrict_def ioa_projections) | |
| 497 | done | |
| 498 | ||
| 499 | lemma cancel_restrict_b: "reachable (restrict ioa acts) s = reachable ioa s" | |
| 500 | apply (rule iffI) | |
| 501 | apply (erule reachable.induct) | |
| 502 | apply (simp add: cancel_restrict_a reachable_0) | |
| 503 | apply (erule reachable_n) | |
| 504 | apply (simp add: cancel_restrict_a) | |
| 505 | (* <-- *) | |
| 506 | apply (erule reachable.induct) | |
| 507 | apply (rule reachable_0) | |
| 508 | apply (simp add: cancel_restrict_a) | |
| 509 | apply (erule reachable_n) | |
| 510 | apply (simp add: cancel_restrict_a) | |
| 511 | done | |
| 512 | ||
| 513 | lemma acts_restrict: "act (restrict A acts) = act A" | |
| 514 | apply (simp (no_asm) add: actions_def asig_internals_def | |
| 515 | asig_outputs_def asig_inputs_def externals_def asig_of_def restrict_def restrict_asig_def) | |
| 516 | apply auto | |
| 517 | done | |
| 518 | ||
| 519 | lemma cancel_restrict: "starts_of(restrict ioa acts) = starts_of(ioa) & | |
| 520 | trans_of(restrict ioa acts) = trans_of(ioa) & | |
| 521 | reachable (restrict ioa acts) s = reachable ioa s & | |
| 522 | act (restrict A acts) = act A" | |
| 523 | apply (simp (no_asm) add: cancel_restrict_a cancel_restrict_b acts_restrict) | |
| 524 | done | |
| 525 | ||
| 526 | ||
| 527 | subsection "rename" | |
| 528 | ||
| 529 | lemma trans_rename: "s -a--(rename C f)-> t ==> (? x. Some(x) = f(a) & s -x--C-> t)" | |
| 530 | apply (simp add: Let_def rename_def trans_of_def) | |
| 531 | done | |
| 532 | ||
| 533 | ||
| 534 | lemma reachable_rename: "[| reachable (rename C g) s |] ==> reachable C s" | |
| 535 | apply (erule reachable.induct) | |
| 536 | apply (rule reachable_0) | |
| 537 | apply (simp add: rename_def ioa_projections) | |
| 538 | apply (drule trans_rename) | |
| 539 | apply (erule exE) | |
| 540 | apply (erule conjE) | |
| 541 | apply (erule reachable_n) | |
| 542 | apply assumption | |
| 543 | done | |
| 544 | ||
| 545 | ||
| 546 | subsection "trans_of(A||B)" | |
| 547 | ||
| 548 | ||
| 549 | lemma trans_A_proj: "[|(s,a,t):trans_of (A||B); a:act A|] | |
| 550 | ==> (fst s,a,fst t):trans_of A" | |
| 551 | apply (simp add: Let_def par_def trans_of_def) | |
| 552 | done | |
| 553 | ||
| 554 | lemma trans_B_proj: "[|(s,a,t):trans_of (A||B); a:act B|] | |
| 555 | ==> (snd s,a,snd t):trans_of B" | |
| 556 | apply (simp add: Let_def par_def trans_of_def) | |
| 557 | done | |
| 558 | ||
| 559 | lemma trans_A_proj2: "[|(s,a,t):trans_of (A||B); a~:act A|] | |
| 560 | ==> fst s = fst t" | |
| 561 | apply (simp add: Let_def par_def trans_of_def) | |
| 562 | done | |
| 563 | ||
| 564 | lemma trans_B_proj2: "[|(s,a,t):trans_of (A||B); a~:act B|] | |
| 565 | ==> snd s = snd t" | |
| 566 | apply (simp add: Let_def par_def trans_of_def) | |
| 567 | done | |
| 568 | ||
| 569 | lemma trans_AB_proj: "(s,a,t):trans_of (A||B) | |
| 570 | ==> a :act A | a :act B" | |
| 571 | apply (simp add: Let_def par_def trans_of_def) | |
| 572 | done | |
| 573 | ||
| 574 | lemma trans_AB: "[|a:act A;a:act B; | |
| 575 | (fst s,a,fst t):trans_of A;(snd s,a,snd t):trans_of B|] | |
| 576 | ==> (s,a,t):trans_of (A||B)" | |
| 577 | apply (simp add: Let_def par_def trans_of_def) | |
| 578 | done | |
| 579 | ||
| 580 | lemma trans_A_notB: "[|a:act A;a~:act B; | |
| 581 | (fst s,a,fst t):trans_of A;snd s=snd t|] | |
| 582 | ==> (s,a,t):trans_of (A||B)" | |
| 583 | apply (simp add: Let_def par_def trans_of_def) | |
| 584 | done | |
| 585 | ||
| 586 | lemma trans_notA_B: "[|a~:act A;a:act B; | |
| 587 | (snd s,a,snd t):trans_of B;fst s=fst t|] | |
| 588 | ==> (s,a,t):trans_of (A||B)" | |
| 589 | apply (simp add: Let_def par_def trans_of_def) | |
| 590 | done | |
| 591 | ||
| 592 | lemmas trans_of_defs1 = trans_AB trans_A_notB trans_notA_B | |
| 593 | and trans_of_defs2 = trans_A_proj trans_B_proj trans_A_proj2 trans_B_proj2 trans_AB_proj | |
| 594 | ||
| 595 | ||
| 596 | lemma trans_of_par4: | |
| 597 | "((s,a,t) : trans_of(A || B || C || D)) = | |
| 598 | ((a:actions(asig_of(A)) | a:actions(asig_of(B)) | a:actions(asig_of(C)) | | |
| 599 | a:actions(asig_of(D))) & | |
| 600 | (if a:actions(asig_of(A)) then (fst(s),a,fst(t)):trans_of(A) | |
| 601 | else fst t=fst s) & | |
| 602 | (if a:actions(asig_of(B)) then (fst(snd(s)),a,fst(snd(t))):trans_of(B) | |
| 603 | else fst(snd(t))=fst(snd(s))) & | |
| 604 | (if a:actions(asig_of(C)) then | |
| 605 | (fst(snd(snd(s))),a,fst(snd(snd(t)))):trans_of(C) | |
| 606 | else fst(snd(snd(t)))=fst(snd(snd(s)))) & | |
| 607 | (if a:actions(asig_of(D)) then | |
| 608 | (snd(snd(snd(s))),a,snd(snd(snd(t)))):trans_of(D) | |
| 609 | else snd(snd(snd(t)))=snd(snd(snd(s)))))" | |
| 610 | apply (simp (no_asm) add: par_def actions_asig_comp Pair_fst_snd_eq Let_def ioa_projections) | |
| 611 | done | |
| 612 | ||
| 613 | ||
| 614 | subsection "proof obligation generator for IOA requirements" | |
| 615 | ||
| 616 | (* without assumptions on A and B because is_trans_of is also incorporated in ||def *) | |
| 617 | lemma is_trans_of_par: "is_trans_of (A||B)" | |
| 618 | apply (unfold is_trans_of_def) | |
| 619 | apply (simp add: Let_def actions_of_par trans_of_par) | |
| 620 | done | |
| 621 | ||
| 622 | lemma is_trans_of_restrict: | |
| 623 | "is_trans_of A ==> is_trans_of (restrict A acts)" | |
| 624 | apply (unfold is_trans_of_def) | |
| 625 | apply (simp add: cancel_restrict acts_restrict) | |
| 626 | done | |
| 627 | ||
| 628 | lemma is_trans_of_rename: | |
| 629 | "is_trans_of A ==> is_trans_of (rename A f)" | |
| 630 | apply (unfold is_trans_of_def restrict_def restrict_asig_def) | |
| 631 | apply (simp add: Let_def actions_def trans_of_def asig_internals_def | |
| 632 | asig_outputs_def asig_inputs_def externals_def asig_of_def rename_def rename_set_def) | |
| 633 | apply blast | |
| 634 | done | |
| 635 | ||
| 636 | lemma is_asig_of_par: "[| is_asig_of A; is_asig_of B; compatible A B|] | |
| 637 | ==> is_asig_of (A||B)" | |
| 638 | apply (simp add: is_asig_of_def asig_of_par asig_comp_def compatible_def | |
| 639 | asig_internals_def asig_outputs_def asig_inputs_def actions_def is_asig_def) | |
| 640 | apply (simp add: asig_of_def) | |
| 641 | apply auto | |
| 642 | done | |
| 643 | ||
| 644 | lemma is_asig_of_restrict: | |
| 645 | "is_asig_of A ==> is_asig_of (restrict A f)" | |
| 646 | apply (unfold is_asig_of_def is_asig_def asig_of_def restrict_def restrict_asig_def | |
| 647 | asig_internals_def asig_outputs_def asig_inputs_def externals_def o_def) | |
| 648 | apply simp | |
| 649 | apply auto | |
| 650 | done | |
| 651 | ||
| 652 | lemma is_asig_of_rename: "is_asig_of A ==> is_asig_of (rename A f)" | |
| 653 | apply (simp add: is_asig_of_def rename_def rename_set_def asig_internals_def | |
| 654 | asig_outputs_def asig_inputs_def actions_def is_asig_def asig_of_def) | |
| 655 | apply auto | |
| 656 | apply (drule_tac [!] s = "Some ?x" in sym) | |
| 657 | apply auto | |
| 658 | done | |
| 659 | ||
| 660 | lemmas [simp] = is_asig_of_par is_asig_of_restrict | |
| 661 | is_asig_of_rename is_trans_of_par is_trans_of_restrict is_trans_of_rename | |
| 662 | ||
| 663 | ||
| 664 | lemma compatible_par: | |
| 665 | "[|compatible A B; compatible A C |]==> compatible A (B||C)" | |
| 666 | apply (unfold compatible_def) | |
| 667 | apply (simp add: internals_of_par outputs_of_par actions_of_par) | |
| 668 | apply auto | |
| 669 | done | |
| 670 | ||
| 671 | (* better derive by previous one and compat_commute *) | |
| 672 | lemma compatible_par2: | |
| 673 | "[|compatible A C; compatible B C |]==> compatible (A||B) C" | |
| 674 | apply (unfold compatible_def) | |
| 675 | apply (simp add: internals_of_par outputs_of_par actions_of_par) | |
| 676 | apply auto | |
| 677 | done | |
| 678 | ||
| 679 | lemma compatible_restrict: | |
| 680 | "[| compatible A B; (ext B - S) Int ext A = {}|]  
 | |
| 681 | ==> compatible A (restrict B S)" | |
| 682 | apply (unfold compatible_def) | |
| 683 | apply (simp add: ioa_triple_proj asig_triple_proj externals_def | |
| 684 | restrict_def restrict_asig_def actions_def) | |
| 685 | apply auto | |
| 686 | done | |
| 687 | ||
| 688 | ||
| 689 | declare split_paired_Ex [simp] | |
| 3071 | 690 | |
| 691 | end |