| author | wenzelm | 
| Sat, 28 Sep 2024 23:23:30 +0200 | |
| changeset 80993 | addebc07f06e | 
| parent 80914 | d97fdabd9e2b | 
| child 81182 | fc5066122e68 | 
| permissions | -rw-r--r-- | 
| 5252 | 1 | (* Title: HOL/UNITY/Union.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 3 | Copyright 1998 University of Cambridge | |
| 4 | ||
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
30304diff
changeset | 5 | Partly from Misra's Chapter 5: Asynchronous Compositions of Programs. | 
| 5252 | 6 | *) | 
| 7 | ||
| 63146 | 8 | section\<open>Unions of Programs\<close> | 
| 13798 | 9 | |
| 16417 | 10 | theory Union imports SubstAx FP begin | 
| 5252 | 11 | |
| 13805 | 12 |   (*FIXME: conjoin Init F \<inter> Init G \<noteq> {} *) 
 | 
| 36866 | 13 | definition | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
80768diff
changeset | 14 | ok :: "['a program, 'a program] => bool" (infixl \<open>ok\<close> 65) | 
| 36866 | 15 | where "F ok G == Acts F \<subseteq> AllowedActs G & | 
| 13805 | 16 | Acts G \<subseteq> AllowedActs F" | 
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 17 | |
| 13805 | 18 |   (*FIXME: conjoin (\<Inter>i \<in> I. Init (F i)) \<noteq> {} *) 
 | 
| 36866 | 19 | definition | 
| 13792 | 20 | OK :: "['a set, 'a => 'b program] => bool" | 
| 36866 | 21 |   where "OK I F = (\<forall>i \<in> I. \<forall>j \<in> I-{i}. Acts (F i) \<subseteq> AllowedActs (F j))"
 | 
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 22 | |
| 36866 | 23 | definition | 
| 13792 | 24 | JOIN :: "['a set, 'a => 'b program] => 'b program" | 
| 36866 | 25 | where "JOIN I F = mk_program (\<Inter>i \<in> I. Init (F i), \<Union>i \<in> I. Acts (F i), | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
30304diff
changeset | 26 | \<Inter>i \<in> I. AllowedActs (F i))" | 
| 5252 | 27 | |
| 36866 | 28 | definition | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
80768diff
changeset | 29 | Join :: "['a program, 'a program] => 'a program" (infixl \<open>\<squnion>\<close> 65) | 
| 60773 | 30 | where "F \<squnion> G = mk_program (Init F \<inter> Init G, Acts F \<union> Acts G, | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
30304diff
changeset | 31 | AllowedActs F \<inter> AllowedActs G)" | 
| 5252 | 32 | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
80768diff
changeset | 33 | definition SKIP :: "'a program" (\<open>\<bottom>\<close>) | 
| 60773 | 34 |   where "\<bottom> = mk_program (UNIV, {}, UNIV)"
 | 
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 35 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 36 | (*Characterizes safety properties. Used with specifying Allowed*) | 
| 36866 | 37 | definition | 
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 38 | safety_prop :: "'a program set => bool" | 
| 69313 | 39 | where "safety_prop X \<longleftrightarrow> SKIP \<in> X \<and> (\<forall>G. Acts G \<subseteq> \<Union>(Acts ` X) \<longrightarrow> G \<in> X)" | 
| 5259 | 40 | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 41 | syntax | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
80768diff
changeset | 42 | "_JOIN1" :: "[pttrns, 'b set] => 'b set" (\<open>(3\<Squnion>_./ _)\<close> 10) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
80768diff
changeset | 43 | "_JOIN" :: "[pttrn, 'a set, 'b set] => 'b set" (\<open>(3\<Squnion>_\<in>_./ _)\<close> 10) | 
| 80768 | 44 | syntax_consts | 
| 45 | "_JOIN1" "_JOIN" == JOIN | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 46 | translations | 
| 60773 | 47 | "\<Squnion>x \<in> A. B" == "CONST JOIN A (\<lambda>x. B)" | 
| 48 | "\<Squnion>x y. B" == "\<Squnion>x. \<Squnion>y. B" | |
| 49 | "\<Squnion>x. B" == "CONST JOIN (CONST UNIV) (\<lambda>x. B)" | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 50 | |
| 13792 | 51 | |
| 63146 | 52 | subsection\<open>SKIP\<close> | 
| 13792 | 53 | |
| 54 | lemma Init_SKIP [simp]: "Init SKIP = UNIV" | |
| 55 | by (simp add: SKIP_def) | |
| 56 | ||
| 57 | lemma Acts_SKIP [simp]: "Acts SKIP = {Id}"
 | |
