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