| author | blanchet | 
| Thu, 29 Oct 2009 12:09:32 +0100 | |
| changeset 33566 | 1c62ac4ef6d1 | 
| parent 32687 | 27530efec97a | 
| child 34233 | 156c42518cfc | 
| permissions | -rw-r--r-- | 
| 13020 | 1 | |
| 2 | header {* \section{The Single Mutator Case} *}
 | |
| 3 | ||
| 16417 | 4 | theory Gar_Coll imports Graph OG_Syntax begin | 
| 13020 | 5 | |
| 13624 | 6 | declare psubsetE [rule del] | 
| 7 | ||
| 13020 | 8 | text {* Declaration of variables: *}
 | 
| 9 | ||
| 10 | record gar_coll_state = | |
| 11 | M :: nodes | |
| 12 | E :: edges | |
| 13 | bc :: "nat set" | |
| 14 | obc :: "nat set" | |
| 15 | Ma :: nodes | |
| 16 | ind :: nat | |
| 17 | k :: nat | |
| 18 | z :: bool | |
| 19 | ||
| 20 | subsection {* The Mutator *}
 | |
| 21 | ||
| 22 | text {* The mutator first redirects an arbitrary edge @{text "R"} from
 | |
| 23 | an arbitrary accessible node towards an arbitrary accessible node | |
| 24 | @{text "T"}.  It then colors the new target @{text "T"} black. 
 | |
| 25 | ||
| 26 | We declare the arbitrarily selected node and edge as constants:*} | |
| 27 | ||
| 28 | consts R :: nat T :: nat | |
| 29 | ||
| 30 | text {* \noindent The following predicate states, given a list of
 | |
| 31 | nodes @{text "m"} and a list of edges @{text "e"}, the conditions
 | |
| 32 | under which the selected edge @{text "R"} and node @{text "T"} are
 | |
| 33 | valid: *} | |
| 34 | ||
| 35 | constdefs | |
| 36 | Mut_init :: "gar_coll_state \<Rightarrow> bool" | |
| 37 | "Mut_init \<equiv> \<guillemotleft> T \<in> Reach \<acute>E \<and> R < length \<acute>E \<and> T < length \<acute>M \<guillemotright>" | |
| 38 | ||
| 39 | text {* \noindent For the mutator we
 | |
| 40 | consider two modules, one for each action. An auxiliary variable | |
| 41 | @{text "\<acute>z"} is set to false if the mutator has already redirected an
 | |
| 42 | edge but has not yet colored the new target. *} | |
| 43 | ||
| 44 | constdefs | |
| 45 | Redirect_Edge :: "gar_coll_state ann_com" | |
| 46 |   "Redirect_Edge \<equiv> .{\<acute>Mut_init \<and> \<acute>z}. \<langle>\<acute>E:=\<acute>E[R:=(fst(\<acute>E!R), T)],, \<acute>z:= (\<not>\<acute>z)\<rangle>"
 | |
| 47 | ||
| 48 | Color_Target :: "gar_coll_state ann_com" | |
| 49 |   "Color_Target \<equiv> .{\<acute>Mut_init \<and> \<not>\<acute>z}. \<langle>\<acute>M:=\<acute>M[T:=Black],, \<acute>z:= (\<not>\<acute>z)\<rangle>"
 | |
| 50 | ||
| 51 | Mutator :: "gar_coll_state ann_com" | |
| 52 | "Mutator \<equiv> | |
| 53 |   .{\<acute>Mut_init \<and> \<acute>z}. 
 | |
| 54 |   WHILE True INV .{\<acute>Mut_init \<and> \<acute>z}. 
 | |
| 55 | DO Redirect_Edge ;; Color_Target OD" | |
| 56 | ||
| 57 | subsubsection {* Correctness of the mutator *}
 | |
| 58 | ||
| 59 | lemmas mutator_defs = Mut_init_def Redirect_Edge_def Color_Target_def | |
| 60 | ||
| 61 | lemma Redirect_Edge: | |
| 62 | "\<turnstile> Redirect_Edge pre(Color_Target)" | |
| 63 | apply (unfold mutator_defs) | |
| 64 | apply annhoare | |
| 65 | apply(simp_all) | |
| 66 | apply(force elim:Graph2) | |
| 67 | done | |
| 68 | ||
| 69 | lemma Color_Target: | |
| 70 |   "\<turnstile> Color_Target .{\<acute>Mut_init \<and> \<acute>z}."
 | |
| 71 | apply (unfold mutator_defs) | |
| 72 | apply annhoare | |
| 73 | apply(simp_all) | |
| 74 | done | |
| 75 | ||
| 76 | lemma Mutator: | |
| 77 |  "\<turnstile> Mutator .{False}."
 | |
| 78 | apply(unfold Mutator_def) | |
| 79 | apply annhoare | |
| 80 | apply(simp_all add:Redirect_Edge Color_Target) | |
| 81 | apply(simp add:mutator_defs Redirect_Edge_def) | |
| 82 | done | |
| 83 | ||
| 84 | subsection {* The Collector *}
 | |
| 85 | ||
| 86 | text {* \noindent A constant @{text "M_init"} is used to give @{text "\<acute>Ma"} a
 | |
| 87 | suitable first value, defined as a list of nodes where only the @{text
 | |
| 88 | "Roots"} are black. *} | |
| 89 | ||
| 90 | consts M_init :: nodes | |
| 91 | ||
| 92 | constdefs | |
| 93 | Proper_M_init :: "gar_coll_state \<Rightarrow> bool" | |
| 94 | "Proper_M_init \<equiv> \<guillemotleft> Blacks M_init=Roots \<and> length M_init=length \<acute>M \<guillemotright>" | |
| 95 | ||
| 96 | Proper :: "gar_coll_state \<Rightarrow> bool" | |
| 97 | "Proper \<equiv> \<guillemotleft> Proper_Roots \<acute>M \<and> Proper_Edges(\<acute>M, \<acute>E) \<and> \<acute>Proper_M_init \<guillemotright>" | |
| 98 | ||
| 99 | Safe :: "gar_coll_state \<Rightarrow> bool" | |
| 100 | "Safe \<equiv> \<guillemotleft> Reach \<acute>E \<subseteq> Blacks \<acute>M \<guillemotright>" | |
| 101 | ||
| 102 | lemmas collector_defs = Proper_M_init_def Proper_def Safe_def | |
| 103 | ||
| 104 | subsubsection {* Blackening the roots *}
 | |
