| 11479 |      1 | (*  Title:      HOL/UNITY/WFair.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Sidi Ehmety, Computer Laboratory
 | 
|  |      4 |     Copyright   1998  University of Cambridge
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
| 15634 |      7 | header{*Progress under Weak Fairness*}
 | 
|  |      8 | 
 | 
|  |      9 | theory WFair
 | 
|  |     10 | imports UNITY Main_ZFC
 | 
|  |     11 | begin
 | 
|  |     12 | 
 | 
|  |     13 | text{*This theory defines the operators transient, ensures and leadsTo,
 | 
|  |     14 | assuming weak fairness. From Misra, "A Logic for Concurrent Programming",
 | 
|  |     15 | 1994.*}
 | 
|  |     16 | 
 | 
| 11479 |     17 | constdefs
 | 
|  |     18 |   
 | 
| 12195 |     19 |   (* This definition specifies weak fairness.  The rest of the theory
 | 
| 11479 |     20 |     is generic to all forms of fairness.*) 
 | 
|  |     21 |  transient :: "i=>i"
 | 
| 12195 |     22 |   "transient(A) =={F:program. (EX act: Acts(F). A<=domain(act) &
 | 
|  |     23 | 			       act``A <= state-A) & st_set(A)}"
 | 
| 11479 |     24 | 
 | 
| 15634 |     25 |   ensures :: "[i,i] => i"       (infixl "ensures" 60)
 | 
| 11479 |     26 |   "A ensures B == ((A-B) co (A Un B)) Int transient(A-B)"
 | 
|  |     27 |   
 | 
|  |     28 | consts
 | 
|  |     29 | 
 | 
|  |     30 |   (*LEADS-TO constant for the inductive definition*)
 | 
| 12195 |     31 |   leads :: "[i, i]=>i"
 | 
| 11479 |     32 | 
 | 
| 12195 |     33 | inductive 
 | 
| 11479 |     34 |   domains
 | 
| 12195 |     35 |      "leads(D, F)" <= "Pow(D)*Pow(D)"
 | 
| 15634 |     36 |   intros 
 | 
|  |     37 |     Basis:  "[| F:A ensures B;  A:Pow(D); B:Pow(D) |] ==> <A,B>:leads(D, F)"
 | 
|  |     38 |     Trans:  "[| <A,B> : leads(D, F); <B,C> : leads(D, F) |] ==>  <A,C>:leads(D, F)"
 | 
|  |     39 |     Union:   "[| S:Pow({A:S. <A, B>:leads(D, F)}); B:Pow(D); S:Pow(Pow(D)) |] ==> 
 | 
| 12195 |     40 | 	      <Union(S),B>:leads(D, F)"
 | 
| 11479 |     41 | 
 | 
|  |     42 |   monos        Pow_mono
 | 
| 15634 |     43 |   type_intros  Union_Pow_iff [THEN iffD2] UnionI PowI
 | 
| 11479 |     44 |  
 | 
|  |     45 | constdefs
 | 
|  |     46 | 
 | 
| 12195 |     47 |   (* The Visible version of the LEADS-TO relation*)
 | 
| 15634 |     48 |   leadsTo :: "[i, i] => i"       (infixl "leadsTo" 60)
 | 
| 12195 |     49 |   "A leadsTo B == {F:program. <A,B>:leads(state, F)}"
 | 
| 11479 |     50 |   
 | 
| 12195 |     51 |   (* wlt(F, B) is the largest set that leads to B*)
 | 
| 11479 |     52 |   wlt :: "[i, i] => i"
 | 
| 12195 |     53 |     "wlt(F, B) == Union({A:Pow(state). F: A leadsTo B})"
 | 
| 11479 |     54 | 
 | 
|  |     55 | syntax (xsymbols)
 | 
| 15634 |     56 |   "leadsTo" :: "[i, i] => i" (infixl "\<longmapsto>" 60)
 | 
|  |     57 | 
 | 
|  |     58 | (** Ad-hoc set-theory rules **)
 | 
|  |     59 | 
 | 
|  |     60 | lemma Int_Union_Union: "Union(B) Int A = (\<Union>b \<in> B. b Int A)"
 | 
|  |     61 | by auto
 | 
|  |     62 | 
 | 
|  |     63 | lemma Int_Union_Union2: "A Int Union(B) = (\<Union>b \<in> B. A Int b)"
 | 
|  |     64 | by auto
 | 
|  |     65 | 
 | 
|  |     66 | (*** transient ***)
 | 
|  |     67 | 
 | 
|  |     68 | lemma transient_type: "transient(A)<=program"
 | 
|  |     69 | by (unfold transient_def, auto)
 | 
|  |     70 | 
 | 
|  |     71 | lemma transientD2: 
 | 
|  |     72 | "F \<in> transient(A) ==> F \<in> program & st_set(A)"
 | 
|  |     73 | apply (unfold transient_def, auto)
 | 
|  |     74 | done
 | 
|  |     75 | 
 | 
|  |     76 | lemma stable_transient_empty: "[| F \<in> stable(A); F \<in> transient(A) |] ==> A = 0"
 | 
|  |     77 | by (simp add: stable_def constrains_def transient_def, fast)
 | 
|  |     78 | 
 | 
|  |     79 | lemma transient_strengthen: "[|F \<in> transient(A); B<=A|] ==> F \<in> transient(B)"
 | 
|  |     80 | apply (simp add: transient_def st_set_def, clarify)
 | 
|  |     81 | apply (blast intro!: rev_bexI)
 | 
|  |     82 | done
 | 
|  |     83 | 
 | 
|  |     84 | lemma transientI: 
 | 
|  |     85 | "[|act \<in> Acts(F); A <= domain(act); act``A <= state-A;  
 | 
|  |     86 |     F \<in> program; st_set(A)|] ==> F \<in> transient(A)"
 | 
|  |     87 | by (simp add: transient_def, blast)
 | 
|  |     88 | 
 | 
|  |     89 | lemma transientE: 
 | 
|  |     90 |      "[| F \<in> transient(A);  
 | 
|  |     91 |          !!act. [| act \<in> Acts(F);  A <= domain(act); act``A <= state-A|]==>P|]
 | 
|  |     92 |       ==>P"
 | 
|  |     93 | by (simp add: transient_def, blast)
 | 
|  |     94 | 
 | 
|  |     95 | lemma transient_state: "transient(state) = 0"
 | 
|  |     96 | apply (simp add: transient_def)
 | 
|  |     97 | apply (rule equalityI, auto) 
 | 
|  |     98 | apply (cut_tac F = x in Acts_type)
 | 
|  |     99 | apply (simp add: Diff_cancel)
 | 
|  |    100 | apply (auto intro: st0_in_state)
 | 
|  |    101 | done
 | 
|  |    102 | 
 | 
|  |    103 | lemma transient_state2: "state<=B ==> transient(B) = 0"
 | 
|  |    104 | apply (simp add: transient_def st_set_def)
 | 
|  |    105 | apply (rule equalityI, auto)
 | 
|  |    106 | apply (cut_tac F = x in Acts_type)
 | 
|  |    107 | apply (subgoal_tac "B=state")
 | 
|  |    108 | apply (auto intro: st0_in_state)
 | 
|  |    109 | done
 | 
|  |    110 | 
 | 
|  |    111 | lemma transient_empty: "transient(0) = program"
 | 
|  |    112 | by (auto simp add: transient_def)
 | 
|  |    113 | 
 | 
|  |    114 | declare transient_empty [simp] transient_state [simp] transient_state2 [simp]
 | 
|  |    115 | 
 | 
|  |    116 | (*** ensures ***)
 | 
|  |    117 | 
 | 
|  |    118 | lemma ensures_type: "A ensures B <=program"
 | 
|  |    119 | by (simp add: ensures_def constrains_def, auto)
 | 
