| author | wenzelm | 
| Wed, 20 Dec 2023 12:50:33 +0100 | |
| changeset 79318 | 74e245625a67 | 
| parent 76217 | 8655344f1cf6 | 
| child 82967 | 73af47bc277c | 
| permissions | -rw-r--r-- | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 1 | (* Title: ZF/UNITY/Constrains.thy | 
| 11479 | 2 | Author: Sidi O Ehmety, Computer Laboratory | 
| 3 | Copyright 2001 University of Cambridge | |
| 4 | *) | |
| 5 | ||
| 60770 | 6 | section\<open>Weak Safety Properties\<close> | 
| 15634 | 7 | |
| 8 | theory Constrains | |
| 9 | imports UNITY | |
| 24893 | 10 | begin | 
| 15634 | 11 | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 12 | consts traces :: "[i, i] \<Rightarrow> i" | 
| 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 13 | (* Initial states and program \<Rightarrow> (final state, reversed trace to it)... | 
| 12195 | 14 | the domain may also be state*list(state) *) | 
| 11479 | 15 | inductive | 
| 16 | domains | |
| 17 | "traces(init, acts)" <= | |
| 46823 | 18 | "(init \<union> (\<Union>act\<in>acts. field(act)))*list(\<Union>act\<in>acts. field(act))" | 
| 15634 | 19 | intros | 
| 11479 | 20 | (*Initial trace is empty*) | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 21 | Init: "s: init \<Longrightarrow> <s,[]> \<in> traces(init,acts)" | 
| 11479 | 22 | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 23 | Acts: "\<lbrakk>act:acts; \<langle>s,evs\<rangle> \<in> traces(init,acts); <s,s'>: act\<rbrakk> | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 24 | \<Longrightarrow> <s', Cons(s,evs)> \<in> traces(init, acts)" | 
| 11479 | 25 | |
| 15634 | 26 | type_intros list.intros UnI1 UnI2 UN_I fieldI2 fieldI1 | 
| 11479 | 27 | |
| 28 | ||
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 29 | consts reachable :: "i\<Rightarrow>i" | 
| 11479 | 30 | inductive | 
| 31 | domains | |
| 46823 | 32 | "reachable(F)" \<subseteq> "Init(F) \<union> (\<Union>act\<in>Acts(F). field(act))" | 
| 15634 | 33 | intros | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 34 | Init: "s:Init(F) \<Longrightarrow> s:reachable(F)" | 
| 11479 | 35 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 36 | Acts: "\<lbrakk>act: Acts(F); s:reachable(F); <s,s'>: act\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 37 | \<Longrightarrow> s':reachable(F)" | 
| 11479 | 38 | |
| 15634 | 39 | type_intros UnI1 UnI2 fieldI2 UN_I | 
| 11479 | 40 | |
| 41 | ||
| 24893 | 42 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 43 | Constrains :: "[i,i] \<Rightarrow> i" (infixl \<open>Co\<close> 60) where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 44 |   "A Co B \<equiv> {F:program. F:(reachable(F) \<inter> A) co B}"
 | 
| 11479 | 45 | |
| 24893 | 46 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 47 | op_Unless :: "[i, i] \<Rightarrow> i" (infixl \<open>Unless\<close> 60) where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 48 | "A Unless B \<equiv> (A-B) Co (A \<union> B)" | 
| 11479 | 49 | |
| 24893 | 50 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 51 | Stable :: "i \<Rightarrow> i" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 52 | "Stable(A) \<equiv> A Co A" | 
| 11479 | 53 | |
| 24893 | 54 | definition | 
| 11479 | 55 | (*Always is the weak form of "invariant"*) | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 56 | Always :: "i \<Rightarrow> i" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 57 | "Always(A) \<equiv> initially(A) \<inter> Stable(A)" | 
| 11479 | 58 | |
| 15634 | 59 | |
| 60 | (*** traces and reachable ***) | |
| 61 | ||
| 46823 | 62 | lemma reachable_type: "reachable(F) \<subseteq> state" | 
| 15634 | 63 | apply (cut_tac F = F in Init_type) | 
| 64 | apply (cut_tac F = F in Acts_type) | |
| 65 | apply (cut_tac F = F in reachable.dom_subset, blast) | |
| 66 | done | |
| 67 | ||
| 68 | lemma st_set_reachable: "st_set(reachable(F))" | |
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 69 | unfolding st_set_def | 
| 15634 | 70 | apply (rule reachable_type) | 
| 71 | done | |
| 72 | declare st_set_reachable [iff] | |
| 73 | ||
| 46823 | 74 | lemma reachable_Int_state: "reachable(F) \<inter> state = reachable(F)" | 
| 15634 | 75 | by (cut_tac reachable_type, auto) | 
| 76 | declare reachable_Int_state [iff] | |
| 77 | ||
| 46823 | 78 | lemma state_Int_reachable: "state \<inter> reachable(F) = reachable(F)" | 
| 15634 | 79 | by (cut_tac reachable_type, auto) | 
| 80 | declare state_Int_reachable [iff] | |
| 81 | ||
| 82 | lemma reachable_equiv_traces: | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 83 | "F \<in> program \<Longrightarrow> reachable(F)={s \<in> state. \<exists>evs. \<langle>s,evs\<rangle>:traces(Init(F), Acts(F))}"
 | 