| 105 | ||
| 106 | constdefs | |
| 107 | Blacken_Roots :: " gar_coll_state ann_com" | |
| 108 | "Blacken_Roots \<equiv> | |
| 109 |   .{\<acute>Proper}.
 | |
| 110 | \<acute>ind:=0;; | |
| 111 |   .{\<acute>Proper \<and> \<acute>ind=0}.
 | |
| 112 | WHILE \<acute>ind<length \<acute>M | |
| 113 |    INV .{\<acute>Proper \<and> (\<forall>i<\<acute>ind. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind\<le>length \<acute>M}.
 | |
| 114 |   DO .{\<acute>Proper \<and> (\<forall>i<\<acute>ind. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M}.
 | |
| 115 | IF \<acute>ind\<in>Roots THEN | |
| 116 |    .{\<acute>Proper \<and> (\<forall>i<\<acute>ind. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M \<and> \<acute>ind\<in>Roots}. 
 | |
| 117 | \<acute>M:=\<acute>M[\<acute>ind:=Black] FI;; | |
| 118 |    .{\<acute>Proper \<and> (\<forall>i<\<acute>ind+1. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M}.
 | |
| 119 | \<acute>ind:=\<acute>ind+1 | |
| 120 | OD" | |
| 121 | ||
| 122 | lemma Blacken_Roots: | |
| 123 |  "\<turnstile> Blacken_Roots .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M}."
 | |
| 124 | apply (unfold Blacken_Roots_def) | |
| 125 | apply annhoare | |
| 126 | apply(simp_all add:collector_defs Graph_defs) | |
| 127 | apply safe | |
| 128 | apply(simp_all add:nth_list_update) | |
| 129 | apply (erule less_SucE) | |
| 130 | apply simp+ | |
| 131 | apply force | |
| 132 | apply force | |
| 133 | done | |
| 134 | ||
| 135 | subsubsection {* Propagating black *}
 | |
| 136 | ||
| 137 | constdefs | |
| 138 | PBInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" | |
| 139 | "PBInv \<equiv> \<guillemotleft> \<lambda>ind. \<acute>obc < Blacks \<acute>M \<or> (\<forall>i <ind. \<not>BtoW (\<acute>E!i, \<acute>M) \<or> | |
| 140 | (\<not>\<acute>z \<and> i=R \<and> (snd(\<acute>E!R)) = T \<and> (\<exists>r. ind \<le> r \<and> r < length \<acute>E \<and> BtoW(\<acute>E!r,\<acute>M))))\<guillemotright>" | |
| 141 | ||
| 142 | constdefs | |
| 143 | Propagate_Black_aux :: "gar_coll_state ann_com" | |
| 144 | "Propagate_Black_aux \<equiv> | |
| 145 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M}.
 | |
| 146 | \<acute>ind:=0;; | |
| 147 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M \<and> \<acute>ind=0}. 
 | |
| 148 | WHILE \<acute>ind<length \<acute>E | |
| 149 |    INV .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
 | |
| 150 | \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind\<le>length \<acute>E}. | |
| 151 |   DO .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
 | |
| 152 | \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E}. | |
| 153 | IF \<acute>M!(fst (\<acute>E!\<acute>ind)) = Black THEN | |
| 154 |     .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
 | |
| 155 | \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E \<and> \<acute>M!fst(\<acute>E!\<acute>ind)=Black}. | |
| 156 | \<acute>M:=\<acute>M[snd(\<acute>E!\<acute>ind):=Black];; | |
| 157 |     .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
 | |
| 158 | \<and> \<acute>PBInv (\<acute>ind + 1) \<and> \<acute>ind<length \<acute>E}. | |
| 159 | \<acute>ind:=\<acute>ind+1 | |
| 160 | FI | |
| 161 | OD" | |
| 162 | ||
| 163 | lemma Propagate_Black_aux: | |
| 164 | "\<turnstile> Propagate_Black_aux | |
| 165 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
 | |
| 166 | \<and> ( \<acute>obc < Blacks \<acute>M \<or> \<acute>Safe)}." | |
| 167 | apply (unfold Propagate_Black_aux_def PBInv_def collector_defs) | |
| 168 | apply annhoare | |
| 169 | apply(simp_all add:Graph6 Graph7 Graph8 Graph12) | |
| 170 | apply force | |
| 171 | apply force | |
| 172 | apply force | |
| 173 | --{* 4 subgoals left *}
 | |
| 174 | apply clarify | |
| 175 | apply(simp add:Proper_Edges_def Proper_Roots_def Graph6 Graph7 Graph8 Graph12) | |
| 176 | apply (erule disjE) | |
| 177 | apply(rule disjI1) | |
| 178 | apply(erule Graph13) | |
| 179 | apply force | |
| 180 | apply (case_tac "M x ! snd (E x ! ind x)=Black") | |
| 181 | apply (simp add: Graph10 BtoW_def) | |
| 182 | apply (rule disjI2) | |
| 183 | apply clarify | |
| 184 | apply (erule less_SucE) | |
| 185 | apply (erule_tac x=i in allE , erule (1) notE impE) | |
| 186 | apply simp | |
| 187 | apply clarify | |
| 26316 
9e9e67e33557
removed redundant less_trans, less_linear, le_imp_less_or_eq, le_less_trans, less_le_trans (cf. Orderings.thy);
 wenzelm parents: 