|  |    120 | 
 | 
|  |    121 | lemma ensuresI: 
 | 
|  |    122 | "[|F:(A-B) co (A Un B); F \<in> transient(A-B)|]==>F \<in> A ensures B"
 | 
|  |    123 | apply (unfold ensures_def)
 | 
|  |    124 | apply (auto simp add: transient_type [THEN subsetD])
 | 
|  |    125 | done
 | 
|  |    126 | 
 | 
|  |    127 | (* Added by Sidi, from Misra's notes, Progress chapter, exercise 4 *)
 | 
|  |    128 | lemma ensuresI2: "[| F \<in> A co A Un B; F \<in> transient(A) |] ==> F \<in> A ensures B"
 | 
|  |    129 | apply (drule_tac B = "A-B" in constrains_weaken_L)
 | 
|  |    130 | apply (drule_tac [2] B = "A-B" in transient_strengthen)
 | 
|  |    131 | apply (auto simp add: ensures_def transient_type [THEN subsetD])
 | 
|  |    132 | done
 | 
|  |    133 | 
 | 
|  |    134 | lemma ensuresD: "F \<in> A ensures B ==> F:(A-B) co (A Un B) & F \<in> transient (A-B)"
 | 
|  |    135 | by (unfold ensures_def, auto)
 | 
|  |    136 | 
 | 
|  |    137 | lemma ensures_weaken_R: "[|F \<in> A ensures A'; A'<=B' |] ==> F \<in> A ensures B'"
 | 
|  |    138 | apply (unfold ensures_def)
 | 
|  |    139 | apply (blast intro: transient_strengthen constrains_weaken)
 | 
|  |    140 | done
 | 
|  |    141 | 
 | 
|  |    142 | (*The L-version (precondition strengthening) fails, but we have this*) 
 | 
|  |    143 | lemma stable_ensures_Int: 
 | 
|  |    144 |      "[| F \<in> stable(C);  F \<in> A ensures B |] ==> F:(C Int A) ensures (C Int B)"
 | 
|  |    145 |  
 | 
|  |    146 | apply (unfold ensures_def)
 | 
|  |    147 | apply (simp (no_asm) add: Int_Un_distrib [symmetric] Diff_Int_distrib [symmetric])
 | 
|  |    148 | apply (blast intro: transient_strengthen stable_constrains_Int constrains_weaken)
 | 
|  |    149 | done
 | 
|  |    150 | 
 | 
|  |    151 | lemma stable_transient_ensures: "[|F \<in> stable(A);  F \<in> transient(C); A<=B Un C|] ==> F \<in> A ensures B"
 | 
|  |    152 | apply (frule stable_type [THEN subsetD])
 | 
|  |    153 | apply (simp add: ensures_def stable_def)
 | 
|  |    154 | apply (blast intro: transient_strengthen constrains_weaken)
 | 
|  |    155 | done
 | 
|  |    156 | 
 | 
|  |    157 | lemma ensures_eq: "(A ensures B) = (A unless B) Int transient (A-B)"
 | 
|  |    158 | by (auto simp add: ensures_def unless_def)
 | 
|  |    159 | 
 | 
|  |    160 | lemma subset_imp_ensures: "[| F \<in> program; A<=B  |] ==> F \<in> A ensures B"
 | 
|  |    161 | by (auto simp add: ensures_def constrains_def transient_def st_set_def)
 | 
|  |    162 | 
 | 
|  |    163 | (*** leadsTo ***)
 | 
|  |    164 | lemmas leads_left = leads.dom_subset [THEN subsetD, THEN SigmaD1]
 | 
|  |    165 | lemmas leads_right = leads.dom_subset [THEN subsetD, THEN SigmaD2]
 | 
|  |    166 | 
 | 
|  |    167 | lemma leadsTo_type: "A leadsTo B <= program"
 | 
|  |    168 | by (unfold leadsTo_def, auto)
 | 
|  |    169 | 
 | 
|  |    170 | lemma leadsToD2: 
 | 
|  |    171 | "F \<in> A leadsTo B ==> F \<in> program & st_set(A) & st_set(B)"
 | 
|  |    172 | apply (unfold leadsTo_def st_set_def)
 | 
|  |    173 | apply (blast dest: leads_left leads_right)
 | 
|  |    174 | done
 | 
|  |    175 | 
 | 
|  |    176 | lemma leadsTo_Basis: 
 | 
|  |    177 |     "[|F \<in> A ensures B; st_set(A); st_set(B)|] ==> F \<in> A leadsTo B"
 | 
|  |    178 | apply (unfold leadsTo_def st_set_def)
 | 
|  |    179 | apply (cut_tac ensures_type)
 | 
|  |    180 | apply (auto intro: leads.Basis)
 | 
|  |    181 | done
 | 
|  |    182 | declare leadsTo_Basis [intro]
 | 
|  |    183 | 
 | 
|  |    184 | (* Added by Sidi, from Misra's notes, Progress chapter, exercise number 4 *)
 | 
|  |    185 | (* [| F \<in> program; A<=B;  st_set(A); st_set(B) |] ==> A leadsTo B *)
 | 
|  |    186 | lemmas subset_imp_leadsTo = subset_imp_ensures [THEN leadsTo_Basis, standard]
 | 
|  |    187 | 
 | 
|  |    188 | lemma leadsTo_Trans: "[|F \<in> A leadsTo B;  F \<in> B leadsTo C |]==>F \<in> A leadsTo C"
 | 
|  |    189 | apply (unfold leadsTo_def)
 | 
|  |    190 | apply (auto intro: leads.Trans)
 | 
|  |    191 | done
 | 
|  |    192 | 
 | 
|  |    193 | (* Better when used in association with leadsTo_weaken_R *)
 | 
|  |    194 | lemma transient_imp_leadsTo: "F \<in> transient(A) ==> F \<in> A leadsTo (state-A)"
 | 
|  |    195 | apply (unfold transient_def)
 | 
|  |    196 | apply (blast intro: ensuresI [THEN leadsTo_Basis] constrains_weaken transientI)
 | 
|  |    197 | done
 | 
|  |    198 | 
 | 
|  |    199 | (*Useful with cancellation, disjunction*)
 | 
|  |    200 | lemma leadsTo_Un_duplicate: "F \<in> A leadsTo (A' Un A') ==> F \<in> A leadsTo A'"
 | 
|  |    201 | by simp
 | 
|  |    202 | 
 | 
|  |    203 | lemma leadsTo_Un_duplicate2:
 | 
|  |    204 |      "F \<in> A leadsTo (A' Un C Un C) ==> F \<in> A leadsTo (A' Un C)"
 | 
|  |    205 | by (simp add: Un_ac)
 | 
|  |    206 | 
 | 
|  |    207 | (*The Union introduction rule as we should have liked to state it*)
 | 
|  |    208 | lemma leadsTo_Union: 
 | 
|  |    209 |     "[|!!A. A \<in> S ==> F \<in> A leadsTo B; F \<in> program; st_set(B)|]
 | 
|  |    210 |      ==> F \<in> Union(S) leadsTo B"
 | 
|  |    211 | apply (unfold leadsTo_def st_set_def)
 | 
|  |    212 | apply (blast intro: leads.Union dest: leads_left)
 | 
|  |    213 | done
 | 
|  |    214 | 
 | 
|  |    215 | lemma leadsTo_Union_Int: 
 | 
|  |    216 |     "[|!!A. A \<in> S ==>F : (A Int C) leadsTo B; F \<in> program; st_set(B)|]  
 | 
|  |    217 |      ==> F : (Union(S)Int C)leadsTo B"
 | 
|  |    218 | apply (unfold leadsTo_def st_set_def)
 | 
|  |    219 | apply (simp only: Int_Union_Union)
 | 
|  |    220 | apply (blast dest: leads_left intro: leads.Union)
 | 
|  |    221 | done
 | 
|  |    222 | 
 | 
|  |    223 | lemma leadsTo_UN: 
 | 
|  |    224 |     "[| !!i. i \<in> I ==> F \<in> A(i) leadsTo B; F \<in> program; st_set(B)|]
 | 
|  |    225 |      ==> F:(\<Union>i \<in> I. A(i)) leadsTo B"
 | 
|  |    226 | apply (simp add: Int_Union_Union leadsTo_def st_set_def)
 | 
|  |    227 | apply (blast dest: leads_left intro: leads.Union)
 | 
|  |    228 | done
 | 
|  |    229 | 
 | 
|  |    230 | (* Binary union introduction rule *)
 | 
|  |    231 | lemma leadsTo_Un:
 | 
|  |    232 |      "[| F \<in> A leadsTo C; F \<in> B leadsTo C |] ==> F \<in> (A Un B) leadsTo C"
 | 
|  |    233 | apply (subst Un_eq_Union)
 | 
|  |    234 | apply (blast intro: leadsTo_Union dest: leadsToD2)
 | 
|  |    235 | done
 | 
|  |    236 | 
 | 
|  |    237 | lemma single_leadsTo_I:
 | 
|  |    238 |     "[|!!x. x \<in> A==> F:{x} leadsTo B; F \<in> program; st_set(B) |] 
 | 
|  |    239 |      ==> F \<in> A leadsTo B"
 | 
|  |    240 | apply (rule_tac b = A in UN_singleton [THEN subst])
 | 
|  |    241 | apply (rule leadsTo_UN, auto) 
 | 
|  |    242 | done
 | 
|  |    243 | 
 | 
|  |    244 | lemma leadsTo_refl: "[| F \<in> program; st_set(A) |] ==> F \<in> A leadsTo A"
 | 
|  |    245 | by (blast intro: subset_imp_leadsTo)
 | 
|  |    246 | 
 | 
|  |    247 | lemma leadsTo_refl_iff: "F \<in> A leadsTo A <-> F \<in> program & st_set(A)"
 | 
|  |    248 | by (auto intro: leadsTo_refl dest: leadsToD2)
 | 
|  |    249 | 
 | 
|  |    250 | lemma empty_leadsTo: "F \<in> 0 leadsTo B <-> (F \<in> program & st_set(B))"
 | 
|  |    251 | by (auto intro: subset_imp_leadsTo dest: leadsToD2)
 | 
|  |    252 | declare empty_leadsTo [iff]
 | 
|  |    253 | 
 | 
|  |    254 | lemma leadsTo_state: "F \<in> A leadsTo state <-> (F \<in> program & st_set(A))"
 | 
|  |    255 | by (auto intro: subset_imp_leadsTo dest: leadsToD2 st_setD)
 | 
|  |    256 | declare leadsTo_state [iff]
 | 
|  |    257 | 
 | 
|  |    258 | lemma leadsTo_weaken_R: "[| F \<in> A leadsTo A'; A'<=B'; st_set(B') |] ==> F \<in> A leadsTo B'"
 | 
|  |    259 | by (blast dest: leadsToD2 intro: subset_imp_leadsTo leadsTo_Trans)
 | 
|  |    260 | 
 | 
|  |    261 | lemma leadsTo_weaken_L: "[| F \<in> A leadsTo A'; B<=A |] ==> F \<in> B leadsTo A'"
 | 
|  |    262 | apply (frule leadsToD2)
 | 
|  |    263 | apply (blast intro: leadsTo_Trans subset_imp_leadsTo st_set_subset)
 | 
|  |    264 | done
 | 
|  |    265 | 
 | 
|  |    266 | lemma leadsTo_weaken: "[| F \<in> A leadsTo A'; B<=A; A'<=B'; st_set(B')|]==> F \<in> B leadsTo B'"
 | 
|  |    267 | apply (frule leadsToD2)
 | 
|  |    268 | apply (blast intro: leadsTo_weaken_R leadsTo_weaken_L leadsTo_Trans leadsTo_refl)
 | 
|  |    269 | done
 | 
|  |    270 | 
 | 
|  |    271 | (* This rule has a nicer conclusion *)
 | 
|  |    272 | lemma transient_imp_leadsTo2: "[| F \<in> transient(A); state-A<=B; st_set(B)|] ==> F \<in> A leadsTo B"
 | 
|  |    273 | apply (frule transientD2)
 | 
|  |    274 | apply (rule leadsTo_weaken_R)
 | 
|  |    275 | apply (auto simp add: transient_imp_leadsTo)
 | 
|  |    276 | done
 | 
|  |    277 | 
 | 
|  |    278 | (*Distributes over binary unions*)
 | 
|  |    279 | lemma leadsTo_Un_distrib: "F:(A Un B) leadsTo C  <->  (F \<in> A leadsTo C & F \<in> B leadsTo C)"
 | 
|  |    280 | by (blast intro: leadsTo_Un leadsTo_weaken_L)
 | 
|  |    281 | 
 | 
|  |    282 | lemma leadsTo_UN_distrib: 
 | 
|  |    283 | "(F:(\<Union>i \<in> I. A(i)) leadsTo B)<-> ((\<forall>i \<in> I. F \<in> A(i) leadsTo B) & F \<in> program & st_set(B))"
 | 
|  |    284 | apply (blast dest: leadsToD2 intro: leadsTo_UN leadsTo_weaken_L)
 | 
|  |    285 | done
 | 
|  |    286 | 
 | 
|  |    287 | lemma leadsTo_Union_distrib: "(F \<in> Union(S) leadsTo B) <->  (\<forall>A \<in> S. F \<in> A leadsTo B) & F \<in> program & st_set(B)"
 | 
|  |    288 | by (blast dest: leadsToD2 intro: leadsTo_Union leadsTo_weaken_L)
 | 
|  |    289 | 
 | 
|  |    290 | text{*Set difference: maybe combine with @{text leadsTo_weaken_L}??*}
 | 
|  |    291 | lemma leadsTo_Diff:
 | 
|  |    292 |      "[| F: (A-B) leadsTo C; F \<in> B leadsTo C; st_set(C) |]
 | 
|  |    293 |       ==> F \<in> A leadsTo C"
 | 
|  |    294 | by (blast intro: leadsTo_Un leadsTo_weaken dest: leadsToD2)
 | 
|  |    295 | 
 | 
|  |    296 | lemma leadsTo_UN_UN:
 | 
|  |    297 |     "[|!!i. i \<in> I ==> F \<in> A(i) leadsTo A'(i); F \<in> program |]  
 | 
|  |    298 |      ==> F: (\<Union>i \<in> I. A(i)) leadsTo (\<Union>i \<in> I. A'(i))"
 | 
|  |    299 | apply (rule leadsTo_Union)
 | 
|  |    300 | apply (auto intro: leadsTo_weaken_R dest: leadsToD2) 
 | 
|  |    301 | done
 | 
|  |    302 | 
 | 
|  |    303 | (*Binary union version*)
 | 
|  |    304 | lemma leadsTo_Un_Un: "[| F \<in> A leadsTo A'; F \<in> B leadsTo B' |] ==> F \<in> (A Un B) leadsTo (A' Un B')"
 | 
|  |    305 | apply (subgoal_tac "st_set (A) & st_set (A') & st_set (B) & st_set (B') ")
 | 
|  |    306 | prefer 2 apply (blast dest: leadsToD2)
 | 
|  |    307 | apply (blast intro: leadsTo_Un leadsTo_weaken_R)
 | 
|  |    308 | done
 | 
|  |    309 | 
 | 
|  |    310 | (** The cancellation law **)
 | 
|  |    311 | lemma leadsTo_cancel2: "[|F \<in> A leadsTo (A' Un B); F \<in> B leadsTo B'|] ==> F \<in> A leadsTo (A' Un B')"
 | 
|  |    312 | apply (subgoal_tac "st_set (A) & st_set (A') & st_set (B) & st_set (B') &F \<in> program")
 | 
|  |    313 | prefer 2 apply (blast dest: leadsToD2)
 | 
|  |    314 | apply (blast intro: leadsTo_Trans leadsTo_Un_Un leadsTo_refl)
 | 
|  |    315 | done
 | 
|  |    316 | 
 | 
|  |    317 | lemma leadsTo_cancel_Diff2: "[|F \<in> A leadsTo (A' Un B); F \<in> (B-A') leadsTo B'|]==> F \<in> A leadsTo (A' Un B')"
 | 
|  |    318 | apply (rule leadsTo_cancel2)
 | 
|  |    319 | prefer 2 apply assumption
 | 
|  |    320 | apply (blast dest: leadsToD2 intro: leadsTo_weaken_R)
 | 
|  |    321 | done
 | 
|  |    322 | 
 | 
|  |    323 | 
 | 
|  |    324 | lemma leadsTo_cancel1: "[| F \<in> A leadsTo (B Un A'); F \<in> B leadsTo B' |] ==> F \<in> A leadsTo (B' Un A')"
 | 
|  |    325 | apply (simp add: Un_commute)
 | 
|  |    326 | apply (blast intro!: leadsTo_cancel2)
 | 
|  |    327 | done
 | 
|  |    328 | 
 | 
|  |    329 | lemma leadsTo_cancel_Diff1:
 | 
|  |    330 |      "[|F \<in> A leadsTo (B Un A'); F: (B-A') leadsTo B'|]==> F \<in> A leadsTo (B' Un A')"
 | 
|  |    331 | apply (rule leadsTo_cancel1)
 | 
|  |    332 | prefer 2 apply assumption
 | 
|  |    333 | apply (blast intro: leadsTo_weaken_R dest: leadsToD2)
 | 
|  |    334 | done
 | 
|  |    335 | 
 | 
|  |    336 | (*The INDUCTION rule as we should have liked to state it*)
 | 
|  |    337 | lemma leadsTo_induct:
 | 
|  |    338 |   assumes major: "F \<in> za leadsTo zb"
 | 
|  |    339 |       and basis: "!!A B. [|F \<in> A ensures B; st_set(A); st_set(B)|] ==> P(A,B)"
 | 
|  |    340 |       and trans: "!!A B C. [| F \<in> A leadsTo B; P(A, B);  
 | 
|  |    341 |                               F \<in> B leadsTo C; P(B, C) |] ==> P(A,C)"
 | 
|  |    342 |       and union: "!!B S. [| \<forall>A \<in> S. F \<in> A leadsTo B; \<forall>A \<in> S. P(A,B); 
 | 
|  |    343 |                            st_set(B); \<forall>A \<in> S. st_set(A)|] ==> P(Union(S), B)"
 | 
|  |    344 |   shows "P(za, zb)"
 | 
|  |    345 | apply (cut_tac major)
 | 
|  |    346 | apply (unfold leadsTo_def, clarify) 
 | 
|  |    347 | apply (erule leads.induct) 
 | 
|  |    348 |   apply (blast intro: basis [unfolded st_set_def])
 | 
|  |    349 |  apply (blast intro: trans [unfolded leadsTo_def]) 
 | 
|  |    350 | apply (force intro: union [unfolded st_set_def leadsTo_def]) 
 | 
|  |    351 | done
 | 
|  |    352 | 
 | 
|  |    353 | 
 | 
|  |    354 | (* Added by Sidi, an induction rule without ensures *)
 | 
|  |    355 | lemma leadsTo_induct2:
 | 
|  |    356 |   assumes major: "F \<in> za leadsTo zb"
 | 
|  |    357 |       and basis1: "!!A B. [| A<=B; st_set(B) |] ==> P(A, B)"
 | 
|  |    358 |       and basis2: "!!A B. [| F \<in> A co A Un B; F \<in> transient(A); st_set(B) |] 
 | 
|  |    359 |                           ==> P(A, B)"
 | 
|  |    360 |       and trans: "!!A B C. [| F \<in> A leadsTo B; P(A, B);  
 | 
|  |    361 |                               F \<in> B leadsTo C; P(B, C) |] ==> P(A,C)"
 | 