| 15634 | 84 | apply (rule equalityI, safe) | 
| 85 | apply (blast dest: reachable_type [THEN subsetD]) | |
| 86 | apply (erule_tac [2] traces.induct) | |
| 87 | apply (erule reachable.induct) | |
| 88 | apply (blast intro: reachable.intros traces.intros)+ | |
| 89 | done | |
| 90 | ||
| 46823 | 91 | lemma Init_into_reachable: "Init(F) \<subseteq> reachable(F)" | 
| 15634 | 92 | by (blast intro: reachable.intros) | 
| 93 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 94 | lemma stable_reachable: "\<lbrakk>F \<in> program; G \<in> program; | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 95 | Acts(G) \<subseteq> Acts(F)\<rbrakk> \<Longrightarrow> G \<in> stable(reachable(F))" | 
| 15634 | 96 | apply (blast intro: stableI constrainsI st_setI | 
| 97 | reachable_type [THEN subsetD] reachable.intros) | |
| 98 | done | |
| 99 | ||
| 100 | declare stable_reachable [intro!] | |
| 101 | declare stable_reachable [simp] | |
| 102 | ||
| 103 | (*The set of all reachable states is an invariant...*) | |
| 104 | lemma invariant_reachable: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 105 | "F \<in> program \<Longrightarrow> F \<in> invariant(reachable(F))" | 
| 76217 | 106 | unfolding invariant_def initially_def | 
| 15634 | 107 | apply (blast intro: reachable_type [THEN subsetD] reachable.intros) | 
| 108 | done | |
| 109 | ||
| 110 | (*...in fact the strongest invariant!*) | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 111 | lemma invariant_includes_reachable: "F \<in> invariant(A) \<Longrightarrow> reachable(F) \<subseteq> A" | 
| 15634 | 112 | apply (cut_tac F = F in Acts_type) | 
| 113 | apply (cut_tac F = F in Init_type) | |
| 114 | apply (cut_tac F = F in reachable_type) | |
| 115 | apply (simp (no_asm_use) add: stable_def constrains_def invariant_def initially_def) | |
| 116 | apply (rule subsetI) | |
| 117 | apply (erule reachable.induct) | |
| 118 | apply (blast intro: reachable.intros)+ | |
| 119 | done | |
| 120 | ||
| 121 | (*** Co ***) | |
| 122 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 123 | lemma constrains_reachable_Int: "F \<in> B co B'\<Longrightarrow>F:(reachable(F) \<inter> B) co (reachable(F) \<inter> B')" | 
| 15634 | 124 | apply (frule constrains_type [THEN subsetD]) | 
| 125 | apply (frule stable_reachable [OF _ _ subset_refl]) | |
| 126 | apply (simp_all add: stable_def constrains_Int) | |
| 127 | done | |
| 128 | ||
| 129 | (*Resembles the previous definition of Constrains*) | |
| 130 | lemma Constrains_eq_constrains: | |
| 46823 | 131 | "A Co B = {F \<in> program. F:(reachable(F) \<inter> A) co (reachable(F)  \<inter>  B)}"
 | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 132 | unfolding Constrains_def | 