24742diff
changeset | 188 | apply (drule_tac y = r in le_imp_less_or_eq) | 
| 13020 | 189 | apply (erule disjE) | 
| 190 | apply (subgoal_tac "Suc (ind x)\<le>r") | |
| 191 | apply fast | |
| 192 | apply arith | |
| 193 | apply fast | |
| 194 | apply fast | |
| 195 | apply(rule disjI1) | |
| 196 | apply(erule subset_psubset_trans) | |
| 197 | apply(erule Graph11) | |
| 198 | apply fast | |
| 199 | --{* 3 subgoals left *}
 | |
| 200 | apply force | |
| 201 | apply force | |
| 202 | --{* last *}
 | |
| 203 | apply clarify | |
| 204 | apply simp | |
| 205 | apply(subgoal_tac "ind x = length (E x)") | |
| 206 | apply (rotate_tac -1) | |
| 13601 | 207 | apply (simp (asm_lr)) | 
| 13020 | 208 | apply(drule Graph1) | 
| 209 | apply simp | |
| 210 | apply clarify | |
| 211 | apply(erule allE, erule impE, assumption) | |
| 212 | apply force | |
| 213 | apply force | |
| 214 | apply arith | |
| 215 | done | |
| 216 | ||
| 217 | subsubsection {* Refining propagating black *}
 | |
| 218 | ||
| 219 | constdefs | |
| 220 | Auxk :: "gar_coll_state \<Rightarrow> bool" | |
| 221 | "Auxk \<equiv> \<guillemotleft>\<acute>k<length \<acute>M \<and> (\<acute>M!\<acute>k\<noteq>Black \<or> \<not>BtoW(\<acute>E!\<acute>ind, \<acute>M) \<or> | |
| 222 | \<acute>obc<Blacks \<acute>M \<or> (\<not>\<acute>z \<and> \<acute>ind=R \<and> snd(\<acute>E!R)=T | |
| 223 | \<and> (\<exists>r. \<acute>ind<r \<and> r<length \<acute>E \<and> BtoW(\<acute>E!r, \<acute>M))))\<guillemotright>" | |
| 224 | ||
| 225 | constdefs | |
| 226 | Propagate_Black :: " gar_coll_state ann_com" | |
| 227 | "Propagate_Black \<equiv> | |
| 228 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M}.
 | |
| 229 | \<acute>ind:=0;; | |
| 230 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M \<and> \<acute>ind=0}.
 | |
| 231 | WHILE \<acute>ind<length \<acute>E | |
| 232 |    INV .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
 | |
| 233 | \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind\<le>length \<acute>E}. | |
| 234 |   DO .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
 | |
| 235 | \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E}. | |
| 236 | IF (\<acute>M!(fst (\<acute>E!\<acute>ind)))=Black THEN | |
| 237 |     .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
 | |
| 238 | \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E \<and> (\<acute>M!fst(\<acute>E!\<acute>ind))=Black}. | |
| 239 | \<acute>k:=(snd(\<acute>E!\<acute>ind));; | |
| 240 |     .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
 | |
| 241 | \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E \<and> (\<acute>M!fst(\<acute>E!\<acute>ind))=Black | |
| 242 | \<and> \<acute>Auxk}. | |
| 243 | \<langle>\<acute>M:=\<acute>M[\<acute>k:=Black],, \<acute>ind:=\<acute>ind+1\<rangle> | |
| 244 |    ELSE .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
 | |
| 245 | \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E}. | |
| 246 | \<langle>IF (\<acute>M!(fst (\<acute>E!\<acute>ind)))\<noteq>Black THEN \<acute>ind:=\<acute>ind+1 FI\<rangle> | |
| 247 | FI | |
| 248 | OD" | |
| 249 | ||
| 250 | lemma Propagate_Black: | |
| 251 | "\<turnstile> Propagate_Black | |
| 252 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
 | |
| 253 | \<and> ( \<acute>obc < Blacks \<acute>M \<or> \<acute>Safe)}." | |
| 254 | apply (unfold Propagate_Black_def PBInv_def Auxk_def collector_defs) | |
| 255 | apply annhoare | |
| 32687 
27530efec97a
tuned proofs; be more cautios wrt. default simp rules
 haftmann parents: 
32621diff
changeset | 256 | apply(simp_all add: Graph6 Graph7 Graph8 Graph12) | 
| 13020 | 257 | apply force | 
| 258 | apply force | |
| 259 | apply force | |
| 260 | --{* 5 subgoals left *}
 | |
| 261 | apply clarify | |
| 262 | apply(simp add:BtoW_def Proper_Edges_def) | |
| 263 | --{* 4 subgoals left *}
 | |
| 264 | apply clarify | |
| 265 | apply(simp add:Proper_Edges_def Graph6 Graph7 Graph8 Graph12) | |
| 266 | apply (erule disjE) | |
| 267 | apply (rule disjI1) | |
| 268 | apply (erule psubset_subset_trans) | |
| 269 | apply (erule Graph9) | |
| 270 | apply (case_tac "M x!k x=Black") | |
| 271 | apply (case_tac "M x ! snd (E x ! ind x)=Black") | |
| 272 | apply (simp add: Graph10 BtoW_def) | |
| 273 | apply (rule disjI2) | |
| 274 | apply clarify | |
| 275 | apply (erule less_SucE) | |
| 276 | apply (erule_tac x=i in allE , erule (1) notE impE) | |
| 277 | apply simp | |
| 278 | apply clarify | |
| 26316 
9e9e67e33557
removed redundant less_trans, less_linear, le_imp_less_or_eq, le_less_trans, less_le_trans (cf. Orderings.thy);
 wenzelm parents: 
24742diff
changeset | 279 | apply (drule_tac y = r in le_imp_less_or_eq) | 
| 13020 | 280 | apply (erule disjE) | 
| 281 | apply (subgoal_tac "Suc (ind x)\<le>r") | |
| 282 | apply fast | |
| 283 | apply arith | |
| 284 | apply fast | |
| 285 | apply fast | |
| 286 | apply (simp add: Graph10 BtoW_def) | |
| 287 | apply (erule disjE) | |
| 288 | apply (erule disjI1) | |
| 289 | apply clarify | |
| 290 | apply (erule less_SucE) | |
| 291 | apply force | |
| 292 | apply simp | |
| 293 | apply (subgoal_tac "Suc R\<le>r") | |
| 294 | apply fast | |
| 295 | apply arith | |
| 296 | apply(rule disjI1) | |
| 297 | apply(erule subset_psubset_trans) | |
| 298 | apply(erule Graph11) | |
| 299 | apply fast | |
| 300 | --{* 2 subgoals left *}
 | |
