| author | wenzelm | 
| Fri, 05 Aug 2022 19:02:38 +0200 | |
| changeset 75769 | 4d27b520622a | 
| parent 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 62008 | 1 | (* Title: HOL/HOLCF/IOA/CompoTraces.thy | 
| 40945 | 2 | Author: Olaf Müller | 
| 62005 | 3 | *) | 
| 3071 | 4 | |
| 62002 | 5 | section \<open>Compositionality on Trace level\<close> | 
| 17233 | 6 | |
| 7 | theory CompoTraces | |
| 8 | imports CompoScheds ShortExecutions | |
| 9 | begin | |
| 3071 | 10 | |
| 62156 | 11 | definition mksch :: | 
| 12 |     "('a, 's) ioa \<Rightarrow> ('a, 't) ioa \<Rightarrow> 'a Seq \<rightarrow> 'a Seq \<rightarrow> 'a Seq \<rightarrow> 'a Seq"
 | |
| 13 | where "mksch A B = | |
| 63549 | 14 | (fix \<cdot> | 
| 62156 | 15 | (LAM h tr schA schB. | 
| 16 | case tr of | |
| 17 | nil \<Rightarrow> nil | |
| 18 | | x ## xs \<Rightarrow> | |
| 19 | (case x of | |
| 20 | UU \<Rightarrow> UU | |
| 21 | | Def y \<Rightarrow> | |
| 22 | (if y \<in> act A then | |
| 23 | (if y \<in> act B then | |
| 63549 | 24 | ((Takewhile (\<lambda>a. a \<in> int A) \<cdot> schA) @@ | 
| 25 | (Takewhile (\<lambda>a. a \<in> int B) \<cdot> schB) @@ | |
| 26 | (y \<leadsto> (h \<cdot> xs \<cdot> (TL \<cdot> (Dropwhile (\<lambda>a. a \<in> int A) \<cdot> schA)) | |
| 27 | \<cdot> (TL \<cdot> (Dropwhile (\<lambda>a. a \<in> int B) \<cdot> schB))))) | |
| 62156 | 28 | else | 
| 63549 | 29 | ((Takewhile (\<lambda>a. a \<in> int A) \<cdot> schA) @@ | 
| 30 | (y \<leadsto> (h \<cdot> xs \<cdot> (TL \<cdot> (Dropwhile (\<lambda>a. a \<in> int A) \<cdot> schA)) \<cdot> schB)))) | |
| 62156 | 31 | else | 
| 32 | (if y \<in> act B then | |
| 63549 | 33 | ((Takewhile (\<lambda>a. a \<in> int B) \<cdot> schB) @@ | 
| 34 | (y \<leadsto> (h \<cdot> xs \<cdot> schA \<cdot> (TL \<cdot> (Dropwhile (\<lambda>a. a \<in> int B) \<cdot> schB))))) | |
| 62156 | 35 | else UU)))))" | 
| 3071 | 36 | |
| 62156 | 37 | definition par_traces :: "'a trace_module \<Rightarrow> 'a trace_module \<Rightarrow> 'a trace_module" | 
| 38 | where "par_traces TracesA TracesB = | |
| 39 | (let | |
| 40 | trA = fst TracesA; sigA = snd TracesA; | |
| 41 | trB = fst TracesB; sigB = snd TracesB | |
| 42 | in | |
| 63549 | 43 |        ({tr. Filter (\<lambda>a. a \<in> actions sigA) \<cdot> tr \<in> trA} \<inter>
 | 
| 44 |         {tr. Filter (\<lambda>a. a \<in> actions sigB) \<cdot> tr \<in> trB} \<inter>
 | |
| 62156 | 45 |         {tr. Forall (\<lambda>x. x \<in> externals sigA \<union> externals sigB) tr},
 | 
| 62005 | 46 | asig_comp sigA sigB))" | 
| 3521 | 47 | |
| 62156 | 48 | axiomatization where | 
| 63549 | 49 | finiteR_mksch: "Finite (mksch A B \<cdot> tr \<cdot> x \<cdot> y) \<Longrightarrow> Finite tr" | 
| 48194 
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
 wenzelm parents: 
47239diff
changeset | 50 | |
| 63549 | 51 | lemma finiteR_mksch': "\<not> Finite tr \<Longrightarrow> \<not> Finite (mksch A B \<cdot> tr \<cdot> x \<cdot> y)" | 
| 48194 
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
 wenzelm parents: 
