author | Lars Hupel <lars.hupel@mytum.de> |
Fri, 25 Aug 2023 08:33:46 +0200 | |
changeset 78521 | 8347ffa1f92c |
parent 76217 | 8655344f1cf6 |
permissions | -rw-r--r-- |
11479 | 1 |
(* Title: ZF/UNITY/UNITY.thy |
2 |
Author: Sidi O Ehmety, Computer Laboratory |
|
3 |
Copyright 2001 University of Cambridge |
|
4 |
*) |
|
5 |
||
60770 | 6 |
section \<open>The Basic UNITY Theory\<close> |
14077 | 7 |
|
16417 | 8 |
theory UNITY imports State begin |
16183
052d9aba392d
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14077
diff
changeset
|
9 |
|
60770 | 10 |
text\<open>The basic UNITY theory (revised version, based upon the "co" operator) |
16183
052d9aba392d
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14077
diff
changeset
|
11 |
From Misra, "A Logic for Concurrent Programming", 1994. |
052d9aba392d
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14077
diff
changeset
|
12 |
|
60770 | 13 |
This ZF theory was ported from its HOL equivalent.\<close> |
16183
052d9aba392d
renamed "constrains" to "safety" to avoid keyword clash
paulson
parents:
14077
diff
changeset
|
14 |
|
24893 | 15 |
definition |
16 |
program :: i where |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
17 |
"program \<equiv> {<init, acts, allowed>: |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
18 |
Pow(state) * Pow(Pow(state*state)) * Pow(Pow(state*state)). |
76214 | 19 |
id(state) \<in> acts \<and> id(state) \<in> allowed}" |
11479 | 20 |
|
24893 | 21 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
22 |
mk_program :: "[i,i,i]\<Rightarrow>i" where |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
61798
diff
changeset
|
23 |
\<comment> \<open>The definition yields a program thanks to the coercions |
60770 | 24 |
init \<inter> state, acts \<inter> Pow(state*state), etc.\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
25 |
"mk_program(init, acts, allowed) \<equiv> |
14077 | 26 |
<init \<inter> state, cons(id(state), acts \<inter> Pow(state*state)), |
27 |
cons(id(state), allowed \<inter> Pow(state*state))>" |
|
11479 | 28 |
|
24893 | 29 |
definition |
69587 | 30 |
SKIP :: i (\<open>\<bottom>\<close>) where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
31 |
"SKIP \<equiv> mk_program(state, 0, Pow(state*state))" |
11479 | 32 |
|
12195 | 33 |
(* Coercion from anything to program *) |
24893 | 34 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
35 |
programify :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
36 |
"programify(F) \<equiv> if F \<in> program then F else SKIP" |
12195 | 37 |
|
24893 | 38 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
39 |
RawInit :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
40 |
"RawInit(F) \<equiv> fst(F)" |
14077 | 41 |
|
24893 | 42 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
43 |
Init :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
44 |
"Init(F) \<equiv> RawInit(programify(F))" |
11479 | 45 |
|
24893 | 46 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
47 |
RawActs :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
48 |
"RawActs(F) \<equiv> cons(id(state), fst(snd(F)))" |
11479 | 49 |
|
24893 | 50 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
51 |
Acts :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
52 |
"Acts(F) \<equiv> RawActs(programify(F))" |
11479 | 53 |
|
24893 | 54 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
55 |
RawAllowedActs :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
56 |
"RawAllowedActs(F) \<equiv> cons(id(state), snd(snd(F)))" |
11479 | 57 |
|
24893 | 58 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
59 |
AllowedActs :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
60 |
"AllowedActs(F) \<equiv> RawAllowedActs(programify(F))" |
11479 | 61 |
|
14077 | 62 |
|
24893 | 63 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
64 |
Allowed :: "i \<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
65 |
"Allowed(F) \<equiv> {G \<in> program. Acts(G) \<subseteq> AllowedActs(F)}" |
11479 | 66 |
|
24893 | 67 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
68 |
initially :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
69 |
"initially(A) \<equiv> {F \<in> program. Init(F)\<subseteq>A}" |
14077 | 70 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
71 |
definition "constrains" :: "[i, i] \<Rightarrow> i" (infixl \<open>co\<close> 60) where |
76214 | 72 |
"A co B \<equiv> {F \<in> program. (\<forall>act \<in> Acts(F). act``A\<subseteq>B) \<and> st_set(A)}" |
69593 | 73 |
\<comment> \<open>the condition \<^term>\<open>st_set(A)\<close> makes the definition slightly |
61083 | 74 |
stronger than the HOL one\<close> |
75 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
76 |
definition unless :: "[i, i] \<Rightarrow> i" (infixl \<open>unless\<close> 60) where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
77 |
"A unless B \<equiv> (A - B) co (A \<union> B)" |
61083 | 78 |
|
24893 | 79 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
80 |
stable :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
81 |
"stable(A) \<equiv> A co A" |
11479 | 82 |
|
24893 | 83 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
84 |
strongest_rhs :: "[i, i] \<Rightarrow> i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
85 |
"strongest_rhs(F, A) \<equiv> \<Inter>({B \<in> Pow(state). F \<in> A co B})" |
11479 | 86 |
|
24893 | 87 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
88 |
invariant :: "i \<Rightarrow> i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
89 |
"invariant(A) \<equiv> initially(A) \<inter> stable(A)" |
14077 | 90 |
|
14046 | 91 |
(* meta-function composition *) |
24893 | 92 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
93 |
metacomp :: "[i\<Rightarrow>i, i\<Rightarrow>i] \<Rightarrow> (i\<Rightarrow>i)" (infixl \<open>comp\<close> 65) where |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
94 |
"f comp g \<equiv> \<lambda>x. f(g(x))" |
11479 | 95 |
|
24893 | 96 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
97 |
pg_compl :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
98 |
"pg_compl(X)\<equiv> program - X" |
14077 | 99 |
|
60770 | 100 |
text\<open>SKIP\<close> |
14077 | 101 |
lemma SKIP_in_program [iff,TC]: "SKIP \<in> program" |
102 |
by (force simp add: SKIP_def program_def mk_program_def) |
|
103 |
||
104 |
||
69593 | 105 |
subsection\<open>The function \<^term>\<open>programify\<close>, the coercion from anything to |
60770 | 106 |
program\<close> |
14077 | 107 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
108 |
lemma programify_program [simp]: "F \<in> program \<Longrightarrow> programify(F)=F" |
14077 | 109 |
by (force simp add: programify_def) |
110 |
||
111 |
lemma programify_in_program [iff,TC]: "programify(F) \<in> program" |
|
112 |
by (force simp add: programify_def) |
|
113 |
||
60770 | 114 |
text\<open>Collapsing rules: to remove programify from expressions\<close> |
14077 | 115 |
lemma programify_idem [simp]: "programify(programify(F))=programify(F)" |
116 |
by (force simp add: programify_def) |
|
117 |
||
118 |
lemma Init_programify [simp]: "Init(programify(F)) = Init(F)" |
|
119 |
by (simp add: Init_def) |
|
120 |
||
121 |
lemma Acts_programify [simp]: "Acts(programify(F)) = Acts(F)" |
|
122 |
by (simp add: Acts_def) |
|
123 |
||
124 |
lemma AllowedActs_programify [simp]: |
|
125 |
"AllowedActs(programify(F)) = AllowedActs(F)" |
|
126 |
by (simp add: AllowedActs_def) |
|
127 |
||
60770 | 128 |
subsection\<open>The Inspectors for Programs\<close> |
14077 | 129 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
130 |
lemma id_in_RawActs: "F \<in> program \<Longrightarrow>id(state) \<in> RawActs(F)" |
14077 | 131 |
by (auto simp add: program_def RawActs_def) |
132 |
||
133 |
lemma id_in_Acts [iff,TC]: "id(state) \<in> Acts(F)" |
|
134 |
by (simp add: id_in_RawActs Acts_def) |
|
135 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
136 |
lemma id_in_RawAllowedActs: "F \<in> program \<Longrightarrow>id(state) \<in> RawAllowedActs(F)" |
14077 | 137 |
by (auto simp add: program_def RawAllowedActs_def) |
138 |
||
139 |
lemma id_in_AllowedActs [iff,TC]: "id(state) \<in> AllowedActs(F)" |
|
140 |
by (simp add: id_in_RawAllowedActs AllowedActs_def) |
|
141 |
||
142 |
lemma cons_id_Acts [simp]: "cons(id(state), Acts(F)) = Acts(F)" |
|
143 |
by (simp add: cons_absorb) |
|
144 |
||
145 |
lemma cons_id_AllowedActs [simp]: |
|
146 |
"cons(id(state), AllowedActs(F)) = AllowedActs(F)" |
|
147 |
by (simp add: cons_absorb) |
|
148 |
||
149 |
||
60770 | 150 |
subsection\<open>Types of the Inspectors\<close> |
14077 | 151 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
152 |
lemma RawInit_type: "F \<in> program \<Longrightarrow> RawInit(F)\<subseteq>state" |
14077 | 153 |
by (auto simp add: program_def RawInit_def) |
154 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
155 |
lemma RawActs_type: "F \<in> program \<Longrightarrow> RawActs(F)\<subseteq>Pow(state*state)" |
14077 | 156 |
by (auto simp add: program_def RawActs_def) |
157 |
||
158 |
lemma RawAllowedActs_type: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
159 |
"F \<in> program \<Longrightarrow> RawAllowedActs(F)\<subseteq>Pow(state*state)" |
14077 | 160 |
by (auto simp add: program_def RawAllowedActs_def) |
161 |
||
162 |
lemma Init_type: "Init(F)\<subseteq>state" |
|
163 |
by (simp add: RawInit_type Init_def) |
|
164 |
||
45602 | 165 |
lemmas InitD = Init_type [THEN subsetD] |
14077 | 166 |
|
167 |
lemma st_set_Init [iff]: "st_set(Init(F))" |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
168 |
unfolding st_set_def |
14077 | 169 |
apply (rule Init_type) |
170 |
done |
|
171 |
||
172 |
lemma Acts_type: "Acts(F)\<subseteq>Pow(state*state)" |
|
173 |
by (simp add: RawActs_type Acts_def) |
|
174 |
||
175 |
lemma AllowedActs_type: "AllowedActs(F) \<subseteq> Pow(state*state)" |
|
176 |
by (simp add: RawAllowedActs_type AllowedActs_def) |
|
177 |
||
60770 | 178 |
text\<open>Needed in Behaviors\<close> |
76214 | 179 |
lemma ActsD: "\<lbrakk>act \<in> Acts(F); <s,s'> \<in> act\<rbrakk> \<Longrightarrow> s \<in> state \<and> s' \<in> state" |
14077 | 180 |
by (blast dest: Acts_type [THEN subsetD]) |
181 |
||
182 |
lemma AllowedActsD: |
|
76214 | 183 |
"\<lbrakk>act \<in> AllowedActs(F); <s,s'> \<in> act\<rbrakk> \<Longrightarrow> s \<in> state \<and> s' \<in> state" |
14077 | 184 |
by (blast dest: AllowedActs_type [THEN subsetD]) |
185 |
||
69593 | 186 |
subsection\<open>Simplification rules involving \<^term>\<open>state\<close>, \<^term>\<open>Init\<close>, |
187 |
\<^term>\<open>Acts\<close>, and \<^term>\<open>AllowedActs\<close>\<close> |
|
14077 | 188 |
|
60770 | 189 |
text\<open>But are they really needed?\<close> |
14077 | 190 |
|
46823 | 191 |
lemma state_subset_is_Init_iff [iff]: "state \<subseteq> Init(F) \<longleftrightarrow> Init(F)=state" |
14077 | 192 |
by (cut_tac F = F in Init_type, auto) |
193 |
||
194 |
lemma Pow_state_times_state_is_subset_Acts_iff [iff]: |
|
46823 | 195 |
"Pow(state*state) \<subseteq> Acts(F) \<longleftrightarrow> Acts(F)=Pow(state*state)" |
14077 | 196 |
by (cut_tac F = F in Acts_type, auto) |
197 |
||
198 |
lemma Pow_state_times_state_is_subset_AllowedActs_iff [iff]: |
|
46823 | 199 |
"Pow(state*state) \<subseteq> AllowedActs(F) \<longleftrightarrow> AllowedActs(F)=Pow(state*state)" |
14077 | 200 |
by (cut_tac F = F in AllowedActs_type, auto) |
201 |
||
61798 | 202 |
subsubsection\<open>Eliminating \<open>\<inter> state\<close> from expressions\<close> |
14077 | 203 |
|
204 |
lemma Init_Int_state [simp]: "Init(F) \<inter> state = Init(F)" |
|
205 |
by (cut_tac F = F in Init_type, blast) |
|
206 |
||
207 |
lemma state_Int_Init [simp]: "state \<inter> Init(F) = Init(F)" |
|
208 |
by (cut_tac F = F in Init_type, blast) |
|
209 |
||
210 |
lemma Acts_Int_Pow_state_times_state [simp]: |
|
211 |
"Acts(F) \<inter> Pow(state*state) = Acts(F)" |
|
212 |
by (cut_tac F = F in Acts_type, blast) |
|
213 |
||
214 |
lemma state_times_state_Int_Acts [simp]: |
|
215 |
"Pow(state*state) \<inter> Acts(F) = Acts(F)" |
|
216 |
by (cut_tac F = F in Acts_type, blast) |
|
217 |
||
218 |
lemma AllowedActs_Int_Pow_state_times_state [simp]: |
|
219 |
"AllowedActs(F) \<inter> Pow(state*state) = AllowedActs(F)" |
|
220 |
by (cut_tac F = F in AllowedActs_type, blast) |
|
221 |
||
222 |
lemma state_times_state_Int_AllowedActs [simp]: |
|
223 |
"Pow(state*state) \<inter> AllowedActs(F) = AllowedActs(F)" |
|
224 |
by (cut_tac F = F in AllowedActs_type, blast) |
|
225 |
||
226 |
||
69593 | 227 |
subsubsection\<open>The Operator \<^term>\<open>mk_program\<close>\<close> |
14077 | 228 |
|
229 |
lemma mk_program_in_program [iff,TC]: |
|
230 |
"mk_program(init, acts, allowed) \<in> program" |
|
231 |
by (auto simp add: mk_program_def program_def) |
|
232 |
||
233 |
lemma RawInit_eq [simp]: |
|
234 |
"RawInit(mk_program(init, acts, allowed)) = init \<inter> state" |
|
235 |
by (auto simp add: mk_program_def RawInit_def) |
|
236 |
||
237 |
lemma RawActs_eq [simp]: |
|
238 |
"RawActs(mk_program(init, acts, allowed)) = |
|
239 |
cons(id(state), acts \<inter> Pow(state*state))" |
|
240 |
by (auto simp add: mk_program_def RawActs_def) |
|
241 |
||
242 |
lemma RawAllowedActs_eq [simp]: |
|
243 |
"RawAllowedActs(mk_program(init, acts, allowed)) = |
|
244 |
cons(id(state), allowed \<inter> Pow(state*state))" |
|
245 |
by (auto simp add: mk_program_def RawAllowedActs_def) |
|
246 |
||
247 |
lemma Init_eq [simp]: "Init(mk_program(init, acts, allowed)) = init \<inter> state" |
|
248 |
by (simp add: Init_def) |
|
249 |
||
250 |
lemma Acts_eq [simp]: |
|
251 |
"Acts(mk_program(init, acts, allowed)) = |
|
252 |
cons(id(state), acts \<inter> Pow(state*state))" |
|
253 |
by (simp add: Acts_def) |
|
254 |
||
255 |
lemma AllowedActs_eq [simp]: |
|
256 |
"AllowedActs(mk_program(init, acts, allowed))= |
|
257 |
cons(id(state), allowed \<inter> Pow(state*state))" |
|
258 |
by (simp add: AllowedActs_def) |
|
259 |
||
60770 | 260 |
text\<open>Init, Acts, and AlowedActs of SKIP\<close> |
14077 | 261 |
|
262 |
lemma RawInit_SKIP [simp]: "RawInit(SKIP) = state" |
|
263 |
by (simp add: SKIP_def) |
|
264 |
||
265 |
lemma RawAllowedActs_SKIP [simp]: "RawAllowedActs(SKIP) = Pow(state*state)" |
|
266 |
by (force simp add: SKIP_def) |
|
267 |
||
268 |
lemma RawActs_SKIP [simp]: "RawActs(SKIP) = {id(state)}" |
|
269 |
by (force simp add: SKIP_def) |
|
270 |
||
271 |
lemma Init_SKIP [simp]: "Init(SKIP) = state" |
|
272 |
by (force simp add: SKIP_def) |
|
273 |
||
274 |
lemma Acts_SKIP [simp]: "Acts(SKIP) = {id(state)}" |
|
275 |
by (force simp add: SKIP_def) |
|
276 |
||
277 |
lemma AllowedActs_SKIP [simp]: "AllowedActs(SKIP) = Pow(state*state)" |
|
278 |
by (force simp add: SKIP_def) |
|
279 |
||
60770 | 280 |
text\<open>Equality of UNITY programs\<close> |
14077 | 281 |
|
282 |
lemma raw_surjective_mk_program: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
283 |
"F \<in> program \<Longrightarrow> mk_program(RawInit(F), RawActs(F), RawAllowedActs(F))=F" |
14077 | 284 |
apply (auto simp add: program_def mk_program_def RawInit_def RawActs_def |
285 |
RawAllowedActs_def, blast+) |
|
286 |
done |
|
287 |
||
288 |
lemma surjective_mk_program [simp]: |
|
289 |
"mk_program(Init(F), Acts(F), AllowedActs(F)) = programify(F)" |
|
290 |
by (auto simp add: raw_surjective_mk_program Init_def Acts_def AllowedActs_def) |
|
291 |
||
292 |
lemma program_equalityI: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
293 |
"\<lbrakk>Init(F) = Init(G); Acts(F) = Acts(G); |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
294 |
AllowedActs(F) = AllowedActs(G); F \<in> program; G \<in> program\<rbrakk> \<Longrightarrow> F = G" |
14077 | 295 |
apply (subgoal_tac "programify(F) = programify(G)") |
296 |
apply simp |
|
297 |
apply (simp only: surjective_mk_program [symmetric]) |
|
298 |
done |
|
299 |
||
300 |
lemma program_equalityE: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
301 |
"\<lbrakk>F = G; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
302 |
\<lbrakk>Init(F) = Init(G); Acts(F) = Acts(G); AllowedActs(F) = AllowedActs(G)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
303 |
\<Longrightarrow> P\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
304 |
\<Longrightarrow> P" |
14077 | 305 |
by force |
306 |
||
307 |
||
308 |
lemma program_equality_iff: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
309 |
"\<lbrakk>F \<in> program; G \<in> program\<rbrakk> \<Longrightarrow>(F=G) \<longleftrightarrow> |
76214 | 310 |
(Init(F) = Init(G) \<and> Acts(F) = Acts(G) \<and> AllowedActs(F) = AllowedActs(G))" |
14077 | 311 |
by (blast intro: program_equalityI program_equalityE) |
312 |
||
60770 | 313 |
subsection\<open>These rules allow "lazy" definition expansion\<close> |
14077 | 314 |
|
315 |
lemma def_prg_Init: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
316 |
"F \<equiv> mk_program (init,acts,allowed) \<Longrightarrow> Init(F) = init \<inter> state" |
14077 | 317 |
by auto |
318 |
||
319 |
lemma def_prg_Acts: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
320 |
"F \<equiv> mk_program (init,acts,allowed) |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
321 |
\<Longrightarrow> Acts(F) = cons(id(state), acts \<inter> Pow(state*state))" |
14077 | 322 |
by auto |
323 |
||
324 |
lemma def_prg_AllowedActs: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
325 |
"F \<equiv> mk_program (init,acts,allowed) |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
326 |
\<Longrightarrow> AllowedActs(F) = cons(id(state), allowed \<inter> Pow(state*state))" |
14077 | 327 |
by auto |
328 |
||
329 |
lemma def_prg_simps: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
330 |
"\<lbrakk>F \<equiv> mk_program (init,acts,allowed)\<rbrakk> |
76214 | 331 |
\<Longrightarrow> Init(F) = init \<inter> state \<and> |
332 |
Acts(F) = cons(id(state), acts \<inter> Pow(state*state)) \<and> |
|
14077 | 333 |
AllowedActs(F) = cons(id(state), allowed \<inter> Pow(state*state))" |
334 |
by auto |
|
335 |
||
336 |
||
60770 | 337 |
text\<open>An action is expanded only if a pair of states is being tested against it\<close> |
14077 | 338 |
lemma def_act_simp: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
339 |
"\<lbrakk>act \<equiv> {<s,s'> \<in> A*B. P(s, s')}\<rbrakk> |
76214 | 340 |
\<Longrightarrow> (<s,s'> \<in> act) \<longleftrightarrow> (<s,s'> \<in> A*B \<and> P(s, s'))" |
14077 | 341 |
by auto |
342 |
||
60770 | 343 |
text\<open>A set is expanded only if an element is being tested against it\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
344 |
lemma def_set_simp: "A \<equiv> B \<Longrightarrow> (x \<in> A) \<longleftrightarrow> (x \<in> B)" |
14077 | 345 |
by auto |
346 |
||
347 |
||
60770 | 348 |
subsection\<open>The Constrains Operator\<close> |
14077 | 349 |
|
350 |
lemma constrains_type: "A co B \<subseteq> program" |
|
351 |
by (force simp add: constrains_def) |
|
352 |
||
353 |
lemma constrainsI: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
354 |
"\<lbrakk>(\<And>act s s'. \<lbrakk>act: Acts(F); <s,s'> \<in> act; s \<in> A\<rbrakk> \<Longrightarrow> s' \<in> A'); |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
355 |
F \<in> program; st_set(A)\<rbrakk> \<Longrightarrow> F \<in> A co A'" |
14077 | 356 |
by (force simp add: constrains_def) |
357 |
||
358 |
lemma constrainsD: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
359 |
"F \<in> A co B \<Longrightarrow> \<forall>act \<in> Acts(F). act``A\<subseteq>B" |
14077 | 360 |
by (force simp add: constrains_def) |
361 |
||
76214 | 362 |
lemma constrainsD2: "F \<in> A co B \<Longrightarrow> F \<in> program \<and> st_set(A)" |
14077 | 363 |
by (force simp add: constrains_def) |
364 |
||
46823 | 365 |
lemma constrains_empty [iff]: "F \<in> 0 co B \<longleftrightarrow> F \<in> program" |
14077 | 366 |
by (force simp add: constrains_def st_set_def) |
367 |
||
76214 | 368 |
lemma constrains_empty2 [iff]: "(F \<in> A co 0) \<longleftrightarrow> (A=0 \<and> F \<in> program)" |
14077 | 369 |
by (force simp add: constrains_def st_set_def) |
370 |
||
76214 | 371 |
lemma constrains_state [iff]: "(F \<in> state co B) \<longleftrightarrow> (state\<subseteq>B \<and> F \<in> program)" |
14077 | 372 |
apply (cut_tac F = F in Acts_type) |
373 |
apply (force simp add: constrains_def st_set_def) |
|
374 |
done |
|
375 |
||
76214 | 376 |
lemma constrains_state2 [iff]: "F \<in> A co state \<longleftrightarrow> (F \<in> program \<and> st_set(A))" |
14077 | 377 |
apply (cut_tac F = F in Acts_type) |
378 |
apply (force simp add: constrains_def st_set_def) |
|
379 |
done |
|
380 |
||
60770 | 381 |
text\<open>monotonic in 2nd argument\<close> |
14077 | 382 |
lemma constrains_weaken_R: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
383 |
"\<lbrakk>F \<in> A co A'; A'\<subseteq>B'\<rbrakk> \<Longrightarrow> F \<in> A co B'" |
14077 | 384 |
apply (unfold constrains_def, blast) |
385 |
done |
|
386 |
||
60770 | 387 |
text\<open>anti-monotonic in 1st argument\<close> |
14077 | 388 |
lemma constrains_weaken_L: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
389 |
"\<lbrakk>F \<in> A co A'; B\<subseteq>A\<rbrakk> \<Longrightarrow> F \<in> B co A'" |
14077 | 390 |
apply (unfold constrains_def st_set_def, blast) |
391 |
done |
|
392 |
||
393 |
lemma constrains_weaken: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
394 |
"\<lbrakk>F \<in> A co A'; B\<subseteq>A; A'\<subseteq>B'\<rbrakk> \<Longrightarrow> F \<in> B co B'" |
14077 | 395 |
apply (drule constrains_weaken_R) |
396 |
apply (drule_tac [2] constrains_weaken_L, blast+) |
|
397 |
done |
|
398 |
||
399 |
||
60770 | 400 |
subsection\<open>Constrains and Union\<close> |
14077 | 401 |
|
402 |
lemma constrains_Un: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
403 |
"\<lbrakk>F \<in> A co A'; F \<in> B co B'\<rbrakk> \<Longrightarrow> F \<in> (A \<union> B) co (A' \<union> B')" |
14077 | 404 |
by (auto simp add: constrains_def st_set_def, force) |
405 |
||
406 |
lemma constrains_UN: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
407 |
"\<lbrakk>\<And>i. i \<in> I \<Longrightarrow> F \<in> A(i) co A'(i); F \<in> program\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
408 |
\<Longrightarrow> F \<in> (\<Union>i \<in> I. A(i)) co (\<Union>i \<in> I. A'(i))" |
14077 | 409 |
by (force simp add: constrains_def st_set_def) |
410 |
||
411 |
lemma constrains_Un_distrib: |
|
46823 | 412 |
"(A \<union> B) co C = (A co C) \<inter> (B co C)" |
14077 | 413 |
by (force simp add: constrains_def st_set_def) |
414 |
||
415 |
lemma constrains_UN_distrib: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
416 |
"i \<in> I \<Longrightarrow> (\<Union>i \<in> I. A(i)) co B = (\<Inter>i \<in> I. A(i) co B)" |
14077 | 417 |
by (force simp add: constrains_def st_set_def) |
418 |
||
419 |
||
60770 | 420 |
subsection\<open>Constrains and Intersection\<close> |
14077 | 421 |
|
422 |
lemma constrains_Int_distrib: "C co (A \<inter> B) = (C co A) \<inter> (C co B)" |
|
423 |
by (force simp add: constrains_def st_set_def) |
|
424 |
||
425 |
lemma constrains_INT_distrib: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
426 |
"x \<in> I \<Longrightarrow> A co (\<Inter>i \<in> I. B(i)) = (\<Inter>i \<in> I. A co B(i))" |
14077 | 427 |
by (force simp add: constrains_def st_set_def) |
428 |
||
429 |
lemma constrains_Int: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
430 |
"\<lbrakk>F \<in> A co A'; F \<in> B co B'\<rbrakk> \<Longrightarrow> F \<in> (A \<inter> B) co (A' \<inter> B')" |
14077 | 431 |
by (force simp add: constrains_def st_set_def) |
432 |
||
433 |
lemma constrains_INT [rule_format]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
434 |
"\<lbrakk>\<forall>i \<in> I. F \<in> A(i) co A'(i); F \<in> program\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
435 |
\<Longrightarrow> F \<in> (\<Inter>i \<in> I. A(i)) co (\<Inter>i \<in> I. A'(i))" |
14077 | 436 |
apply (case_tac "I=0") |
437 |
apply (simp add: Inter_def) |
|
438 |
apply (erule not_emptyE) |
|
439 |
apply (auto simp add: constrains_def st_set_def, blast) |
|
440 |
apply (drule bspec, assumption, force) |
|
441 |
done |
|
442 |
||
443 |
(* The rule below simulates the HOL's one for (\<Inter>z. A i) co (\<Inter>z. B i) *) |
|
444 |
lemma constrains_All: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
445 |
"\<lbrakk>\<forall>z. F:{s \<in> state. P(s, z)} co {s \<in> state. Q(s, z)}; F \<in> program\<rbrakk>\<Longrightarrow> |
14077 | 446 |
F:{s \<in> state. \<forall>z. P(s, z)} co {s \<in> state. \<forall>z. Q(s, z)}" |
447 |
by (unfold constrains_def, blast) |
|
448 |
||
449 |
lemma constrains_imp_subset: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
450 |
"\<lbrakk>F \<in> A co A'\<rbrakk> \<Longrightarrow> A \<subseteq> A'" |
14077 | 451 |
by (unfold constrains_def st_set_def, force) |
452 |
||
60770 | 453 |
text\<open>The reasoning is by subsets since "co" refers to single actions |
454 |
only. So this rule isn't that useful.\<close> |
|
14077 | 455 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
456 |
lemma constrains_trans: "\<lbrakk>F \<in> A co B; F \<in> B co C\<rbrakk> \<Longrightarrow> F \<in> A co C" |
14077 | 457 |
by (unfold constrains_def st_set_def, auto, blast) |
458 |
||
459 |
lemma constrains_cancel: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
460 |
"\<lbrakk>F \<in> A co (A' \<union> B); F \<in> B co B'\<rbrakk> \<Longrightarrow> F \<in> A co (A' \<union> B')" |
14077 | 461 |
apply (drule_tac A = B in constrains_imp_subset) |
462 |
apply (blast intro: constrains_weaken_R) |
|
463 |
done |
|
464 |
||
465 |
||
60770 | 466 |
subsection\<open>The Unless Operator\<close> |
14077 | 467 |
|
468 |
lemma unless_type: "A unless B \<subseteq> program" |
|
469 |
by (force simp add: unless_def constrains_def) |
|
470 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
471 |
lemma unlessI: "\<lbrakk>F \<in> (A-B) co (A \<union> B)\<rbrakk> \<Longrightarrow> F \<in> A unless B" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
472 |
unfolding unless_def |
14077 | 473 |
apply (blast dest: constrainsD2) |
474 |
done |
|
475 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
476 |
lemma unlessD: "F :A unless B \<Longrightarrow> F \<in> (A-B) co (A \<union> B)" |
14077 | 477 |
by (unfold unless_def, auto) |
478 |
||
479 |
||
69593 | 480 |
subsection\<open>The Operator \<^term>\<open>initially\<close>\<close> |
14077 | 481 |
|
482 |
lemma initially_type: "initially(A) \<subseteq> program" |
|
483 |
by (unfold initially_def, blast) |
|
484 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
485 |
lemma initiallyI: "\<lbrakk>F \<in> program; Init(F)\<subseteq>A\<rbrakk> \<Longrightarrow> F \<in> initially(A)" |
14077 | 486 |
by (unfold initially_def, blast) |
487 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
488 |
lemma initiallyD: "F \<in> initially(A) \<Longrightarrow> Init(F)\<subseteq>A" |
14077 | 489 |
by (unfold initially_def, blast) |
490 |
||
491 |
||
69593 | 492 |
subsection\<open>The Operator \<^term>\<open>stable\<close>\<close> |
14077 | 493 |
|
494 |
lemma stable_type: "stable(A)\<subseteq>program" |
|
495 |
by (unfold stable_def constrains_def, blast) |
|
496 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
497 |
lemma stableI: "F \<in> A co A \<Longrightarrow> F \<in> stable(A)" |
14077 | 498 |
by (unfold stable_def, assumption) |
499 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
500 |
lemma stableD: "F \<in> stable(A) \<Longrightarrow> F \<in> A co A" |
14077 | 501 |
by (unfold stable_def, assumption) |
502 |
||
76214 | 503 |
lemma stableD2: "F \<in> stable(A) \<Longrightarrow> F \<in> program \<and> st_set(A)" |
14077 | 504 |
by (unfold stable_def constrains_def, auto) |
505 |
||
506 |
lemma stable_state [simp]: "stable(state) = program" |
|
507 |
by (auto simp add: stable_def constrains_def dest: Acts_type [THEN subsetD]) |
|
508 |
||
509 |
||
510 |
lemma stable_unless: "stable(A)= A unless 0" |
|
511 |
by (auto simp add: unless_def stable_def) |
|
512 |
||
513 |
||
69593 | 514 |
subsection\<open>Union and Intersection with \<^term>\<open>stable\<close>\<close> |
14077 | 515 |
|
516 |
lemma stable_Un: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
517 |
"\<lbrakk>F \<in> stable(A); F \<in> stable(A')\<rbrakk> \<Longrightarrow> F \<in> stable(A \<union> A')" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
518 |
unfolding stable_def |
14077 | 519 |
apply (blast intro: constrains_Un) |
520 |
done |
|
521 |
||
522 |
lemma stable_UN: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
523 |
"\<lbrakk>\<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:
69593
diff
changeset
|
524 |
\<Longrightarrow> F \<in> stable (\<Union>i \<in> I. A(i))" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
525 |
unfolding stable_def |
14077 | 526 |
apply (blast intro: constrains_UN) |
527 |
done |
|
528 |
||
529 |
lemma stable_Int: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
530 |
"\<lbrakk>F \<in> stable(A); F \<in> stable(A')\<rbrakk> \<Longrightarrow> F \<in> stable (A \<inter> A')" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
531 |
unfolding stable_def |
14077 | 532 |
apply (blast intro: constrains_Int) |
533 |
done |
|
534 |
||
535 |
lemma stable_INT: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
536 |
"\<lbrakk>\<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:
69593
diff
changeset
|
537 |
\<Longrightarrow> F \<in> stable (\<Inter>i \<in> I. A(i))" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
538 |
unfolding stable_def |
14077 | 539 |
apply (blast intro: constrains_INT) |
540 |
done |
|
541 |
||
542 |
lemma stable_All: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
543 |
"\<lbrakk>\<forall>z. F \<in> stable({s \<in> state. P(s, z)}); F \<in> program\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
544 |
\<Longrightarrow> F \<in> stable({s \<in> state. \<forall>z. P(s, z)})" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
545 |
unfolding stable_def |
14077 | 546 |
apply (rule constrains_All, auto) |
547 |
done |
|
548 |
||
549 |
lemma stable_constrains_Un: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
550 |
"\<lbrakk>F \<in> stable(C); F \<in> A co (C \<union> A')\<rbrakk> \<Longrightarrow> F \<in> (C \<union> A) co (C \<union> A')" |
14077 | 551 |
apply (unfold stable_def constrains_def st_set_def, auto) |
552 |
apply (blast dest!: bspec) |
|
553 |
done |
|
554 |
||
555 |
lemma stable_constrains_Int: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
556 |
"\<lbrakk>F \<in> stable(C); F \<in> (C \<inter> A) co A'\<rbrakk> \<Longrightarrow> F \<in> (C \<inter> A) co (C \<inter> A')" |
14077 | 557 |
by (unfold stable_def constrains_def st_set_def, blast) |
558 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
559 |
(* \<lbrakk>F \<in> stable(C); F \<in> (C \<inter> A) co A\<rbrakk> \<Longrightarrow> F \<in> stable(C \<inter> A) *) |
45602 | 560 |
lemmas stable_constrains_stable = stable_constrains_Int [THEN stableI] |
14077 | 561 |
|
69593 | 562 |
subsection\<open>The Operator \<^term>\<open>invariant\<close>\<close> |
14077 | 563 |
|
564 |
lemma invariant_type: "invariant(A) \<subseteq> program" |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
565 |
unfolding invariant_def |
14077 | 566 |
apply (blast dest: stable_type [THEN subsetD]) |
567 |
done |
|
568 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
569 |
lemma invariantI: "\<lbrakk>Init(F)\<subseteq>A; F \<in> stable(A)\<rbrakk> \<Longrightarrow> F \<in> invariant(A)" |
76217 | 570 |
unfolding invariant_def initially_def |
14077 | 571 |
apply (frule stable_type [THEN subsetD], auto) |
572 |
done |
|
573 |
||
76214 | 574 |
lemma invariantD: "F \<in> invariant(A) \<Longrightarrow> Init(F)\<subseteq>A \<and> F \<in> stable(A)" |
14077 | 575 |
by (unfold invariant_def initially_def, auto) |
576 |
||
76214 | 577 |
lemma invariantD2: "F \<in> invariant(A) \<Longrightarrow> F \<in> program \<and> st_set(A)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
578 |
unfolding invariant_def |
14077 | 579 |
apply (blast dest: stableD2) |
580 |
done |
|
581 |
||
60770 | 582 |
text\<open>Could also say |
69593 | 583 |
\<^term>\<open>invariant(A) \<inter> invariant(B) \<subseteq> invariant (A \<inter> B)\<close>\<close> |
14077 | 584 |
lemma invariant_Int: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
585 |
"\<lbrakk>F \<in> invariant(A); F \<in> invariant(B)\<rbrakk> \<Longrightarrow> F \<in> invariant(A \<inter> B)" |
76217 | 586 |
unfolding invariant_def initially_def |
14077 | 587 |
apply (simp add: stable_Int, blast) |
588 |
done |
|
589 |
||
590 |
||
60770 | 591 |
subsection\<open>The Elimination Theorem\<close> |
14077 | 592 |
|
593 |
(** The "free" m has become universally quantified! |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
594 |
Should the premise be \<And>m instead of \<forall>m ? Would make it harder |
14077 | 595 |
to use in forward proof. **) |
596 |
||
60770 | 597 |
text\<open>The general case is easier to prove than the special case!\<close> |
14077 | 598 |
lemma "elimination": |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
599 |
"\<lbrakk>\<forall>m \<in> M. F \<in> {s \<in> A. x(s) = m} co B(m); F \<in> program\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
600 |
\<Longrightarrow> F \<in> {s \<in> A. x(s) \<in> M} co (\<Union>m \<in> M. B(m))" |
14077 | 601 |
by (auto simp add: constrains_def st_set_def, blast) |
602 |
||
60770 | 603 |
text\<open>As above, but for the special case of A=state\<close> |
14077 | 604 |
lemma elimination2: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
605 |
"\<lbrakk>\<forall>m \<in> M. F \<in> {s \<in> state. x(s) = m} co B(m); F \<in> program\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
606 |
\<Longrightarrow> F:{s \<in> state. x(s) \<in> M} co (\<Union>m \<in> M. B(m))" |
14077 | 607 |
by (rule UNITY.elimination, auto) |
608 |
||
69593 | 609 |
subsection\<open>The Operator \<^term>\<open>strongest_rhs\<close>\<close> |
14077 | 610 |
|
611 |
lemma constrains_strongest_rhs: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
612 |
"\<lbrakk>F \<in> program; st_set(A)\<rbrakk> \<Longrightarrow> F \<in> A co (strongest_rhs(F,A))" |
14077 | 613 |
by (auto simp add: constrains_def strongest_rhs_def st_set_def |
614 |
dest: Acts_type [THEN subsetD]) |
|
615 |
||
616 |
lemma strongest_rhs_is_strongest: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
617 |
"\<lbrakk>F \<in> A co B; st_set(B)\<rbrakk> \<Longrightarrow> strongest_rhs(F,A) \<subseteq> B" |
14077 | 618 |
by (auto simp add: constrains_def strongest_rhs_def st_set_def) |
619 |
||
60770 | 620 |
ML \<open> |
24893 | 621 |
fun simp_of_act def = def RS @{thm def_act_simp}; |
622 |
fun simp_of_set def = def RS @{thm def_set_simp}; |
|
60770 | 623 |
\<close> |
14077 | 624 |
|
11479 | 625 |
end |