| author | sultana | 
| Tue, 03 Sep 2013 21:46:40 +0100 | |
| changeset 53390 | 51b562922cb1 | 
| parent 45605 | a89b4bc311a5 | 
| child 58889 | 5b7a9633cfa8 | 
| permissions | -rw-r--r-- | 
| 37936 | 1 | (* Title: HOL/UNITY/WFair.thy | 
| 4776 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | Copyright 1998 University of Cambridge | |
| 4 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 5 | Conditional Fairness versions of transient, ensures, leadsTo. | 
| 4776 | 6 | |
| 7 | From Misra, "A Logic for Concurrent Programming", 1994 | |
| 8 | *) | |
| 9 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 10 | header{*Progress*}
 | 
| 13798 | 11 | |
| 16417 | 12 | theory WFair imports UNITY begin | 
| 4776 | 13 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 14 | text{*The original version of this theory was based on weak fairness.  (Thus,
 | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 15 | the entire UNITY development embodied this assumption, until February 2003.) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 16 | Weak fairness states that if a command is enabled continuously, then it is | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 17 | eventually executed. Ernie Cohen suggested that I instead adopt unconditional | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 18 | fairness: every command is executed infinitely often. | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 19 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 20 | In fact, Misra's paper on "Progress" seems to be ambiguous about the correct | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 21 | interpretation, and says that the two forms of fairness are equivalent. They | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 22 | differ only on their treatment of partial transitions, which under | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 23 | unconditional fairness behave magically. That is because if there are partial | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 24 | transitions then there may be no fair executions, making all leads-to | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 25 | properties hold vacuously. | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 26 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 27 | Unconditional fairness has some great advantages. By distinguishing partial | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 28 | transitions from total ones that are the identity on part of their domain, it | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 29 | is more expressive. Also, by simplifying the definition of the transient | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 30 | property, it simplifies many proofs. A drawback is that some laws only hold | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 31 | under the assumption that all transitions are total. The best-known of these | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 32 | is the impossibility law for leads-to. | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 33 | *} | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 34 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32693diff
changeset | 35 | definition | 
| 4776 | 36 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 37 |   --{*This definition specifies conditional fairness.  The rest of the theory
 | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 38 | is generic to all forms of fairness. To get weak fairness, conjoin | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 39 |       the inclusion below with @{term "A \<subseteq> Domain act"}, which specifies 
 | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 40 |       that the action is enabled over all of @{term A}.*}
 | 
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32693diff
changeset | 41 | transient :: "'a set => 'a program set" where | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 42 |     "transient A == {F. \<exists>act\<in>Acts F. act``A \<subseteq> -A}"
 | 
| 4776 | 43 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32693diff
changeset | 44 | definition | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32693diff
changeset | 45 | ensures :: "['a set, 'a set] => 'a program set" (infixl "ensures" 60) where | 
| 13805 | 46 | "A ensures B == (A-B co A \<union> B) \<inter> transient (A-B)" | 
| 8006 | 47 | |
| 6536 | 48 | |
| 23767 | 49 | inductive_set | 
| 6801 
9e0037839d63
renamed the underlying relation of leadsTo from "leadsto"
 paulson parents: 
6536diff
changeset | 50 |   leads :: "'a program => ('a set * 'a set) set"
 | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 51 |     --{*LEADS-TO constant for the inductive definition*}
 | 
| 23767 | 52 | for F :: "'a program" | 
| 53 | where | |
| 4776 | 54 | |
| 13805 | 55 | Basis: "F \<in> A ensures B ==> (A,B) \<in> leads F" | 
| 4776 | 56 | |
| 23767 | 57 | | Trans: "[| (A,B) \<in> leads F; (B,C) \<in> leads F |] ==> (A,C) \<in> leads F" | 
| 4776 | 58 | |
| 23767 | 59 | | Union: "\<forall>A \<in> S. (A,B) \<in> leads F ==> (Union S, B) \<in> leads F" | 
| 4776 | 60 | |
| 5155 | 61 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32693diff
changeset | 62 | definition leadsTo :: "['a set, 'a set] => 'a program set" (infixl "leadsTo" 60) where | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 63 |      --{*visible version of the LEADS-TO relation*}
 | 
| 13805 | 64 |     "A leadsTo B == {F. (A,B) \<in> leads F}"
 | 
| 5648 | 65 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32693diff
changeset | 66 | definition wlt :: "['a program, 'a set] => 'a set" where | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 67 |      --{*predicate transformer: the largest set that leads to @{term B}*}
 | 
| 13805 | 68 |     "wlt F B == Union {A. F \<in> A leadsTo B}"
 | 
| 4776 | 69 | |
| 35355 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
32693diff
changeset | 70 | notation (xsymbols) | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
32693diff
changeset | 71 | leadsTo (infixl "\<longmapsto>" 60) | 
| 13797 | 72 | |
| 73 | ||
| 13798 | 74 | subsection{*transient*}
 | 
| 13797 | 75 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 76 | lemma stable_transient: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 77 | "[| F \<in> stable A; F \<in> transient A |] ==> \<exists>act\<in>Acts F. A \<subseteq> - (Domain act)" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 78 | apply (simp add: stable_def constrains_def transient_def, clarify) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 79 | apply (rule rev_bexI, auto) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 80 | done | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 81 | |
| 13797 | 82 | lemma stable_transient_empty: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 83 |     "[| F \<in> stable A; F \<in> transient A; all_total F |] ==> A = {}"
 | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 84 | apply (drule stable_transient, assumption) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 85 | apply (simp add: all_total_def) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 86 | done | 
| 13797 | 87 | |
| 88 | lemma transient_strengthen: | |
| 13805 | 89 | "[| F \<in> transient A; B \<subseteq> A |] ==> F \<in> transient B" | 
| 13797 | 90 | apply (unfold transient_def, clarify) | 
| 91 | apply (blast intro!: rev_bexI) | |
| 92 | done | |
| 93 | ||
| 94 | lemma transientI: | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 95 | "[| act: Acts F; act``A \<subseteq> -A |] ==> F \<in> transient A" | 
| 13797 | 96 | by (unfold transient_def, blast) | 
| 97 | ||
| 98 | lemma transientE: | |
| 13805 | 99 | "[| F \<in> transient A; | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 100 | !!act. [| act: Acts F; act``A \<subseteq> -A |] ==> P |] | 
| 13797 | 101 | ==> P" | 
| 102 | by (unfold transient_def, blast) | |
| 103 | ||
| 104 | lemma transient_empty [simp]: "transient {} = UNIV"
 | |
| 105 | by (unfold transient_def, auto) | |
| 106 | ||
| 107 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 108 | text{*This equation recovers the notion of weak fairness.  A totalized
 | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 109 | program satisfies a transient assertion just if the original program | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 110 | contains a suitable action that is also enabled.*} | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 111 | lemma totalize_transient_iff: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 112 | "(totalize F \<in> transient A) = (\<exists>act\<in>Acts F. A \<subseteq> Domain act & act``A \<subseteq> -A)" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 113 | apply (simp add: totalize_def totalize_act_def transient_def | 
| 32693 | 114 | Un_Image, safe) | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 115 | apply (blast intro!: rev_bexI)+ | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 116 | done | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 117 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 118 | lemma totalize_transientI: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 119 | "[| act: Acts F; A \<subseteq> Domain act; act``A \<subseteq> -A |] | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 120 | ==> totalize F \<in> transient A" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 121 | by (simp add: totalize_transient_iff, blast) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 122 | |
| 13798 | 123 | subsection{*ensures*}
 | 
| 13797 | 124 | |
| 125 | lemma ensuresI: | |
| 13805 | 126 | "[| F \<in> (A-B) co (A \<union> B); F \<in> transient (A-B) |] ==> F \<in> A ensures B" | 
| 13797 | 127 | by (unfold ensures_def, blast) | 
| 128 | ||
| 129 | lemma ensuresD: | |
| 13805 | 130 | "F \<in> A ensures B ==> F \<in> (A-B) co (A \<union> B) & F \<in> transient (A-B)" | 
| 13797 | 131 | by (unfold ensures_def, blast) | 
| 132 | ||
| 133 | lemma ensures_weaken_R: | |
| 13805 | 134 | "[| F \<in> A ensures A'; A'<=B' |] ==> F \<in> A ensures B'" | 
| 13797 | 135 | apply (unfold ensures_def) | 
| 136 | apply (blast intro: constrains_weaken transient_strengthen) | |
| 137 | done | |
| 138 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 139 | text{*The L-version (precondition strengthening) fails, but we have this*}
 | 
| 13797 | 140 | lemma stable_ensures_Int: | 
| 13805 | 141 | "[| F \<in> stable C; F \<in> A ensures B |] | 
| 142 | ==> F \<in> (C \<inter> A) ensures (C \<inter> B)" | |
| 13797 | 143 | apply (unfold ensures_def) | 
| 144 | apply (auto simp add: ensures_def Int_Un_distrib [symmetric] Diff_Int_distrib [symmetric]) | |
| 145 | prefer 2 apply (blast intro: transient_strengthen) | |
| 146 | apply (blast intro: stable_constrains_Int constrains_weaken) | |
| 147 | done | |
| 148 | ||
| 149 | lemma stable_transient_ensures: | |
| 13805 | 150 | "[| F \<in> stable A; F \<in> transient C; A \<subseteq> B \<union> C |] ==> F \<in> A ensures B" | 
| 13797 | 151 | apply (simp add: ensures_def stable_def) | 
| 152 | apply (blast intro: constrains_weaken transient_strengthen) | |
| 153 | done | |
| 154 | ||
| 13805 | 155 | lemma ensures_eq: "(A ensures B) = (A unless B) \<inter> transient (A-B)" | 
| 13797 | 156 | by (simp (no_asm) add: ensures_def unless_def) | 
| 157 | ||
| 158 | ||
| 13798 | 159 | subsection{*leadsTo*}
 | 
| 13797 | 160 | |
| 13805 | 161 | lemma leadsTo_Basis [intro]: "F \<in> A ensures B ==> F \<in> A leadsTo B" | 
| 13797 | 162 | apply (unfold leadsTo_def) | 
| 163 | apply (blast intro: leads.Basis) | |
| 164 | done | |
| 165 | ||
| 166 | lemma leadsTo_Trans: | |
| 13805 | 167 | "[| F \<in> A leadsTo B; F \<in> B leadsTo C |] ==> F \<in> A leadsTo C" | 
| 13797 | 168 | apply (unfold leadsTo_def) | 
| 169 | apply (blast intro: leads.Trans) | |
| 170 | done | |
| 171 | ||
| 14112 | 172 | lemma leadsTo_Basis': | 
| 173 | "[| F \<in> A co A \<union> B; F \<in> transient A |] ==> F \<in> A leadsTo B" | |
| 174 | apply (drule_tac B = "A-B" in constrains_weaken_L) | |
| 175 | apply (drule_tac [2] B = "A-B" in transient_strengthen) | |
| 176 | apply (rule_tac [3] ensuresI [THEN leadsTo_Basis]) | |
| 177 | apply (blast+) | |
| 178 | done | |
| 179 | ||
| 13805 | 180 | lemma transient_imp_leadsTo: "F \<in> transient A ==> F \<in> A leadsTo (-A)" | 
| 13797 | 181 | by (simp (no_asm_simp) add: leadsTo_Basis ensuresI Compl_partition) | 
| 182 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 183 | text{*Useful with cancellation, disjunction*}
 | 
| 13805 | 184 | lemma leadsTo_Un_duplicate: "F \<in> A leadsTo (A' \<union> A') ==> F \<in> A leadsTo A'" | 
| 13797 | 185 | by (simp add: Un_ac) | 
| 186 | ||
| 187 | lemma leadsTo_Un_duplicate2: | |
| 13805 | 188 | "F \<in> A leadsTo (A' \<union> C \<union> C) ==> F \<in> A leadsTo (A' \<union> C)" | 
| 13797 | 189 | by (simp add: Un_ac) | 
| 190 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 191 | text{*The Union introduction rule as we should have liked to state it*}
 | 
| 13797 | 192 | lemma leadsTo_Union: | 
| 13805 | 193 | "(!!A. A \<in> S ==> F \<in> A leadsTo B) ==> F \<in> (Union S) leadsTo B" | 
| 13797 | 194 | apply (unfold leadsTo_def) | 
| 195 | apply (blast intro: leads.Union) | |
| 196 | done | |
| 197 | ||
| 198 | lemma leadsTo_Union_Int: | |
| 13805 | 199 | "(!!A. A \<in> S ==> F \<in> (A \<inter> C) leadsTo B) ==> F \<in> (Union S \<inter> C) leadsTo B" | 
| 13797 | 200 | apply (unfold leadsTo_def) | 
| 201 | apply (simp only: Int_Union_Union) | |
| 202 | apply (blast intro: leads.Union) | |
| 203 | done | |
| 204 | ||
| 205 | lemma leadsTo_UN: | |
| 13805 | 206 | "(!!i. i \<in> I ==> F \<in> (A i) leadsTo B) ==> F \<in> (\<Union>i \<in> I. A i) leadsTo B" | 
| 13797 | 207 | apply (subst Union_image_eq [symmetric]) | 
| 208 | apply (blast intro: leadsTo_Union) | |
| 209 | done | |
| 210 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 211 | text{*Binary union introduction rule*}
 | 
| 13797 | 212 | lemma leadsTo_Un: | 
| 13805 | 213 | "[| F \<in> A leadsTo C; F \<in> B leadsTo C |] ==> F \<in> (A \<union> B) leadsTo C" | 
| 44106 | 214 |   using leadsTo_Union [of "{A, B}" F C] by auto
 | 
| 13797 | 215 | |
| 216 | lemma single_leadsTo_I: | |
| 13805 | 217 |      "(!!x. x \<in> A ==> F \<in> {x} leadsTo B) ==> F \<in> A leadsTo B"
 | 
| 13797 | 218 | by (subst UN_singleton [symmetric], rule leadsTo_UN, blast) | 
| 219 | ||
| 220 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 221 | text{*The INDUCTION rule as we should have liked to state it*}
 | 
| 13797 | 222 | lemma leadsTo_induct: | 
| 13805 | 223 | "[| F \<in> za leadsTo zb; | 
| 224 | !!A B. F \<in> A ensures B ==> P A B; | |
| 225 | !!A B C. [| F \<in> A leadsTo B; P A B; F \<in> B leadsTo C; P B C |] | |
| 13797 | 226 | ==> P A C; | 
| 13805 | 227 | !!B S. \<forall>A \<in> S. F \<in> A leadsTo B & P A B ==> P (Union S) B | 
| 13797 | 228 | |] ==> P za zb" | 
| 229 | apply (unfold leadsTo_def) | |
| 230 | apply (drule CollectD, erule leads.induct) | |
| 231 | apply (blast+) | |
| 232 | done | |
| 233 | ||
| 234 | ||
| 13805 | 235 | lemma subset_imp_ensures: "A \<subseteq> B ==> F \<in> A ensures B" | 
| 13797 | 236 | by (unfold ensures_def constrains_def transient_def, blast) | 
| 237 | ||
| 45605 | 238 | lemmas subset_imp_leadsTo = subset_imp_ensures [THEN leadsTo_Basis] | 
| 13797 | 239 | |
| 45605 | 240 | lemmas leadsTo_refl = subset_refl [THEN subset_imp_leadsTo] | 
| 13797 | 241 | |
| 45605 | 242 | lemmas empty_leadsTo = empty_subsetI [THEN subset_imp_leadsTo, simp] | 
| 13797 | 243 | |
| 45605 | 244 | lemmas leadsTo_UNIV = subset_UNIV [THEN subset_imp_leadsTo, simp] | 
| 13797 | 245 | |
| 246 | ||
| 247 | ||
| 248 | (** Variant induction rule: on the preconditions for B **) | |
| 249 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 250 | text{*Lemma is the weak version: can't see how to do it in one step*}
 | 
| 13797 | 251 | lemma leadsTo_induct_pre_lemma: | 
| 13805 | 252 | "[| F \<in> za leadsTo zb; | 
| 13797 | 253 | P zb; | 
| 13805 | 254 | !!A B. [| F \<in> A ensures B; P B |] ==> P A; | 
| 255 | !!S. \<forall>A \<in> S. P A ==> P (Union S) | |
| 13797 | 256 | |] ==> P za" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 257 | txt{*by induction on this formula*}
 | 