| 58 | by (simp add: SKIP_def) | |
| 59 | ||
| 60 | lemma AllowedActs_SKIP [simp]: "AllowedActs SKIP = UNIV" | |
| 61 | by (auto simp add: SKIP_def) | |
| 62 | ||
| 63 | lemma reachable_SKIP [simp]: "reachable SKIP = UNIV" | |
| 64 | by (force elim: reachable.induct intro: reachable.intros) | |
| 65 | ||
| 63146 | 66 | subsection\<open>SKIP and safety properties\<close> | 
| 13792 | 67 | |
| 13805 | 68 | lemma SKIP_in_constrains_iff [iff]: "(SKIP \<in> A co B) = (A \<subseteq> B)" | 
| 13792 | 69 | by (unfold constrains_def, auto) | 
| 70 | ||
| 13805 | 71 | lemma SKIP_in_Constrains_iff [iff]: "(SKIP \<in> A Co B) = (A \<subseteq> B)" | 
| 13792 | 72 | by (unfold Constrains_def, auto) | 
| 73 | ||
| 13805 | 74 | lemma SKIP_in_stable [iff]: "SKIP \<in> stable A" | 
| 13792 | 75 | by (unfold stable_def, auto) | 
| 76 | ||
| 77 | declare SKIP_in_stable [THEN stable_imp_Stable, iff] | |
| 78 | ||
| 79 | ||
| 63146 | 80 | subsection\<open>Join\<close> | 
| 13792 | 81 | |
| 13819 | 82 | lemma Init_Join [simp]: "Init (F\<squnion>G) = Init F \<inter> Init G" | 
| 13792 | 83 | by (simp add: Join_def) | 
| 84 | ||
| 13819 | 85 | lemma Acts_Join [simp]: "Acts (F\<squnion>G) = Acts F \<union> Acts G" | 
| 13792 | 86 | by (auto simp add: Join_def) | 
| 87 | ||
| 88 | lemma AllowedActs_Join [simp]: | |
| 13819 | 89 | "AllowedActs (F\<squnion>G) = AllowedActs F \<inter> AllowedActs G" | 
| 13792 | 90 | by (auto simp add: Join_def) | 
| 91 | ||
| 92 | ||
| 63146 | 93 | subsection\<open>JN\<close> | 
| 13792 | 94 | |
| 13805 | 95 | lemma JN_empty [simp]: "(\<Squnion>i\<in>{}. F i) = SKIP"
 | 