| 301 | apply clarify | |
| 302 | apply(simp add:Proper_Edges_def Graph6 Graph7 Graph8 Graph12) | |
| 303 | apply (erule disjE) | |
| 304 | apply fast | |
| 305 | apply clarify | |
| 306 | apply (erule less_SucE) | |
| 307 | apply (erule_tac x=i in allE , erule (1) notE impE) | |
| 308 | apply simp | |
| 309 | apply clarify | |
| 26316 
9e9e67e33557
removed redundant less_trans, less_linear, le_imp_less_or_eq, le_less_trans, less_le_trans (cf. Orderings.thy);
 wenzelm parents: 
24742diff
changeset | 310 | apply (drule_tac y = r in le_imp_less_or_eq) | 
| 13020 | 311 | apply (erule disjE) | 
| 312 | apply (subgoal_tac "Suc (ind x)\<le>r") | |
| 313 | apply fast | |
| 314 | apply arith | |
| 315 | apply (simp add: BtoW_def) | |
| 316 | apply (simp add: BtoW_def) | |
| 317 | --{* last *}
 | |
| 318 | apply clarify | |
| 319 | apply simp | |
| 320 | apply(subgoal_tac "ind x = length (E x)") | |
| 321 | apply (rotate_tac -1) | |
| 13601 | 322 | apply (simp (asm_lr)) | 
| 13020 | 323 | apply(drule Graph1) | 
| 324 | apply simp | |
| 325 | apply clarify | |
| 326 | apply(erule allE, erule impE, assumption) | |
| 327 | apply force | |
| 328 | apply force | |
| 329 | apply arith | |
| 330 | done | |
| 331 | ||
| 332 | subsubsection {* Counting black nodes *}
 | |
| 333 | ||
| 334 | constdefs | |
| 335 | CountInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" | |
| 336 |   "CountInv \<equiv> \<guillemotleft> \<lambda>ind. {i. i<ind \<and> \<acute>Ma!i=Black}\<subseteq>\<acute>bc \<guillemotright>"
 | |
| 337 | ||
| 338 | constdefs | |
| 339 | Count :: " gar_coll_state ann_com" | |
| 340 | "Count \<equiv> | |
| 341 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
 | |
| 342 | \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M | |
| 343 |     \<and> length \<acute>Ma=length \<acute>M \<and> (\<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>bc={}}.
 | |
| 344 | \<acute>ind:=0;; | |
| 345 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
 | |
| 346 | \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M | |
| 347 |    \<and> length \<acute>Ma=length \<acute>M \<and> (\<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>bc={} 
 | |
| 348 | \<and> \<acute>ind=0}. | |
| 349 | WHILE \<acute>ind<length \<acute>M | |
| 350 |      INV .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
 | |
| 351 | \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M | |
| 352 | \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>CountInv \<acute>ind | |
| 353 | \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>ind\<le>length \<acute>M}. | |
| 354 |    DO .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
 | |
| 355 | \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M | |
| 356 | \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>CountInv \<acute>ind | |
| 357 | \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>ind<length \<acute>M}. | |
| 358 | IF \<acute>M!\<acute>ind=Black | |
| 359 |           THEN .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
 | |
| 360 | \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M | |
| 361 | \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>CountInv \<acute>ind | |
| 362 | \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>ind<length \<acute>M \<and> \<acute>M!\<acute>ind=Black}. | |
| 363 | \<acute>bc:=insert \<acute>ind \<acute>bc | |
| 364 | FI;; | |
| 365 |       .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
 | |
| 366 | \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M | |
| 367 | \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>CountInv (\<acute>ind+1) | |
| 368 | \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>ind<length \<acute>M}. | |
| 369 | \<acute>ind:=\<acute>ind+1 | |
| 370 | OD" | |
| 371 | ||
| 372 | lemma Count: | |
| 373 | "\<turnstile> Count | |
| 374 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
 | |
| 375 | \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>\<acute>bc \<and> \<acute>bc\<subseteq>Blacks \<acute>M \<and> length \<acute>Ma=length \<acute>M | |
| 376 | \<and> (\<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe)}." | |
| 377 | apply(unfold Count_def) | |
| 378 | apply annhoare | |
| 379 | apply(simp_all add:CountInv_def Graph6 Graph7 Graph8 Graph12 Blacks_def collector_defs) | |
| 380 | apply force | |
| 381 | apply force | |
| 382 | apply force | |
| 383 | apply clarify | |
| 384 | apply simp | |
| 385 | apply(fast elim:less_SucE) | |
| 386 | apply clarify | |
| 387 | apply simp | |
| 388 | apply(fast elim:less_SucE) | |
| 389 | apply force | |
| 390 | apply force | |
| 391 | done | |
| 392 | ||
| 393 | subsubsection {* Appending garbage nodes to the free list *}
 | |
| 394 | ||
| 395 | consts Append_to_free :: "nat \<times> edges \<Rightarrow> edges" | |
| 396 | ||
| 397 | axioms | |
| 398 | Append_to_free0: "length (Append_to_free (i, e)) = length e" | |
| 399 | Append_to_free1: "Proper_Edges (m, e) | |
| 400 | \<Longrightarrow> Proper_Edges (m, Append_to_free(i, e))" | |
| 401 | Append_to_free2: "i \<notin> Reach e | |
| 402 | \<Longrightarrow> n \<in> Reach (Append_to_free(i, e)) = ( n = i \<or> n \<in> Reach e)" | |
| 403 | ||
| 404 | constdefs | |
| 405 | AppendInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" | |
| 406 | "AppendInv \<equiv> \<guillemotleft>\<lambda>ind. \<forall>i<length \<acute>M. ind\<le>i \<longrightarrow> i\<in>Reach \<acute>E \<longrightarrow> \<acute>M!i=Black\<guillemotright>" | |
| 407 | ||
| 408 | constdefs | |
| 409 | Append :: " gar_coll_state ann_com" | |
| 410 | "Append \<equiv> | |
| 411 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe}.
 | |