| 13797 | 258 | apply (subgoal_tac "P zb --> P za") | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 259 | txt{*now solve first subgoal: this formula is sufficient*}
 | 
| 13797 | 260 | apply (blast intro: leadsTo_refl) | 
| 261 | apply (erule leadsTo_induct) | |
| 262 | apply (blast+) | |
| 263 | done | |
| 264 | ||
| 265 | lemma leadsTo_induct_pre: | |
| 13805 | 266 | "[| F \<in> za leadsTo zb; | 
| 13797 | 267 | P zb; | 
| 13805 | 268 | !!A B. [| F \<in> A ensures B; F \<in> B leadsTo zb; P B |] ==> P A; | 
| 269 | !!S. \<forall>A \<in> S. F \<in> A leadsTo zb & P A ==> P (Union S) | |
| 13797 | 270 | |] ==> P za" | 
| 13805 | 271 | apply (subgoal_tac "F \<in> za leadsTo zb & P za") | 
| 13797 | 272 | apply (erule conjunct2) | 
| 273 | apply (erule leadsTo_induct_pre_lemma) | |
| 274 | prefer 3 apply (blast intro: leadsTo_Union) | |
| 275 | prefer 2 apply (blast intro: leadsTo_Trans) | |
| 276 | apply (blast intro: leadsTo_refl) | |
| 277 | done | |
| 278 | ||
| 279 | ||
| 13805 | 280 | lemma leadsTo_weaken_R: "[| F \<in> A leadsTo A'; A'<=B' |] ==> F \<in> A leadsTo B'" | 
| 13797 | 281 | by (blast intro: subset_imp_leadsTo leadsTo_Trans) | 
| 282 | ||
| 45477 | 283 | lemma leadsTo_weaken_L: | 
| 13805 | 284 | "[| F \<in> A leadsTo A'; B \<subseteq> A |] ==> F \<in> B leadsTo A'" | 
| 13797 | 285 | by (blast intro: leadsTo_Trans subset_imp_leadsTo) | 
| 286 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 287 | text{*Distributes over binary unions*}
 | 