|  |    362 |       and union: "!!B S. [| \<forall>A \<in> S. F \<in> A leadsTo B; \<forall>A \<in> S. P(A,B); 
 | 
|  |    363 |                            st_set(B); \<forall>A \<in> S. st_set(A)|] ==> P(Union(S), B)"
 | 
|  |    364 |   shows "P(za, zb)"
 | 
|  |    365 | apply (cut_tac major)
 | 
|  |    366 | apply (erule leadsTo_induct)
 | 
|  |    367 | apply (auto intro: trans union)
 | 
|  |    368 | apply (simp add: ensures_def, clarify)
 | 
|  |    369 | apply (frule constrainsD2)
 | 
|  |    370 | apply (drule_tac B' = " (A-B) Un B" in constrains_weaken_R)
 | 
|  |    371 | apply blast
 | 
|  |    372 | apply (frule ensuresI2 [THEN leadsTo_Basis])
 | 
|  |    373 | apply (drule_tac [4] basis2, simp_all)
 | 
|  |    374 | apply (frule_tac A1 = A and B = B in Int_lower2 [THEN basis1])
 | 
|  |    375 | apply (subgoal_tac "A=Union ({A - B, A Int B}) ")
 | 
|  |    376 | prefer 2 apply blast
 | 
|  |    377 | apply (erule ssubst)
 | 
|  |    378 | apply (rule union)
 | 
|  |    379 | apply (auto intro: subset_imp_leadsTo)
 | 
|  |    380 | done
 | 
|  |    381 | 
 | 
|  |    382 | 
 | 
|  |    383 | (** Variant induction rule: on the preconditions for B **)
 | 
|  |    384 | (*Lemma is the weak version: can't see how to do it in one step*)
 | 
|  |    385 | lemma leadsTo_induct_pre_aux: 
 | 
|  |    386 |   "[| F \<in> za leadsTo zb;   
 | 
|  |    387 |       P(zb);  
 | 
|  |    388 |       !!A B. [| F \<in> A ensures B;  P(B); st_set(A); st_set(B) |] ==> P(A);  
 | 
|  |    389 |       !!S. [| \<forall>A \<in> S. P(A); \<forall>A \<in> S. st_set(A) |] ==> P(Union(S))  
 | 
|  |    390 |    |] ==> P(za)"
 | 
|  |    391 | txt{*by induction on this formula*}
 | 
|  |    392 | apply (subgoal_tac "P (zb) --> P (za) ")
 | 
|  |    393 | txt{*now solve first subgoal: this formula is sufficient*}
 | 
|  |    394 | apply (blast intro: leadsTo_refl)
 | 
|  |    395 | apply (erule leadsTo_induct)
 | 
|  |    396 | apply (blast+)
 | 
|  |    397 | done
 | 
|  |    398 | 
 | 
|  |    399 | 
 | 
|  |    400 | lemma leadsTo_induct_pre: 
 | 
|  |    401 |   "[| F \<in> za leadsTo zb;   
 | 
|  |    402 |       P(zb);  
 | 
|  |    403 |       !!A B. [| F \<in> A ensures B;  F \<in> B leadsTo zb;  P(B); st_set(A) |] ==> P(A);  
 | 
|  |    404 |       !!S. \<forall>A \<in> S. F \<in> A leadsTo zb & P(A) & st_set(A) ==> P(Union(S))  
 | 
|  |    405 |    |] ==> P(za)"
 | 