| 412 | \<acute>ind:=0;; | |
| 413 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe \<and> \<acute>ind=0}.
 | |
| 414 | WHILE \<acute>ind<length \<acute>M | |
| 415 |       INV .{\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind\<le>length \<acute>M}.
 | |
| 416 |     DO .{\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M}.
 | |
| 417 | IF \<acute>M!\<acute>ind=Black THEN | |
| 418 |           .{\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>M!\<acute>ind=Black}. 
 | |
| 419 | \<acute>M:=\<acute>M[\<acute>ind:=White] | |
| 420 |        ELSE .{\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>ind\<notin>Reach \<acute>E}.
 | |
| 421 | \<acute>E:=Append_to_free(\<acute>ind,\<acute>E) | |
| 422 | FI;; | |
| 423 |      .{\<acute>Proper \<and> \<acute>AppendInv (\<acute>ind+1) \<and> \<acute>ind<length \<acute>M}. 
 | |
| 424 | \<acute>ind:=\<acute>ind+1 | |
| 425 | OD" | |
| 426 | ||
| 427 | lemma Append: | |
| 428 |   "\<turnstile> Append .{\<acute>Proper}."
 | |
| 429 | apply(unfold Append_def AppendInv_def) | |
| 430 | apply annhoare | |
| 431 | apply(simp_all add:collector_defs Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) | |
| 432 | apply(force simp:Blacks_def nth_list_update) | |
| 433 | apply force | |
| 434 | apply force | |
| 435 | apply(force simp add:Graph_defs) | |
| 436 | apply force | |
| 437 | apply clarify | |
| 438 | apply simp | |
| 439 | apply(rule conjI) | |
| 440 | apply (erule Append_to_free1) | |
| 441 | apply clarify | |
| 442 | apply (drule_tac n = "i" in Append_to_free2) | |
| 443 | apply force | |
| 444 | apply force | |
| 445 | apply force | |
| 446 | done | |
| 447 | ||
| 448 | subsubsection {* Correctness of the Collector *}
 | |
| 449 | ||
| 450 | constdefs | |
| 451 | Collector :: " gar_coll_state ann_com" | |
| 452 | "Collector \<equiv> | |
| 453 | .{\<acute>Proper}.  
 | |
| 454 |  WHILE True INV .{\<acute>Proper}. 
 | |
| 455 | DO | |
| 456 | Blacken_Roots;; | |
| 457 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M}.  
 | |
| 458 |    \<acute>obc:={};; 
 | |
| 459 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={}}. 
 | |
| 460 | \<acute>bc:=Roots;; | |
| 461 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots}. 
 | |
| 462 | \<acute>Ma:=M_init;; | |
| 463 |   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots \<and> \<acute>Ma=M_init}. 
 | |
| 464 | WHILE \<acute>obc\<noteq>\<acute>bc | |
| 465 |      INV .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
 | |
| 466 | \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>\<acute>bc \<and> \<acute>bc\<subseteq>Blacks \<acute>M | |
| 467 | \<and> length \<acute>Ma=length \<acute>M \<and> (\<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe)}. | |
| 468 |    DO .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M}.
 | |
| 469 | \<acute>obc:=\<acute>bc;; | |
| 470 | Propagate_Black;; | |
| 471 |       .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
 | |
| 472 | \<and> (\<acute>obc < Blacks \<acute>M \<or> \<acute>Safe)}. | |
| 473 | \<acute>Ma:=\<acute>M;; | |
| 474 |       .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma 
 | |
| 475 | \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M \<and> length \<acute>Ma=length \<acute>M | |
| 476 | \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe)}. | |
| 477 |        \<acute>bc:={};;
 | |
| 478 | Count | |
| 479 | OD;; | |
| 480 | Append | |
| 481 | OD" | |
| 482 | ||
| 483 | lemma Collector: | |
| 484 |   "\<turnstile> Collector .{False}."
 | |
| 485 | apply(unfold Collector_def) | |
| 486 | apply annhoare | |
| 487 | apply(simp_all add: Blacken_Roots Propagate_Black Count Append) | |
| 488 | apply(simp_all add:Blacken_Roots_def Propagate_Black_def Count_def Append_def collector_defs) | |
| 489 | apply (force simp add: Proper_Roots_def) | |
| 490 | apply force | |
| 491 | apply force | |
| 492 | apply clarify | |
| 493 | apply (erule disjE) | |
| 494 | apply(simp add:psubsetI) | |
| 495 | apply(force dest:subset_antisym) | |
| 496 | done | |
| 497 | ||
| 498 | subsection {* Interference Freedom *}
 | |
| 499 | ||
| 500 | lemmas modules = Redirect_Edge_def Color_Target_def Blacken_Roots_def | |
| 501 | Propagate_Black_def Count_def Append_def | |
| 502 | lemmas Invariants = PBInv_def Auxk_def CountInv_def AppendInv_def | |
| 503 | lemmas abbrev = collector_defs mutator_defs Invariants | |
| 504 | ||
| 505 | lemma interfree_Blacken_Roots_Redirect_Edge: | |
| 506 |  "interfree_aux (Some Blacken_Roots, {}, Some Redirect_Edge)"
 | |
| 507 | apply (unfold modules) | |
| 508 | apply interfree_aux | |
| 509 | apply safe | |
| 510 | apply (simp_all add:Graph6 Graph12 abbrev) | |
| 511 | done | |
| 512 | ||
| 513 | lemma interfree_Redirect_Edge_Blacken_Roots: | |
| 514 |   "interfree_aux (Some Redirect_Edge, {}, Some Blacken_Roots)"
 | |