| 13797 | 288 | lemma leadsTo_Un_distrib: | 
| 13805 | 289 | "F \<in> (A \<union> B) leadsTo C = (F \<in> A leadsTo C & F \<in> B leadsTo C)" | 
| 13797 | 290 | by (blast intro: leadsTo_Un leadsTo_weaken_L) | 
| 291 | ||
| 292 | lemma leadsTo_UN_distrib: | |
| 13805 | 293 | "F \<in> (\<Union>i \<in> I. A i) leadsTo B = (\<forall>i \<in> I. F \<in> (A i) leadsTo B)" | 
| 13797 | 294 | by (blast intro: leadsTo_UN leadsTo_weaken_L) | 
| 295 | ||
| 296 | lemma leadsTo_Union_distrib: | |
| 13805 | 297 | "F \<in> (Union S) leadsTo B = (\<forall>A \<in> S. F \<in> A leadsTo B)" | 
| 13797 | 298 | by (blast intro: leadsTo_Union leadsTo_weaken_L) | 
| 299 | ||
| 300 | ||
| 301 | lemma leadsTo_weaken: | |
| 13805 | 302 | "[| F \<in> A leadsTo A'; B \<subseteq> A; A'<=B' |] ==> F \<in> B leadsTo B'" | 
| 13797 | 303 | by (blast intro: leadsTo_weaken_R leadsTo_weaken_L leadsTo_Trans) | 
| 304 | ||
| 305 | ||
| 14150 | 306 | text{*Set difference: maybe combine with @{text leadsTo_weaken_L}??*}
 | 
