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