| 13792 | 96 | by (unfold JOIN_def SKIP_def, auto) | 
| 97 | ||
| 13819 | 98 | lemma JN_insert [simp]: "(\<Squnion>i \<in> insert a I. F i) = (F a)\<squnion>(\<Squnion>i \<in> I. F i)" | 
| 13792 | 99 | apply (rule program_equalityI) | 
| 100 | apply (auto simp add: JOIN_def Join_def) | |
| 101 | done | |
| 102 | ||
| 13805 | 103 | lemma Init_JN [simp]: "Init (\<Squnion>i \<in> I. F i) = (\<Inter>i \<in> I. Init (F i))" | 
| 13792 | 104 | by (simp add: JOIN_def) | 
| 105 | ||
| 13805 | 106 | lemma Acts_JN [simp]: "Acts (\<Squnion>i \<in> I. F i) = insert Id (\<Union>i \<in> I. Acts (F i))" | 
| 13792 | 107 | by (auto simp add: JOIN_def) | 
| 108 | ||
| 109 | lemma AllowedActs_JN [simp]: | |
| 13805 | 110 | "AllowedActs (\<Squnion>i \<in> I. F i) = (\<Inter>i \<in> I. AllowedActs (F i))" | 
| 13792 | 111 | by (auto simp add: JOIN_def) | 
| 112 | ||
| 113 | ||
| 114 | lemma JN_cong [cong]: | |
| 13805 | 115 | "[| I=J; !!i. i \<in> J ==> F i = G i |] ==> (\<Squnion>i \<in> I. F i) = (\<Squnion>i \<in> J. G i)" | 
| 13792 | 116 | by (simp add: JOIN_def) | 
| 117 | ||
| 118 | ||
| 63146 | 119 | subsection\<open>Algebraic laws\<close> | 
| 13792 | 120 | |
| 13819 | 121 | lemma Join_commute: "F\<squnion>G = G\<squnion>F" | 
| 13792 | 122 | by (simp add: Join_def Un_commute Int_commute) | 
| 123 | ||
| 13819 | 124 | lemma Join_assoc: "(F\<squnion>G)\<squnion>H = F\<squnion>(G\<squnion>H)" | 
| 13792 | 125 | by (simp add: Un_ac Join_def Int_assoc insert_absorb) | 
| 126 | ||
| 13819 | 127 | lemma Join_left_commute: "A\<squnion>(B\<squnion>C) = B\<squnion>(A\<squnion>C)" | 
| 13792 | 128 | by (simp add: Un_ac Int_ac Join_def insert_absorb) | 
| 129 | ||
| 13819 | 130 | lemma Join_SKIP_left [simp]: "SKIP\<squnion>F = F" | 
| 13792 | 131 | apply (unfold Join_def SKIP_def) | 
| 132 | apply (rule program_equalityI) | |
| 133 | apply (simp_all (no_asm) add: insert_absorb) | |
| 134 | done | |
| 135 | ||
| 13819 | 136 | lemma Join_SKIP_right [simp]: "F\<squnion>SKIP = F" | 
| 13792 | 137 | apply (unfold Join_def SKIP_def) | 
| 138 | apply (rule program_equalityI) | |
| 139 | apply (simp_all (no_asm) add: insert_absorb) | |
| 140 | done | |
| 141 | ||
| 13819 | 142 | lemma Join_absorb [simp]: "F\<squnion>F = F" | 
| 13792 | 143 | apply (unfold Join_def) | 
| 144 | apply (rule program_equalityI, auto) | |
| 145 | done | |
| 146 | ||
| 13819 | 147 | lemma Join_left_absorb: "F\<squnion>(F\<squnion>G) = F\<squnion>G" | 
| 13792 | 148 | apply (unfold Join_def) | 
| 149 | apply (rule program_equalityI, auto) | |
| 150 | done | |
| 151 | ||
| 152 | (*Join is an AC-operator*) | |
| 153 | lemmas Join_ac = Join_assoc Join_left_absorb Join_commute Join_left_commute | |
| 154 | ||
| 155 | ||
| 63146 | 156 | subsection\<open>Laws Governing \<open>\<Squnion>\<close>\<close> | 
| 13792 | 157 | |
| 158 | (*Also follows by JN_insert and insert_absorb, but the proof is longer*) | |
| 13819 | 159 | lemma JN_absorb: "k \<in> I ==> F k\<squnion>(\<Squnion>i \<in> I. F i) = (\<Squnion>i \<in> I. F i)" | 
| 13792 | 160 | by (auto intro!: program_equalityI) | 
| 161 | ||
| 13819 | 162 | lemma JN_Un: "(\<Squnion>i \<in> I \<union> J. F i) = ((\<Squnion>i \<in> I. F i)\<squnion>(\<Squnion>i \<in> J. F i))" | 
| 13792 | 163 | by (auto intro!: program_equalityI) | 
| 164 | ||
| 13805 | 165 | lemma JN_constant: "(\<Squnion>i \<in> I. c) = (if I={} then SKIP else c)"
 | 
| 13792 | 166 | by (rule program_equalityI, auto) | 
| 167 | ||
| 168 | lemma JN_Join_distrib: | |
| 13819 | 169 | "(\<Squnion>i \<in> I. F i\<squnion>G i) = (\<Squnion>i \<in> I. F i) \<squnion> (\<Squnion>i \<in> I. G i)" | 
| 13792 | 170 | by (auto intro!: program_equalityI) | 
| 171 | ||
| 172 | lemma JN_Join_miniscope: | |
| 13819 | 173 | "i \<in> I ==> (\<Squnion>i \<in> I. F i\<squnion>G) = ((\<Squnion>i \<in> I. F i)\<squnion>G)" | 
| 13792 | 174 | by (auto simp add: JN_Join_distrib JN_constant) | 
| 175 | ||
| 176 | (*Used to prove guarantees_JN_I*) | |
| 13819 | 177 | lemma JN_Join_diff: "i \<in> I ==> F i\<squnion>JOIN (I - {i}) F = JOIN I F"
 | 