| 13797 | 307 | lemma leadsTo_Diff: | 
| 13805 | 308 | "[| F \<in> (A-B) leadsTo C; F \<in> B leadsTo C |] ==> F \<in> A leadsTo C" | 
| 13797 | 309 | by (blast intro: leadsTo_Un leadsTo_weaken) | 
| 310 | ||
| 311 | lemma leadsTo_UN_UN: | |
| 13805 | 312 | "(!! i. i \<in> I ==> F \<in> (A i) leadsTo (A' i)) | 
| 313 | ==> F \<in> (\<Union>i \<in> I. A i) leadsTo (\<Union>i \<in> I. A' i)" | |
| 13797 | 314 | apply (simp only: Union_image_eq [symmetric]) | 
| 315 | apply (blast intro: leadsTo_Union leadsTo_weaken_R) | |
| 316 | done | |
| 317 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 318 | text{*Binary union version*}
 | 
| 13797 | 319 | lemma leadsTo_Un_Un: | 
| 13805 | 320 | "[| F \<in> A leadsTo A'; F \<in> B leadsTo B' |] | 
| 321 | ==> F \<in> (A \<union> B) leadsTo (A' \<union> B')" | |
| 13797 | 322 | by (blast intro: leadsTo_Un leadsTo_weaken_R) | 
| 323 | ||
| 324 | ||
| 325 | (** The cancellation law **) | |
| 326 | ||
| 327 | lemma leadsTo_cancel2: | |
| 13805 | 328 | "[| F \<in> A leadsTo (A' \<union> B); F \<in> B leadsTo B' |] | 
| 329 | ==> F \<in> A leadsTo (A' \<union> B')" | |
| 13797 | 330 | by (blast intro: leadsTo_Un_Un subset_imp_leadsTo leadsTo_Trans) | 
| 331 | ||
| 332 | lemma leadsTo_cancel_Diff2: | |
| 13805 | 333 | "[| F \<in> A leadsTo (A' \<union> B); F \<in> (B-A') leadsTo B' |] | 
| 334 | ==> F \<in> A leadsTo (A' \<union> B')" | |
| 13797 | 335 | apply (rule leadsTo_cancel2) | 
| 336 | prefer 2 apply assumption | |
| 337 | apply (simp_all (no_asm_simp)) | |
| 338 | done | |
| 339 | ||
| 340 | lemma leadsTo_cancel1: | |
| 13805 | 341 | "[| F \<in> A leadsTo (B \<union> A'); F \<in> B leadsTo B' |] | 
| 342 | ==> F \<in> A leadsTo (B' \<union> A')" | |
| 13797 | 343 | apply (simp add: Un_commute) | 
| 344 | apply (blast intro!: leadsTo_cancel2) | |
| 345 | done | |
| 346 | ||
| 347 | lemma leadsTo_cancel_Diff1: | |
| 13805 | 348 | "[| F \<in> A leadsTo (B \<union> A'); F \<in> (B-A') leadsTo B' |] | 
| 349 | ==> F \<in> A leadsTo (B' \<union> A')" | |
| 13797 | 350 | apply (rule leadsTo_cancel1) | 
| 351 | prefer 2 apply assumption | |
| 352 | apply (simp_all (no_asm_simp)) | |
| 353 | done | |
| 354 | ||
| 355 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 356 | text{*The impossibility law*}
 | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 357 | lemma leadsTo_empty: "[|F \<in> A leadsTo {}; all_total F|] ==> A={}"
 | 
| 13797 | 358 | apply (erule leadsTo_induct_pre) | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 359 | apply (simp_all add: ensures_def constrains_def transient_def all_total_def, clarify) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 360 | apply (drule bspec, assumption)+ | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 361 | apply blast | 
| 13797 | 362 | done | 
| 363 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 364 | subsection{*PSP: Progress-Safety-Progress*}
 | 
| 13797 | 365 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 366 | text{*Special case of PSP: Misra's "stable conjunction"*}
 | 
| 13797 | 367 | lemma psp_stable: | 
| 13805 | 368 | "[| F \<in> A leadsTo A'; F \<in> stable B |] | 
| 369 | ==> F \<in> (A \<inter> B) leadsTo (A' \<inter> B)" | |
| 13797 | 370 | apply (unfold stable_def) | 
| 371 | apply (erule leadsTo_induct) | |
| 372 | prefer 3 apply (blast intro: leadsTo_Union_Int) | |
| 373 | prefer 2 apply (blast intro: leadsTo_Trans) | |
| 374 | apply (rule leadsTo_Basis) | |
| 375 | apply (simp add: ensures_def Diff_Int_distrib2 [symmetric] Int_Un_distrib2 [symmetric]) | |
| 376 | apply (blast intro: transient_strengthen constrains_Int) | |
| 377 | done | |
| 378 | ||
| 379 | lemma psp_stable2: | |
| 13805 | 380 | "[| F \<in> A leadsTo A'; F \<in> stable B |] ==> F \<in> (B \<inter> A) leadsTo (B \<inter> A')" | 
| 13797 | 381 | by (simp add: psp_stable Int_ac) | 
| 382 | ||
| 383 | lemma psp_ensures: | |
| 13805 | 384 | "[| F \<in> A ensures A'; F \<in> B co B' |] | 
| 385 | ==> F \<in> (A \<inter> B') ensures ((A' \<inter> B) \<union> (B' - B))" | |
| 13797 | 386 | apply (unfold ensures_def constrains_def, clarify) (*speeds up the proof*) | 
| 387 | apply (blast intro: transient_strengthen) | |
| 388 | done | |
| 389 | ||
| 390 | lemma psp: | |
| 13805 | 391 | "[| F \<in> A leadsTo A'; F \<in> B co B' |] | 
| 392 | ==> F \<in> (A \<inter> B') leadsTo ((A' \<inter> B) \<union> (B' - B))" | |
| 13797 | 393 | apply (erule leadsTo_induct) | 
| 394 | prefer 3 apply (blast intro: leadsTo_Union_Int) | |
| 395 |  txt{*Basis case*}
 | |
| 396 | apply (blast intro: psp_ensures) | |
| 397 | txt{*Transitivity case has a delicate argument involving "cancellation"*}
 | |
| 398 | apply (rule leadsTo_Un_duplicate2) | |
| 399 | apply (erule leadsTo_cancel_Diff1) | |
| 400 | apply (simp add: Int_Diff Diff_triv) | |
| 401 | apply (blast intro: leadsTo_weaken_L dest: constrains_imp_subset) | |
| 402 | done | |
| 403 | ||
| 404 | lemma psp2: | |
| 13805 | 405 | "[| F \<in> A leadsTo A'; F \<in> B co B' |] | 
| 406 | ==> F \<in> (B' \<inter> A) leadsTo ((B \<inter> A') \<union> (B' - B))" | |
| 13797 | 407 | by (simp (no_asm_simp) add: psp Int_ac) | 
| 408 | ||
| 409 | lemma psp_unless: | |
| 13805 | 410 | "[| F \<in> A leadsTo A'; F \<in> B unless B' |] | 
| 411 | ==> F \<in> (A \<inter> B) leadsTo ((A' \<inter> B) \<union> B')" | |
| 13797 | 412 | |
| 413 | apply (unfold unless_def) | |
| 414 | apply (drule psp, assumption) | |
| 415 | apply (blast intro: leadsTo_weaken) | |
| 416 | done | |
| 417 | ||
| 418 | ||
| 13798 | 419 | subsection{*Proving the induction rules*}
 | 
| 13797 | 420 | |
| 421 | (** The most general rule: r is any wf relation; f is any variant function **) | |
| 422 | ||
| 423 | lemma leadsTo_wf_induct_lemma: | |
| 424 | "[| wf r; | |
| 13805 | 425 |          \<forall>m. F \<in> (A \<inter> f-`{m}) leadsTo                      
 | 
| 426 |                     ((A \<inter> f-`(r^-1 `` {m})) \<union> B) |]  
 | |
| 427 |       ==> F \<in> (A \<inter> f-`{m}) leadsTo B"
 | |
| 13797 | 428 | apply (erule_tac a = m in wf_induct) | 
| 13805 | 429 | apply (subgoal_tac "F \<in> (A \<inter> (f -` (r^-1 `` {x}))) leadsTo B")
 | 
| 13797 | 430 | apply (blast intro: leadsTo_cancel1 leadsTo_Un_duplicate) | 
| 431 | apply (subst vimage_eq_UN) | |
| 432 | apply (simp only: UN_simps [symmetric]) | |
| 433 | apply (blast intro: leadsTo_UN) | |
| 434 | done | |
| 435 | ||
| 436 | ||
| 437 | (** Meta or object quantifier ? **) | |
| 438 | lemma leadsTo_wf_induct: | |
| 439 | "[| wf r; | |
| 13805 | 440 |          \<forall>m. F \<in> (A \<inter> f-`{m}) leadsTo                      
 | 
| 441 |                     ((A \<inter> f-`(r^-1 `` {m})) \<union> B) |]  
 | |
| 442 | ==> F \<in> A leadsTo B" | |
| 13797 | 443 | apply (rule_tac t = A in subst) | 
| 444 | defer 1 | |
| 445 | apply (rule leadsTo_UN) | |
| 446 | apply (erule leadsTo_wf_induct_lemma) | |
| 447 | apply assumption | |
| 448 | apply fast (*Blast_tac: Function unknown's argument not a parameter*) | |
| 449 | done | |
| 450 | ||
| 451 | ||
| 452 | lemma bounded_induct: | |
| 453 | "[| wf r; | |
| 13805 | 454 |          \<forall>m \<in> I. F \<in> (A \<inter> f-`{m}) leadsTo                    
 | 
| 455 |                       ((A \<inter> f-`(r^-1 `` {m})) \<union> B) |]  
 | |
| 456 | ==> F \<in> A leadsTo ((A - (f-`I)) \<union> B)" | |
| 13797 | 457 | apply (erule leadsTo_wf_induct, safe) | 
| 13805 | 458 | apply (case_tac "m \<in> I") | 
| 13797 | 459 | apply (blast intro: leadsTo_weaken) | 
| 460 | apply (blast intro: subset_imp_leadsTo) | |
| 461 | done | |
| 462 | ||
| 463 | ||
| 13805 | 464 | (*Alternative proof is via the lemma F \<in> (A \<inter> f-`(lessThan m)) leadsTo B*) | 
| 13797 | 465 | lemma lessThan_induct: | 
| 15045 | 466 |      "[| !!m::nat. F \<in> (A \<inter> f-`{m}) leadsTo ((A \<inter> f-`{..<m}) \<union> B) |]  
 | 
| 13805 | 467 | ==> F \<in> A leadsTo B" | 
| 13797 | 468 | apply (rule wf_less_than [THEN leadsTo_wf_induct]) | 
| 469 | apply (simp (no_asm_simp)) | |
| 470 | apply blast | |
| 471 | done | |
| 472 | ||
| 473 | lemma lessThan_bounded_induct: | |
| 13805 | 474 | "!!l::nat. [| \<forall>m \<in> greaterThan l. | 
| 475 |             F \<in> (A \<inter> f-`{m}) leadsTo ((A \<inter> f-`(lessThan m)) \<union> B) |]  
 | |