|  |    406 | apply (subgoal_tac " (F \<in> za leadsTo zb) & P (za) ")
 | 
|  |    407 | apply (erule conjunct2)
 | 
|  |    408 | apply (frule leadsToD2) 
 | 
|  |    409 | apply (erule leadsTo_induct_pre_aux)
 | 
|  |    410 | prefer 3 apply (blast dest: leadsToD2 intro: leadsTo_Union)
 | 
|  |    411 | prefer 2 apply (blast intro: leadsTo_Trans leadsTo_Basis)
 | 
|  |    412 | apply (blast intro: leadsTo_refl)
 | 
|  |    413 | done
 | 
|  |    414 | 
 | 
|  |    415 | (** The impossibility law **)
 | 
|  |    416 | lemma leadsTo_empty: 
 | 
|  |    417 |    "F \<in> A leadsTo 0 ==> A=0"
 | 
|  |    418 | apply (erule leadsTo_induct_pre)
 | 
|  |    419 | apply (auto simp add: ensures_def constrains_def transient_def st_set_def)
 | 
|  |    420 | apply (drule bspec, assumption)+
 | 
|  |    421 | apply blast
 | 
|  |    422 | done
 | 
|  |    423 | declare leadsTo_empty [simp]
 | 
|  |    424 | 
 | 