| 13792 | 178 | apply (unfold JOIN_def Join_def) | 
| 179 | apply (rule program_equalityI, auto) | |
| 180 | done | |
| 181 | ||
| 182 | ||
| 63146 | 183 | subsection\<open>Safety: co, stable, FP\<close> | 
| 13792 | 184 | |
| 13805 | 185 | (*Fails if I={} because it collapses to SKIP \<in> A co B, i.e. to A \<subseteq> B.  So an
 | 
| 186 | alternative precondition is A \<subseteq> B, but most proofs using this rule require | |
| 13792 | 187 | I to be nonempty for other reasons anyway.*) | 
| 188 | lemma JN_constrains: | |
| 13805 | 189 | "i \<in> I ==> (\<Squnion>i \<in> I. F i) \<in> A co B = (\<forall>i \<in> I. F i \<in> A co B)" | 
| 13792 | 190 | by (simp add: constrains_def JOIN_def, blast) | 
| 191 | ||
| 192 | lemma Join_constrains [simp]: | |
| 13819 | 193 | "(F\<squnion>G \<in> A co B) = (F \<in> A co B & G \<in> A co B)" | 
| 13792 | 194 | by (auto simp add: constrains_def Join_def) | 
| 195 | ||
| 196 | lemma Join_unless [simp]: | |
| 13819 | 197 | "(F\<squnion>G \<in> A unless B) = (F \<in> A unless B & G \<in> A unless B)" | 
| 46577 | 198 | by (simp add: unless_def) | 
| 13792 | 199 | |
| 200 | (*Analogous weak versions FAIL; see Misra [1994] 5.4.1, Substitution Axiom. | |
| 13819 | 201 | reachable (F\<squnion>G) could be much bigger than reachable F, reachable G | 
| 13792 | 202 | *) | 
| 203 | ||
| 204 | ||
| 205 | lemma Join_constrains_weaken: | |
| 13805 | 206 | "[| F \<in> A co A'; G \<in> B co B' |] | 
| 13819 | 207 | ==> F\<squnion>G \<in> (A \<inter> B) co (A' \<union> B')" | 
| 13792 | 208 | by (simp, blast intro: constrains_weaken) | 
| 209 | ||
| 13805 | 210 | (*If I={}, it degenerates to SKIP \<in> UNIV co {}, which is false.*)
 | 
