| author | paulson | 
| Fri, 03 Dec 2004 15:28:12 +0100 | |
| changeset 15370 | 05b03ea0f18d | 
| parent 14092 | 68da54626309 | 
| permissions | -rw-r--r-- | 
| 11479 | 1 | (* Title: ZF/UNITY/Constrains.ML | 
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 2 | ID: $Id \\<in> Constrains.ML,v 1.10 2003/06/20 10:10:45 paulson Exp $ | 
| 11479 | 3 | Author: Sidi O Ehmety, Computer Laboratory | 
| 4 | Copyright 2001 University of Cambridge | |
| 5 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 6 | Safety relations \\<in> restricted to the set of reachable states. | 
| 11479 | 7 | |
| 8 | Proofs ported from HOL. | |
| 9 | *) | |
| 10 | ||
| 11 | (*** traces and reachable ***) | |
| 12 | ||
| 12195 | 13 | Goal "reachable(F) <= state"; | 
| 14 | by (cut_inst_tac [("F", "F")] Init_type 1);
 | |
| 15 | by (cut_inst_tac [("F", "F")] Acts_type 1);
 | |
| 16 | by (cut_inst_tac [("F", "F")] reachable.dom_subset 1);
 | |
| 17 | by (Blast_tac 1); | |
| 11479 | 18 | qed "reachable_type"; | 
| 12195 | 19 | |
| 20 | Goalw [st_set_def] "st_set(reachable(F))"; | |
| 12484 | 21 | by (rtac reachable_type 1); | 
| 12195 | 22 | qed "st_set_reachable"; | 
| 23 | AddIffs [st_set_reachable]; | |
| 11479 | 24 | |
| 12195 | 25 | Goal "reachable(F) Int state = reachable(F)"; | 
| 26 | by (cut_facts_tac [reachable_type] 1); | |
| 27 | by Auto_tac; | |
| 28 | qed "reachable_Int_state"; | |
| 29 | AddIffs [reachable_Int_state]; | |
| 11479 | 30 | |
| 12195 | 31 | Goal "state Int reachable(F) = reachable(F)"; | 
| 32 | by (cut_facts_tac [reachable_type] 1); | |
| 33 | by Auto_tac; | |
| 34 | qed "state_Int_reachable"; | |
| 35 | AddIffs [state_Int_reachable]; | |
| 36 | ||
| 37 | Goal | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 38 | "F \\<in> program ==> reachable(F)={s \\<in> state. \\<exists>evs. <s,evs>:traces(Init(F), Acts(F))}";
 | 
| 11479 | 39 | by (rtac equalityI 1); | 
| 40 | by Safe_tac; | |
| 12195 | 41 | by (blast_tac (claset() addDs [reachable_type RS subsetD]) 1); | 
| 11479 | 42 | by (etac traces.induct 2); | 
| 43 | by (etac reachable.induct 1); | |
| 44 | by (ALLGOALS (blast_tac (claset() addIs reachable.intrs @ traces.intrs))); | |
| 45 | qed "reachable_equiv_traces"; | |
| 46 | ||
| 47 | Goal "Init(F) <= reachable(F)"; | |
| 48 | by (blast_tac (claset() addIs reachable.intrs) 1); | |
| 49 | qed "Init_into_reachable"; | |
| 50 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 51 | Goal "[| F \\<in> program; G \\<in> program; \ | 
| 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 52 | \ Acts(G) <= Acts(F) |] ==> G \\<in> stable(reachable(F))"; | 
| 11479 | 53 | by (blast_tac (claset() | 
| 12195 | 54 | addIs [stableI, constrainsI, st_setI, | 
| 55 | reachable_type RS subsetD] @ reachable.intrs) 1); | |
| 11479 | 56 | qed "stable_reachable"; | 
| 57 | ||
| 58 | AddSIs [stable_reachable]; | |
| 59 | Addsimps [stable_reachable]; | |
| 60 | ||
| 61 | (*The set of all reachable states is an invariant...*) | |
| 62 | Goalw [invariant_def, initially_def] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 63 | "F \\<in> program ==> F \\<in> invariant(reachable(F))"; | 
| 12195 | 64 | by (blast_tac (claset() addIs [reachable_type RS subsetD]@reachable.intrs) 1); | 
| 11479 | 65 | qed "invariant_reachable"; | 
| 66 | ||
| 67 | (*...in fact the strongest invariant!*) | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 68 | Goal "F \\<in> invariant(A) ==> reachable(F) <= A"; | 
| 12195 | 69 | by (cut_inst_tac [("F", "F")] Acts_type 1);
 | 