| 476 | ==> F \<in> A leadsTo ((A \<inter> (f-`(atMost l))) \<union> B)" | |
| 13797 | 477 | apply (simp only: Diff_eq [symmetric] vimage_Compl Compl_greaterThan [symmetric]) | 
| 478 | apply (rule wf_less_than [THEN bounded_induct]) | |
| 479 | apply (simp (no_asm_simp)) | |
| 480 | done | |
| 481 | ||
| 482 | lemma greaterThan_bounded_induct: | |
| 13805 | 483 | "(!!l::nat. \<forall>m \<in> lessThan l. | 
| 484 |                  F \<in> (A \<inter> f-`{m}) leadsTo ((A \<inter> f-`(greaterThan m)) \<union> B))
 | |
| 485 | ==> F \<in> A leadsTo ((A \<inter> (f-`(atLeast l))) \<union> B)" | |
| 13797 | 486 | apply (rule_tac f = f and f1 = "%k. l - k" | 
| 487 | in wf_less_than [THEN wf_inv_image, THEN leadsTo_wf_induct]) | |
| 19769 
c40ce2de2020
Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
 krauss parents: 
16417diff
changeset | 488 | apply (simp (no_asm) add:Image_singleton) | 
| 13797 | 489 | apply clarify | 
| 490 | apply (case_tac "m<l") | |
| 13805 | 491 | apply (blast intro: leadsTo_weaken_R diff_less_mono2) | 
| 492 | apply (blast intro: not_leE subset_imp_leadsTo) | |
| 13797 | 493 | done | 
| 494 | ||
| 495 | ||
| 13798 | 496 | subsection{*wlt*}
 | 
| 13797 | 497 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 498 | text{*Misra's property W3*}
 | 
| 13805 | 499 | lemma wlt_leadsTo: "F \<in> (wlt F B) leadsTo B" | 
| 13797 | 500 | apply (unfold wlt_def) | 
| 501 | apply (blast intro!: leadsTo_Union) | |
| 502 | done | |
| 503 | ||
| 13805 | 504 | lemma leadsTo_subset: "F \<in> A leadsTo B ==> A \<subseteq> wlt F B" | 
| 13797 | 505 | apply (unfold wlt_def) | 
| 506 | apply (blast intro!: leadsTo_Union) | |
| 507 | done | |
| 508 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 509 | text{*Misra's property W2*}
 | 
| 13805 | 510 | lemma leadsTo_eq_subset_wlt: "F \<in> A leadsTo B = (A \<subseteq> wlt F B)" | 
| 13797 | 511 | by (blast intro!: leadsTo_subset wlt_leadsTo [THEN leadsTo_weaken_L]) | 
| 512 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 513 | text{*Misra's property W4*}
 | 
| 13805 | 514 | lemma wlt_increasing: "B \<subseteq> wlt F B" | 
| 13797 | 515 | apply (simp (no_asm_simp) add: leadsTo_eq_subset_wlt [symmetric] subset_imp_leadsTo) | 
| 516 | done | |
| 517 | ||
| 518 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 519 | text{*Used in the Trans case below*}
 | 
| 13797 | 520 | lemma lemma1: | 
| 13805 | 521 | "[| B \<subseteq> A2; | 
| 522 | F \<in> (A1 - B) co (A1 \<union> B); | |
| 523 | F \<in> (A2 - C) co (A2 \<union> C) |] | |
| 524 | ==> F \<in> (A1 \<union> A2 - C) co (A1 \<union> A2 \<union> C)" | |
| 13797 | 525 | by (unfold constrains_def, clarify, blast) | 
| 526 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 527 | text{*Lemma (1,2,3) of Misra's draft book, Chapter 4, "Progress"*}
 | 
| 13797 | 528 | lemma leadsTo_123: | 
| 13805 | 529 | "F \<in> A leadsTo A' | 
| 530 | ==> \<exists>B. A \<subseteq> B & F \<in> B leadsTo A' & F \<in> (B-A') co (B \<union> A')" | |
| 13797 | 531 | apply (erule leadsTo_induct) | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 532 |   txt{*Basis*}
 | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 533 | apply (blast dest: ensuresD) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 534 |  txt{*Trans*}
 | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 535 | apply clarify | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 536 | apply (rule_tac x = "Ba \<union> Bb" in exI) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 537 | apply (blast intro: lemma1 leadsTo_Un_Un leadsTo_cancel1 leadsTo_Un_duplicate) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 538 | txt{*Union*}
 | 
| 13797 | 539 | apply (clarify dest!: ball_conj_distrib [THEN iffD1] bchoice) | 
| 13805 | 540 | apply (rule_tac x = "\<Union>A \<in> S. f A" in exI) | 
| 13797 | 541 | apply (auto intro: leadsTo_UN) | 
| 542 | (*Blast_tac says PROOF FAILED*) | |
| 13805 | 543 | apply (rule_tac I1=S and A1="%i. f i - B" and A'1="%i. f i \<union> B" | 
| 13798 | 544 | in constrains_UN [THEN constrains_weaken], auto) | 
| 13797 | 545 | done | 
| 546 | ||
| 547 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 548 | text{*Misra's property W5*}
 | 
| 13805 | 549 | lemma wlt_constrains_wlt: "F \<in> (wlt F B - B) co (wlt F B)" | 
| 13798 | 550 | proof - | 
| 551 | from wlt_leadsTo [of F B, THEN leadsTo_123] | |
| 552 | show ?thesis | |
| 553 | proof (elim exE conjE) | |
| 554 | (* assumes have to be in exactly the form as in the goal displayed at | |
| 555 | this point. Isar doesn't give you any automation. *) | |
| 556 | fix C | |
| 557 | assume wlt: "wlt F B \<subseteq> C" | |
| 558 | and lt: "F \<in> C leadsTo B" | |
| 559 | and co: "F \<in> C - B co C \<union> B" | |
| 560 | have eq: "C = wlt F B" | |
| 561 | proof - | |
| 562 | from lt and wlt show ?thesis | |
| 563 | by (blast dest: leadsTo_eq_subset_wlt [THEN iffD1]) | |
| 564 | qed | |
| 565 | from co show ?thesis by (simp add: eq wlt_increasing Un_absorb2) | |
| 566 | qed | |
| 567 | qed | |
| 13797 | 568 | |
| 569 | ||
| 13798 | 570 | subsection{*Completion: Binary and General Finite versions*}
 | 
| 13797 | 571 | |
| 572 | lemma completion_lemma : | |
| 13805 | 573 | "[| W = wlt F (B' \<union> C); | 
| 574 | F \<in> A leadsTo (A' \<union> C); F \<in> A' co (A' \<union> C); | |
| 575 | F \<in> B leadsTo (B' \<union> C); F \<in> B' co (B' \<union> C) |] | |
| 576 | ==> F \<in> (A \<inter> B) leadsTo ((A' \<inter> B') \<union> C)" | |
| 577 | apply (subgoal_tac "F \<in> (W-C) co (W \<union> B' \<union> C) ") | |
| 13797 | 578 | prefer 2 | 
| 579 | apply (blast intro: wlt_constrains_wlt [THEN [2] constrains_Un, | |
| 580 | THEN constrains_weaken]) | |
| 13805 | 581 | apply (subgoal_tac "F \<in> (W-C) co W") | 
| 13797 | 582 | prefer 2 | 
| 583 | apply (simp add: wlt_increasing Un_assoc Un_absorb2) | |
| 13805 | 584 | apply (subgoal_tac "F \<in> (A \<inter> W - C) leadsTo (A' \<inter> W \<union> C) ") | 
| 13797 | 585 | prefer 2 apply (blast intro: wlt_leadsTo psp [THEN leadsTo_weaken]) | 
| 586 | (** LEVEL 6 **) | |
| 13805 | 587 | apply (subgoal_tac "F \<in> (A' \<inter> W \<union> C) leadsTo (A' \<inter> B' \<union> C) ") | 
| 13797 | 588 | prefer 2 | 
| 589 | apply (rule leadsTo_Un_duplicate2) | |
| 590 | apply (blast intro: leadsTo_Un_Un wlt_leadsTo | |
| 591 | [THEN psp2, THEN leadsTo_weaken] leadsTo_refl) | |
| 592 | apply (drule leadsTo_Diff) | |
| 593 | apply (blast intro: subset_imp_leadsTo) | |
| 13805 | 594 | apply (subgoal_tac "A \<inter> B \<subseteq> A \<inter> W") | 
| 13797 | 595 | prefer 2 | 
| 596 | apply (blast dest!: leadsTo_subset intro!: subset_refl [THEN Int_mono]) | |
| 597 | apply (blast intro: leadsTo_Trans subset_imp_leadsTo) | |
| 598 | done | |
| 599 | ||
| 600 | lemmas completion = completion_lemma [OF refl] | |
| 601 | ||
| 602 | lemma finite_completion_lemma: | |
| 13805 | 603 | "finite I ==> (\<forall>i \<in> I. F \<in> (A i) leadsTo (A' i \<union> C)) --> | 
| 604 | (\<forall>i \<in> I. F \<in> (A' i) co (A' i \<union> C)) --> | |
| 605 | F \<in> (\<Inter>i \<in> I. A i) leadsTo ((\<Inter>i \<in> I. A' i) \<union> C)" | |
| 13797 | 606 | apply (erule finite_induct, auto) | 
| 607 | apply (rule completion) | |
| 608 | prefer 4 | |
| 609 | apply (simp only: INT_simps [symmetric]) | |
| 610 | apply (rule constrains_INT, auto) | |
| 611 | done | |
| 612 | ||
| 613 | lemma finite_completion: | |
| 614 | "[| finite I; | |
| 13805 | 615 | !!i. i \<in> I ==> F \<in> (A i) leadsTo (A' i \<union> C); | 
| 616 | !!i. i \<in> I ==> F \<in> (A' i) co (A' i \<union> C) |] | |
| 617 | ==> F \<in> (\<Inter>i \<in> I. A i) leadsTo ((\<Inter>i \<in> I. A' i) \<union> C)" | |
| 13797 | 618 | by (blast intro: finite_completion_lemma [THEN mp, THEN mp]) | 
| 619 | ||
| 620 | lemma stable_completion: | |
| 13805 | 621 | "[| F \<in> A leadsTo A'; F \<in> stable A'; | 
| 622 | F \<in> B leadsTo B'; F \<in> stable B' |] | |
| 623 | ==> F \<in> (A \<inter> B) leadsTo (A' \<inter> B')" | |
| 13797 | 624 | apply (unfold stable_def) | 
| 625 | apply (rule_tac C1 = "{}" in completion [THEN leadsTo_weaken_R])
 | |
| 626 | apply (force+) | |
| 627 | done | |
| 628 | ||
| 629 | lemma finite_stable_completion: | |
| 630 | "[| finite I; | |
| 13805 | 631 | !!i. i \<in> I ==> F \<in> (A i) leadsTo (A' i); | 
| 632 | !!i. i \<in> I ==> F \<in> stable (A' i) |] | |
| 633 | ==> F \<in> (\<Inter>i \<in> I. A i) leadsTo (\<Inter>i \<in> I. A' i)" | |
| 13797 | 634 | apply (unfold stable_def) | 
| 635 | apply (rule_tac C1 = "{}" in finite_completion [THEN leadsTo_weaken_R])
 | |
| 636 | apply (simp_all (no_asm_simp)) | |
| 637 | apply blast+ | |
| 638 | done | |
| 9685 | 639 | |
| 35422 | 640 | end |