| 515 | apply (unfold modules) | |
| 516 | apply interfree_aux | |
| 517 | apply safe | |
| 518 | apply(simp add:abbrev)+ | |
| 519 | done | |
| 520 | ||
| 521 | lemma interfree_Blacken_Roots_Color_Target: | |
| 522 |   "interfree_aux (Some Blacken_Roots, {}, Some Color_Target)"
 | |
| 523 | apply (unfold modules) | |
| 524 | apply interfree_aux | |
| 525 | apply safe | |
| 526 | apply(simp_all add:Graph7 Graph8 nth_list_update abbrev) | |
| 527 | done | |
| 528 | ||
| 529 | lemma interfree_Color_Target_Blacken_Roots: | |
| 530 |   "interfree_aux (Some Color_Target, {}, Some Blacken_Roots)"
 | |
| 531 | apply (unfold modules ) | |
| 532 | apply interfree_aux | |
| 533 | apply safe | |
| 534 | apply(simp add:abbrev)+ | |
| 535 | done | |
| 536 | ||
| 537 | lemma interfree_Propagate_Black_Redirect_Edge: | |
| 538 |   "interfree_aux (Some Propagate_Black, {}, Some Redirect_Edge)"
 | |
| 539 | apply (unfold modules ) | |
| 540 | apply interfree_aux | |
| 541 | --{* 11 subgoals left *}
 | |
| 542 | apply(clarify, simp add:abbrev Graph6 Graph12) | |
| 543 | apply(clarify, simp add:abbrev Graph6 Graph12) | |
| 544 | apply(clarify, simp add:abbrev Graph6 Graph12) | |
| 545 | apply(clarify, simp add:abbrev Graph6 Graph12) | |
| 546 | apply(erule conjE)+ | |
| 547 | apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) | |
| 548 | apply(erule Graph4) | |
| 549 | apply(simp)+ | |
| 550 | apply (simp add:BtoW_def) | |
| 551 | apply (simp add:BtoW_def) | |
| 552 | apply(rule conjI) | |
| 553 | apply (force simp add:BtoW_def) | |
| 554 | apply(erule Graph4) | |
| 555 | apply simp+ | |
| 556 | --{* 7 subgoals left *}
 | |
| 557 | apply(clarify, simp add:abbrev Graph6 Graph12) | |
| 558 | apply(erule conjE)+ | |
| 559 | apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) | |
| 560 | apply(erule Graph4) | |
| 561 | apply(simp)+ | |
| 562 | apply (simp add:BtoW_def) | |
| 563 | apply (simp add:BtoW_def) | |
| 564 | apply(rule conjI) | |
| 565 | apply (force simp add:BtoW_def) | |
| 566 | apply(erule Graph4) | |
| 567 | apply simp+ | |
| 568 | --{* 6 subgoals left *}
 | |
| 569 | apply(clarify, simp add:abbrev Graph6 Graph12) | |
| 570 | apply(erule conjE)+ | |
| 571 | apply(rule conjI) | |
| 572 | apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) | |
| 573 | apply(erule Graph4) | |
| 574 | apply(simp)+ | |
| 575 | apply (simp add:BtoW_def) | |
| 576 | apply (simp add:BtoW_def) | |
| 577 | apply(rule conjI) | |
| 578 | apply (force simp add:BtoW_def) | |
| 579 | apply(erule Graph4) | |
| 580 | apply simp+ | |
| 581 | apply(simp add:BtoW_def nth_list_update) | |
| 582 | apply force | |
| 583 | --{* 5 subgoals left *}
 | |
| 584 | apply(clarify, simp add:abbrev Graph6 Graph12) | |
| 585 | --{* 4 subgoals left *}
 | |
| 586 | apply(clarify, simp add:abbrev Graph6 Graph12) | |
| 587 | apply(rule conjI) | |
| 588 | apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) | |
| 589 | apply(erule Graph4) | |
| 590 | apply(simp)+ | |
| 591 | apply (simp add:BtoW_def) | |
| 592 | apply (simp add:BtoW_def) | |
| 593 | apply(rule conjI) | |
| 594 | apply (force simp add:BtoW_def) | |
| 595 | apply(erule Graph4) | |
| 596 | apply simp+ | |
| 597 | apply(rule conjI) | |
| 598 | apply(simp add:nth_list_update) | |
| 599 | apply force | |
| 600 | apply(rule impI, rule impI, erule disjE, erule disjI1, case_tac "R = (ind x)" ,case_tac "M x ! T = Black") | |
| 601 | apply(force simp add:BtoW_def) | |
| 602 | apply(case_tac "M x !snd (E x ! ind x)=Black") | |
| 603 | apply(rule disjI2) | |
| 604 | apply simp | |
| 605 | apply (erule Graph5) | |
| 606 | apply simp+ | |
| 607 | apply(force simp add:BtoW_def) | |
| 608 | apply(force simp add:BtoW_def) | |
| 609 | --{* 3 subgoals left *}
 | |
| 610 | apply(clarify, simp add:abbrev Graph6 Graph12) | |
| 611 | --{* 2 subgoals left *}
 | |
| 612 | apply(clarify, simp add:abbrev Graph6 Graph12) | |
| 613 | apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) | |
| 614 | apply clarify | |
| 615 | apply(erule Graph4) | |
| 616 | apply(simp)+ | |
| 617 | apply (simp add:BtoW_def) | |
| 618 | apply (simp add:BtoW_def) | |
| 619 | apply(rule conjI) | |
| 620 | apply (force simp add:BtoW_def) | |
| 621 | apply(erule Graph4) | |
| 622 | apply simp+ | |
| 623 | done | |
| 624 | ||
| 625 | lemma interfree_Redirect_Edge_Propagate_Black: | |
| 626 |   "interfree_aux (Some Redirect_Edge, {}, Some Propagate_Black)"
 | |
| 627 | apply (unfold modules ) | |
| 628 | apply interfree_aux | |
| 629 | apply(clarify, simp add:abbrev)+ | |
| 630 | done | |
| 631 | ||
| 632 | lemma interfree_Propagate_Black_Color_Target: | |
| 633 |   "interfree_aux (Some Propagate_Black, {}, Some Color_Target)"
 | |