| 15634 | 133 | apply (blast dest: constrains_reachable_Int constrains_type [THEN subsetD] | 
| 134 | intro: constrains_weaken) | |
| 135 | done | |
| 136 | ||
| 137 | lemmas Constrains_def2 = Constrains_eq_constrains [THEN eq_reflection] | |
| 138 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 139 | lemma constrains_imp_Constrains: "F \<in> A co A' \<Longrightarrow> F \<in> A Co A'" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 140 | unfolding Constrains_def | 
| 15634 | 141 | apply (blast intro: constrains_weaken_L dest: constrainsD2) | 
| 142 | done | |
| 143 | ||
| 144 | lemma ConstrainsI: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 145 | "\<lbrakk>\<And>act s s'. \<lbrakk>act \<in> Acts(F); <s,s'>:act; s \<in> A\<rbrakk> \<Longrightarrow> s':A'; | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 146 | F \<in> program\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 147 | \<Longrightarrow> F \<in> A Co A'" | 
| 15634 | 148 | apply (auto simp add: Constrains_def constrains_def st_set_def) | 
| 149 | apply (blast dest: reachable_type [THEN subsetD]) | |
| 150 | done | |
| 151 | ||
| 152 | lemma Constrains_type: | |
| 46823 | 153 | "A Co B \<subseteq> program" | 
| 15634 | 154 | apply (unfold Constrains_def, blast) | 
| 155 | done | |
| 156 | ||
| 46823 | 157 | lemma Constrains_empty: "F \<in> 0 Co B \<longleftrightarrow> F \<in> program" | 
| 15634 | 158 | by (auto dest: Constrains_type [THEN subsetD] | 
| 159 | intro: constrains_imp_Constrains) | |
| 160 | declare Constrains_empty [iff] | |
| 161 | ||
| 46823 | 162 | lemma Constrains_state: "F \<in> A Co state \<longleftrightarrow> F \<in> program" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 163 | unfolding Constrains_def | 
| 15634 | 164 | apply (auto dest: Constrains_type [THEN subsetD] intro: constrains_imp_Constrains) | 
| 165 | done | |
| 166 | declare Constrains_state [iff] | |
| 167 | ||
| 168 | lemma Constrains_weaken_R: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 169 | "\<lbrakk>F \<in> A Co A'; A'<=B'\<rbrakk> \<Longrightarrow> F \<in> A Co B'" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 170 | unfolding Constrains_def2 | 
| 15634 | 171 | apply (blast intro: constrains_weaken_R) | 
| 172 | done | |
| 173 | ||
| 174 | lemma Constrains_weaken_L: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 175 | "\<lbrakk>F \<in> A Co A'; B<=A\<rbrakk> \<Longrightarrow> F \<in> B Co A'" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 176 | unfolding Constrains_def2 | 
| 15634 | 177 | apply (blast intro: constrains_weaken_L st_set_subset) | 
| 178 | done | |
| 179 | ||
| 180 | lemma Constrains_weaken: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 181 | "\<lbrakk>F \<in> A Co A'; B<=A; A'<=B'\<rbrakk> \<Longrightarrow> F \<in> B Co B'" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 182 | unfolding Constrains_def2 | 
| 15634 | 183 | apply (blast intro: constrains_weaken st_set_subset) | 
| 184 | done | |
| 185 | ||
| 186 | (** Union **) | |
| 187 | lemma Constrains_Un: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 188 | "\<lbrakk>F \<in> A Co A'; F \<in> B Co B'\<rbrakk> \<Longrightarrow> F \<in> (A \<union> B) Co (A' \<union> B')" | 
| 15634 | 189 | apply (unfold Constrains_def2, auto) | 
| 190 | apply (simp add: Int_Un_distrib) | |
| 191 | apply (blast intro: constrains_Un) | |
| 192 | done | |
| 193 | ||
| 194 | lemma Constrains_UN: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 195 | "\<lbrakk>(\<And>i. i \<in> I\<Longrightarrow>F \<in> A(i) Co A'(i)); F \<in> program\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 196 | \<Longrightarrow> F:(\<Union>i \<in> I. A(i)) Co (\<Union>i \<in> I. A'(i))" | 
| 15634 | 197 | by (auto intro: constrains_UN simp del: UN_simps | 
| 198 | simp add: Constrains_def2 Int_UN_distrib) | |
| 199 | ||
| 200 | ||
| 201 | (** Intersection **) | |
| 202 | ||
| 203 | lemma Constrains_Int: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 204 | "\<lbrakk>F \<in> A Co A'; F \<in> B Co B'\<rbrakk>\<Longrightarrow> F:(A \<inter> B) Co (A' \<inter> B')" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 205 | unfolding Constrains_def | 
| 46823 | 206 | apply (subgoal_tac "reachable (F) \<inter> (A \<inter> B) = (reachable (F) \<inter> A) \<inter> (reachable (F) \<inter> B) ") | 
| 15634 | 207 | apply (auto intro: constrains_Int) | 
| 208 | done | |
| 209 | ||
| 210 | lemma Constrains_INT: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 211 | "\<lbrakk>(\<And>i. i \<in> I \<Longrightarrow>F \<in> A(i) Co A'(i)); F \<in> program\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 212 | \<Longrightarrow> F:(\<Inter>i \<in> I. A(i)) Co (\<Inter>i \<in> I. A'(i))" | 
| 15634 | 213 | apply (simp (no_asm_simp) del: INT_simps add: Constrains_def INT_extend_simps) | 
| 214 | apply (rule constrains_INT) | |
| 215 | apply (auto simp add: Constrains_def) | |
| 216 | done | |
| 217 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 218 | lemma Constrains_imp_subset: "F \<in> A Co A' \<Longrightarrow> reachable(F) \<inter> A \<subseteq> A'" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 219 | unfolding Constrains_def | 
| 15634 | 220 | apply (blast dest: constrains_imp_subset) | 
| 221 | done | |
| 222 | ||
| 223 | lemma Constrains_trans: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 224 | "\<lbrakk>F \<in> A Co B; F \<in> B Co C\<rbrakk> \<Longrightarrow> F \<in> A Co C" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 225 | unfolding Constrains_def2 | 
| 15634 | 226 | apply (blast intro: constrains_trans constrains_weaken) | 
| 227 | done | |
| 228 | ||
| 229 | lemma Constrains_cancel: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 230 | "\<lbrakk>F \<in> A Co (A' \<union> B); F \<in> B Co B'\<rbrakk> \<Longrightarrow> F \<in> A Co (A' \<union> B')" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 231 | unfolding Constrains_def2 | 
| 15634 | 232 | apply (simp (no_asm_use) add: Int_Un_distrib) | 
| 233 | apply (blast intro: constrains_cancel) | |
| 234 | done | |
| 235 | ||
| 236 | (*** Stable ***) | |
| 237 | (* Useful because there's no Stable_weaken. [Tanja Vos] *) | |
| 238 | ||
| 239 | lemma stable_imp_Stable: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 240 | "F \<in> stable(A) \<Longrightarrow> F \<in> Stable(A)" | 
| 15634 | 241 | |
| 76217 | 242 | unfolding stable_def Stable_def | 
| 15634 | 243 | apply (erule constrains_imp_Constrains) | 
| 244 | done | |
| 245 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 246 | lemma Stable_eq: "\<lbrakk>F \<in> Stable(A); A = B\<rbrakk> \<Longrightarrow> F \<in> Stable(B)" | 
| 15634 | 247 | by blast | 
| 248 | ||
| 249 | lemma Stable_eq_stable: | |
| 46823 | 250 | "F \<in> Stable(A) \<longleftrightarrow> (F \<in> stable(reachable(F) \<inter> A))" | 
| 15634 | 251 | apply (auto dest: constrainsD2 simp add: Stable_def stable_def Constrains_def2) | 
| 252 | done | |
| 253 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 254 | lemma StableI: "F \<in> A Co A \<Longrightarrow> F \<in> Stable(A)" | 
| 15634 | 255 | by (unfold Stable_def, assumption) | 
| 256 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 257 | lemma StableD: "F \<in> Stable(A) \<Longrightarrow> F \<in> A Co A" | 
| 15634 | 258 | by (unfold Stable_def, assumption) | 
| 259 | ||
| 260 | lemma Stable_Un: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 261 | "\<lbrakk>F \<in> Stable(A); F \<in> Stable(A')\<rbrakk> \<Longrightarrow> F \<in> Stable(A \<union> A')" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 262 | unfolding Stable_def | 
| 15634 | 263 | apply (blast intro: Constrains_Un) | 
| 264 | done | |
| 265 | ||
| 266 | lemma Stable_Int: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 267 | "\<lbrakk>F \<in> Stable(A); F \<in> Stable(A')\<rbrakk> \<Longrightarrow> F \<in> Stable (A \<inter> A')" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 268 | unfolding Stable_def | 
| 15634 | 269 | apply (blast intro: Constrains_Int) | 
| 270 | done | |
| 271 | ||
| 272 | lemma Stable_Constrains_Un: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 273 | "\<lbrakk>F \<in> Stable(C); F \<in> A Co (C \<union> A')\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 274 | \<Longrightarrow> F \<in> (C \<union> A) Co (C \<union> A')" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 275 | unfolding Stable_def | 
| 15634 | 276 | apply (blast intro: Constrains_Un [THEN Constrains_weaken_R]) | 
| 277 | done | |
| 278 | ||
| 279 | lemma Stable_Constrains_Int: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 280 | "\<lbrakk>F \<in> Stable(C); F \<in> (C \<inter> A) Co A'\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 281 | \<Longrightarrow> F \<in> (C \<inter> A) Co (C \<inter> A')" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 282 | unfolding Stable_def | 
| 15634 | 283 | apply (blast intro: Constrains_Int [THEN Constrains_weaken]) | 
| 284 | done | |
| 285 | ||
| 286 | lemma Stable_UN: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 287 | "\<lbrakk>(\<And>i. i \<in> I \<Longrightarrow> F \<in> Stable(A(i))); F \<in> program\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 288 | \<Longrightarrow> F \<in> Stable (\<Union>i \<in> I. A(i))" | 
| 15634 | 289 | apply (simp add: Stable_def) | 
| 290 | apply (blast intro: Constrains_UN) | |
| 291 | done | |
| 292 | ||
| 293 | lemma Stable_INT: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 294 | "\<lbrakk>(\<And>i. i \<in> I \<Longrightarrow> F \<in> Stable(A(i))); F \<in> program\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 295 | \<Longrightarrow> F \<in> Stable (\<Inter>i \<in> I. A(i))" | 
| 15634 | 296 | apply (simp add: Stable_def) | 
| 297 | apply (blast intro: Constrains_INT) | |
| 298 | done | |
| 299 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 300 | lemma Stable_reachable: "F \<in> program \<Longrightarrow>F \<in> Stable (reachable(F))" | 
| 15634 | 301 | apply (simp (no_asm_simp) add: Stable_eq_stable Int_absorb) | 
| 302 | done | |
| 303 | ||
| 46823 | 304 | lemma Stable_type: "Stable(A) \<subseteq> program" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 305 | unfolding Stable_def | 
| 15634 | 306 | apply (rule Constrains_type) | 
| 307 | done | |
| 308 | ||
| 309 | (*** The Elimination Theorem. The "free" m has become universally quantified! | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 310 | Should the premise be \<And>m instead of \<forall>m ? Would make it harder to use | 
| 15634 | 311 | in forward proof. ***) | 
| 312 | ||
| 313 | lemma Elimination: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 314 |     "\<lbrakk>\<forall>m \<in> M. F \<in> ({s \<in> A. x(s) = m}) Co (B(m)); F \<in> program\<rbrakk>  
 | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 315 |      \<Longrightarrow> F \<in> ({s \<in> A. x(s):M}) Co (\<Union>m \<in> M. B(m))"
 | 
| 15634 | 316 | apply (unfold Constrains_def, auto) | 
| 46823 | 317 | apply (rule_tac A1 = "reachable (F) \<inter> A" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 318 | in UNITY.elimination [THEN constrains_weaken_L]) | 
| 15634 | 319 | apply (auto intro: constrains_weaken_L) | 
| 320 | done | |
| 321 | ||
| 322 | (* As above, but for the special case of A=state *) | |
| 323 | lemma Elimination2: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 324 |  "\<lbrakk>\<forall>m \<in> M. F \<in> {s \<in> state. x(s) = m} Co B(m); F \<in> program\<rbrakk>  
 | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 325 |      \<Longrightarrow> F \<in> {s \<in> state. x(s):M} Co (\<Union>m \<in> M. B(m))"
 | 
| 15634 | 326 | apply (blast intro: Elimination) | 
| 327 | done | |
| 328 | ||
| 329 | (** Unless **) | |
| 330 | ||
| 331 | lemma Unless_type: "A Unless B <=program" | |
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 332 | unfolding op_Unless_def | 
| 15634 | 333 | apply (rule Constrains_type) | 
| 334 | done | |
| 335 | ||
| 336 | (*** Specialized laws for handling Always ***) | |
| 337 | ||
| 338 | (** Natural deduction rules for "Always A" **) | |
| 339 | ||
| 340 | lemma AlwaysI: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 341 | "\<lbrakk>Init(F)<=A; F \<in> Stable(A)\<rbrakk> \<Longrightarrow> F \<in> Always(A)" | 
| 15634 | 342 | |
| 76217 | 343 | unfolding Always_def initially_def | 
| 15634 | 344 | apply (frule Stable_type [THEN subsetD], auto) | 
| 345 | done | |
| 346 | ||
| 76214 | 347 | lemma AlwaysD: "F \<in> Always(A) \<Longrightarrow> Init(F)<=A \<and> F \<in> Stable(A)" | 
| 15634 | 348 | by (simp add: Always_def initially_def) | 
| 349 | ||
| 45602 | 350 | lemmas AlwaysE = AlwaysD [THEN conjE] | 
| 351 | lemmas Always_imp_Stable = AlwaysD [THEN conjunct2] | |
| 15634 | 352 | |
| 353 | (*The set of all reachable states is Always*) | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 354 | lemma Always_includes_reachable: "F \<in> Always(A) \<Longrightarrow> reachable(F) \<subseteq> A" | 
| 15634 | 355 | apply (simp (no_asm_use) add: Stable_def Constrains_def constrains_def Always_def initially_def) | 
| 356 | apply (rule subsetI) | |
| 357 | apply (erule reachable.induct) | |
| 358 | apply (blast intro: reachable.intros)+ | |
| 359 | done | |
| 360 | ||
| 361 | lemma invariant_imp_Always: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 362 | "F \<in> invariant(A) \<Longrightarrow> F \<in> Always(A)" | 
| 76217 | 363 | unfolding Always_def invariant_def Stable_def stable_def | 
| 15634 | 364 | apply (blast intro: constrains_imp_Constrains) | 
| 365 | done | |
| 366 | ||
| 45602 | 367 | lemmas Always_reachable = invariant_reachable [THEN invariant_imp_Always] | 
| 15634 | 368 | |
| 46823 | 369 | lemma Always_eq_invariant_reachable: "Always(A) = {F \<in> program. F \<in> invariant(reachable(F) \<inter> A)}"
 | 
| 15634 | 370 | apply (simp (no_asm) add: Always_def invariant_def Stable_def Constrains_def2 stable_def initially_def) | 
| 371 | apply (rule equalityI, auto) | |
| 372 | apply (blast intro: reachable.intros reachable_type) | |
| 373 | done | |
| 374 | ||
| 375 | (*the RHS is the traditional definition of the "always" operator*) | |
| 46823 | 376 | lemma Always_eq_includes_reachable: "Always(A) = {F \<in> program. reachable(F) \<subseteq> A}"
 | 
| 15634 | 377 | apply (rule equalityI, safe) | 
| 378 | apply (auto dest: invariant_includes_reachable | |
| 379 | simp add: subset_Int_iff invariant_reachable Always_eq_invariant_reachable) | |
| 380 | done | |
| 381 | ||
| 46823 | 382 | lemma Always_type: "Always(A) \<subseteq> program" | 
| 15634 | 383 | by (unfold Always_def initially_def, auto) | 
| 384 | ||
| 385 | lemma Always_state_eq: "Always(state) = program" | |
| 386 | apply (rule equalityI) | |
| 387 | apply (auto dest: Always_type [THEN subsetD] reachable_type [THEN subsetD] | |
| 388 | simp add: Always_eq_includes_reachable) | |
| 389 | done | |
| 390 | declare Always_state_eq [simp] | |
| 391 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 392 | lemma state_AlwaysI: "F \<in> program \<Longrightarrow> F \<in> Always(state)" | 
| 15634 | 393 | by (auto dest: reachable_type [THEN subsetD] | 
| 394 | simp add: Always_eq_includes_reachable) | |
| 395 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 396 | lemma Always_eq_UN_invariant: "st_set(A) \<Longrightarrow> Always(A) = (\<Union>I \<in> Pow(A). invariant(I))" | 
| 15634 | 397 | apply (simp (no_asm) add: Always_eq_includes_reachable) | 
| 398 | apply (rule equalityI, auto) | |
| 399 | apply (blast intro: invariantI rev_subsetD [OF _ Init_into_reachable] | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32149diff
changeset | 400 | rev_subsetD [OF _ invariant_includes_reachable] | 
| 15634 | 401 | dest: invariant_type [THEN subsetD])+ | 
| 402 | done | |
| 403 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 404 | lemma Always_weaken: "\<lbrakk>F \<in> Always(A); A \<subseteq> B\<rbrakk> \<Longrightarrow> F \<in> Always(B)" | 
| 15634 | 405 | by (auto simp add: Always_eq_includes_reachable) | 
| 406 | ||
| 407 | ||
| 408 | (*** "Co" rules involving Always ***) | |
| 409 | lemmas Int_absorb2 = subset_Int_iff [unfolded iff_def, THEN conjunct1, THEN mp] | |
| 410 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 411 | lemma Always_Constrains_pre: "F \<in> Always(I) \<Longrightarrow> (F:(I \<inter> A) Co A') \<longleftrightarrow> (F \<in> A Co A')" | 
| 15634 | 412 | apply (simp (no_asm_simp) add: Always_includes_reachable [THEN Int_absorb2] Constrains_def Int_assoc [symmetric]) | 
| 413 | done | |
| 414 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 415 | lemma Always_Constrains_post: "F \<in> Always(I) \<Longrightarrow> (F \<in> A Co (I \<inter> A')) \<longleftrightarrow>(F \<in> A Co A')" | 
| 15634 | 416 | apply (simp (no_asm_simp) add: Always_includes_reachable [THEN Int_absorb2] Constrains_eq_constrains Int_assoc [symmetric]) | 
| 417 | done | |
| 418 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 419 | lemma Always_ConstrainsI: "\<lbrakk>F \<in> Always(I); F \<in> (I \<inter> A) Co A'\<rbrakk> \<Longrightarrow> F \<in> A Co A'" | 
| 15634 | 420 | by (blast intro: Always_Constrains_pre [THEN iffD1]) | 
| 421 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 422 | (* \<lbrakk>F \<in> Always(I); F \<in> A Co A'\<rbrakk> \<Longrightarrow> F \<in> A Co (I \<inter> A') *) | 
| 45602 | 423 | lemmas Always_ConstrainsD = Always_Constrains_post [THEN iffD2] | 
| 15634 | 424 | |
| 425 | (*The analogous proof of Always_LeadsTo_weaken doesn't terminate*) | |
| 426 | lemma Always_Constrains_weaken: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 427 | "\<lbrakk>F \<in> Always(C); F \<in> A Co A'; C \<inter> B<=A; C \<inter> A'<=B'\<rbrakk>\<Longrightarrow>F \<in> B Co B'" | 
| 15634 | 428 | apply (rule Always_ConstrainsI) | 
| 429 | apply (drule_tac [2] Always_ConstrainsD, simp_all) | |
| 430 | apply (blast intro: Constrains_weaken) | |
| 431 | done | |
| 432 | ||
| 433 | (** Conjoining Always properties **) | |
| 46823 | 434 | lemma Always_Int_distrib: "Always(A \<inter> B) = Always(A) \<inter> Always(B)" | 
| 15634 | 435 | by (auto simp add: Always_eq_includes_reachable) | 
| 436 | ||
| 437 | (* the premise i \<in> I is need since \<Inter>is formally not defined for I=0 *) | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 438 | lemma Always_INT_distrib: "i \<in> I\<Longrightarrow>Always(\<Inter>i \<in> I. A(i)) = (\<Inter>i \<in> I. Always(A(i)))" | 
| 15634 | 439 | apply (rule equalityI) | 
| 440 | apply (auto simp add: Inter_iff Always_eq_includes_reachable) | |
| 441 | done | |
| 442 | ||
| 443 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 444 | lemma Always_Int_I: "\<lbrakk>F \<in> Always(A); F \<in> Always(B)\<rbrakk> \<Longrightarrow> F \<in> Always(A \<inter> B)" | 
| 15634 | 445 | apply (simp (no_asm_simp) add: Always_Int_distrib) | 
| 446 | done | |
| 447 | ||
| 448 | (*Allows a kind of "implication introduction"*) | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 449 | lemma Always_Diff_Un_eq: "\<lbrakk>F \<in> Always(A)\<rbrakk> \<Longrightarrow> (F \<in> Always(C-A \<union> B)) \<longleftrightarrow> (F \<in> Always(B))" | 
| 15634 | 450 | by (auto simp add: Always_eq_includes_reachable) | 
| 451 | ||
| 452 | (*Delete the nearest invariance assumption (which will be the second one | |
| 453 | used by Always_Int_I) *) | |
| 68233 | 454 | lemmas Always_thin = thin_rl [of "F \<in> Always(A)"] for F A | 
| 15634 | 455 | |
| 57945 
cacb00a569e0
prefer 'named_theorems' over Named_Thms, with subtle change of semantics due to visual order vs. internal reverse order;
 wenzelm parents: 
54742diff
changeset | 456 | (*To allow expansion of the program's definition when appropriate*) | 
| 
cacb00a569e0
prefer 'named_theorems' over Named_Thms, with subtle change of semantics due to visual order vs. internal reverse order;
 wenzelm parents: 
54742diff
changeset | 457 | named_theorems program "program definitions" | 
| 
cacb00a569e0
prefer 'named_theorems' over Named_Thms, with subtle change of semantics due to visual order vs. internal reverse order;
 wenzelm parents: 
54742diff
changeset | 458 | |
| 15634 | 459 | ML | 
| 60770 | 460 | \<open> | 
| 15634 | 461 | (*Combines two invariance ASSUMPTIONS into one. USEFUL??*) | 
| 58963 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 wenzelm parents: 
58871diff
changeset | 462 | fun Always_Int_tac ctxt = | 
| 60754 | 463 |   dresolve_tac ctxt @{thms Always_Int_I} THEN'
 | 
| 464 | assume_tac ctxt THEN' | |
| 465 |   eresolve_tac ctxt @{thms Always_thin};
 | |
| 15634 | 466 | |
| 467 | (*Combines a list of invariance THEOREMS into one.*) | |
| 24893 | 468 | val Always_Int_rule = foldr1 (fn (th1,th2) => [th1,th2] MRS @{thm Always_Int_I});
 | 
| 15634 | 469 | |
| 470 | (*proves "co" properties when the program is specified*) | |
| 471 | ||
| 23894 
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
 wenzelm parents: 
23543diff
changeset | 472 | fun constrains_tac ctxt = | 
| 15634 | 473 | SELECT_GOAL | 
| 58963 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 wenzelm parents: 
58871diff
changeset | 474 | (EVERY [REPEAT (Always_Int_tac ctxt 1), | 
| 60754 | 475 |               REPEAT (eresolve_tac ctxt @{thms Always_ConstrainsI} 1
 | 
| 15634 | 476 | ORELSE | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 477 |                       resolve_tac ctxt [@{thm StableI}, @{thm stableI},
 | 
| 24893 | 478 |                                    @{thm constrains_imp_Constrains}] 1),
 | 
| 60754 | 479 |               resolve_tac ctxt @{thms constrainsI} 1,
 | 
| 15634 | 480 | (* Three subgoals *) | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
51717diff
changeset | 481 |               rewrite_goal_tac ctxt [@{thm st_set_def}] 3,
 | 
| 42793 | 482 | REPEAT (force_tac ctxt 2), | 
| 69593 | 483 | full_simp_tac (ctxt addsimps (Named_Theorems.get ctxt \<^named_theorems>\<open>program\<close>)) 1, | 
| 42793 | 484 | ALLGOALS (clarify_tac ctxt), | 
| 60754 | 485 |               REPEAT (FIRSTGOAL (eresolve_tac ctxt @{thms disjE})),
 | 
| 42793 | 486 | ALLGOALS (clarify_tac ctxt), | 
| 60754 | 487 |               REPEAT (FIRSTGOAL (eresolve_tac ctxt @{thms disjE})),
 | 
| 42793 | 488 | ALLGOALS (clarify_tac ctxt), | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46823diff
changeset | 489 | ALLGOALS (asm_full_simp_tac ctxt), | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46823diff
changeset | 490 | ALLGOALS (clarify_tac ctxt)]); | 
| 15634 | 491 | |
| 492 | (*For proving invariants*) | |
| 42793 | 493 | fun always_tac ctxt i = | 
| 60754 | 494 |   resolve_tac ctxt @{thms AlwaysI} i THEN
 | 
| 495 | force_tac ctxt i | |
| 496 | THEN constrains_tac ctxt i; | |
| 60770 | 497 | \<close> | 
| 15634 | 498 | |
| 60770 | 499 | method_setup safety = \<open> | 
| 500 | Scan.succeed (SIMPLE_METHOD' o constrains_tac)\<close> | |
| 21588 | 501 | "for proving safety properties" | 
| 15634 | 502 | |
| 60770 | 503 | method_setup always = \<open> | 
| 504 | Scan.succeed (SIMPLE_METHOD' o always_tac)\<close> | |
| 23894 
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
 wenzelm parents: 
23543diff
changeset | 505 | "for proving invariants" | 
| 15634 | 506 | |
| 11479 | 507 | end |