47239diff
changeset | 52 | by (blast intro: finiteR_mksch) | 
| 3071 | 53 | |
| 19741 | 54 | |
| 62002 | 55 | declaration \<open>fn _ => Simplifier.map_ss (Simplifier.set_mksym (K (K NONE)))\<close> | 
| 19741 | 56 | |
| 57 | ||
| 58 | subsection "mksch rewrite rules" | |
| 59 | ||
| 60 | lemma mksch_unfold: | |
| 62156 | 61 | "mksch A B = | 
| 62 | (LAM tr schA schB. | |
| 63 | case tr of | |
| 64 | nil \<Rightarrow> nil | |
| 65 | | x ## xs \<Rightarrow> | |
| 66 | (case x of | |
| 67 | UU \<Rightarrow> UU | |
| 68 | | Def y \<Rightarrow> | |
| 69 | (if y \<in> act A then | |
| 70 | (if y \<in> act B then | |
| 63549 | 71 | ((Takewhile (\<lambda>a. a \<in> int A) \<cdot> schA) @@ | 
| 72 | (Takewhile (\<lambda>a. a \<in> int B) \<cdot> schB) @@ | |
| 73 | (y \<leadsto> (mksch A B \<cdot> xs \<cdot>(TL \<cdot> (Dropwhile (\<lambda>a. a \<in> int A) \<cdot> schA)) | |
| 74 | \<cdot>(TL \<cdot> (Dropwhile (\<lambda>a. a \<in> int B) \<cdot> schB))))) | |
| 62156 | 75 | else | 
| 63549 | 76 | ((Takewhile (\<lambda>a. a \<in> int A) \<cdot> schA) @@ | 
| 77 | (y \<leadsto> (mksch A B \<cdot> xs \<cdot> (TL \<cdot> (Dropwhile (\<lambda>a. a \<in> int A) \<cdot> schA)) \<cdot> schB)))) | |
| 62156 | 78 | else | 
| 79 | (if y \<in> act B then | |
| 63549 | 80 | ((Takewhile (\<lambda>a. a \<in> int B) \<cdot> schB) @@ | 
| 81 | (y \<leadsto> (mksch A B \<cdot> xs \<cdot> schA \<cdot> (TL \<cdot> (Dropwhile (\<lambda>a. a \<in> int B) \<cdot> schB))))) | |
| 62156 | 82 | else UU))))" | 
| 83 | apply (rule trans) | |
| 84 | apply (rule fix_eq4) | |
| 85 | apply (rule mksch_def) | |
| 86 | apply (rule beta_cfun) | |
| 87 | apply simp | |
| 88 | done | |
| 89 | ||
| 63549 | 90 | lemma mksch_UU: "mksch A B \<cdot> UU \<cdot> schA \<cdot> schB = UU" | 
| 62156 | 91 | apply (subst mksch_unfold) | 
| 92 | apply simp | |
| 93 | done | |
| 19741 | 94 | |
| 63549 | 95 | lemma mksch_nil: "mksch A B \<cdot> nil \<cdot> schA \<cdot> schB = nil" | 
| 62156 | 96 | apply (subst mksch_unfold) | 
| 97 | apply simp | |
| 98 | done | |
| 19741 | 99 | |
| 62156 | 100 | lemma mksch_cons1: | 
| 101 | "x \<in> act A \<Longrightarrow> x \<notin> act B \<Longrightarrow> | |
| 63549 | 102 | mksch A B \<cdot> (x \<leadsto> tr) \<cdot> schA \<cdot> schB = | 
| 103 | (Takewhile (\<lambda>a. a \<in> int A) \<cdot> schA) @@ | |
| 104 | (x \<leadsto> (mksch A B \<cdot> tr \<cdot> (TL \<cdot> (Dropwhile (\<lambda>a. a \<in> int A) \<cdot> schA)) \<cdot> schB))" | |
| 62156 | 105 | apply (rule trans) | 
| 106 | apply (subst mksch_unfold) | |
| 107 | apply (simp add: Consq_def If_and_if) | |
| 108 | apply (simp add: Consq_def) | |
| 109 | done | |
| 19741 | 110 | |
| 62156 | 111 | lemma mksch_cons2: | 
| 112 | "x \<notin> act A \<Longrightarrow> x \<in> act B \<Longrightarrow> | |
| 63549 | 113 | mksch A B \<cdot> (x \<leadsto> tr) \<cdot> schA \<cdot> schB = | 
| 114 | (Takewhile (\<lambda>a. a \<in> int B) \<cdot> schB) @@ | |
| 115 | (x \<leadsto> (mksch A B \<cdot> tr \<cdot> schA \<cdot> (TL \<cdot> (Dropwhile (\<lambda>a. a \<in> int B) \<cdot> schB))))" | |
| 62156 | 116 | apply (rule trans) | 
| 117 | apply (subst mksch_unfold) | |
| 118 | apply (simp add: Consq_def If_and_if) | |
| 119 | apply (simp add: Consq_def) | |
| 120 | done | |
| 19741 | 121 | |
| 62156 | 122 | lemma mksch_cons3: | 
| 123 | "x \<in> act A \<Longrightarrow> x \<in> act B \<Longrightarrow> | |
| 63549 | 124 | mksch A B \<cdot> (x \<leadsto> tr) \<cdot> schA \<cdot> schB = | 
| 125 | (Takewhile (\<lambda>a. a \<in> int A) \<cdot> schA) @@ | |
| 126 | ((Takewhile (\<lambda>a. a \<in> int B) \<cdot> schB) @@ | |
| 127 | (x \<leadsto> (mksch A B \<cdot> tr \<cdot> (TL \<cdot> (Dropwhile (\<lambda>a. a \<in> int A) \<cdot> schA)) | |
| 128 | \<cdot> (TL \<cdot> (Dropwhile (\<lambda>a. a \<in> int B) \<cdot> schB)))))" | |
| 62156 | 129 | apply (rule trans) | 
| 130 | apply (subst mksch_unfold) | |
| 131 | apply (simp add: Consq_def If_and_if) | |
| 132 | apply (simp add: Consq_def) | |
| 133 | done | |
| 19741 | 134 | |
| 135 | lemmas compotr_simps = mksch_UU mksch_nil mksch_cons1 mksch_cons2 mksch_cons3 | |
| 136 | ||
| 137 | declare compotr_simps [simp] | |
| 138 | ||
| 139 | ||
| 62002 | 140 | subsection \<open>COMPOSITIONALITY on TRACE Level\<close> | 
| 19741 | 141 | |
| 62156 | 142 | subsubsection "Lemmata for \<open>\<Longrightarrow>\<close>" | 
| 19741 | 143 | |
| 62156 | 144 | text \<open> | 
| 145 | Consequence out of \<open>ext1_ext2_is_not_act1(2)\<close>, which in turn are | |
| 146 | consequences out of the compatibility of IOA, in particular out of the | |
| 147 | condition that internals are really hidden. | |
| 148 | \<close> | |
| 19741 | 149 | |
| 62156 | 150 | lemma compatibility_consequence1: "(eB \<and> \<not> eA \<longrightarrow> \<not> A) \<longrightarrow> A \<and> (eA \<or> eB) \<longleftrightarrow> eA \<and> A" | 
| 151 | by fast | |
| 19741 | 152 | |
| 153 | ||
| 154 | (* very similar to above, only the commutativity of | is used to make a slight change *) | |
| 62156 | 155 | lemma compatibility_consequence2: "(eB \<and> \<not> eA \<longrightarrow> \<not> A) \<longrightarrow> A \<and> (eB \<or> eA) \<longleftrightarrow> eA \<and> A" | 
| 156 | by fast | |
| 19741 | 157 | |
| 158 | ||
| 62156 | 159 | subsubsection \<open>Lemmata for \<open>\<Longleftarrow>\<close>\<close> | 
| 19741 | 160 | |
| 161 | (* Lemma for substitution of looping assumption in another specific assumption *) | |
| 62156 | 162 | lemma subst_lemma1: "f \<sqsubseteq> g x \<Longrightarrow> x = h x \<Longrightarrow> f \<sqsubseteq> g (h x)" | 
| 163 | by (erule subst) | |
| 19741 | 164 | |
| 165 | (* Lemma for substitution of looping assumption in another specific assumption *) | |
| 62156 | 166 | lemma subst_lemma2: "(f x) = y \<leadsto> g \<Longrightarrow> x = h x \<Longrightarrow> f (h x) = y \<leadsto> g" | 
| 167 | by (erule subst) | |
| 19741 | 168 | |
| 169 | lemma ForallAorB_mksch [rule_format]: | |
| 62156 | 170 | "compatible A B \<Longrightarrow> | 
| 171 | \<forall>schA schB. Forall (\<lambda>x. x \<in> act (A \<parallel> B)) tr \<longrightarrow> | |
| 63549 | 172 | Forall (\<lambda>x. x \<in> act (A \<parallel> B)) (mksch A B \<cdot> tr \<cdot> schA \<cdot> schB)" | 
| 62195 | 173 | apply (Seq_induct tr simp: Forall_def sforall_def mksch_def) | 
| 62156 | 174 | apply auto | 
| 175 | apply (simp add: actions_of_par) | |
| 176 | apply (case_tac "a \<in> act A") | |
| 177 | apply (case_tac "a \<in> act B") | |
| 178 | text \<open>\<open>a \<in> A\<close>, \<open>a \<in> B\<close>\<close> | |
| 179 | apply simp | |
| 180 | apply (rule Forall_Conc_impl [THEN mp]) | |
| 181 | apply (simp add: intA_is_not_actB int_is_act) | |
| 182 | apply (rule Forall_Conc_impl [THEN mp]) | |
| 183 | apply (simp add: intA_is_not_actB int_is_act) | |
| 184 | text \<open>\<open>a \<in> A\<close>, \<open>a \<notin> B\<close>\<close> | |
| 185 | apply simp | |
| 186 | apply (rule Forall_Conc_impl [THEN mp]) | |
| 187 | apply (simp add: intA_is_not_actB int_is_act) | |
| 67613 | 188 | apply (case_tac "a\<in>act B") | 
| 62156 | 189 | text \<open>\<open>a \<notin> A\<close>, \<open>a \<in> B\<close>\<close> | 
| 190 | apply simp | |
| 191 | apply (rule Forall_Conc_impl [THEN mp]) | |
| 192 | apply (simp add: intA_is_not_actB int_is_act) | |
| 193 | text \<open>\<open>a \<notin> A\<close>, \<open>a \<notin> B\<close>\<close> | |
| 194 | apply auto | |
| 195 | done | |
| 19741 | 196 | |
| 62156 | 197 | lemma ForallBnAmksch [rule_format]: | 
| 198 | "compatible B A \<Longrightarrow> | |
| 199 | \<forall>schA schB. Forall (\<lambda>x. x \<in> act B \<and> x \<notin> act A) tr \<longrightarrow> | |
| 63549 | 200 | Forall (\<lambda>x. x \<in> act B \<and> x \<notin> act A) (mksch A B \<cdot> tr \<cdot> schA \<cdot> schB)" | 
| 62195 | 201 | apply (Seq_induct tr simp: Forall_def sforall_def mksch_def) | 
| 62156 | 202 | apply auto | 
| 203 | apply (rule Forall_Conc_impl [THEN mp]) | |
| 204 | apply (simp add: intA_is_not_actB int_is_act) | |
| 205 | done | |
| 19741 | 206 | |
| 62156 | 207 | lemma ForallAnBmksch [rule_format]: | 
| 208 | "compatible A B \<Longrightarrow> | |
| 209 | \<forall>schA schB. Forall (\<lambda>x. x \<in> act A \<and> x \<notin> act B) tr \<longrightarrow> | |
| 63549 | 210 | Forall (\<lambda>x. x \<in> act A \<and> x \<notin> act B) (mksch A B \<cdot> tr \<cdot> schA \<cdot> schB)" | 
| 62195 | 211 | apply (Seq_induct tr simp: Forall_def sforall_def mksch_def) | 
| 62156 | 212 | apply auto | 
| 213 | apply (rule Forall_Conc_impl [THEN mp]) | |
| 214 | apply (simp add: intA_is_not_actB int_is_act) | |
| 215 | done | |
| 19741 | 216 | |
| 62156 | 217 | (* safe_tac makes too many case distinctions with this lemma in the next proof *) | 
| 19741 | 218 | declare FiniteConc [simp del] | 
| 219 | ||
| 62156 | 220 | lemma FiniteL_mksch [rule_format]: | 
| 221 | "Finite tr \<Longrightarrow> is_asig (asig_of A) \<Longrightarrow> is_asig (asig_of B) \<Longrightarrow> | |
| 222 | \<forall>x y. | |
| 223 | Forall (\<lambda>x. x \<in> act A) x \<and> Forall (\<lambda>x. x \<in> act B) y \<and> | |
| 63549 | 224 | Filter (\<lambda>a. a \<in> ext A) \<cdot> x = Filter (\<lambda>a. a \<in> act A) \<cdot> tr \<and> | 
| 225 | Filter (\<lambda>a. a \<in> ext B) \<cdot> y = Filter (\<lambda>a. a \<in> act B) \<cdot> tr \<and> | |
| 226 | Forall (\<lambda>x. x \<in> ext (A \<parallel> B)) tr \<longrightarrow> Finite (mksch A B \<cdot> tr \<cdot> x \<cdot> y)" | |
| 62156 | 227 | apply (erule Seq_Finite_ind) | 
| 228 | apply simp | |
| 229 | text \<open>main case\<close> | |
| 230 | apply simp | |
| 231 | apply auto | |
| 19741 | 232 | |
| 62156 | 233 | text \<open>\<open>a \<in> act A\<close>; \<open>a \<in> act B\<close>\<close> | 
| 234 | apply (frule sym [THEN eq_imp_below, THEN divide_Seq]) | |
| 235 | apply (frule sym [THEN eq_imp_below, THEN divide_Seq]) | |
| 236 | back | |
| 237 | apply (erule conjE)+ | |
| 238 | text \<open>\<open>Finite (tw iA x)\<close> and \<open>Finite (tw iB y)\<close>\<close> | |
| 239 | apply (simp add: not_ext_is_int_or_not_act FiniteConc) | |
| 240 | text \<open>now for conclusion IH applicable, but assumptions have to be transformed\<close> | |
| 63549 | 241 | apply (drule_tac x = "x" and g = "Filter (\<lambda>a. a \<in> act A) \<cdot> s" in subst_lemma2) | 
| 62156 | 242 | apply assumption | 
| 63549 | 243 | apply (drule_tac x = "y" and g = "Filter (\<lambda>a. a \<in> act B) \<cdot> s" in subst_lemma2) | 
| 62156 | 244 | apply assumption | 
| 245 | text \<open>IH\<close> | |
| 246 | apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile) | |
| 19741 | 247 | |
| 62156 | 248 | text \<open>\<open>a \<in> act B\<close>; \<open>a \<notin> act A\<close>\<close> | 
| 249 | apply (frule sym [THEN eq_imp_below, THEN divide_Seq]) | |
| 19741 | 250 | |
| 62156 | 251 | apply (erule conjE)+ | 
| 252 | text \<open>\<open>Finite (tw iB y)\<close>\<close> | |
| 253 | apply (simp add: not_ext_is_int_or_not_act FiniteConc) | |
| 254 | text \<open>now for conclusion IH applicable, but assumptions have to be transformed\<close> | |
| 63549 | 255 | apply (drule_tac x = "y" and g = "Filter (\<lambda>a. a \<in> act B) \<cdot> s" in subst_lemma2) | 
| 62156 | 256 | apply assumption | 
| 257 | text \<open>IH\<close> | |
| 258 | apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile) | |
| 19741 | 259 | |
| 62156 | 260 | text \<open>\<open>a \<notin> act B\<close>; \<open>a \<in> act A\<close>\<close> | 
| 261 | apply (frule sym [THEN eq_imp_below, THEN divide_Seq]) | |
| 19741 | 262 | |
| 62156 | 263 | apply (erule conjE)+ | 
| 264 | text \<open>\<open>Finite (tw iA x)\<close>\<close> | |
| 265 | apply (simp add: not_ext_is_int_or_not_act FiniteConc) | |
| 266 | text \<open>now for conclusion IH applicable, but assumptions have to be transformed\<close> | |
| 63549 | 267 | apply (drule_tac x = "x" and g = "Filter (\<lambda>a. a \<in> act A) \<cdot> s" in subst_lemma2) | 
| 62156 | 268 | apply assumption | 
| 269 | text \<open>IH\<close> | |
| 270 | apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile) | |
| 19741 | 271 | |
| 62156 | 272 | text \<open>\<open>a \<notin> act B\<close>; \<open>a \<notin> act A\<close>\<close> | 
| 273 | apply (fastforce intro!: ext_is_act simp: externals_of_par) | |
| 274 | done | |
| 19741 | 275 | |
| 276 | declare FiniteConc [simp] | |
| 277 | ||
| 278 | declare FilterConc [simp del] | |
| 279 | ||
| 62156 | 280 | lemma reduceA_mksch1 [rule_format]: | 
| 281 | "Finite bs \<Longrightarrow> is_asig (asig_of A) \<Longrightarrow> is_asig (asig_of B) \<Longrightarrow> compatible A B \<Longrightarrow> | |
| 282 | \<forall>y. Forall (\<lambda>x. x \<in> act B) y \<and> Forall (\<lambda>x. x \<in> act B \<and> x \<notin> act A) bs \<and> | |
| 63549 | 283 | Filter (\<lambda>a. a \<in> ext B) \<cdot> y = Filter (\<lambda>a. a \<in> act B) \<cdot> (bs @@ z) \<longrightarrow> | 
| 62156 | 284 | (\<exists>y1 y2. | 
| 63549 | 285 | (mksch A B \<cdot> (bs @@ z) \<cdot> x \<cdot> y) = (y1 @@ (mksch A B \<cdot> z \<cdot> x \<cdot> y2)) \<and> | 
| 62156 | 286 | Forall (\<lambda>x. x \<in> act B \<and> x \<notin> act A) y1 \<and> | 
| 287 | Finite y1 \<and> y = (y1 @@ y2) \<and> | |
| 63549 | 288 | Filter (\<lambda>a. a \<in> ext B) \<cdot> y1 = bs)" | 
| 62156 | 289 | apply (frule_tac A1 = "A" in compat_commute [THEN iffD1]) | 
| 290 | apply (erule Seq_Finite_ind) | |
| 291 | apply (rule allI)+ | |
| 292 | apply (rule impI) | |
| 293 | apply (rule_tac x = "nil" in exI) | |
| 294 | apply (rule_tac x = "y" in exI) | |
| 295 | apply simp | |
| 296 | text \<open>main case\<close> | |
| 297 | apply (rule allI)+ | |
| 298 | apply (rule impI) | |
| 299 | apply simp | |
| 300 | apply (erule conjE)+ | |
| 301 | apply simp | |
| 302 | text \<open>\<open>divide_Seq\<close> on \<open>s\<close>\<close> | |
| 303 | apply (frule sym [THEN eq_imp_below, THEN divide_Seq]) | |
| 304 | apply (erule conjE)+ | |
| 305 | text \<open>transform assumption \<open>f eB y = f B (s @ z)\<close>\<close> | |
| 63549 | 306 | apply (drule_tac x = "y" and g = "Filter (\<lambda>a. a \<in> act B) \<cdot> (s @@ z) " in subst_lemma2) | 
| 62156 | 307 | apply assumption | 
| 308 | apply (simp add: not_ext_is_int_or_not_act FilterConc) | |
| 309 | text \<open>apply IH\<close> | |
| 63549 | 310 | apply (erule_tac x = "TL \<cdot> (Dropwhile (\<lambda>a. a \<in> int B) \<cdot> y)" in allE) | 
| 62156 | 311 | apply (simp add: ForallTL ForallDropwhile FilterConc) | 
| 312 | apply (erule exE)+ | |
| 313 | apply (erule conjE)+ | |
| 314 | apply (simp add: FilterConc) | |
| 315 | text \<open>for replacing IH in conclusion\<close> | |
| 316 | apply (rotate_tac -2) | |
| 317 | text \<open>instantiate \<open>y1a\<close> and \<open>y2a\<close>\<close> | |
| 63549 | 318 | apply (rule_tac x = "Takewhile (\<lambda>a. a \<in> int B) \<cdot> y @@ a \<leadsto> y1" in exI) | 
| 62156 | 319 | apply (rule_tac x = "y2" in exI) | 
| 320 | text \<open>elminate all obligations up to two depending on \<open>Conc_assoc\<close>\<close> | |
| 321 | apply (simp add: intA_is_not_actB int_is_act int_is_not_ext FilterConc) | |
| 322 | apply (simp add: Conc_assoc FilterConc) | |
| 323 | done | |
| 19741 | 324 | |
| 325 | lemmas reduceA_mksch = conjI [THEN [6] conjI [THEN [5] reduceA_mksch1]] | |
| 326 | ||
| 327 | lemma reduceB_mksch1 [rule_format]: | |
| 62156 | 328 | "Finite a_s \<Longrightarrow> is_asig (asig_of A) \<Longrightarrow> is_asig (asig_of B) \<Longrightarrow> compatible A B \<Longrightarrow> | 
| 329 | \<forall>x. Forall (\<lambda>x. x \<in> act A) x \<and> Forall (\<lambda>x. x \<in> act A \<and> x \<notin> act B) a_s \<and> | |
| 63549 | 330 | Filter (\<lambda>a. a \<in> ext A) \<cdot> x = Filter (\<lambda>a. a \<in> act A) \<cdot> (a_s @@ z) \<longrightarrow> | 
| 62156 | 331 | (\<exists>x1 x2. | 
| 63549 | 332 | (mksch A B \<cdot> (a_s @@ z) \<cdot> x \<cdot> y) = (x1 @@ (mksch A B \<cdot> z \<cdot> x2 \<cdot> y)) \<and> | 
| 62156 | 333 | Forall (\<lambda>x. x \<in> act A \<and> x \<notin> act B) x1 \<and> | 
| 334 | Finite x1 \<and> x = (x1 @@ x2) \<and> | |
| 63549 | 335 | Filter (\<lambda>a. a \<in> ext A) \<cdot> x1 = a_s)" | 
| 62156 | 336 | apply (frule_tac A1 = "A" in compat_commute [THEN iffD1]) | 
| 337 | apply (erule Seq_Finite_ind) | |
| 338 | apply (rule allI)+ | |
| 339 | apply (rule impI) | |
| 340 | apply (rule_tac x = "nil" in exI) | |
| 341 | apply (rule_tac x = "x" in exI) | |
| 342 | apply simp | |
| 343 | text \<open>main case\<close> | |
| 344 | apply (rule allI)+ | |
| 345 | apply (rule impI) | |
| 346 | apply simp | |
| 347 | apply (erule conjE)+ | |
| 348 | apply simp | |
| 349 | text \<open>\<open>divide_Seq\<close> on \<open>s\<close>\<close> | |
| 350 | apply (frule sym [THEN eq_imp_below, THEN divide_Seq]) | |
| 351 | apply (erule conjE)+ | |
| 352 | text \<open>transform assumption \<open>f eA x = f A (s @ z)\<close>\<close> | |
| 63549 | 353 | apply (drule_tac x = "x" and g = "Filter (\<lambda>a. a \<in> act A) \<cdot> (s @@ z)" in subst_lemma2) | 
| 62156 | 354 | apply assumption | 
| 355 | apply (simp add: not_ext_is_int_or_not_act FilterConc) | |
| 356 | text \<open>apply IH\<close> | |
| 63549 | 357 | apply (erule_tac x = "TL \<cdot> (Dropwhile (\<lambda>a. a \<in> int A) \<cdot> x)" in allE) | 
| 62156 | 358 | apply (simp add: ForallTL ForallDropwhile FilterConc) | 
| 359 | apply (erule exE)+ | |
| 360 | apply (erule conjE)+ | |
| 361 | apply (simp add: FilterConc) | |
| 362 | text \<open>for replacing IH in conclusion\<close> | |
| 363 | apply (rotate_tac -2) | |
| 364 | text \<open>instantiate \<open>y1a\<close> and \<open>y2a\<close>\<close> | |
| 63549 | 365 | apply (rule_tac x = "Takewhile (\<lambda>a. a \<in> int A) \<cdot> x @@ a \<leadsto> x1" in exI) | 
| 62156 | 366 | apply (rule_tac x = "x2" in exI) | 
| 367 | text \<open>eliminate all obligations up to two depending on \<open>Conc_assoc\<close>\<close> | |
| 368 | apply (simp add: intA_is_not_actB int_is_act int_is_not_ext FilterConc) | |
| 369 | apply (simp add: Conc_assoc FilterConc) | |
| 370 | done | |
| 19741 | 371 | |
| 372 | lemmas reduceB_mksch = conjI [THEN [6] conjI [THEN [5] reduceB_mksch1]] | |
| 373 | ||
| 374 | declare FilterConc [simp] | |
| 375 | ||
| 376 | ||
| 62156 | 377 | subsection \<open>Filtering external actions out of \<open>mksch (tr, schA, schB)\<close> yields the oracle \<open>tr\<close>\<close> | 
| 19741 | 378 | |
| 62005 | 379 | lemma FilterA_mksch_is_tr: | 
| 62156 | 380 | "compatible A B \<Longrightarrow> compatible B A \<Longrightarrow> is_asig (asig_of A) \<Longrightarrow> is_asig (asig_of B) \<Longrightarrow> | 
| 381 | \<forall>schA schB. | |
| 382 | Forall (\<lambda>x. x \<in> act A) schA \<and> | |
| 383 | Forall (\<lambda>x. x \<in> act B) schB \<and> | |
| 384 | Forall (\<lambda>x. x \<in> ext (A \<parallel> B)) tr \<and> | |
| 63549 | 385 | Filter (\<lambda>a. a \<in> act A) \<cdot> tr \<sqsubseteq> Filter (\<lambda>a. a \<in> ext A) \<cdot> schA \<and> | 
| 386 | Filter (\<lambda>a. a \<in> act B) \<cdot> tr \<sqsubseteq> Filter (\<lambda>a. a \<in> ext B) \<cdot> schB | |
| 387 | \<longrightarrow> Filter (\<lambda>a. a \<in> ext (A \<parallel> B)) \<cdot> (mksch A B \<cdot> tr \<cdot> schA \<cdot> schB) = tr" | |
| 62195 | 388 | apply (Seq_induct tr simp: Forall_def sforall_def mksch_def) | 
| 62156 | 389 | text \<open>main case\<close> | 
| 390 | text \<open>splitting into 4 cases according to \<open>a \<in> A\<close>, \<open>a \<in> B\<close>\<close> | |
| 391 | apply auto | |
| 19741 | 392 | |
| 62156 | 393 | text \<open>Case \<open>a \<in> A\<close>, \<open>a \<in> B\<close>\<close> | 
| 394 | apply (frule divide_Seq) | |
| 395 | apply (frule divide_Seq) | |
| 396 | back | |
| 397 | apply (erule conjE)+ | |
| 398 | text \<open>filtering internals of \<open>A\<close> in \<open>schA\<close> and of \<open>B\<close> in \<open>schB\<close> is \<open>nil\<close>\<close> | |
| 399 | apply (simp add: not_ext_is_int_or_not_act externals_of_par intA_is_not_extB int_is_not_ext) | |
| 400 | text \<open>conclusion of IH ok, but assumptions of IH have to be transformed\<close> | |
| 401 | apply (drule_tac x = "schA" in subst_lemma1) | |
| 402 | apply assumption | |
| 403 | apply (drule_tac x = "schB" in subst_lemma1) | |
| 404 | apply assumption | |
| 405 | text \<open>IH\<close> | |
| 406 | apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile) | |
| 19741 | 407 | |
| 62156 | 408 | text \<open>Case \<open>a \<in> A\<close>, \<open>a \<notin> B\<close>\<close> | 
| 409 | apply (frule divide_Seq) | |
| 410 | apply (erule conjE)+ | |
| 411 | text \<open>filtering internals of \<open>A\<close> is \<open>nil\<close>\<close> | |
| 412 | apply (simp add: not_ext_is_int_or_not_act externals_of_par intA_is_not_extB int_is_not_ext) | |
| 413 | apply (drule_tac x = "schA" in subst_lemma1) | |
| 414 | apply assumption | |
| 415 | apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile) | |
| 19741 | 416 | |
| 62156 | 417 | text \<open>Case \<open>a \<in> B\<close>, \<open>a \<notin> A\<close>\<close> | 
| 418 | apply (frule divide_Seq) | |
| 419 | apply (erule conjE)+ | |
| 420 | text \<open>filtering internals of \<open>A\<close> is \<open>nil\<close>\<close> | |
| 421 | apply (simp add: not_ext_is_int_or_not_act externals_of_par intA_is_not_extB int_is_not_ext) | |
| 422 | apply (drule_tac x = "schB" in subst_lemma1) | |
| 423 | back | |
| 424 | apply assumption | |
| 425 | apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile) | |
| 19741 | 426 | |
| 62156 | 427 | text \<open>Case \<open>a \<notin> A\<close>, \<open>a \<notin> B\<close>\<close> | 
| 428 | apply (fastforce intro!: ext_is_act simp: externals_of_par) | |
| 429 | done | |
| 19741 | 430 | |
| 431 | ||
| 432 | subsection" Filter of mksch(tr,schA,schB) to A is schA -- take lemma proof" | |
| 433 | ||
| 62156 | 434 | lemma FilterAmksch_is_schA: | 
| 435 | "compatible A B \<Longrightarrow> compatible B A \<Longrightarrow> is_asig (asig_of A) \<Longrightarrow> is_asig (asig_of B) \<Longrightarrow> | |
| 436 | Forall (\<lambda>x. x \<in> ext (A \<parallel> B)) tr \<and> | |
| 437 | Forall (\<lambda>x. x \<in> act A) schA \<and> | |
| 438 | Forall (\<lambda>x. x \<in> act B) schB \<and> | |
| 63549 | 439 | Filter (\<lambda>a. a \<in> ext A) \<cdot> schA = Filter (\<lambda>a. a \<in> act A) \<cdot> tr \<and> | 
| 440 | Filter (\<lambda>a. a \<in> ext B) \<cdot> schB = Filter (\<lambda>a. a \<in> act B) \<cdot> tr \<and> | |
| 62156 | 441 | LastActExtsch A schA \<and> LastActExtsch B schB | 
| 63549 | 442 | \<longrightarrow> Filter (\<lambda>a. a \<in> act A) \<cdot> (mksch A B \<cdot> tr \<cdot> schA \<cdot> schB) = schA" | 
| 62156 | 443 | apply (intro strip) | 
| 444 | apply (rule seq.take_lemma) | |
| 445 | apply (rule mp) | |
| 446 | prefer 2 apply assumption | |
| 447 | back back back back | |
| 448 | apply (rule_tac x = "schA" in spec) | |
| 449 | apply (rule_tac x = "schB" in spec) | |
| 450 | apply (rule_tac x = "tr" in spec) | |
| 69597 | 451 | apply (tactic "thin_tac' \<^context> 5 1") | 
| 62156 | 452 | apply (rule nat_less_induct) | 
| 453 | apply (rule allI)+ | |
| 454 | apply (rename_tac tr schB schA) | |
| 455 | apply (intro strip) | |
| 456 | apply (erule conjE)+ | |
| 19741 | 457 | |
| 62156 | 458 | apply (case_tac "Forall (\<lambda>x. x \<in> act B \<and> x \<notin> act A) tr") | 
| 19741 | 459 | |
| 62156 | 460 | apply (rule seq_take_lemma [THEN iffD2, THEN spec]) | 
| 69597 | 461 | apply (tactic "thin_tac' \<^context> 5 1") | 
| 19741 | 462 | |
| 463 | ||
| 62156 | 464 | apply (case_tac "Finite tr") | 
| 19741 | 465 | |
| 62156 | 466 | text \<open>both sides of this equation are \<open>nil\<close>\<close> | 
| 467 | apply (subgoal_tac "schA = nil") | |
| 468 | apply simp | |
| 469 | text \<open>first side: \<open>mksch = nil\<close>\<close> | |
| 470 | apply (auto intro!: ForallQFilterPnil ForallBnAmksch FiniteL_mksch)[1] | |
| 471 | text \<open>second side: \<open>schA = nil\<close>\<close> | |
| 472 | apply (erule_tac A = "A" in LastActExtimplnil) | |
| 473 | apply simp | |
| 474 | apply (erule_tac Q = "\<lambda>x. x \<in> act B \<and> x \<notin> act A" in ForallQFilterPnil) | |
| 475 | apply assumption | |
| 476 | apply fast | |
| 19741 | 477 | |
| 62156 | 478 | text \<open>case \<open>\<not> Finite s\<close>\<close> | 
| 19741 | 479 | |
| 62156 | 480 | text \<open>both sides of this equation are \<open>UU\<close>\<close> | 
| 481 | apply (subgoal_tac "schA = UU") | |
| 482 | apply simp | |
| 483 | text \<open>first side: \<open>mksch = UU\<close>\<close> | |
| 484 | apply (auto intro!: ForallQFilterPUU finiteR_mksch' ForallBnAmksch)[1] | |
| 485 | text \<open>\<open>schA = UU\<close>\<close> | |
| 486 | apply (erule_tac A = "A" in LastActExtimplUU) | |
| 487 | apply simp | |
| 488 | apply (erule_tac Q = "\<lambda>x. x \<in> act B \<and> x \<notin> act A" in ForallQFilterPUU) | |
| 489 | apply assumption | |
| 490 | apply fast | |
| 19741 | 491 | |
| 62156 | 492 | text \<open>case \<open>\<not> Forall (\<lambda>x. x \<in> act B \<and> x \<notin> act A) s\<close>\<close> | 
| 19741 | 493 | |
| 62156 | 494 | apply (drule divide_Seq3) | 
| 19741 | 495 | |
| 62156 | 496 | apply (erule exE)+ | 
| 497 | apply (erule conjE)+ | |
| 498 | apply hypsubst | |
| 19741 | 499 | |
| 62156 | 500 | text \<open>bring in lemma \<open>reduceA_mksch\<close>\<close> | 
| 501 | apply (frule_tac x = "schA" and y = "schB" and A = "A" and B = "B" in reduceA_mksch) | |
| 502 | apply assumption+ | |
| 503 | apply (erule exE)+ | |
| 504 | apply (erule conjE)+ | |
| 19741 | 505 | |
| 62156 | 506 | text \<open>use \<open>reduceA_mksch\<close> to rewrite conclusion\<close> | 
| 507 | apply hypsubst | |
| 508 | apply simp | |
| 19741 | 509 | |
| 62156 | 510 | text \<open>eliminate the \<open>B\<close>-only prefix\<close> | 
| 19741 | 511 | |
| 63549 | 512 | apply (subgoal_tac "(Filter (\<lambda>a. a \<in> act A) \<cdot> y1) = nil") | 
| 62156 | 513 | apply (erule_tac [2] ForallQFilterPnil) | 
| 514 | prefer 2 apply assumption | |
| 515 | prefer 2 apply fast | |
| 19741 | 516 | |
| 62156 | 517 | text \<open>Now real recursive step follows (in \<open>y\<close>)\<close> | 
| 19741 | 518 | |
| 62156 | 519 | apply simp | 
| 520 | apply (case_tac "x \<in> act A") | |
| 521 | apply (case_tac "x \<notin> act B") | |
| 522 | apply (rotate_tac -2) | |
| 523 | apply simp | |
| 19741 | 524 | |
| 63549 | 525 | apply (subgoal_tac "Filter (\<lambda>a. a \<in> act A \<and> a \<in> ext B) \<cdot> y1 = nil") | 
| 62156 | 526 | apply (rotate_tac -1) | 
| 527 | apply simp | |
| 528 | text \<open>eliminate introduced subgoal 2\<close> | |
| 529 | apply (erule_tac [2] ForallQFilterPnil) | |
| 530 | prefer 2 apply assumption | |
| 531 | prefer 2 apply fast | |
| 19741 | 532 | |
| 62156 | 533 | text \<open>bring in \<open>divide_Seq\<close> for \<open>s\<close>\<close> | 
| 534 | apply (frule sym [THEN eq_imp_below, THEN divide_Seq]) | |
| 535 | apply (erule conjE)+ | |
| 19741 | 536 | |
| 62156 | 537 | text \<open>subst \<open>divide_Seq\<close> in conclusion, but only at the rightest occurrence\<close> | 
| 538 | apply (rule_tac t = "schA" in ssubst) | |
| 539 | back | |
| 540 | back | |
| 541 | back | |
| 542 | apply assumption | |
| 19741 | 543 | |
| 62156 | 544 | text \<open>reduce trace_takes from \<open>n\<close> to strictly smaller \<open>k\<close>\<close> | 
| 545 | apply (rule take_reduction) | |
| 19741 | 546 | |
| 62156 | 547 | text \<open>\<open>f A (tw iA) = tw \<not> eA\<close>\<close> | 
| 548 | apply (simp add: int_is_act not_ext_is_int_or_not_act) | |
| 549 | apply (rule refl) | |
| 550 | apply (simp add: int_is_act not_ext_is_int_or_not_act) | |
| 551 | apply (rotate_tac -11) | |
| 19741 | 552 | |
| 62156 | 553 | text \<open>now conclusion fulfills induction hypothesis, but assumptions are not ready\<close> | 
| 19741 | 554 | |
| 62156 | 555 | text \<open>assumption \<open>Forall tr\<close>\<close> | 
| 556 | text \<open>assumption \<open>schB\<close>\<close> | |
| 557 | apply (simp add: ext_and_act) | |
| 558 | text \<open>assumption \<open>schA\<close>\<close> | |
| 63549 | 559 | apply (drule_tac x = "schA" and g = "Filter (\<lambda>a. a \<in> act A) \<cdot> rs" in subst_lemma2) | 
| 62156 | 560 | apply assumption | 
| 561 | apply (simp add: int_is_not_ext) | |
| 562 | text \<open>assumptions concerning \<open>LastActExtsch\<close>, cannot be rewritten, as \<open>LastActExtsmalli\<close> are looping\<close> | |
| 563 | apply (drule_tac sch = "schA" and P = "\<lambda>a. a \<in> int A" in LastActExtsmall1) | |
| 564 | apply (frule_tac ?sch1.0 = "y1" in LastActExtsmall2) | |
| 565 | apply assumption | |
| 19741 | 566 | |
| 62156 | 567 | text \<open>assumption \<open>Forall schA\<close>\<close> | 
| 568 | apply (drule_tac s = "schA" and P = "Forall (\<lambda>x. x \<in> act A) " in subst) | |
| 569 | apply assumption | |
| 570 | apply (simp add: int_is_act) | |
| 19741 | 571 | |
| 62156 | 572 | text \<open>case \<open>x \<in> actions (asig_of A) \<and> x \<in> actions (asig_of B)\<close>\<close> | 
| 19741 | 573 | |
| 62156 | 574 | apply (rotate_tac -2) | 
| 575 | apply simp | |
| 19741 | 576 | |
| 63549 | 577 | apply (subgoal_tac "Filter (\<lambda>a. a \<in> act A \<and> a \<in> ext B) \<cdot> y1 = nil") | 
| 62156 | 578 | apply (rotate_tac -1) | 
| 579 | apply simp | |
| 580 | text \<open>eliminate introduced subgoal 2\<close> | |
| 581 | apply (erule_tac [2] ForallQFilterPnil) | |
| 582 | prefer 2 apply assumption | |
| 583 | prefer 2 apply fast | |
| 19741 | 584 | |
| 62156 | 585 | text \<open>bring in \<open>divide_Seq\<close> for \<open>s\<close>\<close> | 
| 586 | apply (frule sym [THEN eq_imp_below, THEN divide_Seq]) | |
| 587 | apply (erule conjE)+ | |
| 19741 | 588 | |
| 62156 | 589 | text \<open>subst \<open>divide_Seq\<close> in conclusion, but only at the rightmost occurrence\<close> | 
| 590 | apply (rule_tac t = "schA" in ssubst) | |
| 591 | back | |
| 592 | back | |
| 593 | back | |
| 594 | apply assumption | |
| 19741 | 595 | |
| 62156 | 596 | text \<open>\<open>f A (tw iA) = tw \<not> eA\<close>\<close> | 
| 597 | apply (simp add: int_is_act not_ext_is_int_or_not_act) | |
| 19741 | 598 | |
| 62156 | 599 | text \<open>rewrite assumption forall and \<open>schB\<close>\<close> | 
| 600 | apply (rotate_tac 13) | |
| 601 | apply (simp add: ext_and_act) | |
| 19741 | 602 | |
| 62156 | 603 | text \<open>\<open>divide_Seq\<close> for \<open>schB2\<close>\<close> | 
| 604 | apply (frule_tac y = "y2" in sym [THEN eq_imp_below, THEN divide_Seq]) | |
| 605 | apply (erule conjE)+ | |
| 606 | text \<open>assumption \<open>schA\<close>\<close> | |
| 67613 | 607 | apply (drule_tac x = "schA" and g = "Filter (\<lambda>a. a\<in>act A) \<cdot>rs" in subst_lemma2) | 
| 62156 | 608 | apply assumption | 
| 609 | apply (simp add: int_is_not_ext) | |
| 19741 | 610 | |
| 62156 | 611 | text \<open>\<open>f A (tw iB schB2) = nil\<close>\<close> | 
| 612 | apply (simp add: int_is_not_ext not_ext_is_int_or_not_act intA_is_not_actB) | |
| 19741 | 613 | |
| 614 | ||
| 62156 | 615 | text \<open>reduce trace_takes from \<open>n\<close> to strictly smaller \<open>k\<close>\<close> | 
| 616 | apply (rule take_reduction) | |
| 617 | apply (rule refl) | |
| 618 | apply (rule refl) | |
| 19741 | 619 | |
| 62156 | 620 | text \<open>now conclusion fulfills induction hypothesis, but assumptions are not all ready\<close> | 
| 19741 | 621 | |
| 62156 | 622 | text \<open>assumption \<open>schB\<close>\<close> | 
| 63549 | 623 | apply (drule_tac x = "y2" and g = "Filter (\<lambda>a. a \<in> act B) \<cdot> rs" in subst_lemma2) | 
| 62156 | 624 | apply assumption | 
| 625 | apply (simp add: intA_is_not_actB int_is_not_ext) | |
| 19741 | 626 | |
| 62156 | 627 | text \<open>conclusions concerning \<open>LastActExtsch\<close>, cannot be rewritten, as \<open>LastActExtsmalli\<close> are looping\<close> | 
| 628 | apply (drule_tac sch = "schA" and P = "\<lambda>a. a \<in> int A" in LastActExtsmall1) | |
| 629 | apply (frule_tac ?sch1.0 = "y1" in LastActExtsmall2) | |
| 630 | apply assumption | |
| 631 | apply (drule_tac sch = "y2" and P = "\<lambda>a. a \<in> int B" in LastActExtsmall1) | |
| 19741 | 632 | |
| 62156 | 633 | text \<open>assumption \<open>Forall schA\<close>, and \<open>Forall schA\<close> are performed by \<open>ForallTL\<close>, \<open>ForallDropwhile\<close>\<close> | 
| 634 | apply (simp add: ForallTL ForallDropwhile) | |
| 19741 | 635 | |
| 62156 | 636 | text \<open>case \<open>x \<notin> A \<and> x \<in> B\<close>\<close> | 
| 637 | text \<open>cannot occur, as just this case has been scheduled out before as the \<open>B\<close>-only prefix\<close> | |
| 638 | apply (case_tac "x \<in> act B") | |
| 639 | apply fast | |
| 19741 | 640 | |
| 62156 | 641 | text \<open>case \<open>x \<notin> A \<and> x \<notin> B\<close>\<close> | 
| 642 | text \<open>cannot occur because of assumption: \<open>Forall (a \<in> ext A \<or> a \<in> ext B)\<close>\<close> | |
| 643 | apply (rotate_tac -9) | |
| 644 | text \<open>reduce forall assumption from \<open>tr\<close> to \<open>x \<leadsto> rs\<close>\<close> | |
| 645 | apply (simp add: externals_of_par) | |
| 646 | apply (fast intro!: ext_is_act) | |
| 647 | done | |
| 19741 | 648 | |
| 649 | ||
| 62156 | 650 | subsection \<open>Filter of \<open>mksch (tr, schA, schB)\<close> to \<open>B\<close> is \<open>schB\<close> -- take lemma proof\<close> | 
| 19741 | 651 | |
| 62156 | 652 | lemma FilterBmksch_is_schB: | 
| 653 | "compatible A B \<Longrightarrow> compatible B A \<Longrightarrow> is_asig (asig_of A) \<Longrightarrow> is_asig (asig_of B) \<Longrightarrow> | |
| 654 | Forall (\<lambda>x. x \<in> ext (A \<parallel> B)) tr \<and> | |
| 655 | Forall (\<lambda>x. x \<in> act A) schA \<and> | |
| 656 | Forall (\<lambda>x. x \<in> act B) schB \<and> | |
| 63549 | 657 | Filter (\<lambda>a. a \<in> ext A) \<cdot> schA = Filter (\<lambda>a. a \<in> act A) \<cdot> tr \<and> | 
| 658 | Filter (\<lambda>a. a \<in> ext B) \<cdot> schB = Filter (\<lambda>a. a \<in> act B) \<cdot> tr \<and> | |
| 62156 | 659 | LastActExtsch A schA \<and> LastActExtsch B schB | 
| 63549 | 660 | \<longrightarrow> Filter (\<lambda>a. a \<in> act B) \<cdot> (mksch A B \<cdot> tr \<cdot> schA \<cdot> schB) = schB" | 
| 62156 | 661 | apply (intro strip) | 
| 662 | apply (rule seq.take_lemma) | |
| 663 | apply (rule mp) | |
| 664 | prefer 2 apply assumption | |
| 665 | back back back back | |
| 666 | apply (rule_tac x = "schA" in spec) | |
| 667 | apply (rule_tac x = "schB" in spec) | |
| 668 | apply (rule_tac x = "tr" in spec) | |
| 69597 | 669 | apply (tactic "thin_tac' \<^context> 5 1") | 
| 62156 | 670 | apply (rule nat_less_induct) | 
| 671 | apply (rule allI)+ | |
| 672 | apply (rename_tac tr schB schA) | |
| 673 | apply (intro strip) | |
| 674 | apply (erule conjE)+ | |
| 19741 | 675 | |
| 62156 | 676 | apply (case_tac "Forall (\<lambda>x. x \<in> act A \<and> x \<notin> act B) tr") | 
| 19741 | 677 | |
| 62156 | 678 | apply (rule seq_take_lemma [THEN iffD2, THEN spec]) | 
| 69597 | 679 | apply (tactic "thin_tac' \<^context> 5 1") | 
| 19741 | 680 | |
| 62156 | 681 | apply (case_tac "Finite tr") | 
| 19741 | 682 | |
| 62156 | 683 | text \<open>both sides of this equation are \<open>nil\<close>\<close> | 
| 684 | apply (subgoal_tac "schB = nil") | |
| 685 | apply simp | |
| 686 | text \<open>first side: \<open>mksch = nil\<close>\<close> | |
| 687 | apply (auto intro!: ForallQFilterPnil ForallAnBmksch FiniteL_mksch)[1] | |
| 688 | text \<open>second side: \<open>schA = nil\<close>\<close> | |
| 689 | apply (erule_tac A = "B" in LastActExtimplnil) | |
| 690 | apply (simp (no_asm_simp)) | |
| 691 | apply (erule_tac Q = "\<lambda>x. x \<in> act A \<and> x \<notin> act B" in ForallQFilterPnil) | |
| 692 | apply assumption | |
| 693 | apply fast | |
| 19741 | 694 | |
| 62156 | 695 | text \<open>case \<open>\<not> Finite tr\<close>\<close> | 
| 19741 | 696 | |
| 62156 | 697 | text \<open>both sides of this equation are \<open>UU\<close>\<close> | 
| 698 | apply (subgoal_tac "schB = UU") | |
| 699 | apply simp | |
| 700 | text \<open>first side: \<open>mksch = UU\<close>\<close> | |
| 701 | apply (force intro!: ForallQFilterPUU finiteR_mksch' ForallAnBmksch) | |
| 702 | text \<open>\<open>schA = UU\<close>\<close> | |
| 703 | apply (erule_tac A = "B" in LastActExtimplUU) | |
| 704 | apply simp | |
| 705 | apply (erule_tac Q = "\<lambda>x. x \<in> act A \<and> x \<notin> act B" in ForallQFilterPUU) | |
| 706 | apply assumption | |
| 707 | apply fast | |
| 19741 | 708 | |
| 62156 | 709 | text \<open>case \<open>\<not> Forall (\<lambda>x. x \<in> act B \<and> x \<notin> act A) s\<close>\<close> | 
| 19741 | 710 | |
| 62156 | 711 | apply (drule divide_Seq3) | 
| 19741 | 712 | |
| 62156 | 713 | apply (erule exE)+ | 
| 714 | apply (erule conjE)+ | |
| 715 | apply hypsubst | |
| 19741 | 716 | |
| 62156 | 717 | text \<open>bring in lemma \<open>reduceB_mksch\<close>\<close> | 
| 718 | apply (frule_tac y = "schB" and x = "schA" and A = "A" and B = "B" in reduceB_mksch) | |
| 719 | apply assumption+ | |
| 720 | apply (erule exE)+ | |
| 721 | apply (erule conjE)+ | |
| 19741 | 722 | |
| 62156 | 723 | text \<open>use \<open>reduceB_mksch\<close> to rewrite conclusion\<close> | 
| 724 | apply hypsubst | |
| 725 | apply simp | |
| 19741 | 726 | |
| 62156 | 727 | text \<open>eliminate the \<open>A\<close>-only prefix\<close> | 
| 19741 | 728 | |
| 63549 | 729 | apply (subgoal_tac "(Filter (\<lambda>a. a \<in> act B) \<cdot> x1) = nil") | 
| 62156 | 730 | apply (erule_tac [2] ForallQFilterPnil) | 
| 731 | prefer 2 apply (assumption) | |
| 732 | prefer 2 apply (fast) | |
| 19741 | 733 | |
| 62156 | 734 | text \<open>Now real recursive step follows (in \<open>x\<close>)\<close> | 
| 19741 | 735 | |
| 62156 | 736 | apply simp | 
| 737 | apply (case_tac "x \<in> act B") | |
| 738 | apply (case_tac "x \<notin> act A") | |
| 739 | apply (rotate_tac -2) | |
| 740 | apply simp | |
| 19741 | 741 | |
| 63549 | 742 | apply (subgoal_tac "Filter (\<lambda>a. a \<in> act B \<and> a \<in> ext A) \<cdot> x1 = nil") | 
| 62156 | 743 | apply (rotate_tac -1) | 
| 744 | apply simp | |
| 745 | text \<open>eliminate introduced subgoal 2\<close> | |
| 746 | apply (erule_tac [2] ForallQFilterPnil) | |
| 747 | prefer 2 apply assumption | |
| 748 | prefer 2 apply fast | |
| 19741 | 749 | |
| 62156 | 750 | text \<open>bring in \<open>divide_Seq\<close> for \<open>s\<close>\<close> | 
| 751 | apply (frule sym [THEN eq_imp_below, THEN divide_Seq]) | |
| 752 | apply (erule conjE)+ | |
| 19741 | 753 | |
| 62156 | 754 | text \<open>subst \<open>divide_Seq\<close> in conclusion, but only at the rightmost occurrence\<close> | 
| 755 | apply (rule_tac t = "schB" in ssubst) | |
| 756 | back | |
| 757 | back | |
| 758 | back | |
| 759 | apply assumption | |
| 19741 | 760 | |
| 62156 | 761 | text \<open>reduce \<open>trace_takes\<close> from \<open>n\<close> to strictly smaller \<open>k\<close>\<close> | 
| 762 | apply (rule take_reduction) | |
| 19741 | 763 | |
| 62156 | 764 | text \<open>\<open>f B (tw iB) = tw \<not> eB\<close>\<close> | 
| 765 | apply (simp add: int_is_act not_ext_is_int_or_not_act) | |
| 766 | apply (rule refl) | |
| 767 | apply (simp add: int_is_act not_ext_is_int_or_not_act) | |
| 768 | apply (rotate_tac -11) | |
| 19741 | 769 | |
| 62156 | 770 | text \<open>now conclusion fulfills induction hypothesis, but assumptions are not ready\<close> | 
| 19741 | 771 | |
| 62156 | 772 | text \<open>assumption \<open>schA\<close>\<close> | 
| 773 | apply (simp add: ext_and_act) | |
| 774 | text \<open>assumption \<open>schB\<close>\<close> | |
| 63549 | 775 | apply (drule_tac x = "schB" and g = "Filter (\<lambda>a. a \<in> act B) \<cdot> rs" in subst_lemma2) | 
| 62156 | 776 | apply assumption | 
| 777 | apply (simp add: int_is_not_ext) | |
| 778 | text \<open>assumptions concerning \<open>LastActExtsch\<close>, cannot be rewritten, as \<open>LastActExtsmalli\<close> are looping\<close> | |
| 779 | apply (drule_tac sch = "schB" and P = "\<lambda>a. a \<in> int B" in LastActExtsmall1) | |
| 780 | apply (frule_tac ?sch1.0 = "x1" in LastActExtsmall2) | |
| 781 | apply assumption | |
| 19741 | 782 | |
| 62156 | 783 | text \<open>assumption \<open>Forall schB\<close>\<close> | 
| 784 | apply (drule_tac s = "schB" and P = "Forall (\<lambda>x. x \<in> act B)" in subst) | |
| 785 | apply assumption | |
| 786 | apply (simp add: int_is_act) | |
| 19741 | 787 | |
| 62156 | 788 | text \<open>case \<open>x \<in> actions (asig_of A) \<and> x \<in> actions (asig_of B)\<close>\<close> | 
| 19741 | 789 | |
| 62156 | 790 | apply (rotate_tac -2) | 
| 791 | apply simp | |
| 19741 | 792 | |
| 63549 | 793 | apply (subgoal_tac "Filter (\<lambda>a. a \<in> act B \<and> a \<in> ext A) \<cdot> x1 = nil") | 
| 62156 | 794 | apply (rotate_tac -1) | 
| 795 | apply simp | |
| 796 | text \<open>eliminate introduced subgoal 2\<close> | |
| 797 | apply (erule_tac [2] ForallQFilterPnil) | |
| 798 | prefer 2 apply assumption | |
| 799 | prefer 2 apply fast | |
| 19741 | 800 | |
| 62156 | 801 | text \<open>bring in \<open>divide_Seq\<close> for \<open>s\<close>\<close> | 
| 802 | apply (frule sym [THEN eq_imp_below, THEN divide_Seq]) | |
| 803 | apply (erule conjE)+ | |
| 19741 | 804 | |
| 62156 | 805 | text \<open>subst \<open>divide_Seq\<close> in conclusion, but only at the rightmost occurrence\<close> | 
| 806 | apply (rule_tac t = "schB" in ssubst) | |
| 807 | back | |
| 808 | back | |
| 809 | back | |
| 810 | apply assumption | |
| 19741 | 811 | |
| 62156 | 812 | text \<open>\<open>f B (tw iB) = tw \<not> eB\<close>\<close> | 
| 813 | apply (simp add: int_is_act not_ext_is_int_or_not_act) | |
| 19741 | 814 | |
| 62156 | 815 | text \<open>rewrite assumption forall and \<open>schB\<close>\<close> | 
| 816 | apply (rotate_tac 13) | |
| 817 | apply (simp add: ext_and_act) | |
| 19741 | 818 | |
| 62156 | 819 | text \<open>\<open>divide_Seq\<close> for \<open>schB2\<close>\<close> | 
| 820 | apply (frule_tac y = "x2" in sym [THEN eq_imp_below, THEN divide_Seq]) | |
| 821 | apply (erule conjE)+ | |
| 822 | text \<open>assumption \<open>schA\<close>\<close> | |
| 63549 | 823 | apply (drule_tac x = "schB" and g = "Filter (\<lambda>a. a \<in> act B) \<cdot> rs" in subst_lemma2) | 
| 62156 | 824 | apply assumption | 
| 825 | apply (simp add: int_is_not_ext) | |
| 19741 | 826 | |
| 62156 | 827 | text \<open>\<open>f B (tw iA schA2) = nil\<close>\<close> | 
| 828 | apply (simp add: int_is_not_ext not_ext_is_int_or_not_act intA_is_not_actB) | |
| 19741 | 829 | |
| 830 | ||
| 62156 | 831 | text \<open>reduce \<open>trace_takes from \<open>n\<close> to strictly smaller \<open>k\<close>\<close>\<close> | 
| 832 | apply (rule take_reduction) | |
| 833 | apply (rule refl) | |
| 834 | apply (rule refl) | |
| 19741 | 835 | |
| 62156 | 836 | text \<open>now conclusion fulfills induction hypothesis, but assumptions are not all ready\<close> | 
| 19741 | 837 | |
| 62156 | 838 | text \<open>assumption \<open>schA\<close>\<close> | 
| 63549 | 839 | apply (drule_tac x = "x2" and g = "Filter (\<lambda>a. a \<in> act A) \<cdot> rs" in subst_lemma2) | 
| 62156 | 840 | apply assumption | 
| 841 | apply (simp add: intA_is_not_actB int_is_not_ext) | |
| 19741 | 842 | |
| 62156 | 843 | text \<open>conclusions concerning \<open>LastActExtsch\<close>, cannot be rewritten, as \<open>LastActExtsmalli\<close> are looping\<close> | 
| 67613 | 844 | apply (drule_tac sch = "schB" and P = "\<lambda>a. a\<in>int B" in LastActExtsmall1) | 
| 62156 | 845 | apply (frule_tac ?sch1.0 = "x1" in LastActExtsmall2) | 
| 846 | apply assumption | |
| 67613 | 847 | apply (drule_tac sch = "x2" and P = "\<lambda>a. a\<in>int A" in LastActExtsmall1) | 
| 19741 | 848 | |
| 62156 | 849 | text \<open>assumption \<open>Forall schA\<close>, and \<open>Forall schA\<close> are performed by \<open>ForallTL\<close>, \<open>ForallDropwhile\<close>\<close> | 
| 850 | apply (simp add: ForallTL ForallDropwhile) | |
| 19741 | 851 | |
| 62156 | 852 | text \<open>case \<open>x \<notin> B \<and> x \<in> A\<close>\<close> | 
| 853 | text \<open>cannot occur, as just this case has been scheduled out before as the \<open>B\<close>-only prefix\<close> | |
| 854 | apply (case_tac "x \<in> act A") | |
| 855 | apply fast | |
| 19741 | 856 | |
| 62156 | 857 | text \<open>case \<open>x \<notin> B \<and> x \<notin> A\<close>\<close> | 
| 858 | text \<open>cannot occur because of assumption: \<open>Forall (a \<in> ext A \<or> a \<in> ext B)\<close>\<close> | |
| 859 | apply (rotate_tac -9) | |
| 860 | text \<open>reduce forall assumption from \<open>tr\<close> to \<open>x \<leadsto> rs\<close>\<close> | |
| 861 | apply (simp add: externals_of_par) | |
| 862 | apply (fast intro!: ext_is_act) | |
| 863 | done | |
| 19741 | 864 | |
| 865 | ||
| 866 | subsection "COMPOSITIONALITY on TRACE Level -- Main Theorem" | |
| 867 | ||
| 62156 | 868 | theorem compositionality_tr: | 
| 869 | "is_trans_of A \<Longrightarrow> is_trans_of B \<Longrightarrow> compatible A B \<Longrightarrow> compatible B A \<Longrightarrow> | |
| 870 | is_asig (asig_of A) \<Longrightarrow> is_asig (asig_of B) \<Longrightarrow> | |
| 871 | tr \<in> traces (A \<parallel> B) \<longleftrightarrow> | |
| 63549 | 872 | Filter (\<lambda>a. a \<in> act A) \<cdot> tr \<in> traces A \<and> | 
| 873 | Filter (\<lambda>a. a \<in> act B) \<cdot> tr \<in> traces B \<and> | |
| 62156 | 874 | Forall (\<lambda>x. x \<in> ext(A \<parallel> B)) tr" | 
| 875 | apply (simp add: traces_def has_trace_def) | |
| 876 | apply auto | |
| 19741 | 877 | |
| 62156 | 878 | text \<open>\<open>\<Longrightarrow>\<close>\<close> | 
| 879 | text \<open>There is a schedule of \<open>A\<close>\<close> | |
| 63549 | 880 | apply (rule_tac x = "Filter (\<lambda>a. a \<in> act A) \<cdot> sch" in bexI) | 
| 62156 | 881 | prefer 2 | 
| 882 | apply (simp add: compositionality_sch) | |
| 883 | apply (simp add: compatibility_consequence1 externals_of_par ext1_ext2_is_not_act1) | |
| 884 | text \<open>There is a schedule of \<open>B\<close>\<close> | |
| 63549 | 885 | apply (rule_tac x = "Filter (\<lambda>a. a \<in> act B) \<cdot> sch" in bexI) | 
| 62156 | 886 | prefer 2 | 
| 887 | apply (simp add: compositionality_sch) | |
| 888 | apply (simp add: compatibility_consequence2 externals_of_par ext1_ext2_is_not_act2) | |
| 889 | text \<open>Traces of \<open>A \<parallel> B\<close> have only external actions from \<open>A\<close> or \<open>B\<close>\<close> | |
| 890 | apply (rule ForallPFilterP) | |
| 19741 | 891 | |
| 62156 | 892 | text \<open>\<open>\<Longleftarrow>\<close>\<close> | 
| 19741 | 893 | |
| 62156 | 894 | text \<open>replace \<open>schA\<close> and \<open>schB\<close> by \<open>Cut schA\<close> and \<open>Cut schB\<close>\<close> | 
| 895 | apply (drule exists_LastActExtsch) | |
| 896 | apply assumption | |
| 897 | apply (drule exists_LastActExtsch) | |
| 898 | apply assumption | |
| 899 | apply (erule exE)+ | |
| 900 | apply (erule conjE)+ | |
| 901 | text \<open>Schedules of A(B) have only actions of A(B)\<close> | |
| 902 | apply (drule scheds_in_sig) | |
| 903 | apply assumption | |
| 904 | apply (drule scheds_in_sig) | |
| 905 | apply assumption | |
| 19741 | 906 | |
| 62156 | 907 | apply (rename_tac h1 h2 schA schB) | 
| 908 | text \<open>\<open>mksch\<close> is exactly the construction of \<open>trA\<parallel>B\<close> out of \<open>schA\<close>, \<open>schB\<close>, and the oracle \<open>tr\<close>, | |
| 909 | we need here\<close> | |
| 63549 | 910 | apply (rule_tac x = "mksch A B \<cdot> tr \<cdot> schA \<cdot> schB" in bexI) | 
| 19741 | 911 | |
| 62156 | 912 | text \<open>External actions of mksch are just the oracle\<close> | 
| 913 | apply (simp add: FilterA_mksch_is_tr) | |
| 19741 | 914 | |
| 62156 | 915 | text \<open>\<open>mksch\<close> is a schedule -- use compositionality on sch-level\<close> | 
| 916 | apply (simp add: compositionality_sch) | |
| 917 | apply (simp add: FilterAmksch_is_schA FilterBmksch_is_schB) | |
| 918 | apply (erule ForallAorB_mksch) | |
| 919 | apply (erule ForallPForallQ) | |
| 920 | apply (erule ext_is_act) | |
| 921 | done | |
| 19741 | 922 | |
| 923 | ||
| 924 | ||
| 62002 | 925 | subsection \<open>COMPOSITIONALITY on TRACE Level -- for Modules\<close> | 
| 19741 | 926 | |
| 62005 | 927 | lemma compositionality_tr_modules: | 
| 62156 | 928 | "is_trans_of A \<Longrightarrow> is_trans_of B \<Longrightarrow> compatible A B \<Longrightarrow> compatible B A \<Longrightarrow> | 
| 929 | is_asig(asig_of A) \<Longrightarrow> is_asig(asig_of B) \<Longrightarrow> | |
| 930 | Traces (A\<parallel>B) = par_traces (Traces A) (Traces B)" | |
| 931 | apply (unfold Traces_def par_traces_def) | |
| 932 | apply (simp add: asig_of_par) | |
| 933 | apply (rule set_eqI) | |
| 934 | apply (simp add: compositionality_tr externals_of_par) | |
| 935 | done | |
| 19741 | 936 | |
| 937 | ||
| 62002 | 938 | declaration \<open>fn _ => Simplifier.map_ss (Simplifier.set_mksym Simplifier.default_mk_sym)\<close> | 
| 3071 | 939 | |
| 940 | end |