| 634 | apply (unfold modules ) | |
| 635 | apply interfree_aux | |
| 636 | --{* 11 subgoals left *}
 | |
| 637 | apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)+ | |
| 638 | apply(erule conjE)+ | |
| 639 | apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, | |
| 640 | case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, | |
| 641 | erule allE, erule impE, assumption, erule impE, assumption, | |
| 642 | simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) | |
| 643 | --{* 7 subgoals left *}
 | |
| 644 | apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) | |
| 645 | apply(erule conjE)+ | |
| 646 | apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, | |
| 647 | case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, | |
| 648 | erule allE, erule impE, assumption, erule impE, assumption, | |
| 649 | simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) | |
| 650 | --{* 6 subgoals left *}
 | |
| 651 | apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) | |
| 652 | apply clarify | |
| 653 | apply (rule conjI) | |
| 654 | apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, | |
| 655 | case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, | |
| 656 | erule allE, erule impE, assumption, erule impE, assumption, | |
| 657 | simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) | |
| 658 | apply(simp add:nth_list_update) | |
| 659 | --{* 5 subgoals left *}
 | |
| 660 | apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) | |
| 661 | --{* 4 subgoals left *}
 | |
| 662 | apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) | |
| 663 | apply (rule conjI) | |
| 664 | apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, | |
| 665 | case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, | |
| 666 | erule allE, erule impE, assumption, erule impE, assumption, | |
| 667 | simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) | |
| 668 | apply(rule conjI) | |
| 669 | apply(simp add:nth_list_update) | |
| 670 | apply(rule impI,rule impI, case_tac "M x!T=Black",rotate_tac -1, force simp add: BtoW_def Graph10, | |
| 671 | erule subset_psubset_trans, erule Graph11, force) | |
| 672 | --{* 3 subgoals left *}
 | |
| 673 | apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) | |
| 674 | --{* 2 subgoals left *}
 | |
| 675 | apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) | |
| 676 | apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, | |
| 677 | case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, | |
| 678 | erule allE, erule impE, assumption, erule impE, assumption, | |
| 679 | simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) | |
| 680 | --{* 3 subgoals left *}
 | |
| 681 | apply(simp add:abbrev) | |
| 682 | done | |
| 683 | ||
| 684 | lemma interfree_Color_Target_Propagate_Black: | |
| 685 |   "interfree_aux (Some Color_Target, {}, Some Propagate_Black)"
 | |
| 686 | apply (unfold modules ) | |
| 687 | apply interfree_aux | |
| 688 | apply(clarify, simp add:abbrev)+ | |
| 689 | done | |
| 690 | ||
| 691 | lemma interfree_Count_Redirect_Edge: | |
| 692 |   "interfree_aux (Some Count, {}, Some Redirect_Edge)"
 | |
| 693 | apply (unfold modules) | |
| 694 | apply interfree_aux | |
| 695 | --{* 9 subgoals left *}
 | |
| 696 | apply(simp_all add:abbrev Graph6 Graph12) | |
| 697 | --{* 6 subgoals left *}
 | |
| 698 | apply(clarify, simp add:abbrev Graph6 Graph12, | |
| 699 | erule disjE,erule disjI1,rule disjI2,rule subset_trans, erule Graph3,force,force)+ | |
| 700 | done | |
| 701 | ||
| 702 | lemma interfree_Redirect_Edge_Count: | |
| 703 |   "interfree_aux (Some Redirect_Edge, {}, Some Count)"
 | |
| 704 | apply (unfold modules ) | |
| 705 | apply interfree_aux | |
| 706 | apply(clarify,simp add:abbrev)+ | |
| 707 | apply(simp add:abbrev) | |
| 708 | done | |
| 709 | ||
| 710 | lemma interfree_Count_Color_Target: | |
| 711 |   "interfree_aux (Some Count, {}, Some Color_Target)"
 | |
| 712 | apply (unfold modules ) | |
| 713 | apply interfree_aux | |
| 714 | --{* 9 subgoals left *}
 | |
| 715 | apply(simp_all add:abbrev Graph7 Graph8 Graph12) | |
| 716 | --{* 6 subgoals left *}
 | |
| 717 | apply(clarify,simp add:abbrev Graph7 Graph8 Graph12, | |
| 718 | erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9)+ | |
| 719 | --{* 2 subgoals left *}
 | |
| 720 | apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) | |
| 721 | apply(rule conjI) | |
| 722 | apply(erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9) | |
| 723 | apply(simp add:nth_list_update) | |
| 13022 
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
 prensani parents: 
13020diff
changeset | 724 | --{* 1 subgoal left *}
 | 
| 13020 | 725 | apply(clarify, simp add:abbrev Graph7 Graph8 Graph12, | 
| 726 | erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9) | |
| 727 | done | |
| 728 | ||
| 729 | lemma interfree_Color_Target_Count: | |
| 730 |   "interfree_aux (Some Color_Target, {}, Some Count)"
 | |
| 731 | apply (unfold modules ) | |
| 732 | apply interfree_aux | |
| 733 | apply(clarify, simp add:abbrev)+ | |
| 734 | apply(simp add:abbrev) | |
| 735 | done | |
| 736 | ||
| 737 | lemma interfree_Append_Redirect_Edge: | |
| 738 |   "interfree_aux (Some Append, {}, Some Redirect_Edge)"
 | |
| 739 | apply (unfold modules ) | |
| 740 | apply interfree_aux | |
| 741 | apply( simp_all add:abbrev Graph6 Append_to_free0 Append_to_free1 Graph12) | |
| 742 | apply(clarify, simp add:abbrev Graph6 Append_to_free0 Append_to_free1 Graph12, force dest:Graph3)+ | |
| 743 | done | |
| 744 | ||
| 745 | lemma interfree_Redirect_Edge_Append: | |
| 746 |   "interfree_aux (Some Redirect_Edge, {}, Some Append)"
 | |