| 70 | by (cut_inst_tac [("F", "F")] Init_type 1);
 | |
| 71 | by (cut_inst_tac [("F", "F")] reachable_type 1);
 | |
| 72 | by (full_simp_tac (simpset() addsimps [stable_def, constrains_def, | |
| 73 | invariant_def, initially_def]) 1); | |
| 11479 | 74 | by (rtac subsetI 1); | 
| 75 | by (etac reachable.induct 1); | |
| 76 | by (REPEAT (blast_tac (claset() addIs reachable.intrs) 1)); | |
| 77 | qed "invariant_includes_reachable"; | |
| 78 | ||
| 79 | (*** Co ***) | |
| 80 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 81 | Goal "F \\<in> B co B'==>F:(reachable(F) Int B) co (reachable(F) Int B')"; | 
| 12195 | 82 | by (forward_tac [constrains_type RS subsetD] 1); | 
| 83 | by (forward_tac [[asm_rl, asm_rl, subset_refl] MRS stable_reachable] 1); | |
| 84 | by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [stable_def, constrains_Int]))); | |
| 11479 | 85 | qed "constrains_reachable_Int"; | 
| 86 | ||
| 87 | (*Resembles the previous definition of Constrains*) | |
| 88 | Goalw [Constrains_def] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 89 | "A Co B = {F \\<in> program. F:(reachable(F) Int A) co (reachable(F)  Int  B)}";
 | 
| 12195 | 90 | by (blast_tac (claset() addDs [constrains_reachable_Int, | 
| 91 | constrains_type RS subsetD] | |
| 11479 | 92 | addIs [constrains_weaken]) 1); | 
| 93 | qed "Constrains_eq_constrains"; | |
| 12195 | 94 | val Constrains_def2 = Constrains_eq_constrains RS eq_reflection; | 
| 95 | ||
| 96 | Goalw [Constrains_def] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 97 | "F \\<in> A co A' ==> F \\<in> A Co A'"; | 
| 12195 | 98 | by (blast_tac (claset() addIs [constrains_weaken_L] addDs [constrainsD2]) 1); | 
| 99 | qed "constrains_imp_Constrains"; | |
| 100 | ||
| 101 | val prems = Goalw [Constrains_def, constrains_def, st_set_def] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 102 | "[|(!!act s s'. [| act \\<in> Acts(F); <s,s'>:act; s \\<in> A |] ==> s':A'); F \\<in> program|]==>F \\<in> A Co A'"; | 
| 12195 | 103 | by (auto_tac (claset(), simpset() addsimps prems)); | 
| 104 | by (blast_tac (claset() addDs [reachable_type RS subsetD]) 1); | |
| 105 | qed "ConstrainsI"; | |
| 11479 | 106 | |
| 107 | Goalw [Constrains_def] | |
| 12195 | 108 | "A Co B <= program"; | 
| 109 | by (Blast_tac 1); | |
| 110 | qed "Constrains_type"; | |
| 111 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 112 | Goal "F \\<in> 0 Co B <-> F \\<in> program"; | 
| 12195 | 113 | by (auto_tac (claset() addDs [Constrains_type RS subsetD] | 
| 114 | addIs [constrains_imp_Constrains], simpset())); | |
| 115 | qed "Constrains_empty"; | |
| 116 | AddIffs [Constrains_empty]; | |
| 117 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 118 | Goalw [Constrains_def] "F \\<in> A Co state <-> F \\<in> program"; | 
| 12195 | 119 | by (auto_tac (claset() addDs [Constrains_type RS subsetD] | 
| 120 | addIs [constrains_imp_Constrains], simpset())); | |
| 121 | qed "Constrains_state"; | |
| 122 | AddIffs [Constrains_state]; | |
| 123 | ||
| 124 | Goalw [Constrains_def2] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 125 | "[| F \\<in> A Co A'; A'<=B' |] ==> F \\<in> A Co B'"; | 
| 12195 | 126 | by (blast_tac (claset() addIs [constrains_weaken_R]) 1); | 
| 127 | qed "Constrains_weaken_R"; | |
| 128 | ||
| 129 | Goalw [Constrains_def2] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 130 | "[| F \\<in> A Co A'; B<=A |] ==> F \\<in> B Co A'"; | 
| 12195 | 131 | by (blast_tac (claset() addIs [constrains_weaken_L, st_set_subset]) 1); | 
| 132 | qed "Constrains_weaken_L"; | |
| 133 | ||
| 134 | Goalw [Constrains_def2] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 135 | "[| F \\<in> A Co A'; B<=A; A'<=B' |] ==> F \\<in> B Co B'"; | 
| 12195 | 136 | by (blast_tac (claset() addIs [constrains_weaken, st_set_subset]) 1); | 
| 137 | qed "Constrains_weaken"; | |
| 138 | ||
| 139 | (** Union **) | |
| 12215 
55c084921240
Modified to make the files build with the new changes in ZF
 ehmety parents: 