| 13792 | 211 | lemma JN_constrains_weaken: | 
| 13805 | 212 | "[| \<forall>i \<in> I. F i \<in> A i co A' i; i \<in> I |] | 
| 213 | ==> (\<Squnion>i \<in> I. F i) \<in> (\<Inter>i \<in> I. A i) co (\<Union>i \<in> I. A' i)" | |
| 13792 | 214 | apply (simp (no_asm_simp) add: JN_constrains) | 
| 215 | apply (blast intro: constrains_weaken) | |
| 216 | done | |
| 217 | ||
| 13805 | 218 | lemma JN_stable: "(\<Squnion>i \<in> I. F i) \<in> stable A = (\<forall>i \<in> I. F i \<in> stable A)" | 
| 13792 | 219 | by (simp add: stable_def constrains_def JOIN_def) | 
| 220 | ||
| 221 | lemma invariant_JN_I: | |
| 13805 | 222 | "[| !!i. i \<in> I ==> F i \<in> invariant A; i \<in> I |] | 
| 223 | ==> (\<Squnion>i \<in> I. F i) \<in> invariant A" | |
| 13792 | 224 | by (simp add: invariant_def JN_stable, blast) | 
| 225 | ||
| 226 | lemma Join_stable [simp]: | |
| 13819 | 227 | "(F\<squnion>G \<in> stable A) = | 
| 13805 | 228 | (F \<in> stable A & G \<in> stable A)" | 
| 13792 | 229 | by (simp add: stable_def) | 
| 230 | ||
| 231 | lemma Join_increasing [simp]: | |
| 13819 | 232 | "(F\<squnion>G \<in> increasing f) = | 
| 13805 | 233 | (F \<in> increasing f & G \<in> increasing f)" | 
| 46577 | 234 | by (auto simp add: increasing_def) | 
| 13792 | 235 | |
| 236 | lemma invariant_JoinI: | |
| 13805 | 237 | "[| F \<in> invariant A; G \<in> invariant A |] | 
| 13819 | 238 | ==> F\<squnion>G \<in> invariant A" | 
| 46577 | 239 | by (auto simp add: invariant_def) | 
| 13792 | 240 | |
| 13805 | 241 | lemma FP_JN: "FP (\<Squnion>i \<in> I. F i) = (\<Inter>i \<in> I. FP (F i))" | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
36866diff
changeset | 242 | by (simp add: FP_def JN_stable INTER_eq) | 
| 13792 | 243 | |
| 244 | ||
| 63146 | 245 | subsection\<open>Progress: transient, ensures\<close> | 
| 13792 | 246 | |
| 247 | lemma JN_transient: | |
| 13805 | 248 | "i \<in> I ==> | 
| 249 | (\<Squnion>i \<in> I. F i) \<in> transient A = (\<exists>i \<in> I. F i \<in> transient A)" | |
| 13792 | 250 | by (auto simp add: transient_def JOIN_def) | 
| 251 | ||
| 252 | lemma Join_transient [simp]: | |
| 13819 | 253 | "F\<squnion>G \<in> transient A = | 
| 13805 | 254 | (F \<in> transient A | G \<in> transient A)" | 
| 13792 | 255 | by (auto simp add: bex_Un transient_def Join_def) | 
| 256 | ||
| 13819 | 257 | lemma Join_transient_I1: "F \<in> transient A ==> F\<squnion>G \<in> transient A" | 
| 46577 | 258 | by simp | 
| 13792 | 259 | |
| 13819 | 260 | lemma Join_transient_I2: "G \<in> transient A ==> F\<squnion>G \<in> transient A" | 
| 46577 | 261 | by simp | 
| 13792 | 262 | |
| 13805 | 263 | (*If I={} it degenerates to (SKIP \<in> A ensures B) = False, i.e. to ~(A \<subseteq> B) *)
 | 
| 13792 | 264 | lemma JN_ensures: | 
| 13805 | 265 | "i \<in> I ==> | 
| 266 | (\<Squnion>i \<in> I. F i) \<in> A ensures B = | |
| 267 | ((\<forall>i \<in> I. F i \<in> (A-B) co (A \<union> B)) & (\<exists>i \<in> I. F i \<in> A ensures B))" | |
| 13792 | 268 | by (auto simp add: ensures_def JN_constrains JN_transient) | 
| 269 | ||
| 270 | lemma Join_ensures: | |
| 13819 | 271 | "F\<squnion>G \<in> A ensures B = | 
| 13805 | 272 | (F \<in> (A-B) co (A \<union> B) & G \<in> (A-B) co (A \<union> B) & | 
| 273 | (F \<in> transient (A-B) | G \<in> transient (A-B)))" | |
| 46577 | 274 | by (auto simp add: ensures_def) | 
| 13792 | 275 | |
| 276 | lemma stable_Join_constrains: | |
| 13805 | 277 | "[| F \<in> stable A; G \<in> A co A' |] | 
| 13819 | 278 | ==> F\<squnion>G \<in> A co A'" | 
| 13792 | 279 | apply (unfold stable_def constrains_def Join_def) | 
| 280 | apply (simp add: ball_Un, blast) | |
| 281 | done | |
| 282 | ||
| 13805 | 283 | (*Premise for G cannot use Always because F \<in> Stable A is weaker than | 
| 284 | G \<in> stable A *) | |
| 13792 | 285 | lemma stable_Join_Always1: | 
| 13819 | 286 | "[| F \<in> stable A; G \<in> invariant A |] ==> F\<squnion>G \<in> Always A" | 
| 13792 | 287 | apply (simp (no_asm_use) add: Always_def invariant_def Stable_eq_stable) | 
| 288 | apply (force intro: stable_Int) | |
| 289 | done | |
| 290 | ||
| 291 | (*As above, but exchanging the roles of F and G*) | |
| 292 | lemma stable_Join_Always2: | |
| 13819 | 293 | "[| F \<in> invariant A; G \<in> stable A |] ==> F\<squnion>G \<in> Always A" | 
| 13792 | 294 | apply (subst Join_commute) | 
| 295 | apply (blast intro: stable_Join_Always1) | |
| 296 | done | |
| 297 | ||
| 298 | lemma stable_Join_ensures1: | |
| 13819 | 299 | "[| F \<in> stable A; G \<in> A ensures B |] ==> F\<squnion>G \<in> A ensures B" | 
| 13792 | 300 | apply (simp (no_asm_simp) add: Join_ensures) | 
| 301 | apply (simp add: stable_def ensures_def) | |
| 302 | apply (erule constrains_weaken, auto) | |
| 303 | done | |
| 304 | ||
| 305 | (*As above, but exchanging the roles of F and G*) | |
| 306 | lemma stable_Join_ensures2: | |
| 13819 | 307 | "[| F \<in> A ensures B; G \<in> stable A |] ==> F\<squnion>G \<in> A ensures B" | 
| 13792 | 308 | apply (subst Join_commute) | 
| 309 | apply (blast intro: stable_Join_ensures1) | |
| 310 | done | |
| 311 | ||
| 312 | ||
| 63146 | 313 | subsection\<open>the ok and OK relations\<close> | 
| 13792 | 314 | |
| 315 | lemma ok_SKIP1 [iff]: "SKIP ok F" | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 316 | by (simp add: ok_def) | 
| 13792 | 317 | |
| 318 | lemma ok_SKIP2 [iff]: "F ok SKIP" | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 319 | by (simp add: ok_def) | 
| 13792 | 320 | |
| 321 | lemma ok_Join_commute: | |
| 13819 | 322 | "(F ok G & (F\<squnion>G) ok H) = (G ok H & F ok (G\<squnion>H))" | 
| 13792 | 323 | by (auto simp add: ok_def) | 
| 324 | ||
| 325 | lemma ok_commute: "(F ok G) = (G ok F)" | |
| 326 | by (auto simp add: ok_def) | |
| 327 | ||
| 45605 | 328 | lemmas ok_sym = ok_commute [THEN iffD1] | 
| 13792 | 329 | |
| 330 | lemma ok_iff_OK: | |
| 13819 | 331 |      "OK {(0::int,F),(1,G),(2,H)} snd = (F ok G & (F\<squnion>G) ok H)"
 | 
| 16977 | 332 | apply (simp add: Ball_def conj_disj_distribR ok_def Join_def OK_def insert_absorb | 
| 333 | all_conj_distrib) | |
| 334 | apply blast | |
| 335 | done | |
| 13792 | 336 | |
| 13819 | 337 | lemma ok_Join_iff1 [iff]: "F ok (G\<squnion>H) = (F ok G & F ok H)" | 
| 13792 | 338 | by (auto simp add: ok_def) | 
| 339 | ||
| 13819 | 340 | lemma ok_Join_iff2 [iff]: "(G\<squnion>H) ok F = (G ok F & H ok F)" | 
| 13792 | 341 | by (auto simp add: ok_def) | 
| 342 | ||
| 343 | (*useful? Not with the previous two around*) | |
| 13819 | 344 | lemma ok_Join_commute_I: "[| F ok G; (F\<squnion>G) ok H |] ==> F ok (G\<squnion>H)" | 
| 13792 | 345 | by (auto simp add: ok_def) | 
| 346 | ||
| 13805 | 347 | lemma ok_JN_iff1 [iff]: "F ok (JOIN I G) = (\<forall>i \<in> I. F ok G i)" | 
| 13792 | 348 | by (auto simp add: ok_def) | 
| 349 | ||
| 13805 | 350 | lemma ok_JN_iff2 [iff]: "(JOIN I G) ok F = (\<forall>i \<in> I. G i ok F)" | 
| 13792 | 351 | by (auto simp add: ok_def) | 
| 352 | ||
| 13805 | 353 | lemma OK_iff_ok: "OK I F = (\<forall>i \<in> I. \<forall>j \<in> I-{i}. (F i) ok (F j))"
 | 
| 13792 | 354 | by (auto simp add: ok_def OK_def) | 
| 355 | ||
| 13805 | 356 | lemma OK_imp_ok: "[| OK I F; i \<in> I; j \<in> I; i \<noteq> j|] ==> (F i) ok (F j)" | 
| 13792 | 357 | by (auto simp add: OK_iff_ok) | 
| 358 | ||
| 359 | ||
| 63146 | 360 | subsection\<open>Allowed\<close> | 
| 13792 | 361 | |
| 362 | lemma Allowed_SKIP [simp]: "Allowed SKIP = UNIV" | |
| 363 | by (auto simp add: Allowed_def) | |
| 364 | ||
| 13819 | 365 | lemma Allowed_Join [simp]: "Allowed (F\<squnion>G) = Allowed F \<inter> Allowed G" | 
| 13792 | 366 | by (auto simp add: Allowed_def) | 
| 367 | ||
| 13805 | 368 | lemma Allowed_JN [simp]: "Allowed (JOIN I F) = (\<Inter>i \<in> I. Allowed (F i))" | 
| 13792 | 369 | by (auto simp add: Allowed_def) | 
| 370 | ||
| 13805 | 371 | lemma ok_iff_Allowed: "F ok G = (F \<in> Allowed G & G \<in> Allowed F)" | 
| 13792 | 372 | by (simp add: ok_def Allowed_def) | 
| 373 | ||
| 13805 | 374 | lemma OK_iff_Allowed: "OK I F = (\<forall>i \<in> I. \<forall>j \<in> I-{i}. F i \<in> Allowed(F j))"
 | 
| 13792 | 375 | by (auto simp add: OK_iff_ok ok_iff_Allowed) | 
| 376 | ||
| 69597 | 377 | subsection\<open>\<^term>\<open>safety_prop\<close>, for reasoning about | 
| 63146 | 378 | given instances of "ok"\<close> | 
| 13792 | 379 | |
| 380 | lemma safety_prop_Acts_iff: | |
| 69313 | 381 | "safety_prop X ==> (Acts G \<subseteq> insert Id (\<Union>(Acts ` X))) = (G \<in> X)" | 
| 13792 | 382 | by (auto simp add: safety_prop_def) | 
| 383 | ||
| 384 | lemma safety_prop_AllowedActs_iff_Allowed: | |
| 69313 | 385 | "safety_prop X ==> (\<Union>(Acts ` X) \<subseteq> AllowedActs F) = (X \<subseteq> Allowed F)" | 
| 13792 | 386 | by (auto simp add: Allowed_def safety_prop_Acts_iff [symmetric]) | 
| 387 | ||
| 388 | lemma Allowed_eq: | |
| 69313 | 389 | "safety_prop X ==> Allowed (mk_program (init, acts, \<Union>(Acts ` X))) = X" | 
| 13792 | 390 | by (simp add: Allowed_def safety_prop_Acts_iff) | 
| 391 | ||
| 392 | (*For safety_prop to hold, the property must be satisfiable!*) | |
| 13805 | 393 | lemma safety_prop_constrains [iff]: "safety_prop (A co B) = (A \<subseteq> B)" | 
| 13792 | 394 | by (simp add: safety_prop_def constrains_def, blast) | 
| 395 | ||
| 396 | lemma safety_prop_stable [iff]: "safety_prop (stable A)" | |
| 397 | by (simp add: stable_def) | |
| 398 | ||
| 399 | lemma safety_prop_Int [simp]: | |
| 56248 
67dc9549fa15
generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
 haftmann parents: 
46577diff
changeset | 400 | "safety_prop X \<Longrightarrow> safety_prop Y \<Longrightarrow> safety_prop (X \<inter> Y)" | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 401 | proof (clarsimp simp add: safety_prop_def) | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 402 | fix G | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 403 | assume "\<forall>G. Acts G \<subseteq> (\<Union>x\<in>X. Acts x) \<longrightarrow> G \<in> X" | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 404 | then have X: "Acts G \<subseteq> (\<Union>x\<in>X. Acts x) \<Longrightarrow> G \<in> X" by blast | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 405 | assume "\<forall>G. Acts G \<subseteq> (\<Union>x\<in>Y. Acts x) \<longrightarrow> G \<in> Y" | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 406 | then have Y: "Acts G \<subseteq> (\<Union>x\<in>Y. Acts x) \<Longrightarrow> G \<in> Y" by blast | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 407 | assume Acts: "Acts G \<subseteq> (\<Union>x\<in>X \<inter> Y. Acts x)" | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 408 | with X and Y show "G \<in> X \<and> G \<in> Y" by auto | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 409 | qed | 
| 56248 
67dc9549fa15
generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
 haftmann parents: 
46577diff
changeset | 410 | |
| 
67dc9549fa15
generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
 haftmann parents: 
46577diff
changeset | 411 | lemma safety_prop_INTER [simp]: | 
| 
67dc9549fa15
generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
 haftmann parents: 
46577diff
changeset | 412 | "(\<And>i. i \<in> I \<Longrightarrow> safety_prop (X i)) \<Longrightarrow> safety_prop (\<Inter>i\<in>I. X i)" | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 413 | proof (clarsimp simp add: safety_prop_def) | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 414 | fix G and i | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 415 | assume "\<And>i. i \<in> I \<Longrightarrow> \<bottom> \<in> X i \<and> | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 416 | (\<forall>G. Acts G \<subseteq> (\<Union>x\<in>X i. Acts x) \<longrightarrow> G \<in> X i)" | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 417 | then have *: "i \<in> I \<Longrightarrow> Acts G \<subseteq> (\<Union>x\<in>X i. Acts x) \<Longrightarrow> G \<in> X i" | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 418 | by blast | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 419 | assume "i \<in> I" | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 420 | moreover assume "Acts G \<subseteq> (\<Union>j\<in>\<Inter>i\<in>I. X i. Acts j)" | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 421 | ultimately have "Acts G \<subseteq> (\<Union>i\<in>X i. Acts i)" | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 422 | by auto | 
| 63146 | 423 | with * \<open>i \<in> I\<close> show "G \<in> X i" by blast | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61941diff
changeset | 424 | qed | 
| 13792 | 425 | |
| 426 | lemma safety_prop_INTER1 [simp]: | |
| 56248 
67dc9549fa15
generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
 haftmann parents: 
46577diff
changeset | 427 | "(\<And>i. safety_prop (X i)) \<Longrightarrow> safety_prop (\<Inter>i. X i)" | 
| 
67dc9549fa15
generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
 haftmann parents: 
46577diff
changeset | 428 | by (rule safety_prop_INTER) simp | 
| 13792 | 429 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 430 | lemma def_prg_Allowed: | 
| 69313 | 431 | "[| F == mk_program (init, acts, \<Union>(Acts ` X)) ; safety_prop X |] | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 432 | ==> Allowed F = X" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 433 | by (simp add: Allowed_eq) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 434 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 435 | lemma Allowed_totalize [simp]: "Allowed (totalize F) = Allowed F" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 436 | by (simp add: Allowed_def) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 437 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 438 | lemma def_total_prg_Allowed: | 
| 69313 | 439 | "[| F = mk_total_program (init, acts, \<Union>(Acts ` X)) ; safety_prop X |] | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 440 | ==> Allowed F = X" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 441 | by (simp add: mk_total_program_def def_prg_Allowed) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 442 | |
| 13792 | 443 | lemma def_UNION_ok_iff: | 
| 69313 | 444 | "[| F = mk_program(init,acts,\<Union>(Acts ` X)); safety_prop X |] | 
| 13805 | 445 | ==> F ok G = (G \<in> X & acts \<subseteq> AllowedActs G)" | 
| 13792 | 446 | by (auto simp add: ok_def safety_prop_Acts_iff) | 
| 9685 | 447 | |
| 63146 | 448 | text\<open>The union of two total programs is total.\<close> | 
| 13819 | 449 | lemma totalize_Join: "totalize F\<squnion>totalize G = totalize (F\<squnion>G)" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 450 | by (simp add: program_equalityI totalize_def Join_def image_Un) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 451 | |
| 13819 | 452 | lemma all_total_Join: "[|all_total F; all_total G|] ==> all_total (F\<squnion>G)" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 453 | by (simp add: all_total_def, blast) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 454 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 455 | lemma totalize_JN: "(\<Squnion>i \<in> I. totalize (F i)) = totalize(\<Squnion>i \<in> I. F i)" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 456 | by (simp add: program_equalityI totalize_def JOIN_def image_UN) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 457 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 458 | lemma all_total_JN: "(!!i. i\<in>I ==> all_total (F i)) ==> all_total(\<Squnion>i\<in>I. F i)" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 459 | by (simp add: all_total_iff_totalize totalize_JN [symmetric]) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 460 | |
| 5252 | 461 | end |