|  |    425 | subsection{*PSP: Progress-Safety-Progress*}
 | 
|  |    426 | 
 | 
|  |    427 | text{*Special case of PSP: Misra's "stable conjunction"*}
 | 
|  |    428 | 
 | 
|  |    429 | lemma psp_stable: 
 | 
|  |    430 |    "[| F \<in> A leadsTo A'; F \<in> stable(B) |] ==> F:(A Int B) leadsTo (A' Int B)"
 | 
|  |    431 | apply (unfold stable_def)
 | 
|  |    432 | apply (frule leadsToD2) 
 | 
|  |    433 | apply (erule leadsTo_induct)
 | 
|  |    434 | prefer 3 apply (blast intro: leadsTo_Union_Int)
 | 
|  |    435 | prefer 2 apply (blast intro: leadsTo_Trans)
 | 
|  |    436 | apply (rule leadsTo_Basis)
 | 
|  |    437 | apply (simp add: ensures_def Diff_Int_distrib2 [symmetric] Int_Un_distrib2 [symmetric])
 | 
|  |    438 | apply (auto intro: transient_strengthen constrains_Int)
 | 
|  |    439 | done
 | 
|  |    440 | 
 | 
|  |    441 | 
 | 
|  |    442 | lemma psp_stable2: "[|F \<in> A leadsTo A'; F \<in> stable(B) |]==>F: (B Int A) leadsTo (B Int A')"
 | 
|  |    443 | apply (simp (no_asm_simp) add: psp_stable Int_ac)
 | 
|  |    444 | done
 | 
|  |    445 | 
 | 
|  |    446 | lemma psp_ensures: 
 | 
|  |    447 | "[| F \<in> A ensures A'; F \<in> B co B' |]==> F: (A Int B') ensures ((A' Int B) Un (B' - B))"
 | 
|  |    448 | apply (unfold ensures_def constrains_def st_set_def)
 | 
|  |    449 | (*speeds up the proof*)
 | 
|  |    450 | apply clarify
 | 
|  |    451 | apply (blast intro: transient_strengthen)
 | 
|  |    452 | done
 | 
|  |    453 | 
 | 
|  |    454 | lemma psp: 
 | 
|  |    455 | "[|F \<in> A leadsTo A'; F \<in> B co B'; st_set(B')|]==> F:(A Int B') leadsTo ((A' Int B) Un (B' - B))"
 | 
|  |    456 | apply (subgoal_tac "F \<in> program & st_set (A) & st_set (A') & st_set (B) ")
 | 
|  |    457 | prefer 2 apply (blast dest!: constrainsD2 leadsToD2)
 | 
|  |    458 | apply (erule leadsTo_induct)
 | 
|  |    459 | prefer 3 apply (blast intro: leadsTo_Union_Int)
 | 
|  |    460 |  txt{*Basis case*}
 | 
|  |    461 |  apply (blast intro: psp_ensures leadsTo_Basis)
 | 
|  |    462 | txt{*Transitivity case has a delicate argument involving "cancellation"*}
 | 
|  |    463 | apply (rule leadsTo_Un_duplicate2)
 | 
|  |    464 | apply (erule leadsTo_cancel_Diff1)
 | 
|  |    465 | apply (simp add: Int_Diff Diff_triv)
 | 
|  |    466 | apply (blast intro: leadsTo_weaken_L dest: constrains_imp_subset)
 | 
|  |    467 | done
 | 
|  |    468 | 
 | 
|  |    469 | 
 | 
|  |    470 | lemma psp2: "[| F \<in> A leadsTo A'; F \<in> B co B'; st_set(B') |]  
 | 
|  |    471 |     ==> F \<in> (B' Int A) leadsTo ((B Int A') Un (B' - B))"
 | 
|  |    472 | by (simp (no_asm_simp) add: psp Int_ac)
 | 
|  |    473 | 
 | 
|  |    474 | lemma psp_unless: 
 | 
|  |    475 |    "[| F \<in> A leadsTo A';  F \<in> B unless B'; st_set(B); st_set(B') |]  
 | 
|  |    476 |     ==> F \<in> (A Int B) leadsTo ((A' Int B) Un B')"
 | 
|  |    477 | apply (unfold unless_def)
 | 
|  |    478 | apply (subgoal_tac "st_set (A) &st_set (A') ")
 | 
|  |    479 | prefer 2 apply (blast dest: leadsToD2)
 | 
|  |    480 | apply (drule psp, assumption, blast)
 | 
|  |    481 | apply (blast intro: leadsTo_weaken)
 | 
|  |    482 | done
 | 
|  |    483 | 
 | 
|  |    484 | 
 | 
|  |    485 | subsection{*Proving the induction rules*}
 | 
|  |    486 | 
 | 
|  |    487 | (** The most general rule \<in> r is any wf relation; f is any variant function **)
 | 
|  |    488 | lemma leadsTo_wf_induct_aux: "[| wf(r);  
 | 
|  |    489 |          m \<in> I;  
 | 
|  |    490 |          field(r)<=I;  
 | 
|  |    491 |          F \<in> program; st_set(B); 
 | 
|  |    492 |          \<forall>m \<in> I. F \<in> (A Int f-``{m}) leadsTo                      
 | 
|  |    493 |                     ((A Int f-``(converse(r)``{m})) Un B) |]  
 | 
|  |    494 |       ==> F \<in> (A Int f-``{m}) leadsTo B"
 | 
|  |    495 | apply (erule_tac a = m in wf_induct2, simp_all)
 | 
|  |    496 | apply (subgoal_tac "F \<in> (A Int (f-`` (converse (r) ``{x}))) leadsTo B")
 | 
|  |    497 |  apply (blast intro: leadsTo_cancel1 leadsTo_Un_duplicate)
 | 
|  |    498 | apply (subst vimage_eq_UN)
 | 
|  |    499 | apply (simp del: UN_simps add: Int_UN_distrib)
 | 
|  |    500 | apply (auto intro: leadsTo_UN simp del: UN_simps simp add: Int_UN_distrib)
 | 
|  |    501 | done
 | 
|  |    502 | 
 | 
|  |    503 | (** Meta or object quantifier ? **)
 | 
|  |    504 | lemma leadsTo_wf_induct: "[| wf(r);  
 | 
|  |    505 |          field(r)<=I;  
 | 
|  |    506 |          A<=f-``I;  
 | 
|  |    507 |          F \<in> program; st_set(A); st_set(B);  
 | 
|  |    508 |          \<forall>m \<in> I. F \<in> (A Int f-``{m}) leadsTo                      
 | 
|  |    509 |                     ((A Int f-``(converse(r)``{m})) Un B) |]  
 | 
|  |    510 |       ==> F \<in> A leadsTo B"
 | 
|  |    511 | apply (rule_tac b = A in subst)
 | 
|  |    512 |  defer 1
 | 
|  |    513 |  apply (rule_tac I = I in leadsTo_UN)
 | 
|  |    514 |  apply (erule_tac I = I in leadsTo_wf_induct_aux, assumption+, best) 
 | 
|  |    515 | done
 | 
|  |    516 | 
 | 
|  |    517 | lemma nat_measure_field: "field(measure(nat, %x. x)) = nat"
 | 
|  |    518 | apply (unfold field_def)
 | 
|  |    519 | apply (simp add: measure_def)
 | 
|  |    520 | apply (rule equalityI, force, clarify)
 | 
|  |    521 | apply (erule_tac V = "x\<notin>range (?y) " in thin_rl)
 | 
|  |    522 | apply (erule nat_induct)
 | 
|  |    523 | apply (rule_tac [2] b = "succ (succ (xa))" in domainI)
 | 
|  |    524 | apply (rule_tac b = "succ (0) " in domainI)
 | 
|  |    525 | apply simp_all
 | 
|  |    526 | done
 | 
|  |    527 | 
 | 
|  |    528 | 
 | 
|  |    529 | lemma Image_inverse_lessThan: "k<A ==> measure(A, %x. x) -`` {k} = k"
 | 
|  |    530 | apply (rule equalityI)
 | 
|  |    531 | apply (auto simp add: measure_def)
 | 
|  |    532 | apply (blast intro: ltD)
 | 
|  |    533 | apply (rule vimageI)
 | 
|  |    534 | prefer 2 apply blast
 | 
|  |    535 | apply (simp add: lt_Ord lt_Ord2 Ord_mem_iff_lt)
 | 
|  |    536 | apply (blast intro: lt_trans)
 | 
|  |    537 | done
 | 
|  |    538 | 
 | 
|  |    539 | (*Alternative proof is via the lemma F \<in> (A Int f-`(lessThan m)) leadsTo B*)
 | 
|  |    540 | lemma lessThan_induct: 
 | 
|  |    541 |  "[| A<=f-``nat;  
 | 
|  |    542 |      F \<in> program; st_set(A); st_set(B);  
 | 
|  |    543 |      \<forall>m \<in> nat. F:(A Int f-``{m}) leadsTo ((A Int f -`` m) Un B) |]  
 | 
|  |    544 |       ==> F \<in> A leadsTo B"
 | 
|  |    545 | apply (rule_tac A1 = nat and f1 = "%x. x" in wf_measure [THEN leadsTo_wf_induct]) 
 | 
|  |    546 | apply (simp_all add: nat_measure_field)
 | 
|  |    547 | apply (simp add: ltI Image_inverse_lessThan vimage_def [symmetric])
 | 
|  |    548 | done
 | 
|  |    549 | 
 | 
|  |    550 | 
 | 
|  |    551 | (*** wlt ****)
 | 
|  |    552 | 
 | 
|  |    553 | (*Misra's property W3*)
 | 
|  |    554 | lemma wlt_type: "wlt(F,B) <=state"
 | 
|  |    555 | by (unfold wlt_def, auto)
 | 
|  |    556 | 
 | 
|  |    557 | lemma wlt_st_set: "st_set(wlt(F, B))"
 | 
|  |    558 | apply (unfold st_set_def)
 | 
|  |    559 | apply (rule wlt_type)
 | 
|  |    560 | done
 | 
|  |    561 | declare wlt_st_set [iff]
 | 
|  |    562 | 
 | 
|  |    563 | lemma wlt_leadsTo_iff: "F \<in> wlt(F, B) leadsTo B <-> (F \<in> program & st_set(B))"
 | 
|  |    564 | apply (unfold wlt_def)
 | 
|  |    565 | apply (blast dest: leadsToD2 intro!: leadsTo_Union)
 | 
|  |    566 | done
 | 
|  |    567 | 
 | 
|  |    568 | (* [| F \<in> program;  st_set(B) |] ==> F \<in> wlt(F, B) leadsTo B  *)
 | 
|  |    569 | lemmas wlt_leadsTo = conjI [THEN wlt_leadsTo_iff [THEN iffD2], standard]
 | 
|  |    570 | 
 | 
|  |    571 | lemma leadsTo_subset: "F \<in> A leadsTo B ==> A <= wlt(F, B)"
 | 
|  |    572 | apply (unfold wlt_def)
 | 
|  |    573 | apply (frule leadsToD2)
 | 
|  |    574 | apply (auto simp add: st_set_def)
 | 
|  |    575 | done
 | 
|  |    576 | 
 | 
|  |    577 | (*Misra's property W2*)
 | 
|  |    578 | lemma leadsTo_eq_subset_wlt: "F \<in> A leadsTo B <-> (A <= wlt(F,B) & F \<in> program & st_set(B))"
 | 
|  |    579 | apply auto
 | 
|  |    580 | apply (blast dest: leadsToD2 leadsTo_subset intro: leadsTo_weaken_L wlt_leadsTo)+
 | 
|  |    581 | done
 | 
|  |    582 | 
 | 
|  |    583 | (*Misra's property W4*)
 | 
|  |    584 | lemma wlt_increasing: "[| F \<in> program; st_set(B) |] ==> B <= wlt(F,B)"
 | 
|  |    585 | apply (rule leadsTo_subset)
 | 
|  |    586 | apply (simp (no_asm_simp) add: leadsTo_eq_subset_wlt [THEN iff_sym] subset_imp_leadsTo)
 | 
|  |    587 | done
 | 
|  |    588 | 
 | 
|  |    589 | (*Used in the Trans case below*)
 | 
|  |    590 | lemma leadsTo_123_aux: 
 | 
|  |    591 |    "[| B <= A2;  
 | 
|  |    592 |        F \<in> (A1 - B) co (A1 Un B);  
 | 
|  |    593 |        F \<in> (A2 - C) co (A2 Un C) |]  
 | 
|  |    594 |     ==> F \<in> (A1 Un A2 - C) co (A1 Un A2 Un C)"
 | 
|  |    595 | apply (unfold constrains_def st_set_def, blast)
 | 
|  |    596 | done
 | 
|  |    597 | 
 | 
|  |    598 | (*Lemma (1,2,3) of Misra's draft book, Chapter 4, "Progress"*)
 | 
|  |    599 | (* slightly different from the HOL one \<in> B here is bounded *)
 | 
|  |    600 | lemma leadsTo_123: "F \<in> A leadsTo A'  
 | 
|  |    601 |       ==> \<exists>B \<in> Pow(state). A<=B & F \<in> B leadsTo A' & F \<in> (B-A') co (B Un A')"
 | 
|  |    602 | apply (frule leadsToD2)
 | 
|  |    603 | apply (erule leadsTo_induct)
 | 
|  |    604 |   txt{*Basis*}
 | 
|  |    605 |   apply (blast dest: ensuresD constrainsD2 st_setD)
 | 
|  |    606 |  txt{*Trans*}
 | 
|  |    607 |  apply clarify
 | 
|  |    608 |  apply (rule_tac x = "Ba Un Bb" in bexI)
 | 
|  |    609 |  apply (blast intro: leadsTo_123_aux leadsTo_Un_Un leadsTo_cancel1 leadsTo_Un_duplicate, blast)
 | 
|  |    610 | txt{*Union*}
 | 
|  |    611 | apply (clarify dest!: ball_conj_distrib [THEN iffD1])
 | 
|  |    612 | apply (subgoal_tac "\<exists>y. y \<in> Pi (S, %A. {Ba \<in> Pow (state) . A<=Ba & F \<in> Ba leadsTo B & F \<in> Ba - B co Ba Un B}) ")
 | 
|  |    613 | defer 1
 | 
|  |    614 | apply (rule AC_ball_Pi, safe)
 | 
|  |    615 | apply (rotate_tac 1)
 | 
|  |    616 | apply (drule_tac x = x in bspec, blast, blast) 
 | 
|  |    617 | apply (rule_tac x = "\<Union>A \<in> S. y`A" in bexI, safe)
 | 
|  |    618 | apply (rule_tac [3] I1 = S in constrains_UN [THEN constrains_weaken])
 | 
|  |    619 | apply (rule_tac [2] leadsTo_Union)
 | 
|  |    620 | prefer 5 apply (blast dest!: apply_type, simp_all)
 | 
|  |    621 | apply (force dest!: apply_type)+
 | 
|  |    622 | done
 | 
|  |    623 | 
 | 
|  |    624 | 
 | 
|  |    625 | (*Misra's property W5*)
 | 
|  |    626 | lemma wlt_constrains_wlt: "[| F \<in> program; st_set(B) |] ==>F \<in> (wlt(F, B) - B) co (wlt(F,B))"
 | 
|  |    627 | apply (cut_tac F = F in wlt_leadsTo [THEN leadsTo_123], assumption, blast)
 | 
|  |    628 | apply clarify
 | 
|  |    629 | apply (subgoal_tac "Ba = wlt (F,B) ")
 | 
|  |    630 | prefer 2 apply (blast dest: leadsTo_eq_subset_wlt [THEN iffD1], clarify)
 | 
|  |    631 | apply (simp add: wlt_increasing [THEN subset_Un_iff2 [THEN iffD1]])
 | 
|  |    632 | done
 | 
|  |    633 | 
 | 
|  |    634 | 
 | 
|  |    635 | subsection{*Completion: Binary and General Finite versions*}
 | 
|  |    636 | 
 | 
|  |    637 | lemma completion_aux: "[| W = wlt(F, (B' Un C));      
 | 
|  |    638 |        F \<in> A leadsTo (A' Un C);  F \<in> A' co (A' Un C);    
 | 
|  |    639 |        F \<in> B leadsTo (B' Un C);  F \<in> B' co (B' Un C) |]  
 | 
|  |    640 |     ==> F \<in> (A Int B) leadsTo ((A' Int B') Un C)"
 | 
|  |    641 | apply (subgoal_tac "st_set (C) &st_set (W) &st_set (W-C) &st_set (A') &st_set (A) & st_set (B) & st_set (B') & F \<in> program")
 | 
|  |    642 |  prefer 2 
 | 
|  |    643 |  apply simp 
 | 
|  |    644 |  apply (blast dest!: leadsToD2)
 | 
|  |    645 | apply (subgoal_tac "F \<in> (W-C) co (W Un B' Un C) ")
 | 
|  |    646 |  prefer 2
 | 
|  |    647 |  apply (blast intro!: constrains_weaken [OF constrains_Un [OF _ wlt_constrains_wlt]])
 | 
|  |    648 | apply (subgoal_tac "F \<in> (W-C) co W")
 | 
|  |    649 |  prefer 2
 | 
|  |    650 |  apply (simp add: wlt_increasing [THEN subset_Un_iff2 [THEN iffD1]] Un_assoc)
 | 
|  |    651 | apply (subgoal_tac "F \<in> (A Int W - C) leadsTo (A' Int W Un C) ")
 | 
|  |    652 |  prefer 2 apply (blast intro: wlt_leadsTo psp [THEN leadsTo_weaken])
 | 
|  |    653 | (** step 13 **)
 | 
|  |    654 | apply (subgoal_tac "F \<in> (A' Int W Un C) leadsTo (A' Int B' Un C) ")
 | 
|  |    655 | apply (drule leadsTo_Diff)
 | 
|  |    656 | apply (blast intro: subset_imp_leadsTo dest: leadsToD2 constrainsD2)
 | 
|  |    657 | apply (force simp add: st_set_def)
 | 
|  |    658 | apply (subgoal_tac "A Int B <= A Int W")
 | 
|  |    659 | prefer 2 apply (blast dest!: leadsTo_subset intro!: subset_refl [THEN Int_mono])
 | 
|  |    660 | apply (blast intro: leadsTo_Trans subset_imp_leadsTo)
 | 
|  |    661 | txt{*last subgoal*}
 | 
|  |    662 | apply (rule_tac leadsTo_Un_duplicate2)
 | 
|  |    663 | apply (rule_tac leadsTo_Un_Un)
 | 
|  |    664 |  prefer 2 apply (blast intro: leadsTo_refl)
 | 
|  |    665 | apply (rule_tac A'1 = "B' Un C" in wlt_leadsTo[THEN psp2, THEN leadsTo_weaken])
 | 
|  |    666 | apply blast+
 | 
|  |    667 | done
 | 
|  |    668 | 
 | 
|  |    669 | lemmas completion = refl [THEN completion_aux, standard]
 | 
|  |    670 | 
 | 
|  |    671 | lemma finite_completion_aux:
 | 
|  |    672 |      "[| I \<in> Fin(X); F \<in> program; st_set(C) |] ==>  
 | 
|  |    673 |        (\<forall>i \<in> I. F \<in> (A(i)) leadsTo (A'(i) Un C)) -->   
 | 
|  |    674 |                      (\<forall>i \<in> I. F \<in> (A'(i)) co (A'(i) Un C)) -->  
 | 
|  |    675 |                    F \<in> (\<Inter>i \<in> I. A(i)) leadsTo ((\<Inter>i \<in> I. A'(i)) Un C)"
 | 
|  |    676 | apply (erule Fin_induct)
 | 
|  |    677 | apply (auto simp add: Inter_0)
 | 
|  |    678 | apply (rule completion)
 | 
|  |    679 | apply (auto simp del: INT_simps simp add: INT_extend_simps)
 | 
|  |    680 | apply (blast intro: constrains_INT)
 | 
|  |    681 | done
 | 
|  |    682 | 
 | 
|  |    683 | lemma finite_completion: 
 | 
|  |    684 |      "[| I \<in> Fin(X);   
 | 
|  |    685 |          !!i. i \<in> I ==> F \<in> A(i) leadsTo (A'(i) Un C);  
 | 
|  |    686 |          !!i. i \<in> I ==> F \<in> A'(i) co (A'(i) Un C); F \<in> program; st_set(C)|]    
 | 
|  |    687 |       ==> F \<in> (\<Inter>i \<in> I. A(i)) leadsTo ((\<Inter>i \<in> I. A'(i)) Un C)"
 | 
|  |    688 | by (blast intro: finite_completion_aux [THEN mp, THEN mp])
 | 
|  |    689 | 
 | 
|  |    690 | lemma stable_completion: 
 | 
|  |    691 |      "[| F \<in> A leadsTo A';  F \<in> stable(A');    
 | 
|  |    692 |          F \<in> B leadsTo B';  F \<in> stable(B') |]  
 | 
|  |    693 |     ==> F \<in> (A Int B) leadsTo (A' Int B')"
 | 
|  |    694 | apply (unfold stable_def)
 | 
|  |    695 | apply (rule_tac C1 = 0 in completion [THEN leadsTo_weaken_R], simp+)
 | 
|  |    696 | apply (blast dest: leadsToD2)
 | 
|  |    697 | done
 | 
|  |    698 | 
 | 
|  |    699 | 
 | 
|  |    700 | lemma finite_stable_completion: 
 | 
|  |    701 |      "[| I \<in> Fin(X);  
 | 
|  |    702 |          (!!i. i \<in> I ==> F \<in> A(i) leadsTo A'(i));  
 | 
|  |    703 |          (!!i. i \<in> I ==> F \<in> stable(A'(i)));  F \<in> program |]  
 | 
|  |    704 |       ==> F \<in> (\<Inter>i \<in> I. A(i)) leadsTo (\<Inter>i \<in> I. A'(i))"
 | 
|  |    705 | apply (unfold stable_def)
 | 
|  |    706 | apply (subgoal_tac "st_set (\<Inter>i \<in> I. A' (i))")
 | 
|  |    707 | prefer 2 apply (blast dest: leadsToD2)
 | 
|  |    708 | apply (rule_tac C1 = 0 in finite_completion [THEN leadsTo_weaken_R], auto) 
 | 
|  |    709 | done
 | 
|  |    710 | 
 | 
|  |    711 | ML
 | 
|  |    712 | {*
 | 
|  |    713 | val Int_Union_Union = thm "Int_Union_Union";
 | 
|  |    714 | val Int_Union_Union2 = thm "Int_Union_Union2";
 | 
|  |    715 | val transient_type = thm "transient_type";
 | 
|  |    716 | val transientD2 = thm "transientD2";
 | 
|  |    717 | val stable_transient_empty = thm "stable_transient_empty";
 | 
|  |    718 | val transient_strengthen = thm "transient_strengthen";
 | 
|  |    719 | val transientI = thm "transientI";
 | 
|  |    720 | val transientE = thm "transientE";
 | 
|  |    721 | val transient_state = thm "transient_state";
 | 
|  |    722 | val transient_state2 = thm "transient_state2";
 | 
|  |    723 | val transient_empty = thm "transient_empty";
 | 
|  |    724 | val ensures_type = thm "ensures_type";
 | 
|  |    725 | val ensuresI = thm "ensuresI";
 | 
|  |    726 | val ensuresI2 = thm "ensuresI2";
 | 
|  |    727 | val ensuresD = thm "ensuresD";
 | 
|  |    728 | val ensures_weaken_R = thm "ensures_weaken_R";
 | 
|  |    729 | val stable_ensures_Int = thm "stable_ensures_Int";
 | 
|  |    730 | val stable_transient_ensures = thm "stable_transient_ensures";
 | 
|  |    731 | val ensures_eq = thm "ensures_eq";
 | 
|  |    732 | val subset_imp_ensures = thm "subset_imp_ensures";
 | 
|  |    733 | val leads_left = thm "leads_left";
 | 
|  |    734 | val leads_right = thm "leads_right";
 | 
|  |    735 | val leadsTo_type = thm "leadsTo_type";
 | 
|  |    736 | val leadsToD2 = thm "leadsToD2";
 | 
|  |    737 | val leadsTo_Basis = thm "leadsTo_Basis";
 | 
|  |    738 | val subset_imp_leadsTo = thm "subset_imp_leadsTo";
 | 
|  |    739 | val leadsTo_Trans = thm "leadsTo_Trans";
 | 
|  |    740 | val transient_imp_leadsTo = thm "transient_imp_leadsTo";
 | 
|  |    741 | val leadsTo_Un_duplicate = thm "leadsTo_Un_duplicate";
 | 
|  |    742 | val leadsTo_Un_duplicate2 = thm "leadsTo_Un_duplicate2";
 | 
|  |    743 | val leadsTo_Union = thm "leadsTo_Union";
 | 
|  |    744 | val leadsTo_Union_Int = thm "leadsTo_Union_Int";
 | 
|  |    745 | val leadsTo_UN = thm "leadsTo_UN";
 | 
|  |    746 | val leadsTo_Un = thm "leadsTo_Un";
 | 
|  |    747 | val single_leadsTo_I = thm "single_leadsTo_I";
 | 
|  |    748 | val leadsTo_refl = thm "leadsTo_refl";
 | 
|  |    749 | val leadsTo_refl_iff = thm "leadsTo_refl_iff";
 | 
|  |    750 | val empty_leadsTo = thm "empty_leadsTo";
 | 
|  |    751 | val leadsTo_state = thm "leadsTo_state";
 | 
|  |    752 | val leadsTo_weaken_R = thm "leadsTo_weaken_R";
 | 
|  |    753 | val leadsTo_weaken_L = thm "leadsTo_weaken_L";
 | 
|  |    754 | val leadsTo_weaken = thm "leadsTo_weaken";
 | 
|  |    755 | val transient_imp_leadsTo2 = thm "transient_imp_leadsTo2";
 | 
|  |    756 | val leadsTo_Un_distrib = thm "leadsTo_Un_distrib";
 | 
|  |    757 | val leadsTo_UN_distrib = thm "leadsTo_UN_distrib";
 | 
|  |    758 | val leadsTo_Union_distrib = thm "leadsTo_Union_distrib";
 | 
|  |    759 | val leadsTo_Diff = thm "leadsTo_Diff";
 | 
|  |    760 | val leadsTo_UN_UN = thm "leadsTo_UN_UN";
 | 
|  |    761 | val leadsTo_Un_Un = thm "leadsTo_Un_Un";
 | 
|  |    762 | val leadsTo_cancel2 = thm "leadsTo_cancel2";
 | 
|  |    763 | val leadsTo_cancel_Diff2 = thm "leadsTo_cancel_Diff2";
 | 
|  |    764 | val leadsTo_cancel1 = thm "leadsTo_cancel1";
 | 
|  |    765 | val leadsTo_cancel_Diff1 = thm "leadsTo_cancel_Diff1";
 | 
|  |    766 | val leadsTo_induct = thm "leadsTo_induct";
 | 
|  |    767 | val leadsTo_induct2 = thm "leadsTo_induct2";
 | 
|  |    768 | val leadsTo_induct_pre_aux = thm "leadsTo_induct_pre_aux";
 | 
|  |    769 | val leadsTo_induct_pre = thm "leadsTo_induct_pre";
 | 
|  |    770 | val leadsTo_empty = thm "leadsTo_empty";
 | 
|  |    771 | val psp_stable = thm "psp_stable";
 | 
|  |    772 | val psp_stable2 = thm "psp_stable2";
 | 
|  |    773 | val psp_ensures = thm "psp_ensures";
 | 
|  |    774 | val psp = thm "psp";
 | 
|  |    775 | val psp2 = thm "psp2";
 | 
|  |    776 | val psp_unless = thm "psp_unless";
 | 
|  |    777 | val leadsTo_wf_induct_aux = thm "leadsTo_wf_induct_aux";
 | 
|  |    778 | val leadsTo_wf_induct = thm "leadsTo_wf_induct";
 | 
|  |    779 | val nat_measure_field = thm "nat_measure_field";
 | 
|  |    780 | val Image_inverse_lessThan = thm "Image_inverse_lessThan";
 | 
|  |    781 | val lessThan_induct = thm "lessThan_induct";
 | 
|  |    782 | val wlt_type = thm "wlt_type";
 | 
|  |    783 | val wlt_st_set = thm "wlt_st_set";
 | 
|  |    784 | val wlt_leadsTo_iff = thm "wlt_leadsTo_iff";
 | 
|  |    785 | val wlt_leadsTo = thm "wlt_leadsTo";
 | 
|  |    786 | val leadsTo_subset = thm "leadsTo_subset";
 | 
|  |    787 | val leadsTo_eq_subset_wlt = thm "leadsTo_eq_subset_wlt";
 | 
|  |    788 | val wlt_increasing = thm "wlt_increasing";
 | 
|  |    789 | val leadsTo_123_aux = thm "leadsTo_123_aux";
 | 
|  |    790 | val leadsTo_123 = thm "leadsTo_123";
 | 
|  |    791 | val wlt_constrains_wlt = thm "wlt_constrains_wlt";
 | 
|  |    792 | val completion_aux = thm "completion_aux";
 | 
|  |    793 | val completion = thm "completion";
 | 
|  |    794 | val finite_completion_aux = thm "finite_completion_aux";
 | 
|  |    795 | val finite_completion = thm "finite_completion";
 | 
|  |    796 | val stable_completion = thm "stable_completion";
 | 
|  |    797 | val finite_stable_completion = thm "finite_stable_completion";
 | 
|  |    798 | *}
 | 
| 11479 |    799 | 
 | 
|  |    800 | end
 |