author | huffman |
Wed, 22 Jun 2011 13:45:32 -0700 | |
changeset 43525 | 8f28a91ea135 |
parent 36866 | 426d5781bb25 |
child 45605 | a89b4bc311a5 |
permissions | -rw-r--r-- |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30198
diff
changeset
|
1 |
(* Title: HOL/UNITY/UNITY.thy |
4776 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
3 |
Copyright 1998 University of Cambridge |
|
4 |
||
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30198
diff
changeset
|
5 |
The basic UNITY theory (revised version, based upon the "co" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30198
diff
changeset
|
6 |
operator). |
4776 | 7 |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30198
diff
changeset
|
8 |
From Misra, "A Logic for Concurrent Programming", 1994. |
4776 | 9 |
*) |
10 |
||
13798 | 11 |
header {*The Basic UNITY Theory*} |
12 |
||
16417 | 13 |
theory UNITY imports Main begin |
6535 | 14 |
|
15 |
typedef (Program) |
|
10064
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8948
diff
changeset
|
16 |
'a program = "{(init:: 'a set, acts :: ('a * 'a)set set, |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30198
diff
changeset
|
17 |
allowed :: ('a * 'a)set set). Id \<in> acts & Id: allowed}" |
13797 | 18 |
by blast |
6536 | 19 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
20 |
definition Acts :: "'a program => ('a * 'a)set set" where |
14653 | 21 |
"Acts F == (%(init, acts, allowed). acts) (Rep_Program F)" |
22 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
23 |
definition "constrains" :: "['a set, 'a set] => 'a program set" (infixl "co" 60) where |
13805 | 24 |
"A co B == {F. \<forall>act \<in> Acts F. act``A \<subseteq> B}" |
13797 | 25 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
26 |
definition unless :: "['a set, 'a set] => 'a program set" (infixl "unless" 60) where |
13805 | 27 |
"A unless B == (A-B) co (A \<union> B)" |
13797 | 28 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
29 |
definition mk_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set) |
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
30 |
=> 'a program" where |
10064
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8948
diff
changeset
|
31 |
"mk_program == %(init, acts, allowed). |
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8948
diff
changeset
|
32 |
Abs_Program (init, insert Id acts, insert Id allowed)" |
6535 | 33 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
34 |
definition Init :: "'a program => 'a set" where |
10064
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8948
diff
changeset
|
35 |
"Init F == (%(init, acts, allowed). init) (Rep_Program F)" |
6535 | 36 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
37 |
definition AllowedActs :: "'a program => ('a * 'a)set set" where |
10064
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8948
diff
changeset
|
38 |
"AllowedActs F == (%(init, acts, allowed). allowed) (Rep_Program F)" |
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8948
diff
changeset
|
39 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
40 |
definition Allowed :: "'a program => 'a program set" where |
13805 | 41 |
"Allowed F == {G. Acts G \<subseteq> AllowedActs F}" |
4776 | 42 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
43 |
definition stable :: "'a set => 'a program set" where |
6536 | 44 |
"stable A == A co A" |
4776 | 45 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
46 |
definition strongest_rhs :: "['a program, 'a set] => 'a set" where |
13805 | 47 |
"strongest_rhs F A == Inter {B. F \<in> A co B}" |
4776 | 48 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
49 |
definition invariant :: "'a set => 'a program set" where |
13805 | 50 |
"invariant A == {F. Init F \<subseteq> A} \<inter> stable A" |
4776 | 51 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
52 |
definition increasing :: "['a => 'b::{order}] => 'a program set" where |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
53 |
--{*Polymorphic in both states and the meaning of @{text "\<le>"}*} |
13805 | 54 |
"increasing f == \<Inter>z. stable {s. z \<le> f s}" |
4776 | 55 |
|
6536 | 56 |
|
24147 | 57 |
text{*Perhaps HOL shouldn't add this in the first place!*} |
13797 | 58 |
declare image_Collect [simp del] |
59 |
||
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
60 |
subsubsection{*The abstract type of programs*} |
13797 | 61 |
|
62 |
lemmas program_typedef = |
|
63 |
Rep_Program Rep_Program_inverse Abs_Program_inverse |
|
64 |
Program_def Init_def Acts_def AllowedActs_def mk_program_def |
|
65 |
||
13805 | 66 |
lemma Id_in_Acts [iff]: "Id \<in> Acts F" |
13797 | 67 |
apply (cut_tac x = F in Rep_Program) |
68 |
apply (auto simp add: program_typedef) |
|
69 |
done |
|
70 |
||
71 |
lemma insert_Id_Acts [iff]: "insert Id (Acts F) = Acts F" |
|
72 |
by (simp add: insert_absorb Id_in_Acts) |
|
73 |
||
13861 | 74 |
lemma Acts_nonempty [simp]: "Acts F \<noteq> {}" |
75 |
by auto |
|
76 |
||
13805 | 77 |
lemma Id_in_AllowedActs [iff]: "Id \<in> AllowedActs F" |
13797 | 78 |
apply (cut_tac x = F in Rep_Program) |
79 |
apply (auto simp add: program_typedef) |
|
80 |
done |
|
81 |
||
82 |
lemma insert_Id_AllowedActs [iff]: "insert Id (AllowedActs F) = AllowedActs F" |
|
83 |
by (simp add: insert_absorb Id_in_AllowedActs) |
|
84 |
||
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
85 |
subsubsection{*Inspectors for type "program"*} |
13797 | 86 |
|
87 |
lemma Init_eq [simp]: "Init (mk_program (init,acts,allowed)) = init" |
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
88 |
by (simp add: program_typedef) |
13797 | 89 |
|
90 |
lemma Acts_eq [simp]: "Acts (mk_program (init,acts,allowed)) = insert Id acts" |
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
91 |
by (simp add: program_typedef) |
13797 | 92 |
|
93 |
lemma AllowedActs_eq [simp]: |
|
94 |
"AllowedActs (mk_program (init,acts,allowed)) = insert Id allowed" |
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
95 |
by (simp add: program_typedef) |
13797 | 96 |
|
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
97 |
subsubsection{*Equality for UNITY programs*} |
13797 | 98 |
|
99 |
lemma surjective_mk_program [simp]: |
|
100 |
"mk_program (Init F, Acts F, AllowedActs F) = F" |
|
101 |
apply (cut_tac x = F in Rep_Program) |
|
102 |
apply (auto simp add: program_typedef) |
|
103 |
apply (drule_tac f = Abs_Program in arg_cong)+ |
|
104 |
apply (simp add: program_typedef insert_absorb) |
|
105 |
done |
|
106 |
||
107 |
lemma program_equalityI: |
|
108 |
"[| Init F = Init G; Acts F = Acts G; AllowedActs F = AllowedActs G |] |
|
109 |
==> F = G" |
|
110 |
apply (rule_tac t = F in surjective_mk_program [THEN subst]) |
|
111 |
apply (rule_tac t = G in surjective_mk_program [THEN subst], simp) |
|
112 |
done |
|
113 |
||
114 |
lemma program_equalityE: |
|
115 |
"[| F = G; |
|
116 |
[| Init F = Init G; Acts F = Acts G; AllowedActs F = AllowedActs G |] |
|
117 |
==> P |] ==> P" |
|
118 |
by simp |
|
119 |
||
120 |
lemma program_equality_iff: |
|
121 |
"(F=G) = |
|
122 |
(Init F = Init G & Acts F = Acts G &AllowedActs F = AllowedActs G)" |
|
123 |
by (blast intro: program_equalityI program_equalityE) |
|
124 |
||
125 |
||
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
126 |
subsubsection{*co*} |
13797 | 127 |
|
128 |
lemma constrainsI: |
|
13805 | 129 |
"(!!act s s'. [| act: Acts F; (s,s') \<in> act; s \<in> A |] ==> s': A') |
130 |
==> F \<in> A co A'" |
|
13797 | 131 |
by (simp add: constrains_def, blast) |
132 |
||
133 |
lemma constrainsD: |
|
13805 | 134 |
"[| F \<in> A co A'; act: Acts F; (s,s'): act; s \<in> A |] ==> s': A'" |
13797 | 135 |
by (unfold constrains_def, blast) |
136 |
||
13805 | 137 |
lemma constrains_empty [iff]: "F \<in> {} co B" |
13797 | 138 |
by (unfold constrains_def, blast) |
139 |
||
13805 | 140 |
lemma constrains_empty2 [iff]: "(F \<in> A co {}) = (A={})" |
13797 | 141 |
by (unfold constrains_def, blast) |
142 |
||
13805 | 143 |
lemma constrains_UNIV [iff]: "(F \<in> UNIV co B) = (B = UNIV)" |
13797 | 144 |
by (unfold constrains_def, blast) |
145 |
||
13805 | 146 |
lemma constrains_UNIV2 [iff]: "F \<in> A co UNIV" |
13797 | 147 |
by (unfold constrains_def, blast) |
148 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
149 |
text{*monotonic in 2nd argument*} |
13797 | 150 |
lemma constrains_weaken_R: |
13805 | 151 |
"[| F \<in> A co A'; A'<=B' |] ==> F \<in> A co B'" |
13797 | 152 |
by (unfold constrains_def, blast) |
153 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
154 |
text{*anti-monotonic in 1st argument*} |
13797 | 155 |
lemma constrains_weaken_L: |
13805 | 156 |
"[| F \<in> A co A'; B \<subseteq> A |] ==> F \<in> B co A'" |
13797 | 157 |
by (unfold constrains_def, blast) |
158 |
||
159 |
lemma constrains_weaken: |
|
13805 | 160 |
"[| F \<in> A co A'; B \<subseteq> A; A'<=B' |] ==> F \<in> B co B'" |
13797 | 161 |
by (unfold constrains_def, blast) |
162 |
||
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
163 |
subsubsection{*Union*} |
13797 | 164 |
|
165 |
lemma constrains_Un: |
|
13805 | 166 |
"[| F \<in> A co A'; F \<in> B co B' |] ==> F \<in> (A \<union> B) co (A' \<union> B')" |
13797 | 167 |
by (unfold constrains_def, blast) |
168 |
||
169 |
lemma constrains_UN: |
|
13805 | 170 |
"(!!i. i \<in> I ==> F \<in> (A i) co (A' i)) |
171 |
==> F \<in> (\<Union>i \<in> I. A i) co (\<Union>i \<in> I. A' i)" |
|
13797 | 172 |
by (unfold constrains_def, blast) |
173 |
||
13805 | 174 |
lemma constrains_Un_distrib: "(A \<union> B) co C = (A co C) \<inter> (B co C)" |
13797 | 175 |
by (unfold constrains_def, blast) |
176 |
||
13805 | 177 |
lemma constrains_UN_distrib: "(\<Union>i \<in> I. A i) co B = (\<Inter>i \<in> I. A i co B)" |
13797 | 178 |
by (unfold constrains_def, blast) |
179 |
||
13805 | 180 |
lemma constrains_Int_distrib: "C co (A \<inter> B) = (C co A) \<inter> (C co B)" |
13797 | 181 |
by (unfold constrains_def, blast) |
182 |
||
13805 | 183 |
lemma constrains_INT_distrib: "A co (\<Inter>i \<in> I. B i) = (\<Inter>i \<in> I. A co B i)" |
13797 | 184 |
by (unfold constrains_def, blast) |
185 |
||
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
186 |
subsubsection{*Intersection*} |
6536 | 187 |
|
13797 | 188 |
lemma constrains_Int: |
13805 | 189 |
"[| F \<in> A co A'; F \<in> B co B' |] ==> F \<in> (A \<inter> B) co (A' \<inter> B')" |
13797 | 190 |
by (unfold constrains_def, blast) |
191 |
||
192 |
lemma constrains_INT: |
|
13805 | 193 |
"(!!i. i \<in> I ==> F \<in> (A i) co (A' i)) |
194 |
==> F \<in> (\<Inter>i \<in> I. A i) co (\<Inter>i \<in> I. A' i)" |
|
13797 | 195 |
by (unfold constrains_def, blast) |
196 |
||
13805 | 197 |
lemma constrains_imp_subset: "F \<in> A co A' ==> A \<subseteq> A'" |
13797 | 198 |
by (unfold constrains_def, auto) |
199 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
200 |
text{*The reasoning is by subsets since "co" refers to single actions |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
201 |
only. So this rule isn't that useful.*} |
13797 | 202 |
lemma constrains_trans: |
13805 | 203 |
"[| F \<in> A co B; F \<in> B co C |] ==> F \<in> A co C" |
13797 | 204 |
by (unfold constrains_def, blast) |
205 |
||
206 |
lemma constrains_cancel: |
|
13805 | 207 |
"[| F \<in> A co (A' \<union> B); F \<in> B co B' |] ==> F \<in> A co (A' \<union> B')" |
13797 | 208 |
by (unfold constrains_def, clarify, blast) |
209 |
||
210 |
||
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
211 |
subsubsection{*unless*} |
13797 | 212 |
|
13805 | 213 |
lemma unlessI: "F \<in> (A-B) co (A \<union> B) ==> F \<in> A unless B" |
13797 | 214 |
by (unfold unless_def, assumption) |
215 |
||
13805 | 216 |
lemma unlessD: "F \<in> A unless B ==> F \<in> (A-B) co (A \<union> B)" |
13797 | 217 |
by (unfold unless_def, assumption) |
218 |
||
219 |
||
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
220 |
subsubsection{*stable*} |
13797 | 221 |
|
13805 | 222 |
lemma stableI: "F \<in> A co A ==> F \<in> stable A" |
13797 | 223 |
by (unfold stable_def, assumption) |
224 |
||
13805 | 225 |
lemma stableD: "F \<in> stable A ==> F \<in> A co A" |
13797 | 226 |
by (unfold stable_def, assumption) |
227 |
||
228 |
lemma stable_UNIV [simp]: "stable UNIV = UNIV" |
|
229 |
by (unfold stable_def constrains_def, auto) |
|
230 |
||
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
231 |
subsubsection{*Union*} |
13797 | 232 |
|
233 |
lemma stable_Un: |
|
13805 | 234 |
"[| F \<in> stable A; F \<in> stable A' |] ==> F \<in> stable (A \<union> A')" |
13797 | 235 |
|
236 |
apply (unfold stable_def) |
|
237 |
apply (blast intro: constrains_Un) |
|
238 |
done |
|
239 |
||
240 |
lemma stable_UN: |
|
13805 | 241 |
"(!!i. i \<in> I ==> F \<in> stable (A i)) ==> F \<in> stable (\<Union>i \<in> I. A i)" |
13797 | 242 |
apply (unfold stable_def) |
243 |
apply (blast intro: constrains_UN) |
|
244 |
done |
|
245 |
||
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
13861
diff
changeset
|
246 |
lemma stable_Union: |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
13861
diff
changeset
|
247 |
"(!!A. A \<in> X ==> F \<in> stable A) ==> F \<in> stable (\<Union>X)" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
13861
diff
changeset
|
248 |
by (unfold stable_def constrains_def, blast) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
13861
diff
changeset
|
249 |
|
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
250 |
subsubsection{*Intersection*} |
13797 | 251 |
|
252 |
lemma stable_Int: |
|
13805 | 253 |
"[| F \<in> stable A; F \<in> stable A' |] ==> F \<in> stable (A \<inter> A')" |
13797 | 254 |
apply (unfold stable_def) |
255 |
apply (blast intro: constrains_Int) |
|
256 |
done |
|
257 |
||
258 |
lemma stable_INT: |
|
13805 | 259 |
"(!!i. i \<in> I ==> F \<in> stable (A i)) ==> F \<in> stable (\<Inter>i \<in> I. A i)" |
13797 | 260 |
apply (unfold stable_def) |
261 |
apply (blast intro: constrains_INT) |
|
262 |
done |
|
263 |
||
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
13861
diff
changeset
|
264 |
lemma stable_Inter: |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
13861
diff
changeset
|
265 |
"(!!A. A \<in> X ==> F \<in> stable A) ==> F \<in> stable (\<Inter>X)" |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
13861
diff
changeset
|
266 |
by (unfold stable_def constrains_def, blast) |
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
13861
diff
changeset
|
267 |
|
13797 | 268 |
lemma stable_constrains_Un: |
13805 | 269 |
"[| F \<in> stable C; F \<in> A co (C \<union> A') |] ==> F \<in> (C \<union> A) co (C \<union> A')" |
13797 | 270 |
by (unfold stable_def constrains_def, blast) |
271 |
||
272 |
lemma stable_constrains_Int: |
|
13805 | 273 |
"[| F \<in> stable C; F \<in> (C \<inter> A) co A' |] ==> F \<in> (C \<inter> A) co (C \<inter> A')" |
13797 | 274 |
by (unfold stable_def constrains_def, blast) |
275 |
||
13805 | 276 |
(*[| F \<in> stable C; F \<in> (C \<inter> A) co A |] ==> F \<in> stable (C \<inter> A) *) |
13870
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
13861
diff
changeset
|
277 |
lemmas stable_constrains_stable = stable_constrains_Int[THEN stableI, standard] |
13797 | 278 |
|
279 |
||
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
280 |
subsubsection{*invariant*} |
13797 | 281 |
|
13805 | 282 |
lemma invariantI: "[| Init F \<subseteq> A; F \<in> stable A |] ==> F \<in> invariant A" |
13797 | 283 |
by (simp add: invariant_def) |
284 |
||
14150 | 285 |
text{*Could also say @{term "invariant A \<inter> invariant B \<subseteq> invariant(A \<inter> B)"}*} |
13797 | 286 |
lemma invariant_Int: |
13805 | 287 |
"[| F \<in> invariant A; F \<in> invariant B |] ==> F \<in> invariant (A \<inter> B)" |
13797 | 288 |
by (auto simp add: invariant_def stable_Int) |
289 |
||
290 |
||
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
291 |
subsubsection{*increasing*} |
13797 | 292 |
|
293 |
lemma increasingD: |
|
13805 | 294 |
"F \<in> increasing f ==> F \<in> stable {s. z \<subseteq> f s}" |
13797 | 295 |
by (unfold increasing_def, blast) |
296 |
||
13805 | 297 |
lemma increasing_constant [iff]: "F \<in> increasing (%s. c)" |
13797 | 298 |
by (unfold increasing_def stable_def, auto) |
299 |
||
300 |
lemma mono_increasing_o: |
|
13805 | 301 |
"mono g ==> increasing f \<subseteq> increasing (g o f)" |
13797 | 302 |
apply (unfold increasing_def stable_def constrains_def, auto) |
303 |
apply (blast intro: monoD order_trans) |
|
304 |
done |
|
305 |
||
13805 | 306 |
(*Holds by the theorem (Suc m \<subseteq> n) = (m < n) *) |
13797 | 307 |
lemma strict_increasingD: |
13805 | 308 |
"!!z::nat. F \<in> increasing f ==> F \<in> stable {s. z < f s}" |
13797 | 309 |
by (simp add: increasing_def Suc_le_eq [symmetric]) |
310 |
||
311 |
||
312 |
(** The Elimination Theorem. The "free" m has become universally quantified! |
|
13805 | 313 |
Should the premise be !!m instead of \<forall>m ? Would make it harder to use |
13797 | 314 |
in forward proof. **) |
315 |
||
316 |
lemma elimination: |
|
13805 | 317 |
"[| \<forall>m \<in> M. F \<in> {s. s x = m} co (B m) |] |
318 |
==> F \<in> {s. s x \<in> M} co (\<Union>m \<in> M. B m)" |
|
13797 | 319 |
by (unfold constrains_def, blast) |
320 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
321 |
text{*As above, but for the trivial case of a one-variable state, in which the |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
322 |
state is identified with its one variable.*} |
13797 | 323 |
lemma elimination_sing: |
13805 | 324 |
"(\<forall>m \<in> M. F \<in> {m} co (B m)) ==> F \<in> M co (\<Union>m \<in> M. B m)" |
13797 | 325 |
by (unfold constrains_def, blast) |
326 |
||
327 |
||
328 |
||
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
329 |
subsubsection{*Theoretical Results from Section 6*} |
13797 | 330 |
|
331 |
lemma constrains_strongest_rhs: |
|
13805 | 332 |
"F \<in> A co (strongest_rhs F A )" |
13797 | 333 |
by (unfold constrains_def strongest_rhs_def, blast) |
334 |
||
335 |
lemma strongest_rhs_is_strongest: |
|
13805 | 336 |
"F \<in> A co B ==> strongest_rhs F A \<subseteq> B" |
13797 | 337 |
by (unfold constrains_def strongest_rhs_def, blast) |
338 |
||
339 |
||
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
340 |
subsubsection{*Ad-hoc set-theory rules*} |
13797 | 341 |
|
13805 | 342 |
lemma Un_Diff_Diff [simp]: "A \<union> B - (A - B) = B" |
13797 | 343 |
by blast |
344 |
||
13805 | 345 |
lemma Int_Union_Union: "Union(B) \<inter> A = Union((%C. C \<inter> A)`B)" |
13797 | 346 |
by blast |
347 |
||
24147 | 348 |
text{*Needed for WF reasoning in WFair.thy*} |
13797 | 349 |
|
350 |
lemma Image_less_than [simp]: "less_than `` {k} = greaterThan k" |
|
351 |
by blast |
|
352 |
||
353 |
lemma Image_inverse_less_than [simp]: "less_than^-1 `` {k} = lessThan k" |
|
354 |
by blast |
|
6536 | 355 |
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
356 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
357 |
subsection{*Partial versus Total Transitions*} |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
358 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
359 |
definition totalize_act :: "('a * 'a)set => ('a * 'a)set" where |
30198 | 360 |
"totalize_act act == act \<union> Id_on (-(Domain act))" |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
361 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
362 |
definition totalize :: "'a program => 'a program" where |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
363 |
"totalize F == mk_program (Init F, |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30198
diff
changeset
|
364 |
totalize_act ` Acts F, |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30198
diff
changeset
|
365 |
AllowedActs F)" |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
366 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
367 |
definition mk_total_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set) |
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
368 |
=> 'a program" where |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
369 |
"mk_total_program args == totalize (mk_program args)" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
370 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
371 |
definition all_total :: "'a program => bool" where |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
372 |
"all_total F == \<forall>act \<in> Acts F. Domain act = UNIV" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
373 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
374 |
lemma insert_Id_image_Acts: "f Id = Id ==> insert Id (f`Acts F) = f ` Acts F" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
375 |
by (blast intro: sym [THEN image_eqI]) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
376 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
377 |
|
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
378 |
subsubsection{*Basic properties*} |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
379 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
380 |
lemma totalize_act_Id [simp]: "totalize_act Id = Id" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
381 |
by (simp add: totalize_act_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
382 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
383 |
lemma Domain_totalize_act [simp]: "Domain (totalize_act act) = UNIV" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
384 |
by (auto simp add: totalize_act_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
385 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
386 |
lemma Init_totalize [simp]: "Init (totalize F) = Init F" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
387 |
by (unfold totalize_def, auto) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
388 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
389 |
lemma Acts_totalize [simp]: "Acts (totalize F) = (totalize_act ` Acts F)" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
390 |
by (simp add: totalize_def insert_Id_image_Acts) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
391 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
392 |
lemma AllowedActs_totalize [simp]: "AllowedActs (totalize F) = AllowedActs F" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
393 |
by (simp add: totalize_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
394 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
395 |
lemma totalize_constrains_iff [simp]: "(totalize F \<in> A co B) = (F \<in> A co B)" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
396 |
by (simp add: totalize_def totalize_act_def constrains_def, blast) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
397 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
398 |
lemma totalize_stable_iff [simp]: "(totalize F \<in> stable A) = (F \<in> stable A)" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
399 |
by (simp add: stable_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
400 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
401 |
lemma totalize_invariant_iff [simp]: |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
402 |
"(totalize F \<in> invariant A) = (F \<in> invariant A)" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
403 |
by (simp add: invariant_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
404 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
405 |
lemma all_total_totalize: "all_total (totalize F)" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
406 |
by (simp add: totalize_def all_total_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
407 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
408 |
lemma Domain_iff_totalize_act: "(Domain act = UNIV) = (totalize_act act = act)" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
409 |
by (force simp add: totalize_act_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
410 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
411 |
lemma all_total_imp_totalize: "all_total F ==> (totalize F = F)" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
412 |
apply (simp add: all_total_def totalize_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
413 |
apply (rule program_equalityI) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
414 |
apply (simp_all add: Domain_iff_totalize_act image_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
415 |
done |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
416 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
417 |
lemma all_total_iff_totalize: "all_total F = (totalize F = F)" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
418 |
apply (rule iffI) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
419 |
apply (erule all_total_imp_totalize) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
420 |
apply (erule subst) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
421 |
apply (rule all_total_totalize) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
422 |
done |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
423 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
424 |
lemma mk_total_program_constrains_iff [simp]: |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
425 |
"(mk_total_program args \<in> A co B) = (mk_program args \<in> A co B)" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
426 |
by (simp add: mk_total_program_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
427 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
428 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
429 |
subsection{*Rules for Lazy Definition Expansion*} |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
430 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
431 |
text{*They avoid expanding the full program, which is a large expression*} |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
432 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
433 |
lemma def_prg_Init: |
36866 | 434 |
"F = mk_total_program (init,acts,allowed) ==> Init F = init" |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
435 |
by (simp add: mk_total_program_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
436 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
437 |
lemma def_prg_Acts: |
36866 | 438 |
"F = mk_total_program (init,acts,allowed) |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
439 |
==> Acts F = insert Id (totalize_act ` acts)" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
440 |
by (simp add: mk_total_program_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
441 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
442 |
lemma def_prg_AllowedActs: |
36866 | 443 |
"F = mk_total_program (init,acts,allowed) |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
444 |
==> AllowedActs F = insert Id allowed" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
445 |
by (simp add: mk_total_program_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
446 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
447 |
text{*An action is expanded if a pair of states is being tested against it*} |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
448 |
lemma def_act_simp: |
36866 | 449 |
"act = {(s,s'). P s s'} ==> ((s,s') \<in> act) = P s s'" |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
450 |
by (simp add: mk_total_program_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
451 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
452 |
text{*A set is expanded only if an element is being tested against it*} |
36866 | 453 |
lemma def_set_simp: "A = B ==> (x \<in> A) = (x \<in> B)" |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
454 |
by (simp add: mk_total_program_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
455 |
|
16184
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14653
diff
changeset
|
456 |
subsubsection{*Inspectors for type "program"*} |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
457 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
458 |
lemma Init_total_eq [simp]: |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
459 |
"Init (mk_total_program (init,acts,allowed)) = init" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
460 |
by (simp add: mk_total_program_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
461 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
462 |
lemma Acts_total_eq [simp]: |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
463 |
"Acts(mk_total_program(init,acts,allowed)) = insert Id (totalize_act`acts)" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
464 |
by (simp add: mk_total_program_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
465 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
466 |
lemma AllowedActs_total_eq [simp]: |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
467 |
"AllowedActs (mk_total_program (init,acts,allowed)) = insert Id allowed" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
468 |
by (auto simp add: mk_total_program_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
469 |
|
4776 | 470 |
end |