12195diff
changeset | 140 | Goalw [Constrains_def2] | 
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 141 | "[| F \\<in> A Co A'; F \\<in> B Co B' |] ==> F \\<in> (A Un B) Co (A' Un B')"; | 
| 12195 | 142 | by Auto_tac; | 
| 12215 
55c084921240
Modified to make the files build with the new changes in ZF
 ehmety parents: 
12195diff
changeset | 143 | by (asm_full_simp_tac (simpset() addsimps [Int_Un_distrib]) 1); | 
| 12195 | 144 | by (blast_tac (claset() addIs [constrains_Un]) 1); | 
| 145 | qed "Constrains_Un"; | |
| 146 | ||
| 147 | val [major, minor] = Goalw [Constrains_def2] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 148 | "[|(!!i. i \\<in> I==>F \\<in> A(i) Co A'(i)); F \\<in> program|] ==> F:(\\<Union>i \\<in> I. A(i)) Co (\\<Union>i \\<in> I. A'(i))"; | 
| 12195 | 149 | by (cut_facts_tac [minor] 1); | 
| 150 | by (auto_tac (claset() addDs [major] | |
| 151 | addIs [constrains_UN], | |
| 12215 
55c084921240
Modified to make the files build with the new changes in ZF
 ehmety parents: 
12195diff
changeset | 152 | simpset() delsimps UN_simps addsimps [Int_UN_distrib])); | 
| 12195 | 153 | qed "Constrains_UN"; | 
| 154 | ||
| 155 | (** Intersection **) | |
| 156 | ||
| 157 | Goalw [Constrains_def] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 158 | "[| F \\<in> A Co A'; F \\<in> B Co B'|]==> F:(A Int B) Co (A' Int B')"; | 
| 12195 | 159 | by (subgoal_tac "reachable(F) Int (A Int B) = \ | 
| 160 | \ (reachable(F) Int A) Int (reachable(F) Int B)" 1); | |
| 161 | by (ALLGOALS(force_tac (claset() addIs [constrains_Int], simpset()))); | |
| 162 | qed "Constrains_Int"; | |
| 163 | ||
| 164 | val [major,minor] = Goal | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 165 | "[| (!!i. i \\<in> I ==>F \\<in> A(i) Co A'(i)); F \\<in> program |] \ | 
| 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 166 | \ ==> F:(\\<Inter>i \\<in> I. A(i)) Co (\\<Inter>i \\<in> I. A'(i))"; | 
| 12195 | 167 | by (cut_facts_tac [minor] 1); | 
| 12220 | 168 | by (asm_simp_tac (simpset() delsimps INT_simps | 
| 169 | addsimps [Constrains_def]@INT_extend_simps) 1); | |
| 12195 | 170 | by (rtac constrains_INT 1); | 
| 12484 | 171 | by (dtac major 1); | 
| 12220 | 172 | by (auto_tac (claset(), simpset() addsimps [Constrains_def])); | 
| 12195 | 173 | qed "Constrains_INT"; | 
| 174 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 175 | Goalw [Constrains_def] "F \\<in> A Co A' ==> reachable(F) Int A <= A'"; | 
| 12195 | 176 | by (blast_tac (claset() addDs [constrains_imp_subset]) 1); | 
| 177 | qed "Constrains_imp_subset"; | |
| 178 | ||
| 179 | Goalw [Constrains_def2] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 180 | "[| F \\<in> A Co B; F \\<in> B Co C |] ==> F \\<in> A Co C"; | 
| 12195 | 181 | by (blast_tac (claset() addIs [constrains_trans, constrains_weaken]) 1); | 
| 182 | qed "Constrains_trans"; | |
| 183 | ||
| 184 | Goalw [Constrains_def2] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 185 | "[| F \\<in> A Co (A' Un B); F \\<in> B Co B' |] ==> F \\<in> A Co (A' Un B')"; | 
| 12215 
55c084921240
Modified to make the files build with the new changes in ZF
 ehmety parents: 
12195diff
changeset | 186 | by (full_simp_tac (simpset() addsimps [Int_Un_distrib]) 1); | 
| 12195 | 187 | by (blast_tac (claset() addIs [constrains_cancel]) 1); | 
| 188 | qed "Constrains_cancel"; | |
| 189 | ||
| 190 | (*** Stable ***) | |
| 191 | (* Useful because there's no Stable_weaken. [Tanja Vos] *) | |
| 11479 | 192 | |
| 193 | Goalw [stable_def, Stable_def] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 194 | "F \\<in> stable(A) ==> F \\<in> Stable(A)"; | 
| 11479 | 195 | by (etac constrains_imp_Constrains 1); | 
| 196 | qed "stable_imp_Stable"; | |
| 197 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 198 | Goal "[| F \\<in> Stable(A); A = B |] ==> F \\<in> Stable(B)"; | 
| 11479 | 199 | by (Blast_tac 1); | 
| 200 | qed "Stable_eq"; | |
| 201 | ||
| 12195 | 202 | Goal | 
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 203 | "F \\<in> Stable(A) <-> (F \\<in> stable(reachable(F) Int A))"; | 
| 12195 | 204 | by (auto_tac (claset() addDs [constrainsD2], | 
| 205 | simpset() addsimps [Stable_def, stable_def, Constrains_def2])); | |
| 11479 | 206 | qed "Stable_eq_stable"; | 
| 207 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 208 | Goalw [Stable_def] "F \\<in> A Co A ==> F \\<in> Stable(A)"; | 
| 11479 | 209 | by (assume_tac 1); | 
| 210 | qed "StableI"; | |
| 211 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 212 | Goalw [Stable_def] "F \\<in> Stable(A) ==> F \\<in> A Co A"; | 
| 11479 | 213 | by (assume_tac 1); | 
| 214 | qed "StableD"; | |
| 215 | ||
| 216 | Goalw [Stable_def] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 217 | "[| F \\<in> Stable(A); F \\<in> Stable(A') |] ==> F \\<in> Stable(A Un A')"; | 
| 11479 | 218 | by (blast_tac (claset() addIs [Constrains_Un]) 1); | 
| 219 | qed "Stable_Un"; | |
| 220 | ||
| 221 | Goalw [Stable_def] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 222 | "[| F \\<in> Stable(A); F \\<in> Stable(A') |] ==> F \\<in> Stable (A Int A')"; | 
| 11479 | 223 | by (blast_tac (claset() addIs [Constrains_Int]) 1); | 
| 224 | qed "Stable_Int"; | |
| 225 | ||
| 226 | Goalw [Stable_def] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 227 | "[| F \\<in> Stable(C); F \\<in> A Co (C Un A') |] \ | 
| 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 228 | \ ==> F \\<in> (C Un A) Co (C Un A')"; | 
| 11479 | 229 | by (blast_tac (claset() addIs [Constrains_Un RS Constrains_weaken_R]) 1); | 
| 230 | qed "Stable_Constrains_Un"; | |
| 231 | ||
| 232 | Goalw [Stable_def] | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 233 | "[| F \\<in> Stable(C); F \\<in> (C Int A) Co A' |] \ | 
| 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 234 | \ ==> F \\<in> (C Int A) Co (C Int A')"; | 
| 12195 | 235 | by (blast_tac (claset() addIs [Constrains_Int RS Constrains_weaken]) 1); | 
| 11479 | 236 | qed "Stable_Constrains_Int"; | 
| 237 | ||
| 12195 | 238 | val [major,minor] = Goalw [Stable_def] | 
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 239 | "[| (!!i. i \\<in> I ==> F \\<in> Stable(A(i))); F \\<in> program |]==> F \\<in> Stable (\\<Union>i \\<in> I. A(i))"; | 
| 12195 | 240 | by (cut_facts_tac [minor] 1); | 
| 241 | by (blast_tac (claset() addIs [Constrains_UN,major]) 1); | |
| 11479 | 242 | qed "Stable_UN"; | 
| 243 | ||
| 12195 | 244 | val [major,minor] = Goalw [Stable_def] | 
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 245 | "[|(!!i. i \\<in> I ==> F \\<in> Stable(A(i))); F \\<in> program |]==> F \\<in> Stable (\\<Inter>i \\<in> I. A(i))"; | 
| 12195 | 246 | by (cut_facts_tac [minor] 1); | 
| 247 | by (blast_tac (claset() addIs [Constrains_INT, major]) 1); | |
| 11479 | 248 | qed "Stable_INT"; | 
| 249 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 250 | Goal "F \\<in> program ==>F \\<in> Stable (reachable(F))"; | 
| 11479 | 251 | by (asm_simp_tac (simpset() | 
| 12195 | 252 | addsimps [Stable_eq_stable, Int_absorb]) 1); | 
| 11479 | 253 | qed "Stable_reachable"; | 
| 254 | ||
| 255 | Goalw [Stable_def] | |
| 12195 | 256 | "Stable(A) <= program"; | 
| 257 | by (rtac Constrains_type 1); | |
| 258 | qed "Stable_type"; | |
| 11479 | 259 | |
| 260 | (*** The Elimination Theorem. The "free" m has become universally quantified! | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 261 | Should the premise be !!m instead of \\<forall>m ? Would make it harder to use | 
| 11479 | 262 | in forward proof. ***) | 
| 263 | ||
| 12195 | 264 | Goalw [Constrains_def] | 
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 265 | "[| \\<forall>m \\<in> M. F \\<in> ({s \\<in> A. x(s) = m}) Co (B(m)); F \\<in> program |] \
 | 
| 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 266 | \    ==> F \\<in> ({s \\<in> A. x(s):M}) Co (\\<Union>m \\<in> M. B(m))";
 | 
| 11479 | 267 | by Auto_tac; | 
| 12195 | 268 | by (res_inst_tac [("A1","reachable(F)Int A")] (elimination RS constrains_weaken_L) 1);
 | 
| 269 | by (auto_tac (claset() addIs [constrains_weaken_L], simpset())); | |
| 11479 | 270 | qed "Elimination"; | 
| 271 | ||
| 12195 | 272 | (* As above, but for the special case of A=state *) | 
| 11479 | 273 | Goal | 
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 274 |  "[| \\<forall>m \\<in> M. F \\<in> {s \\<in> state. x(s) = m} Co B(m); F \\<in> program |] \
 | 
| 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 275 | \    ==> F \\<in> {s \\<in> state. x(s):M} Co (\\<Union>m \\<in> M. B(m))";
 | 
| 11479 | 276 | by (blast_tac (claset() addIs [Elimination]) 1); | 
| 277 | qed "Elimination2"; | |
| 278 | ||
| 279 | (** Unless **) | |
| 280 | ||
| 12195 | 281 | Goalw [Unless_def] "A Unless B <=program"; | 
| 282 | by (rtac Constrains_type 1); | |
| 283 | qed "Unless_type"; | |
| 11479 | 284 | |
| 285 | (*** Specialized laws for handling Always ***) | |
| 286 | ||
| 287 | (** Natural deduction rules for "Always A" **) | |
| 12195 | 288 | |
| 11479 | 289 | Goalw [Always_def, initially_def] | 
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 290 | "[| Init(F)<=A; F \\<in> Stable(A) |] ==> F \\<in> Always(A)"; | 
| 12195 | 291 | by (forward_tac [Stable_type RS subsetD] 1); | 
| 292 | by Auto_tac; | |
| 11479 | 293 | qed "AlwaysI"; | 
| 294 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 295 | Goal "F \\<in> Always(A) ==> Init(F)<=A & F \\<in> Stable(A)"; | 
| 12195 | 296 | by (asm_full_simp_tac (simpset() addsimps [Always_def, initially_def]) 1); | 
| 11479 | 297 | qed "AlwaysD"; | 
| 298 | ||
| 299 | bind_thm ("AlwaysE", AlwaysD RS conjE);
 | |
| 300 | bind_thm ("Always_imp_Stable", AlwaysD RS conjunct2);
 | |
| 301 | ||
| 302 | (*The set of all reachable states is Always*) | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 303 | Goal "F \\<in> Always(A) ==> reachable(F) <= A"; | 
| 12195 | 304 | by (full_simp_tac (simpset() addsimps | 
| 305 | [Stable_def, Constrains_def, constrains_def, Always_def, initially_def]) 1); | |
| 11479 | 306 | by (rtac subsetI 1); | 
| 307 | by (etac reachable.induct 1); | |
| 308 | by (REPEAT (blast_tac (claset() addIs reachable.intrs) 1)); | |
| 309 | qed "Always_includes_reachable"; | |
| 310 | ||
| 12195 | 311 | Goalw [Always_def, invariant_def, Stable_def, stable_def] | 
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 312 | "F \\<in> invariant(A) ==> F \\<in> Always(A)"; | 
| 11479 | 313 | by (blast_tac (claset() addIs [constrains_imp_Constrains]) 1); | 
| 314 | qed "invariant_imp_Always"; | |
| 315 | ||
| 316 | bind_thm ("Always_reachable", invariant_reachable RS invariant_imp_Always);
 | |
| 317 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 318 | Goal "Always(A) = {F \\<in> program. F \\<in> invariant(reachable(F) Int A)}";
 | 
| 11479 | 319 | by (simp_tac (simpset() addsimps [Always_def, invariant_def, Stable_def, | 
| 12195 | 320 | Constrains_def2, stable_def, initially_def]) 1); | 
| 11479 | 321 | by (rtac equalityI 1); | 
| 322 | by (ALLGOALS(Clarify_tac)); | |
| 12195 | 323 | by (REPEAT(blast_tac (claset() addIs reachable.intrs@[reachable_type]) 1)); | 
| 11479 | 324 | qed "Always_eq_invariant_reachable"; | 
| 325 | ||
| 326 | (*the RHS is the traditional definition of the "always" operator*) | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 327 | Goal "Always(A) = {F \\<in> program. reachable(F) <= A}";
 | 
| 12484 | 328 | by (rtac equalityI 1); | 
| 11479 | 329 | by (ALLGOALS(Clarify_tac)); | 
| 330 | by (auto_tac (claset() addDs [invariant_includes_reachable], | |
| 331 | simpset() addsimps [subset_Int_iff, invariant_reachable, | |
| 332 | Always_eq_invariant_reachable])); | |
| 333 | qed "Always_eq_includes_reachable"; | |
| 334 | ||
| 12195 | 335 | Goalw [Always_def, initially_def] "Always(A) <= program"; | 
| 336 | by Auto_tac; | |
| 337 | qed "Always_type"; | |
| 11479 | 338 | |
| 339 | Goal "Always(state) = program"; | |
| 12484 | 340 | by (rtac equalityI 1); | 
| 12195 | 341 | by (auto_tac (claset() addDs [Always_type RS subsetD, | 
| 342 | reachable_type RS subsetD], | |
| 343 | simpset() addsimps [Always_eq_includes_reachable])); | |
| 11479 | 344 | qed "Always_state_eq"; | 
| 345 | Addsimps [Always_state_eq]; | |
| 346 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 347 | Goal "F \\<in> program ==> F \\<in> Always(state)"; | 
| 12195 | 348 | by (auto_tac (claset() addDs [reachable_type RS subsetD], simpset() | 
| 11479 | 349 | addsimps [Always_eq_includes_reachable])); | 
| 350 | qed "state_AlwaysI"; | |
| 351 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 352 | Goal "st_set(A) ==> Always(A) = (\\<Union>I \\<in> Pow(A). invariant(I))"; | 
| 11479 | 353 | by (simp_tac (simpset() addsimps [Always_eq_includes_reachable]) 1); | 
| 354 | by (rtac equalityI 1); | |
| 355 | by (ALLGOALS(Clarify_tac)); | |
| 356 | by (REPEAT(blast_tac (claset() | |
| 357 | addIs [invariantI, impOfSubs Init_into_reachable, | |
| 358 | impOfSubs invariant_includes_reachable] | |
| 12195 | 359 | addDs [invariant_type RS subsetD]) 1)); | 
| 11479 | 360 | qed "Always_eq_UN_invariant"; | 
| 361 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 362 | Goal "[| F \\<in> Always(A); A <= B |] ==> F \\<in> Always(B)"; | 
| 11479 | 363 | by (auto_tac (claset(), simpset() addsimps [Always_eq_includes_reachable])); | 
| 364 | qed "Always_weaken"; | |
| 365 | ||
| 366 | ||
| 367 | (*** "Co" rules involving Always ***) | |
| 368 | val Int_absorb2 = rewrite_rule [iff_def] subset_Int_iff RS conjunct1 RS mp; | |
| 369 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 370 | Goal "F \\<in> Always(I) ==> (F:(I Int A) Co A') <-> (F \\<in> A Co A')"; | 
| 11479 | 371 | by (asm_simp_tac | 
| 372 | (simpset() addsimps [Always_includes_reachable RS Int_absorb2, | |
| 373 | Constrains_def, Int_assoc RS sym]) 1); | |
| 374 | qed "Always_Constrains_pre"; | |
| 375 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 376 | Goal "F \\<in> Always(I) ==> (F \\<in> A Co (I Int A')) <->(F \\<in> A Co A')"; | 
| 11479 | 377 | by (asm_simp_tac | 
| 378 | (simpset() addsimps [Always_includes_reachable RS Int_absorb2, | |
| 379 | Constrains_eq_constrains, Int_assoc RS sym]) 1); | |
| 380 | qed "Always_Constrains_post"; | |
| 381 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 382 | Goal "[| F \\<in> Always(I); F \\<in> (I Int A) Co A' |] ==> F \\<in> A Co A'"; | 
| 12195 | 383 | by (blast_tac (claset() addIs [Always_Constrains_pre RS iffD1]) 1); | 
| 384 | qed "Always_ConstrainsI"; | |
| 11479 | 385 | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 386 | (* [| F \\<in> Always(I); F \\<in> A Co A' |] ==> F \\<in> A Co (I Int A') *) | 
| 11479 | 387 | bind_thm ("Always_ConstrainsD", Always_Constrains_post RS iffD2);
 | 
| 388 | ||
| 389 | (*The analogous proof of Always_LeadsTo_weaken doesn't terminate*) | |
| 12195 | 390 | Goal | 
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 391 | "[|F \\<in> Always(C); F \\<in> A Co A'; C Int B<=A; C Int A'<=B'|]==>F \\<in> B Co B'"; | 
| 11479 | 392 | by (rtac Always_ConstrainsI 1); | 
| 12195 | 393 | by (dtac Always_ConstrainsD 2); | 
| 394 | by (ALLGOALS(Asm_simp_tac)); | |
| 11479 | 395 | by (blast_tac (claset() addIs [Constrains_weaken]) 1); | 
| 396 | qed "Always_Constrains_weaken"; | |
| 397 | ||
| 398 | (** Conjoining Always properties **) | |
| 12195 | 399 | Goal "Always(A Int B) = Always(A) Int Always(B)"; | 
| 11479 | 400 | by (auto_tac (claset(), simpset() addsimps [Always_eq_includes_reachable])); | 
| 401 | qed "Always_Int_distrib"; | |
| 402 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 403 | (* the premise i \\<in> I is need since \\<Inter>is formally not defined for I=0 *) | 
| 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 404 | Goal "i \\<in> I==>Always(\\<Inter>i \\<in> I. A(i)) = (\\<Inter>i \\<in> I. Always(A(i)))"; | 
| 11479 | 405 | by (rtac equalityI 1); | 
| 12195 | 406 | by (auto_tac (claset(), simpset() addsimps | 
| 407 | [Inter_iff, Always_eq_includes_reachable])); | |
| 11479 | 408 | qed "Always_INT_distrib"; | 
| 409 | ||
| 410 | ||
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 411 | Goal "[| F \\<in> Always(A); F \\<in> Always(B) |] ==> F \\<in> Always(A Int B)"; | 
| 12195 | 412 | by (asm_simp_tac (simpset() addsimps [Always_Int_distrib]) 1); | 
| 11479 | 413 | qed "Always_Int_I"; | 
| 414 | ||
| 415 | (*Allows a kind of "implication introduction"*) | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 416 | Goal "[| F \\<in> Always(A) |] ==> (F \\<in> Always(C-A Un B)) <-> (F \\<in> Always(B))"; | 
| 11479 | 417 | by (auto_tac (claset(), simpset() addsimps [Always_eq_includes_reachable])); | 
| 12195 | 418 | qed "Always_Diff_Un_eq"; | 
| 11479 | 419 | |
| 420 | (*Delete the nearest invariance assumption (which will be the second one | |
| 421 | used by Always_Int_I) *) | |
| 422 | val Always_thin = | |
| 423 | read_instantiate_sg (sign_of thy) | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 424 |                 [("V", "?F \\<in> Always(?A)")] thin_rl;
 | 
| 11479 | 425 | |
| 426 | (*Combines two invariance ASSUMPTIONS into one. USEFUL??*) | |
| 427 | val Always_Int_tac = dtac Always_Int_I THEN' assume_tac THEN' etac Always_thin; | |
| 428 | ||
| 429 | (*Combines a list of invariance THEOREMS into one.*) | |
| 430 | val Always_Int_rule = foldr1 (fn (th1,th2) => [th1,th2] MRS Always_Int_I); | |
| 431 | ||
| 432 | (*To allow expansion of the program's definition when appropriate*) | |
| 14092 
68da54626309
Conversion of ZF/UNITY/{FP,Union} to Isar script.
 paulson parents: 
14060diff
changeset | 433 | val program_defs_ref = ref ([]: thm list); | 
| 11479 | 434 | |
| 435 | (*proves "co" properties when the program is specified*) | |
| 436 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
14046diff
changeset | 437 | fun gen_constrains_tac(cs,ss) i = | 
| 11479 | 438 | SELECT_GOAL | 
| 439 | (EVERY [REPEAT (Always_Int_tac 1), | |
| 440 | REPEAT (etac Always_ConstrainsI 1 | |
| 441 | ORELSE | |
| 442 | resolve_tac [StableI, stableI, | |
| 443 | constrains_imp_Constrains] 1), | |
| 444 | rtac constrainsI 1, | |
| 12195 | 445 | (* Three subgoals *) | 
| 446 | rewrite_goal_tac [st_set_def] 3, | |
| 447 | REPEAT (Force_tac 2), | |
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
14046diff
changeset | 448 | full_simp_tac (ss addsimps !program_defs_ref) 1, | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
14046diff
changeset | 449 | ALLGOALS (clarify_tac cs), | 
| 11479 | 450 | REPEAT (FIRSTGOAL (etac disjE)), | 
| 451 | ALLGOALS Clarify_tac, | |
| 452 | REPEAT (FIRSTGOAL (etac disjE)), | |
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
14046diff
changeset | 453 | ALLGOALS (clarify_tac cs), | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
14046diff
changeset | 454 | ALLGOALS (asm_full_simp_tac ss), | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
14046diff
changeset | 455 | ALLGOALS (clarify_tac cs)]) i; | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
14046diff
changeset | 456 | |
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
14046diff
changeset | 457 | fun constrains_tac st = gen_constrains_tac (claset(), simpset()) st; | 
| 11479 | 458 | |
| 459 | (*For proving invariants*) | |
| 460 | fun always_tac i = | |
| 461 | rtac AlwaysI i THEN Force_tac i THEN constrains_tac i; |