| 747 | apply (unfold modules ) | |
| 748 | apply interfree_aux | |
| 749 | apply(clarify, simp add:abbrev Append_to_free0)+ | |
| 750 | apply (force simp add: Append_to_free2) | |
| 751 | apply(clarify, simp add:abbrev Append_to_free0)+ | |
| 752 | done | |
| 753 | ||
| 754 | lemma interfree_Append_Color_Target: | |
| 755 |   "interfree_aux (Some Append, {}, Some Color_Target)"
 | |
| 756 | apply (unfold modules ) | |
| 757 | apply interfree_aux | |
| 758 | apply(clarify, simp add:abbrev Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12 nth_list_update)+ | |
| 759 | apply(simp add:abbrev Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12 nth_list_update) | |
| 760 | done | |
| 761 | ||
| 762 | lemma interfree_Color_Target_Append: | |
| 763 |   "interfree_aux (Some Color_Target, {}, Some Append)"
 | |
| 764 | apply (unfold modules ) | |
| 765 | apply interfree_aux | |
| 766 | apply(clarify, simp add:abbrev Append_to_free0)+ | |
| 767 | apply (force simp add: Append_to_free2) | |
| 768 | apply(clarify,simp add:abbrev Append_to_free0)+ | |
| 769 | done | |
| 770 | ||
| 771 | lemmas collector_mutator_interfree = | |
| 772 | interfree_Blacken_Roots_Redirect_Edge interfree_Blacken_Roots_Color_Target | |
| 773 | interfree_Propagate_Black_Redirect_Edge interfree_Propagate_Black_Color_Target | |
| 774 | interfree_Count_Redirect_Edge interfree_Count_Color_Target | |
| 775 | interfree_Append_Redirect_Edge interfree_Append_Color_Target | |
| 776 | interfree_Redirect_Edge_Blacken_Roots interfree_Color_Target_Blacken_Roots | |
| 777 | interfree_Redirect_Edge_Propagate_Black interfree_Color_Target_Propagate_Black | |
| 778 | interfree_Redirect_Edge_Count interfree_Color_Target_Count | |
| 779 | interfree_Redirect_Edge_Append interfree_Color_Target_Append | |
| 780 | ||
| 781 | subsubsection {* Interference freedom Collector-Mutator *}
 | |
| 782 | ||
| 783 | lemma interfree_Collector_Mutator: | |
| 784 |  "interfree_aux (Some Collector, {}, Some Mutator)"
 | |
| 785 | apply(unfold Collector_def Mutator_def) | |
| 786 | apply interfree_aux | |
| 787 | apply(simp_all add:collector_mutator_interfree) | |
| 788 | apply(unfold modules collector_defs mutator_defs) | |
| 789 | apply(tactic  {* TRYALL (interfree_aux_tac) *})
 | |
| 790 | --{* 32 subgoals left *}
 | |
| 791 | apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) | |
| 792 | --{* 20 subgoals left *}
 | |
| 23894 
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
 wenzelm parents: 
21669diff
changeset | 793 | apply(tactic{* TRYALL (clarify_tac @{claset}) *})
 | 
| 13020 | 794 | apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) | 
| 795 | apply(tactic {* TRYALL (etac disjE) *})
 | |
| 796 | apply simp_all | |
| 26342 | 797 | apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac subset_trans,etac @{thm Graph3},force_tac @{clasimpset}, assume_tac]) *})
 | 
| 798 | apply(tactic {* TRYALL(EVERY'[rtac disjI2,etac subset_trans,rtac @{thm Graph9},force_tac @{clasimpset}]) *})
 | |
| 799 | apply(tactic {* TRYALL(EVERY'[rtac disjI1,etac @{thm psubset_subset_trans},rtac @{thm Graph9},force_tac @{clasimpset}]) *})
 | |
| 13020 | 800 | done | 
| 801 | ||
| 802 | subsubsection {* Interference freedom Mutator-Collector *}
 | |
| 803 | ||
| 804 | lemma interfree_Mutator_Collector: | |
| 805 |  "interfree_aux (Some Mutator, {}, Some Collector)"
 | |
| 806 | apply(unfold Collector_def Mutator_def) | |
| 807 | apply interfree_aux | |
| 808 | apply(simp_all add:collector_mutator_interfree) | |
| 809 | apply(unfold modules collector_defs mutator_defs) | |
| 810 | apply(tactic  {* TRYALL (interfree_aux_tac) *})
 | |
| 811 | --{* 64 subgoals left *}
 | |
| 812 | apply(simp_all add:nth_list_update Invariants Append_to_free0)+ | |
| 23894 
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
 wenzelm parents: 
21669diff
changeset | 813 | apply(tactic{* TRYALL (clarify_tac @{claset}) *})
 | 
| 13020 | 814 | --{* 4 subgoals left *}
 | 
| 815 | apply force | |
| 816 | apply(simp add:Append_to_free2) | |
| 817 | apply force | |
| 818 | apply(simp add:Append_to_free2) | |
| 819 | done | |
| 820 | ||
| 821 | subsubsection {* The Garbage Collection algorithm *}
 | |
| 822 | ||
| 823 | text {* In total there are 289 verification conditions.  *}
 | |
| 824 | ||
| 825 | lemma Gar_Coll: | |
| 826 |   "\<parallel>- .{\<acute>Proper \<and> \<acute>Mut_init \<and> \<acute>z}.  
 | |
| 827 | COBEGIN | |
| 828 | Collector | |
| 829 |   .{False}.
 | |
| 830 | \<parallel> | |
| 831 | Mutator | |
| 832 |   .{False}. 
 | |
| 833 | COEND | |
| 834 |   .{False}."
 | |
| 835 | apply oghoare | |
| 836 | apply(force simp add: Mutator_def Collector_def modules) | |
| 837 | apply(rule Collector) | |
| 838 | apply(rule Mutator) | |
| 839 | apply(simp add:interfree_Collector_Mutator) | |
| 840 | apply(simp add:interfree_Mutator_Collector) | |
| 841 | apply force | |
| 842 | done | |
| 843 | ||
| 844 | end |