author | wenzelm |
Mon, 12 Oct 2020 16:19:11 +0200 | |
changeset 72455 | 7bf67a58f54a |
parent 69597 | ff784d5a5bfb |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
35108 | 1 |
(* Title: HOL/TLA/TLA.thy |
2 |
Author: Stephan Merz |
|
3 |
Copyright: 1998 University of Munich |
|
21624 | 4 |
*) |
3807 | 5 |
|
60592 | 6 |
section \<open>The temporal level of TLA\<close> |
3807 | 7 |
|
17309 | 8 |
theory TLA |
9 |
imports Init |
|
10 |
begin |
|
3807 | 11 |
|
12 |
consts |
|
6255 | 13 |
(** abstract syntax **) |
60588 | 14 |
Box :: "('w::world) form \<Rightarrow> temporal" |
15 |
Dmd :: "('w::world) form \<Rightarrow> temporal" |
|
16 |
leadsto :: "['w::world form, 'v::world form] \<Rightarrow> temporal" |
|
17 |
Stable :: "stpred \<Rightarrow> temporal" |
|
18 |
WF :: "[action, 'a stfun] \<Rightarrow> temporal" |
|
19 |
SF :: "[action, 'a stfun] \<Rightarrow> temporal" |
|
3807 | 20 |
|
21 |
(* Quantification over (flexible) state variables *) |
|
60588 | 22 |
EEx :: "('a stfun \<Rightarrow> temporal) \<Rightarrow> temporal" (binder "Eex " 10) |
23 |
AAll :: "('a stfun \<Rightarrow> temporal) \<Rightarrow> temporal" (binder "Aall " 10) |
|
6255 | 24 |
|
25 |
(** concrete syntax **) |
|
26 |
syntax |
|
60591 | 27 |
"_Box" :: "lift \<Rightarrow> lift" ("(\<box>_)" [40] 40) |
28 |
"_Dmd" :: "lift \<Rightarrow> lift" ("(\<diamond>_)" [40] 40) |
|
29 |
"_leadsto" :: "[lift,lift] \<Rightarrow> lift" ("(_ \<leadsto> _)" [23,22] 22) |
|
60588 | 30 |
"_stable" :: "lift \<Rightarrow> lift" ("(stable/ _)") |
31 |
"_WF" :: "[lift,lift] \<Rightarrow> lift" ("(WF'(_')'_(_))" [0,60] 55) |
|
32 |
"_SF" :: "[lift,lift] \<Rightarrow> lift" ("(SF'(_')'_(_))" [0,60] 55) |
|
6255 | 33 |
|
60591 | 34 |
"_EEx" :: "[idts, lift] \<Rightarrow> lift" ("(3\<exists>\<exists> _./ _)" [0,10] 10) |
35 |
"_AAll" :: "[idts, lift] \<Rightarrow> lift" ("(3\<forall>\<forall> _./ _)" [0,10] 10) |
|
3807 | 36 |
|
37 |
translations |
|
35068 | 38 |
"_Box" == "CONST Box" |
39 |
"_Dmd" == "CONST Dmd" |
|
40 |
"_leadsto" == "CONST leadsto" |
|
41 |
"_stable" == "CONST Stable" |
|
42 |
"_WF" == "CONST WF" |
|
43 |
"_SF" == "CONST SF" |
|
6255 | 44 |
"_EEx v A" == "Eex v. A" |
45 |
"_AAll v A" == "Aall v. A" |
|
46 |
||
60591 | 47 |
"sigma \<Turnstile> \<box>F" <= "_Box F sigma" |
48 |
"sigma \<Turnstile> \<diamond>F" <= "_Dmd F sigma" |
|
49 |
"sigma \<Turnstile> F \<leadsto> G" <= "_leadsto F G sigma" |
|
50 |
"sigma \<Turnstile> stable P" <= "_stable P sigma" |
|
51 |
"sigma \<Turnstile> WF(A)_v" <= "_WF A v sigma" |
|
52 |
"sigma \<Turnstile> SF(A)_v" <= "_SF A v sigma" |
|
53 |
"sigma \<Turnstile> \<exists>\<exists>x. F" <= "_EEx x F sigma" |
|
54 |
"sigma \<Turnstile> \<forall>\<forall>x. F" <= "_AAll x F sigma" |
|
3808 | 55 |
|
47968 | 56 |
axiomatization where |
6255 | 57 |
(* Definitions of derived operators *) |
60587 | 58 |
dmd_def: "\<And>F. TEMP \<diamond>F == TEMP \<not>\<box>\<not>F" |
47968 | 59 |
|
60 |
axiomatization where |
|
60587 | 61 |
boxInit: "\<And>F. TEMP \<box>F == TEMP \<box>Init F" and |
60588 | 62 |
leadsto_def: "\<And>F G. TEMP F \<leadsto> G == TEMP \<box>(Init F \<longrightarrow> \<diamond>G)" and |
63 |
stable_def: "\<And>P. TEMP stable P == TEMP \<box>($P \<longrightarrow> P$)" and |
|
64 |
WF_def: "TEMP WF(A)_v == TEMP \<diamond>\<box> Enabled(<A>_v) \<longrightarrow> \<box>\<diamond><A>_v" and |
|
65 |
SF_def: "TEMP SF(A)_v == TEMP \<box>\<diamond> Enabled(<A>_v) \<longrightarrow> \<box>\<diamond><A>_v" and |
|
60587 | 66 |
aall_def: "TEMP (\<forall>\<forall>x. F x) == TEMP \<not> (\<exists>\<exists>x. \<not> F x)" |
3807 | 67 |
|
47968 | 68 |
axiomatization where |
6255 | 69 |
(* Base axioms for raw TLA. *) |
60588 | 70 |
normalT: "\<And>F G. \<turnstile> \<box>(F \<longrightarrow> G) \<longrightarrow> (\<box>F \<longrightarrow> \<box>G)" and (* polymorphic *) |
71 |
reflT: "\<And>F. \<turnstile> \<box>F \<longrightarrow> F" and (* F::temporal *) |
|
72 |
transT: "\<And>F. \<turnstile> \<box>F \<longrightarrow> \<box>\<box>F" and (* polymorphic *) |
|
60591 | 73 |
linT: "\<And>F G. \<turnstile> \<diamond>F \<and> \<diamond>G \<longrightarrow> (\<diamond>(F \<and> \<diamond>G)) \<or> (\<diamond>(G \<and> \<diamond>F))" and |
74 |
discT: "\<And>F. \<turnstile> \<box>(F \<longrightarrow> \<diamond>(\<not>F \<and> \<diamond>F)) \<longrightarrow> (F \<longrightarrow> \<box>\<diamond>F)" and |
|
60588 | 75 |
primeI: "\<And>P. \<turnstile> \<box>P \<longrightarrow> Init P`" and |
76 |
primeE: "\<And>P F. \<turnstile> \<box>(Init P \<longrightarrow> \<box>F) \<longrightarrow> Init P` \<longrightarrow> (F \<longrightarrow> \<box>F)" and |
|
60591 | 77 |
indT: "\<And>P F. \<turnstile> \<box>(Init P \<and> \<not>\<box>F \<longrightarrow> Init P` \<and> F) \<longrightarrow> Init P \<longrightarrow> \<box>F" and |
60588 | 78 |
allT: "\<And>F. \<turnstile> (\<forall>x. \<box>(F x)) = (\<box>(\<forall> x. F x))" |
3807 | 79 |
|
47968 | 80 |
axiomatization where |
60588 | 81 |
necT: "\<And>F. \<turnstile> F \<Longrightarrow> \<turnstile> \<box>F" (* polymorphic *) |
3807 | 82 |
|
47968 | 83 |
axiomatization where |
3807 | 84 |
(* Flexible quantification: refinement mappings, history variables *) |
60588 | 85 |
eexI: "\<turnstile> F x \<longrightarrow> (\<exists>\<exists>x. F x)" and |
86 |
eexE: "\<lbrakk> sigma \<Turnstile> (\<exists>\<exists>x. F x); basevars vs; |
|
87 |
(\<And>x. \<lbrakk> basevars (x, vs); sigma \<Turnstile> F x \<rbrakk> \<Longrightarrow> (G sigma)::bool) |
|
88 |
\<rbrakk> \<Longrightarrow> G sigma" and |
|
60591 | 89 |
history: "\<turnstile> \<exists>\<exists>h. Init(h = ha) \<and> \<box>(\<forall>x. $h = #x \<longrightarrow> h` = hb x)" |
17309 | 90 |
|
21624 | 91 |
|
92 |
(* Specialize intensional introduction/elimination rules for temporal formulas *) |
|
93 |
||
60588 | 94 |
lemma tempI [intro!]: "(\<And>sigma. sigma \<Turnstile> (F::temporal)) \<Longrightarrow> \<turnstile> F" |
21624 | 95 |
apply (rule intI) |
96 |
apply (erule meta_spec) |
|
97 |
done |
|
98 |
||
60588 | 99 |
lemma tempD [dest]: "\<turnstile> (F::temporal) \<Longrightarrow> sigma \<Turnstile> F" |
21624 | 100 |
by (erule intD) |
101 |
||
102 |
||
103 |
(* ======== Functions to "unlift" temporal theorems ====== *) |
|
104 |
||
60592 | 105 |
ML \<open> |
21624 | 106 |
(* The following functions are specialized versions of the corresponding |
107 |
functions defined in theory Intensional in that they introduce a |
|
108 |
"world" parameter of type "behavior". |
|
109 |
*) |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
110 |
fun temp_unlift ctxt th = |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
111 |
(rewrite_rule ctxt @{thms action_rews} (th RS @{thm tempD})) |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
112 |
handle THM _ => action_unlift ctxt th; |
21624 | 113 |
|
60588 | 114 |
(* Turn \<turnstile> F = G into meta-level rewrite rule F == G *) |
21624 | 115 |
val temp_rewrite = int_rewrite |
116 |
||
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
117 |
fun temp_use ctxt th = |
59582 | 118 |
case Thm.concl_of th of |
69597 | 119 |
Const _ $ (Const (\<^const_name>\<open>Intensional.Valid\<close>, _) $ _) => |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
120 |
((flatten (temp_unlift ctxt th)) handle THM _ => th) |
21624 | 121 |
| _ => th; |
122 |
||
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
123 |
fun try_rewrite ctxt th = temp_rewrite ctxt th handle THM _ => temp_use ctxt th; |
60592 | 124 |
\<close> |
21624 | 125 |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
126 |
attribute_setup temp_unlift = |
61853
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents:
60754
diff
changeset
|
127 |
\<open>Scan.succeed (Thm.rule_attribute [] (temp_unlift o Context.proof_of))\<close> |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
128 |
attribute_setup temp_rewrite = |
61853
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents:
60754
diff
changeset
|
129 |
\<open>Scan.succeed (Thm.rule_attribute [] (temp_rewrite o Context.proof_of))\<close> |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
130 |
attribute_setup temp_use = |
61853
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents:
60754
diff
changeset
|
131 |
\<open>Scan.succeed (Thm.rule_attribute [] (temp_use o Context.proof_of))\<close> |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
132 |
attribute_setup try_rewrite = |
61853
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents:
60754
diff
changeset
|
133 |
\<open>Scan.succeed (Thm.rule_attribute [] (try_rewrite o Context.proof_of))\<close> |
30528 | 134 |
|
21624 | 135 |
|
136 |
(* ------------------------------------------------------------------------- *) |
|
60587 | 137 |
(*** "Simple temporal logic": only \<box> and \<diamond> ***) |
21624 | 138 |
(* ------------------------------------------------------------------------- *) |
139 |
section "Simple temporal logic" |
|
140 |
||
60587 | 141 |
(* \<box>\<not>F == \<box>\<not>Init F *) |
142 |
lemmas boxNotInit = boxInit [of "LIFT \<not>F", unfolded Init_simps] for F |
|
21624 | 143 |
|
60587 | 144 |
lemma dmdInit: "TEMP \<diamond>F == TEMP \<diamond> Init F" |
21624 | 145 |
apply (unfold dmd_def) |
60587 | 146 |
apply (unfold boxInit [of "LIFT \<not>F"]) |
21624 | 147 |
apply (simp (no_asm) add: Init_simps) |
148 |
done |
|
149 |
||
60587 | 150 |
lemmas dmdNotInit = dmdInit [of "LIFT \<not>F", unfolded Init_simps] for F |
21624 | 151 |
|
152 |
(* boxInit and dmdInit cannot be used as rewrites, because they loop. |
|
153 |
Non-looping instances for state predicates and actions are occasionally useful. |
|
154 |
*) |
|
45605 | 155 |
lemmas boxInit_stp = boxInit [where 'a = state] |
156 |
lemmas boxInit_act = boxInit [where 'a = "state * state"] |
|
157 |
lemmas dmdInit_stp = dmdInit [where 'a = state] |
|
158 |
lemmas dmdInit_act = dmdInit [where 'a = "state * state"] |
|
21624 | 159 |
|
160 |
(* The symmetric equations can be used to get rid of Init *) |
|
161 |
lemmas boxInitD = boxInit [symmetric] |
|
162 |
lemmas dmdInitD = dmdInit [symmetric] |
|
163 |
lemmas boxNotInitD = boxNotInit [symmetric] |
|
164 |
lemmas dmdNotInitD = dmdNotInit [symmetric] |
|
165 |
||
166 |
lemmas Init_simps = Init_simps boxInitD dmdInitD boxNotInitD dmdNotInitD |
|
167 |
||
168 |
(* ------------------------ STL2 ------------------------------------------- *) |
|
169 |
lemmas STL2 = reflT |
|
170 |
||
171 |
(* The "polymorphic" (generic) variant *) |
|
60588 | 172 |
lemma STL2_gen: "\<turnstile> \<box>F \<longrightarrow> Init F" |
21624 | 173 |
apply (unfold boxInit [of F]) |
174 |
apply (rule STL2) |
|
175 |
done |
|
176 |
||
60588 | 177 |
(* see also STL2_pr below: "\<turnstile> \<box>P \<longrightarrow> Init P & Init (P`)" *) |
21624 | 178 |
|
179 |
||
60587 | 180 |
(* Dual versions for \<diamond> *) |
60588 | 181 |
lemma InitDmd: "\<turnstile> F \<longrightarrow> \<diamond> F" |
21624 | 182 |
apply (unfold dmd_def) |
183 |
apply (auto dest!: STL2 [temp_use]) |
|
184 |
done |
|
185 |
||
60588 | 186 |
lemma InitDmd_gen: "\<turnstile> Init F \<longrightarrow> \<diamond>F" |
21624 | 187 |
apply clarsimp |
188 |
apply (drule InitDmd [temp_use]) |
|
189 |
apply (simp add: dmdInitD) |
|
190 |
done |
|
191 |
||
192 |
||
193 |
(* ------------------------ STL3 ------------------------------------------- *) |
|
60588 | 194 |
lemma STL3: "\<turnstile> (\<box>\<box>F) = (\<box>F)" |
21624 | 195 |
by (auto elim: transT [temp_use] STL2 [temp_use]) |
196 |
||
197 |
(* corresponding elimination rule introduces double boxes: |
|
60588 | 198 |
\<lbrakk> (sigma \<Turnstile> \<box>F); (sigma \<Turnstile> \<box>\<box>F) \<Longrightarrow> PROP W \<rbrakk> \<Longrightarrow> PROP W |
21624 | 199 |
*) |
200 |
lemmas dup_boxE = STL3 [temp_unlift, THEN iffD2, elim_format] |
|
45605 | 201 |
lemmas dup_boxD = STL3 [temp_unlift, THEN iffD1] |
21624 | 202 |
|
60587 | 203 |
(* dual versions for \<diamond> *) |
60588 | 204 |
lemma DmdDmd: "\<turnstile> (\<diamond>\<diamond>F) = (\<diamond>F)" |
21624 | 205 |
by (auto simp add: dmd_def [try_rewrite] STL3 [try_rewrite]) |
206 |
||
207 |
lemmas dup_dmdE = DmdDmd [temp_unlift, THEN iffD2, elim_format] |
|
45605 | 208 |
lemmas dup_dmdD = DmdDmd [temp_unlift, THEN iffD1] |
21624 | 209 |
|
210 |
||
211 |
(* ------------------------ STL4 ------------------------------------------- *) |
|
212 |
lemma STL4: |
|
60588 | 213 |
assumes "\<turnstile> F \<longrightarrow> G" |
214 |
shows "\<turnstile> \<box>F \<longrightarrow> \<box>G" |
|
21624 | 215 |
apply clarsimp |
216 |
apply (rule normalT [temp_use]) |
|
217 |
apply (rule assms [THEN necT, temp_use]) |
|
218 |
apply assumption |
|
219 |
done |
|
220 |
||
221 |
(* Unlifted version as an elimination rule *) |
|
60588 | 222 |
lemma STL4E: "\<lbrakk> sigma \<Turnstile> \<box>F; \<turnstile> F \<longrightarrow> G \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<box>G" |
21624 | 223 |
by (erule (1) STL4 [temp_use]) |
224 |
||
60588 | 225 |
lemma STL4_gen: "\<turnstile> Init F \<longrightarrow> Init G \<Longrightarrow> \<turnstile> \<box>F \<longrightarrow> \<box>G" |
21624 | 226 |
apply (drule STL4) |
227 |
apply (simp add: boxInitD) |
|
228 |
done |
|
229 |
||
60588 | 230 |
lemma STL4E_gen: "\<lbrakk> sigma \<Turnstile> \<box>F; \<turnstile> Init F \<longrightarrow> Init G \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<box>G" |
21624 | 231 |
by (erule (1) STL4_gen [temp_use]) |
232 |
||
233 |
(* see also STL4Edup below, which allows an auxiliary boxed formula: |
|
60587 | 234 |
\<box>A /\ F => G |
21624 | 235 |
----------------- |
60587 | 236 |
\<box>A /\ \<box>F => \<box>G |
21624 | 237 |
*) |
238 |
||
60587 | 239 |
(* The dual versions for \<diamond> *) |
21624 | 240 |
lemma DmdImpl: |
60588 | 241 |
assumes prem: "\<turnstile> F \<longrightarrow> G" |
242 |
shows "\<turnstile> \<diamond>F \<longrightarrow> \<diamond>G" |
|
21624 | 243 |
apply (unfold dmd_def) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
244 |
apply (fastforce intro!: prem [temp_use] elim!: STL4E [temp_use]) |
21624 | 245 |
done |
246 |
||
60588 | 247 |
lemma DmdImplE: "\<lbrakk> sigma \<Turnstile> \<diamond>F; \<turnstile> F \<longrightarrow> G \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<diamond>G" |
21624 | 248 |
by (erule (1) DmdImpl [temp_use]) |
249 |
||
250 |
(* ------------------------ STL5 ------------------------------------------- *) |
|
60591 | 251 |
lemma STL5: "\<turnstile> (\<box>F \<and> \<box>G) = (\<box>(F \<and> G))" |
21624 | 252 |
apply auto |
60591 | 253 |
apply (subgoal_tac "sigma \<Turnstile> \<box> (G \<longrightarrow> (F \<and> G))") |
21624 | 254 |
apply (erule normalT [temp_use]) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
255 |
apply (fastforce elim!: STL4E [temp_use])+ |
21624 | 256 |
done |
257 |
||
258 |
(* rewrite rule to split conjunctions under boxes *) |
|
45605 | 259 |
lemmas split_box_conj = STL5 [temp_unlift, symmetric] |
21624 | 260 |
|
261 |
||
262 |
(* the corresponding elimination rule allows to combine boxes in the hypotheses |
|
263 |
(NB: F and G must have the same type, i.e., both actions or temporals.) |
|
264 |
Use "addSE2" etc. if you want to add this to a claset, otherwise it will loop! |
|
265 |
*) |
|
266 |
lemma box_conjE: |
|
60588 | 267 |
assumes "sigma \<Turnstile> \<box>F" |
268 |
and "sigma \<Turnstile> \<box>G" |
|
60591 | 269 |
and "sigma \<Turnstile> \<box>(F\<and>G) \<Longrightarrow> PROP R" |
21624 | 270 |
shows "PROP R" |
271 |
by (rule assms STL5 [temp_unlift, THEN iffD1] conjI)+ |
|
272 |
||
273 |
(* Instances of box_conjE for state predicates, actions, and temporals |
|
274 |
in case the general rule is "too polymorphic". |
|
275 |
*) |
|
45605 | 276 |
lemmas box_conjE_temp = box_conjE [where 'a = behavior] |
277 |
lemmas box_conjE_stp = box_conjE [where 'a = state] |
|
278 |
lemmas box_conjE_act = box_conjE [where 'a = "state * state"] |
|
21624 | 279 |
|
280 |
(* Define a tactic that tries to merge all boxes in an antecedent. The definition is |
|
281 |
a bit kludgy in order to simulate "double elim-resolution". |
|
282 |
*) |
|
283 |
||
60588 | 284 |
lemma box_thin: "\<lbrakk> sigma \<Turnstile> \<box>F; PROP W \<rbrakk> \<Longrightarrow> PROP W" . |
21624 | 285 |
|
60592 | 286 |
ML \<open> |
60754 | 287 |
fun merge_box_tac ctxt i = |
288 |
REPEAT_DETERM (EVERY [eresolve_tac ctxt @{thms box_conjE} i, assume_tac ctxt i, |
|
289 |
eresolve_tac ctxt @{thms box_thin} i]) |
|
21624 | 290 |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
26305
diff
changeset
|
291 |
fun merge_temp_box_tac ctxt i = |
60754 | 292 |
REPEAT_DETERM (EVERY [eresolve_tac ctxt @{thms box_conjE_temp} i, assume_tac ctxt i, |
59780 | 293 |
Rule_Insts.eres_inst_tac ctxt [((("'a", 0), Position.none), "behavior")] [] @{thm box_thin} i]) |
21624 | 294 |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
26305
diff
changeset
|
295 |
fun merge_stp_box_tac ctxt i = |
60754 | 296 |
REPEAT_DETERM (EVERY [eresolve_tac ctxt @{thms box_conjE_stp} i, assume_tac ctxt i, |
59780 | 297 |
Rule_Insts.eres_inst_tac ctxt [((("'a", 0), Position.none), "state")] [] @{thm box_thin} i]) |
21624 | 298 |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
26305
diff
changeset
|
299 |
fun merge_act_box_tac ctxt i = |
60754 | 300 |
REPEAT_DETERM (EVERY [eresolve_tac ctxt @{thms box_conjE_act} i, assume_tac ctxt i, |
59780 | 301 |
Rule_Insts.eres_inst_tac ctxt [((("'a", 0), Position.none), "state * state")] [] @{thm box_thin} i]) |
60592 | 302 |
\<close> |
21624 | 303 |
|
60754 | 304 |
method_setup merge_box = \<open>Scan.succeed (SIMPLE_METHOD' o merge_box_tac)\<close> |
60592 | 305 |
method_setup merge_temp_box = \<open>Scan.succeed (SIMPLE_METHOD' o merge_temp_box_tac)\<close> |
306 |
method_setup merge_stp_box = \<open>Scan.succeed (SIMPLE_METHOD' o merge_stp_box_tac)\<close> |
|
307 |
method_setup merge_act_box = \<open>Scan.succeed (SIMPLE_METHOD' o merge_act_box_tac)\<close> |
|
42787 | 308 |
|
21624 | 309 |
(* rewrite rule to push universal quantification through box: |
60588 | 310 |
(sigma \<Turnstile> \<box>(\<forall>x. F x)) = (\<forall>x. (sigma \<Turnstile> \<box>F x)) |
21624 | 311 |
*) |
45605 | 312 |
lemmas all_box = allT [temp_unlift, symmetric] |
21624 | 313 |
|
60591 | 314 |
lemma DmdOr: "\<turnstile> (\<diamond>(F \<or> G)) = (\<diamond>F \<or> \<diamond>G)" |
21624 | 315 |
apply (auto simp add: dmd_def split_box_conj [try_rewrite]) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
316 |
apply (erule contrapos_np, merge_box, fastforce elim!: STL4E [temp_use])+ |
21624 | 317 |
done |
318 |
||
60588 | 319 |
lemma exT: "\<turnstile> (\<exists>x. \<diamond>(F x)) = (\<diamond>(\<exists>x. F x))" |
21624 | 320 |
by (auto simp: dmd_def Not_Rex [try_rewrite] all_box [try_rewrite]) |
321 |
||
45605 | 322 |
lemmas ex_dmd = exT [temp_unlift, symmetric] |
21624 | 323 |
|
60591 | 324 |
lemma STL4Edup: "\<And>sigma. \<lbrakk> sigma \<Turnstile> \<box>A; sigma \<Turnstile> \<box>F; \<turnstile> F \<and> \<box>A \<longrightarrow> G \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<box>G" |
21624 | 325 |
apply (erule dup_boxE) |
42787 | 326 |
apply merge_box |
21624 | 327 |
apply (erule STL4E) |
328 |
apply assumption |
|
329 |
done |
|
330 |
||
60587 | 331 |
lemma DmdImpl2: |
60588 | 332 |
"\<And>sigma. \<lbrakk> sigma \<Turnstile> \<diamond>F; sigma \<Turnstile> \<box>(F \<longrightarrow> G) \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<diamond>G" |
21624 | 333 |
apply (unfold dmd_def) |
334 |
apply auto |
|
335 |
apply (erule notE) |
|
42787 | 336 |
apply merge_box |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
337 |
apply (fastforce elim!: STL4E [temp_use]) |
21624 | 338 |
done |
339 |
||
340 |
lemma InfImpl: |
|
60588 | 341 |
assumes 1: "sigma \<Turnstile> \<box>\<diamond>F" |
342 |
and 2: "sigma \<Turnstile> \<box>G" |
|
60591 | 343 |
and 3: "\<turnstile> F \<and> G \<longrightarrow> H" |
60588 | 344 |
shows "sigma \<Turnstile> \<box>\<diamond>H" |
21624 | 345 |
apply (insert 1 2) |
346 |
apply (erule_tac F = G in dup_boxE) |
|
42787 | 347 |
apply merge_box |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
348 |
apply (fastforce elim!: STL4E [temp_use] DmdImpl2 [temp_use] intro!: 3 [temp_use]) |
21624 | 349 |
done |
350 |
||
351 |
(* ------------------------ STL6 ------------------------------------------- *) |
|
352 |
(* Used in the proof of STL6, but useful in itself. *) |
|
60591 | 353 |
lemma BoxDmd: "\<turnstile> \<box>F \<and> \<diamond>G \<longrightarrow> \<diamond>(\<box>F \<and> G)" |
21624 | 354 |
apply (unfold dmd_def) |
355 |
apply clarsimp |
|
356 |
apply (erule dup_boxE) |
|
42787 | 357 |
apply merge_box |
21624 | 358 |
apply (erule contrapos_np) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
359 |
apply (fastforce elim!: STL4E [temp_use]) |
21624 | 360 |
done |
361 |
||
362 |
(* weaker than BoxDmd, but more polymorphic (and often just right) *) |
|
60591 | 363 |
lemma BoxDmd_simple: "\<turnstile> \<box>F \<and> \<diamond>G \<longrightarrow> \<diamond>(F \<and> G)" |
21624 | 364 |
apply (unfold dmd_def) |
365 |
apply clarsimp |
|
42787 | 366 |
apply merge_box |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
367 |
apply (fastforce elim!: notE STL4E [temp_use]) |
21624 | 368 |
done |
369 |
||
60591 | 370 |
lemma BoxDmd2_simple: "\<turnstile> \<box>F \<and> \<diamond>G \<longrightarrow> \<diamond>(G \<and> F)" |
21624 | 371 |
apply (unfold dmd_def) |
372 |
apply clarsimp |
|
42787 | 373 |
apply merge_box |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
374 |
apply (fastforce elim!: notE STL4E [temp_use]) |
21624 | 375 |
done |
376 |
||
377 |
lemma DmdImpldup: |
|
60588 | 378 |
assumes 1: "sigma \<Turnstile> \<box>A" |
379 |
and 2: "sigma \<Turnstile> \<diamond>F" |
|
60591 | 380 |
and 3: "\<turnstile> \<box>A \<and> F \<longrightarrow> G" |
60588 | 381 |
shows "sigma \<Turnstile> \<diamond>G" |
21624 | 382 |
apply (rule 2 [THEN 1 [THEN BoxDmd [temp_use]], THEN DmdImplE]) |
383 |
apply (rule 3) |
|
384 |
done |
|
385 |
||
60591 | 386 |
lemma STL6: "\<turnstile> \<diamond>\<box>F \<and> \<diamond>\<box>G \<longrightarrow> \<diamond>\<box>(F \<and> G)" |
21624 | 387 |
apply (auto simp: STL5 [temp_rewrite, symmetric]) |
388 |
apply (drule linT [temp_use]) |
|
389 |
apply assumption |
|
390 |
apply (erule thin_rl) |
|
391 |
apply (rule DmdDmd [temp_unlift, THEN iffD1]) |
|
392 |
apply (erule disjE) |
|
393 |
apply (erule DmdImplE) |
|
394 |
apply (rule BoxDmd) |
|
395 |
apply (erule DmdImplE) |
|
396 |
apply auto |
|
397 |
apply (drule BoxDmd [temp_use]) |
|
398 |
apply assumption |
|
399 |
apply (erule thin_rl) |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
400 |
apply (fastforce elim!: DmdImplE [temp_use]) |
21624 | 401 |
done |
402 |
||
403 |
||
404 |
(* ------------------------ True / False ----------------------------------------- *) |
|
405 |
section "Simplification of constants" |
|
406 |
||
60588 | 407 |
lemma BoxConst: "\<turnstile> (\<box>#P) = #P" |
21624 | 408 |
apply (rule tempI) |
409 |
apply (cases P) |
|
410 |
apply (auto intro!: necT [temp_use] dest: STL2_gen [temp_use] simp: Init_simps) |
|
411 |
done |
|
412 |
||
60588 | 413 |
lemma DmdConst: "\<turnstile> (\<diamond>#P) = #P" |
21624 | 414 |
apply (unfold dmd_def) |
415 |
apply (cases P) |
|
416 |
apply (simp_all add: BoxConst [try_rewrite]) |
|
417 |
done |
|
418 |
||
419 |
lemmas temp_simps [temp_rewrite, simp] = BoxConst DmdConst |
|
420 |
||
421 |
||
422 |
(* ------------------------ Further rewrites ----------------------------------------- *) |
|
423 |
section "Further rewrites" |
|
424 |
||
60588 | 425 |
lemma NotBox: "\<turnstile> (\<not>\<box>F) = (\<diamond>\<not>F)" |
21624 | 426 |
by (simp add: dmd_def) |
427 |
||
60588 | 428 |
lemma NotDmd: "\<turnstile> (\<not>\<diamond>F) = (\<box>\<not>F)" |
21624 | 429 |
by (simp add: dmd_def) |
430 |
||
431 |
(* These are not declared by default, because they could be harmful, |
|
60587 | 432 |
e.g. \<box>F & \<not>\<box>F becomes \<box>F & \<diamond>\<not>F !! *) |
26305 | 433 |
lemmas more_temp_simps1 = |
21624 | 434 |
STL3 [temp_rewrite] DmdDmd [temp_rewrite] NotBox [temp_rewrite] NotDmd [temp_rewrite] |
435 |
NotBox [temp_unlift, THEN eq_reflection] |
|
436 |
NotDmd [temp_unlift, THEN eq_reflection] |
|
437 |
||
60588 | 438 |
lemma BoxDmdBox: "\<turnstile> (\<box>\<diamond>\<box>F) = (\<diamond>\<box>F)" |
21624 | 439 |
apply (auto dest!: STL2 [temp_use]) |
440 |
apply (rule ccontr) |
|
60591 | 441 |
apply (subgoal_tac "sigma \<Turnstile> \<diamond>\<box>\<box>F \<and> \<diamond>\<box>\<not>\<box>F") |
21624 | 442 |
apply (erule thin_rl) |
443 |
apply auto |
|
444 |
apply (drule STL6 [temp_use]) |
|
445 |
apply assumption |
|
446 |
apply simp |
|
26305 | 447 |
apply (simp_all add: more_temp_simps1) |
21624 | 448 |
done |
449 |
||
60588 | 450 |
lemma DmdBoxDmd: "\<turnstile> (\<diamond>\<box>\<diamond>F) = (\<box>\<diamond>F)" |
21624 | 451 |
apply (unfold dmd_def) |
452 |
apply (auto simp: BoxDmdBox [unfolded dmd_def, try_rewrite]) |
|
453 |
done |
|
454 |
||
26305 | 455 |
lemmas more_temp_simps2 = more_temp_simps1 BoxDmdBox [temp_rewrite] DmdBoxDmd [temp_rewrite] |
21624 | 456 |
|
457 |
||
458 |
(* ------------------------ Miscellaneous ----------------------------------- *) |
|
459 |
||
60591 | 460 |
lemma BoxOr: "\<And>sigma. \<lbrakk> sigma \<Turnstile> \<box>F \<or> \<box>G \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<box>(F \<or> G)" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
461 |
by (fastforce elim!: STL4E [temp_use]) |
21624 | 462 |
|
463 |
(* "persistently implies infinitely often" *) |
|
60588 | 464 |
lemma DBImplBD: "\<turnstile> \<diamond>\<box>F \<longrightarrow> \<box>\<diamond>F" |
21624 | 465 |
apply clarsimp |
466 |
apply (rule ccontr) |
|
26305 | 467 |
apply (simp add: more_temp_simps2) |
21624 | 468 |
apply (drule STL6 [temp_use]) |
469 |
apply assumption |
|
470 |
apply simp |
|
471 |
done |
|
472 |
||
60591 | 473 |
lemma BoxDmdDmdBox: "\<turnstile> \<box>\<diamond>F \<and> \<diamond>\<box>G \<longrightarrow> \<box>\<diamond>(F \<and> G)" |
21624 | 474 |
apply clarsimp |
475 |
apply (rule ccontr) |
|
26305 | 476 |
apply (unfold more_temp_simps2) |
21624 | 477 |
apply (drule STL6 [temp_use]) |
478 |
apply assumption |
|
60588 | 479 |
apply (subgoal_tac "sigma \<Turnstile> \<diamond>\<box>\<not>F") |
21624 | 480 |
apply (force simp: dmd_def) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
481 |
apply (fastforce elim: DmdImplE [temp_use] STL4E [temp_use]) |
21624 | 482 |
done |
483 |
||
484 |
||
485 |
(* ------------------------------------------------------------------------- *) |
|
486 |
(*** TLA-specific theorems: primed formulas ***) |
|
487 |
(* ------------------------------------------------------------------------- *) |
|
488 |
section "priming" |
|
489 |
||
490 |
(* ------------------------ TLA2 ------------------------------------------- *) |
|
60591 | 491 |
lemma STL2_pr: "\<turnstile> \<box>P \<longrightarrow> Init P \<and> Init P`" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
492 |
by (fastforce intro!: STL2_gen [temp_use] primeI [temp_use]) |
21624 | 493 |
|
494 |
(* Auxiliary lemma allows priming of boxed actions *) |
|
60591 | 495 |
lemma BoxPrime: "\<turnstile> \<box>P \<longrightarrow> \<box>($P \<and> P$)" |
21624 | 496 |
apply clarsimp |
497 |
apply (erule dup_boxE) |
|
498 |
apply (unfold boxInit_act) |
|
499 |
apply (erule STL4E) |
|
500 |
apply (auto simp: Init_simps dest!: STL2_pr [temp_use]) |
|
501 |
done |
|
502 |
||
503 |
lemma TLA2: |
|
60591 | 504 |
assumes "\<turnstile> $P \<and> P$ \<longrightarrow> A" |
60588 | 505 |
shows "\<turnstile> \<box>P \<longrightarrow> \<box>A" |
21624 | 506 |
apply clarsimp |
507 |
apply (drule BoxPrime [temp_use]) |
|
41529 | 508 |
apply (auto simp: Init_stp_act_rev [try_rewrite] intro!: assms [temp_use] |
21624 | 509 |
elim!: STL4E [temp_use]) |
510 |
done |
|
511 |
||
60591 | 512 |
lemma TLA2E: "\<lbrakk> sigma \<Turnstile> \<box>P; \<turnstile> $P \<and> P$ \<longrightarrow> A \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<box>A" |
21624 | 513 |
by (erule (1) TLA2 [temp_use]) |
514 |
||
60588 | 515 |
lemma DmdPrime: "\<turnstile> (\<diamond>P`) \<longrightarrow> (\<diamond>P)" |
21624 | 516 |
apply (unfold dmd_def) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
517 |
apply (fastforce elim!: TLA2E [temp_use]) |
21624 | 518 |
done |
519 |
||
45605 | 520 |
lemmas PrimeDmd = InitDmd_gen [temp_use, THEN DmdPrime [temp_use]] |
21624 | 521 |
|
522 |
(* ------------------------ INV1, stable --------------------------------------- *) |
|
523 |
section "stable, invariant" |
|
524 |
||
525 |
lemma ind_rule: |
|
60591 | 526 |
"\<lbrakk> sigma \<Turnstile> \<box>H; sigma \<Turnstile> Init P; \<turnstile> H \<longrightarrow> (Init P \<and> \<not>\<box>F \<longrightarrow> Init(P`) \<and> F) \<rbrakk> |
60588 | 527 |
\<Longrightarrow> sigma \<Turnstile> \<box>F" |
21624 | 528 |
apply (rule indT [temp_use]) |
529 |
apply (erule (2) STL4E) |
|
530 |
done |
|
531 |
||
60588 | 532 |
lemma box_stp_act: "\<turnstile> (\<box>$P) = (\<box>P)" |
21624 | 533 |
by (simp add: boxInit_act Init_simps) |
534 |
||
45605 | 535 |
lemmas box_stp_actI = box_stp_act [temp_use, THEN iffD2] |
536 |
lemmas box_stp_actD = box_stp_act [temp_use, THEN iffD1] |
|
21624 | 537 |
|
26305 | 538 |
lemmas more_temp_simps3 = box_stp_act [temp_rewrite] more_temp_simps2 |
21624 | 539 |
|
60587 | 540 |
lemma INV1: |
60588 | 541 |
"\<turnstile> (Init P) \<longrightarrow> (stable P) \<longrightarrow> \<box>P" |
21624 | 542 |
apply (unfold stable_def boxInit_stp boxInit_act) |
543 |
apply clarsimp |
|
544 |
apply (erule ind_rule) |
|
545 |
apply (auto simp: Init_simps elim: ind_rule) |
|
546 |
done |
|
547 |
||
60587 | 548 |
lemma StableT: |
60591 | 549 |
"\<And>P. \<turnstile> $P \<and> A \<longrightarrow> P` \<Longrightarrow> \<turnstile> \<box>A \<longrightarrow> stable P" |
21624 | 550 |
apply (unfold stable_def) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
551 |
apply (fastforce elim!: STL4E [temp_use]) |
21624 | 552 |
done |
553 |
||
60591 | 554 |
lemma Stable: "\<lbrakk> sigma \<Turnstile> \<box>A; \<turnstile> $P \<and> A \<longrightarrow> P` \<rbrakk> \<Longrightarrow> sigma \<Turnstile> stable P" |
21624 | 555 |
by (erule (1) StableT [temp_use]) |
556 |
||
557 |
(* Generalization of INV1 *) |
|
60588 | 558 |
lemma StableBox: "\<turnstile> (stable P) \<longrightarrow> \<box>(Init P \<longrightarrow> \<box>P)" |
21624 | 559 |
apply (unfold stable_def) |
560 |
apply clarsimp |
|
561 |
apply (erule dup_boxE) |
|
562 |
apply (force simp: stable_def elim: STL4E [temp_use] INV1 [temp_use]) |
|
563 |
done |
|
564 |
||
60591 | 565 |
lemma DmdStable: "\<turnstile> (stable P) \<and> \<diamond>P \<longrightarrow> \<diamond>\<box>P" |
21624 | 566 |
apply clarsimp |
567 |
apply (rule DmdImpl2) |
|
568 |
prefer 2 |
|
569 |
apply (erule StableBox [temp_use]) |
|
570 |
apply (simp add: dmdInitD) |
|
571 |
done |
|
572 |
||
573 |
(* ---------------- (Semi-)automatic invariant tactics ---------------------- *) |
|
574 |
||
60592 | 575 |
ML \<open> |
60588 | 576 |
(* inv_tac reduces goals of the form ... \<Longrightarrow> sigma \<Turnstile> \<box>P *) |
42793 | 577 |
fun inv_tac ctxt = |
578 |
SELECT_GOAL |
|
579 |
(EVERY |
|
580 |
[auto_tac ctxt, |
|
60754 | 581 |
TRY (merge_box_tac ctxt 1), |
582 |
resolve_tac ctxt [temp_use ctxt @{thm INV1}] 1, (* fail if the goal is not a box *) |
|
583 |
TRYALL (eresolve_tac ctxt @{thms Stable})]); |
|
21624 | 584 |
|
585 |
(* auto_inv_tac applies inv_tac and then tries to attack the subgoals |
|
60588 | 586 |
in simple cases it may be able to handle goals like \<turnstile> MyProg \<longrightarrow> \<box>Inv. |
21624 | 587 |
In these simple cases the simplifier seems to be more useful than the |
588 |
auto-tactic, which applies too much propositional logic and simplifies |
|
589 |
too late. |
|
590 |
*) |
|
42803 | 591 |
fun auto_inv_tac ctxt = |
42793 | 592 |
SELECT_GOAL |
42803 | 593 |
(inv_tac ctxt 1 THEN |
42793 | 594 |
(TRYALL (action_simp_tac |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51668
diff
changeset
|
595 |
(ctxt addsimps [@{thm Init_stp}, @{thm Init_act}]) [] [@{thm squareE}]))); |
60592 | 596 |
\<close> |
21624 | 597 |
|
60592 | 598 |
method_setup invariant = \<open> |
42793 | 599 |
Method.sections Clasimp.clasimp_modifiers >> (K (SIMPLE_METHOD' o inv_tac)) |
60592 | 600 |
\<close> |
42769 | 601 |
|
60592 | 602 |
method_setup auto_invariant = \<open> |
42803 | 603 |
Method.sections Clasimp.clasimp_modifiers >> (K (SIMPLE_METHOD' o auto_inv_tac)) |
60592 | 604 |
\<close> |
42769 | 605 |
|
60591 | 606 |
lemma unless: "\<turnstile> \<box>($P \<longrightarrow> P` \<or> Q`) \<longrightarrow> (stable P) \<or> \<diamond>Q" |
21624 | 607 |
apply (unfold dmd_def) |
608 |
apply (clarsimp dest!: BoxPrime [temp_use]) |
|
42787 | 609 |
apply merge_box |
21624 | 610 |
apply (erule contrapos_np) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
611 |
apply (fastforce elim!: Stable [temp_use]) |
21624 | 612 |
done |
613 |
||
614 |
||
615 |
(* --------------------- Recursive expansions --------------------------------------- *) |
|
616 |
section "recursive expansions" |
|
617 |
||
60587 | 618 |
(* Recursive expansions of \<box> and \<diamond> for state predicates *) |
60591 | 619 |
lemma BoxRec: "\<turnstile> (\<box>P) = (Init P \<and> \<box>P`)" |
21624 | 620 |
apply (auto intro!: STL2_gen [temp_use]) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
621 |
apply (fastforce elim!: TLA2E [temp_use]) |
21624 | 622 |
apply (auto simp: stable_def elim!: INV1 [temp_use] STL4E [temp_use]) |
623 |
done |
|
624 |
||
60591 | 625 |
lemma DmdRec: "\<turnstile> (\<diamond>P) = (Init P \<or> \<diamond>P`)" |
21624 | 626 |
apply (unfold dmd_def BoxRec [temp_rewrite]) |
627 |
apply (auto simp: Init_simps) |
|
628 |
done |
|
629 |
||
60588 | 630 |
lemma DmdRec2: "\<And>sigma. \<lbrakk> sigma \<Turnstile> \<diamond>P; sigma \<Turnstile> \<box>\<not>P` \<rbrakk> \<Longrightarrow> sigma \<Turnstile> Init P" |
21624 | 631 |
apply (force simp: DmdRec [temp_rewrite] dmd_def) |
632 |
done |
|
633 |
||
60588 | 634 |
lemma InfinitePrime: "\<turnstile> (\<box>\<diamond>P) = (\<box>\<diamond>P`)" |
21624 | 635 |
apply auto |
636 |
apply (rule classical) |
|
637 |
apply (rule DBImplBD [temp_use]) |
|
60588 | 638 |
apply (subgoal_tac "sigma \<Turnstile> \<diamond>\<box>P") |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
639 |
apply (fastforce elim!: DmdImplE [temp_use] TLA2E [temp_use]) |
60591 | 640 |
apply (subgoal_tac "sigma \<Turnstile> \<diamond>\<box> (\<diamond>P \<and> \<box>\<not>P`)") |
21624 | 641 |
apply (force simp: boxInit_stp [temp_use] |
642 |
elim!: DmdImplE [temp_use] STL4E [temp_use] DmdRec2 [temp_use]) |
|
26305 | 643 |
apply (force intro!: STL6 [temp_use] simp: more_temp_simps3) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
644 |
apply (fastforce intro: DmdPrime [temp_use] elim!: STL4E [temp_use]) |
21624 | 645 |
done |
646 |
||
647 |
lemma InfiniteEnsures: |
|
60591 | 648 |
"\<lbrakk> sigma \<Turnstile> \<box>N; sigma \<Turnstile> \<box>\<diamond>A; \<turnstile> A \<and> N \<longrightarrow> P` \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<box>\<diamond>P" |
21624 | 649 |
apply (unfold InfinitePrime [temp_rewrite]) |
650 |
apply (rule InfImpl) |
|
651 |
apply assumption+ |
|
652 |
done |
|
653 |
||
654 |
(* ------------------------ fairness ------------------------------------------- *) |
|
655 |
section "fairness" |
|
656 |
||
657 |
(* alternative definitions of fairness *) |
|
60591 | 658 |
lemma WF_alt: "\<turnstile> WF(A)_v = (\<box>\<diamond>\<not>Enabled(<A>_v) \<or> \<box>\<diamond><A>_v)" |
21624 | 659 |
apply (unfold WF_def dmd_def) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
660 |
apply fastforce |
21624 | 661 |
done |
662 |
||
60591 | 663 |
lemma SF_alt: "\<turnstile> SF(A)_v = (\<diamond>\<box>\<not>Enabled(<A>_v) \<or> \<box>\<diamond><A>_v)" |
21624 | 664 |
apply (unfold SF_def dmd_def) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
665 |
apply fastforce |
21624 | 666 |
done |
667 |
||
668 |
(* theorems to "box" fairness conditions *) |
|
60588 | 669 |
lemma BoxWFI: "\<turnstile> WF(A)_v \<longrightarrow> \<box>WF(A)_v" |
26305 | 670 |
by (auto simp: WF_alt [try_rewrite] more_temp_simps3 intro!: BoxOr [temp_use]) |
21624 | 671 |
|
60588 | 672 |
lemma WF_Box: "\<turnstile> (\<box>WF(A)_v) = WF(A)_v" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
673 |
by (fastforce intro!: BoxWFI [temp_use] dest!: STL2 [temp_use]) |
21624 | 674 |
|
60588 | 675 |
lemma BoxSFI: "\<turnstile> SF(A)_v \<longrightarrow> \<box>SF(A)_v" |
26305 | 676 |
by (auto simp: SF_alt [try_rewrite] more_temp_simps3 intro!: BoxOr [temp_use]) |
21624 | 677 |
|
60588 | 678 |
lemma SF_Box: "\<turnstile> (\<box>SF(A)_v) = SF(A)_v" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
679 |
by (fastforce intro!: BoxSFI [temp_use] dest!: STL2 [temp_use]) |
21624 | 680 |
|
26305 | 681 |
lemmas more_temp_simps = more_temp_simps3 WF_Box [temp_rewrite] SF_Box [temp_rewrite] |
21624 | 682 |
|
60588 | 683 |
lemma SFImplWF: "\<turnstile> SF(A)_v \<longrightarrow> WF(A)_v" |
21624 | 684 |
apply (unfold SF_def WF_def) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
685 |
apply (fastforce dest!: DBImplBD [temp_use]) |
21624 | 686 |
done |
687 |
||
688 |
(* A tactic that "boxes" all fairness conditions. Apply more_temp_simps to "unbox". *) |
|
60592 | 689 |
ML \<open> |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58889
diff
changeset
|
690 |
fun box_fair_tac ctxt = |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58889
diff
changeset
|
691 |
SELECT_GOAL (REPEAT (dresolve_tac ctxt [@{thm BoxWFI}, @{thm BoxSFI}] 1)) |
60592 | 692 |
\<close> |
21624 | 693 |
|
694 |
||
695 |
(* ------------------------------ leads-to ------------------------------ *) |
|
696 |
||
60587 | 697 |
section "\<leadsto>" |
21624 | 698 |
|
60591 | 699 |
lemma leadsto_init: "\<turnstile> (Init F) \<and> (F \<leadsto> G) \<longrightarrow> \<diamond>G" |
21624 | 700 |
apply (unfold leadsto_def) |
701 |
apply (auto dest!: STL2 [temp_use]) |
|
702 |
done |
|
703 |
||
60588 | 704 |
(* \<turnstile> F & (F \<leadsto> G) \<longrightarrow> \<diamond>G *) |
45605 | 705 |
lemmas leadsto_init_temp = leadsto_init [where 'a = behavior, unfolded Init_simps] |
21624 | 706 |
|
60588 | 707 |
lemma streett_leadsto: "\<turnstile> (\<box>\<diamond>Init F \<longrightarrow> \<box>\<diamond>G) = (\<diamond>(F \<leadsto> G))" |
21624 | 708 |
apply (unfold leadsto_def) |
709 |
apply auto |
|
710 |
apply (simp add: more_temp_simps) |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
711 |
apply (fastforce elim!: DmdImplE [temp_use] STL4E [temp_use]) |
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
712 |
apply (fastforce intro!: InitDmd [temp_use] elim!: STL4E [temp_use]) |
60588 | 713 |
apply (subgoal_tac "sigma \<Turnstile> \<box>\<diamond>\<diamond>G") |
21624 | 714 |
apply (simp add: more_temp_simps) |
715 |
apply (drule BoxDmdDmdBox [temp_use]) |
|
716 |
apply assumption |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
717 |
apply (fastforce elim!: DmdImplE [temp_use] STL4E [temp_use]) |
21624 | 718 |
done |
719 |
||
60591 | 720 |
lemma leadsto_infinite: "\<turnstile> \<box>\<diamond>F \<and> (F \<leadsto> G) \<longrightarrow> \<box>\<diamond>G" |
21624 | 721 |
apply clarsimp |
722 |
apply (erule InitDmd [temp_use, THEN streett_leadsto [temp_unlift, THEN iffD2, THEN mp]]) |
|
723 |
apply (simp add: dmdInitD) |
|
724 |
done |
|
725 |
||
726 |
(* In particular, strong fairness is a Streett condition. The following |
|
727 |
rules are sometimes easier to use than WF2 or SF2 below. |
|
728 |
*) |
|
60588 | 729 |
lemma leadsto_SF: "\<turnstile> (Enabled(<A>_v) \<leadsto> <A>_v) \<longrightarrow> SF(A)_v" |
21624 | 730 |
apply (unfold SF_def) |
731 |
apply (clarsimp elim!: leadsto_infinite [temp_use]) |
|
732 |
done |
|
733 |
||
60588 | 734 |
lemma leadsto_WF: "\<turnstile> (Enabled(<A>_v) \<leadsto> <A>_v) \<longrightarrow> WF(A)_v" |
21624 | 735 |
by (clarsimp intro!: SFImplWF [temp_use] leadsto_SF [temp_use]) |
736 |
||
737 |
(* introduce an invariant into the proof of a leadsto assertion. |
|
60588 | 738 |
\<box>I \<longrightarrow> ((P \<leadsto> Q) = (P /\ I \<leadsto> Q)) |
21624 | 739 |
*) |
60591 | 740 |
lemma INV_leadsto: "\<turnstile> \<box>I \<and> (P \<and> I \<leadsto> Q) \<longrightarrow> (P \<leadsto> Q)" |
21624 | 741 |
apply (unfold leadsto_def) |
742 |
apply clarsimp |
|
743 |
apply (erule STL4Edup) |
|
744 |
apply assumption |
|
745 |
apply (auto simp: Init_simps dest!: STL2_gen [temp_use]) |
|
746 |
done |
|
747 |
||
60591 | 748 |
lemma leadsto_classical: "\<turnstile> (Init F \<and> \<box>\<not>G \<leadsto> G) \<longrightarrow> (F \<leadsto> G)" |
21624 | 749 |
apply (unfold leadsto_def dmd_def) |
750 |
apply (force simp: Init_simps elim!: STL4E [temp_use]) |
|
751 |
done |
|
752 |
||
60591 | 753 |
lemma leadsto_false: "\<turnstile> (F \<leadsto> #False) = (\<box>\<not>F)" |
21624 | 754 |
apply (unfold leadsto_def) |
755 |
apply (simp add: boxNotInitD) |
|
756 |
done |
|
757 |
||
60588 | 758 |
lemma leadsto_exists: "\<turnstile> ((\<exists>x. F x) \<leadsto> G) = (\<forall>x. (F x \<leadsto> G))" |
21624 | 759 |
apply (unfold leadsto_def) |
760 |
apply (auto simp: allT [try_rewrite] Init_simps elim!: STL4E [temp_use]) |
|
761 |
done |
|
762 |
||
763 |
(* basic leadsto properties, cf. Unity *) |
|
764 |
||
60588 | 765 |
lemma ImplLeadsto_gen: "\<turnstile> \<box>(Init F \<longrightarrow> Init G) \<longrightarrow> (F \<leadsto> G)" |
21624 | 766 |
apply (unfold leadsto_def) |
767 |
apply (auto intro!: InitDmd_gen [temp_use] |
|
768 |
elim!: STL4E_gen [temp_use] simp: Init_simps) |
|
769 |
done |
|
770 |
||
45605 | 771 |
lemmas ImplLeadsto = |
772 |
ImplLeadsto_gen [where 'a = behavior and 'b = behavior, unfolded Init_simps] |
|
21624 | 773 |
|
60588 | 774 |
lemma ImplLeadsto_simple: "\<And>F G. \<turnstile> F \<longrightarrow> G \<Longrightarrow> \<turnstile> F \<leadsto> G" |
21624 | 775 |
by (auto simp: Init_def intro!: ImplLeadsto_gen [temp_use] necT [temp_use]) |
776 |
||
777 |
lemma EnsuresLeadsto: |
|
60591 | 778 |
assumes "\<turnstile> A \<and> $P \<longrightarrow> Q`" |
60588 | 779 |
shows "\<turnstile> \<box>A \<longrightarrow> (P \<leadsto> Q)" |
21624 | 780 |
apply (unfold leadsto_def) |
781 |
apply (clarsimp elim!: INV_leadsto [temp_use]) |
|
782 |
apply (erule STL4E_gen) |
|
783 |
apply (auto simp: Init_defs intro!: PrimeDmd [temp_use] assms [temp_use]) |
|
784 |
done |
|
785 |
||
60588 | 786 |
lemma EnsuresLeadsto2: "\<turnstile> \<box>($P \<longrightarrow> Q`) \<longrightarrow> (P \<leadsto> Q)" |
21624 | 787 |
apply (unfold leadsto_def) |
788 |
apply clarsimp |
|
789 |
apply (erule STL4E_gen) |
|
790 |
apply (auto simp: Init_simps intro!: PrimeDmd [temp_use]) |
|
791 |
done |
|
792 |
||
793 |
lemma ensures: |
|
60591 | 794 |
assumes 1: "\<turnstile> $P \<and> N \<longrightarrow> P` \<or> Q`" |
795 |
and 2: "\<turnstile> ($P \<and> N) \<and> A \<longrightarrow> Q`" |
|
796 |
shows "\<turnstile> \<box>N \<and> \<box>(\<box>P \<longrightarrow> \<diamond>A) \<longrightarrow> (P \<leadsto> Q)" |
|
21624 | 797 |
apply (unfold leadsto_def) |
798 |
apply clarsimp |
|
799 |
apply (erule STL4Edup) |
|
800 |
apply assumption |
|
801 |
apply clarsimp |
|
60591 | 802 |
apply (subgoal_tac "sigmaa \<Turnstile> \<box>($P \<longrightarrow> P` \<or> Q`) ") |
21624 | 803 |
apply (drule unless [temp_use]) |
804 |
apply (clarsimp dest!: INV1 [temp_use]) |
|
805 |
apply (rule 2 [THEN DmdImpl, temp_use, THEN DmdPrime [temp_use]]) |
|
806 |
apply (force intro!: BoxDmd_simple [temp_use] |
|
807 |
simp: split_box_conj [try_rewrite] box_stp_act [try_rewrite]) |
|
808 |
apply (force elim: STL4E [temp_use] dest: 1 [temp_use]) |
|
809 |
done |
|
810 |
||
811 |
lemma ensures_simple: |
|
60591 | 812 |
"\<lbrakk> \<turnstile> $P \<and> N \<longrightarrow> P` \<or> Q`; |
813 |
\<turnstile> ($P \<and> N) \<and> A \<longrightarrow> Q` |
|
814 |
\<rbrakk> \<Longrightarrow> \<turnstile> \<box>N \<and> \<box>\<diamond>A \<longrightarrow> (P \<leadsto> Q)" |
|
21624 | 815 |
apply clarsimp |
816 |
apply (erule (2) ensures [temp_use]) |
|
817 |
apply (force elim!: STL4E [temp_use]) |
|
818 |
done |
|
819 |
||
820 |
lemma EnsuresInfinite: |
|
60591 | 821 |
"\<lbrakk> sigma \<Turnstile> \<box>\<diamond>P; sigma \<Turnstile> \<box>A; \<turnstile> A \<and> $P \<longrightarrow> Q` \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<box>\<diamond>Q" |
21624 | 822 |
apply (erule leadsto_infinite [temp_use]) |
823 |
apply (erule EnsuresLeadsto [temp_use]) |
|
824 |
apply assumption |
|
825 |
done |
|
826 |
||
827 |
||
828 |
(*** Gronning's lattice rules (taken from TLP) ***) |
|
829 |
section "Lattice rules" |
|
830 |
||
60588 | 831 |
lemma LatticeReflexivity: "\<turnstile> F \<leadsto> F" |
21624 | 832 |
apply (unfold leadsto_def) |
833 |
apply (rule necT InitDmd_gen)+ |
|
834 |
done |
|
835 |
||
60591 | 836 |
lemma LatticeTransitivity: "\<turnstile> (G \<leadsto> H) \<and> (F \<leadsto> G) \<longrightarrow> (F \<leadsto> H)" |
21624 | 837 |
apply (unfold leadsto_def) |
838 |
apply clarsimp |
|
60588 | 839 |
apply (erule dup_boxE) (* \<box>\<box>(Init G \<longrightarrow> H) *) |
42787 | 840 |
apply merge_box |
21624 | 841 |
apply (clarsimp elim!: STL4E [temp_use]) |
842 |
apply (rule dup_dmdD) |
|
60588 | 843 |
apply (subgoal_tac "sigmaa \<Turnstile> \<diamond>Init G") |
21624 | 844 |
apply (erule DmdImpl2) |
845 |
apply assumption |
|
846 |
apply (simp add: dmdInitD) |
|
847 |
done |
|
848 |
||
60591 | 849 |
lemma LatticeDisjunctionElim1: "\<turnstile> (F \<or> G \<leadsto> H) \<longrightarrow> (F \<leadsto> H)" |
21624 | 850 |
apply (unfold leadsto_def) |
851 |
apply (auto simp: Init_simps elim!: STL4E [temp_use]) |
|
852 |
done |
|
853 |
||
60591 | 854 |
lemma LatticeDisjunctionElim2: "\<turnstile> (F \<or> G \<leadsto> H) \<longrightarrow> (G \<leadsto> H)" |
21624 | 855 |
apply (unfold leadsto_def) |
856 |
apply (auto simp: Init_simps elim!: STL4E [temp_use]) |
|
857 |
done |
|
858 |
||
60591 | 859 |
lemma LatticeDisjunctionIntro: "\<turnstile> (F \<leadsto> H) \<and> (G \<leadsto> H) \<longrightarrow> (F \<or> G \<leadsto> H)" |
21624 | 860 |
apply (unfold leadsto_def) |
861 |
apply clarsimp |
|
42787 | 862 |
apply merge_box |
21624 | 863 |
apply (auto simp: Init_simps elim!: STL4E [temp_use]) |
864 |
done |
|
865 |
||
60591 | 866 |
lemma LatticeDisjunction: "\<turnstile> (F \<or> G \<leadsto> H) = ((F \<leadsto> H) \<and> (G \<leadsto> H))" |
21624 | 867 |
by (auto intro: LatticeDisjunctionIntro [temp_use] |
868 |
LatticeDisjunctionElim1 [temp_use] |
|
869 |
LatticeDisjunctionElim2 [temp_use]) |
|
870 |
||
60591 | 871 |
lemma LatticeDiamond: "\<turnstile> (A \<leadsto> B \<or> C) \<and> (B \<leadsto> D) \<and> (C \<leadsto> D) \<longrightarrow> (A \<leadsto> D)" |
21624 | 872 |
apply clarsimp |
60591 | 873 |
apply (subgoal_tac "sigma \<Turnstile> (B \<or> C) \<leadsto> D") |
874 |
apply (erule_tac G = "LIFT (B \<or> C)" in LatticeTransitivity [temp_use]) |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42814
diff
changeset
|
875 |
apply (fastforce intro!: LatticeDisjunctionIntro [temp_use])+ |
21624 | 876 |
done |
877 |
||
60591 | 878 |
lemma LatticeTriangle: "\<turnstile> (A \<leadsto> D \<or> B) \<and> (B \<leadsto> D) \<longrightarrow> (A \<leadsto> D)" |
21624 | 879 |
apply clarsimp |
60591 | 880 |
apply (subgoal_tac "sigma \<Turnstile> (D \<or> B) \<leadsto> D") |
881 |
apply (erule_tac G = "LIFT (D \<or> B)" in LatticeTransitivity [temp_use]) |
|
21624 | 882 |
apply assumption |
883 |
apply (auto intro: LatticeDisjunctionIntro [temp_use] LatticeReflexivity [temp_use]) |
|
884 |
done |
|
885 |
||
60591 | 886 |
lemma LatticeTriangle2: "\<turnstile> (A \<leadsto> B \<or> D) \<and> (B \<leadsto> D) \<longrightarrow> (A \<leadsto> D)" |
21624 | 887 |
apply clarsimp |
60591 | 888 |
apply (subgoal_tac "sigma \<Turnstile> B \<or> D \<leadsto> D") |
889 |
apply (erule_tac G = "LIFT (B \<or> D)" in LatticeTransitivity [temp_use]) |
|
21624 | 890 |
apply assumption |
891 |
apply (auto intro: LatticeDisjunctionIntro [temp_use] LatticeReflexivity [temp_use]) |
|
892 |
done |
|
893 |
||
894 |
(*** Lamport's fairness rules ***) |
|
895 |
section "Fairness rules" |
|
896 |
||
897 |
lemma WF1: |
|
60591 | 898 |
"\<lbrakk> \<turnstile> $P \<and> N \<longrightarrow> P` \<or> Q`; |
899 |
\<turnstile> ($P \<and> N) \<and> <A>_v \<longrightarrow> Q`; |
|
900 |
\<turnstile> $P \<and> N \<longrightarrow> $(Enabled(<A>_v)) \<rbrakk> |
|
901 |
\<Longrightarrow> \<turnstile> \<box>N \<and> WF(A)_v \<longrightarrow> (P \<leadsto> Q)" |
|
21624 | 902 |
apply (clarsimp dest!: BoxWFI [temp_use]) |
903 |
apply (erule (2) ensures [temp_use]) |
|
904 |
apply (erule (1) STL4Edup) |
|
905 |
apply (clarsimp simp: WF_def) |
|
906 |
apply (rule STL2 [temp_use]) |
|
907 |
apply (clarsimp elim!: mp intro!: InitDmd [temp_use]) |
|
908 |
apply (erule STL4 [temp_use, THEN box_stp_actD [temp_use]]) |
|
909 |
apply (simp add: split_box_conj box_stp_actI) |
|
910 |
done |
|
911 |
||
912 |
(* Sometimes easier to use; designed for action B rather than state predicate Q *) |
|
913 |
lemma WF_leadsto: |
|
60591 | 914 |
assumes 1: "\<turnstile> N \<and> $P \<longrightarrow> $Enabled (<A>_v)" |
915 |
and 2: "\<turnstile> N \<and> <A>_v \<longrightarrow> B" |
|
916 |
and 3: "\<turnstile> \<box>(N \<and> [\<not>A]_v) \<longrightarrow> stable P" |
|
917 |
shows "\<turnstile> \<box>N \<and> WF(A)_v \<longrightarrow> (P \<leadsto> B)" |
|
21624 | 918 |
apply (unfold leadsto_def) |
919 |
apply (clarsimp dest!: BoxWFI [temp_use]) |
|
920 |
apply (erule (1) STL4Edup) |
|
921 |
apply clarsimp |
|
922 |
apply (rule 2 [THEN DmdImpl, temp_use]) |
|
923 |
apply (rule BoxDmd_simple [temp_use]) |
|
924 |
apply assumption |
|
925 |
apply (rule classical) |
|
926 |
apply (rule STL2 [temp_use]) |
|
927 |
apply (clarsimp simp: WF_def elim!: mp intro!: InitDmd [temp_use]) |
|
928 |
apply (rule 1 [THEN STL4, temp_use, THEN box_stp_actD]) |
|
929 |
apply (simp (no_asm_simp) add: split_box_conj [try_rewrite] box_stp_act [try_rewrite]) |
|
930 |
apply (erule INV1 [temp_use]) |
|
931 |
apply (rule 3 [temp_use]) |
|
932 |
apply (simp add: split_box_conj [try_rewrite] NotDmd [temp_use] not_angle [try_rewrite]) |
|
933 |
done |
|
934 |
||
935 |
lemma SF1: |
|
60591 | 936 |
"\<lbrakk> \<turnstile> $P \<and> N \<longrightarrow> P` \<or> Q`; |
937 |
\<turnstile> ($P \<and> N) \<and> <A>_v \<longrightarrow> Q`; |
|
938 |
\<turnstile> \<box>P \<and> \<box>N \<and> \<box>F \<longrightarrow> \<diamond>Enabled(<A>_v) \<rbrakk> |
|
939 |
\<Longrightarrow> \<turnstile> \<box>N \<and> SF(A)_v \<and> \<box>F \<longrightarrow> (P \<leadsto> Q)" |
|
21624 | 940 |
apply (clarsimp dest!: BoxSFI [temp_use]) |
941 |
apply (erule (2) ensures [temp_use]) |
|
942 |
apply (erule_tac F = F in dup_boxE) |
|
42787 | 943 |
apply merge_temp_box |
21624 | 944 |
apply (erule STL4Edup) |
945 |
apply assumption |
|
946 |
apply (clarsimp simp: SF_def) |
|
947 |
apply (rule STL2 [temp_use]) |
|
948 |
apply (erule mp) |
|
949 |
apply (erule STL4 [temp_use]) |
|
950 |
apply (simp add: split_box_conj [try_rewrite] STL3 [try_rewrite]) |
|
951 |
done |
|
952 |
||
953 |
lemma WF2: |
|
60591 | 954 |
assumes 1: "\<turnstile> N \<and> <B>_f \<longrightarrow> <M>_g" |
955 |
and 2: "\<turnstile> $P \<and> P` \<and> <N \<and> A>_f \<longrightarrow> B" |
|
956 |
and 3: "\<turnstile> P \<and> Enabled(<M>_g) \<longrightarrow> Enabled(<A>_f)" |
|
957 |
and 4: "\<turnstile> \<box>(N \<and> [\<not>B]_f) \<and> WF(A)_f \<and> \<box>F \<and> \<diamond>\<box>Enabled(<M>_g) \<longrightarrow> \<diamond>\<box>P" |
|
958 |
shows "\<turnstile> \<box>N \<and> WF(A)_f \<and> \<box>F \<longrightarrow> WF(M)_g" |
|
21624 | 959 |
apply (clarsimp dest!: BoxWFI [temp_use] BoxDmdBox [temp_use, THEN iffD2] |
960 |
simp: WF_def [where A = M]) |
|
961 |
apply (erule_tac F = F in dup_boxE) |
|
42787 | 962 |
apply merge_temp_box |
21624 | 963 |
apply (erule STL4Edup) |
964 |
apply assumption |
|
965 |
apply (clarsimp intro!: BoxDmd_simple [temp_use, THEN 1 [THEN DmdImpl, temp_use]]) |
|
966 |
apply (rule classical) |
|
60591 | 967 |
apply (subgoal_tac "sigmaa \<Turnstile> \<diamond> (($P \<and> P` \<and> N) \<and> <A>_f)") |
21624 | 968 |
apply (force simp: angle_def intro!: 2 [temp_use] elim!: DmdImplE [temp_use]) |
969 |
apply (rule BoxDmd_simple [THEN DmdImpl, unfolded DmdDmd [temp_rewrite], temp_use]) |
|
970 |
apply (simp add: NotDmd [temp_use] not_angle [try_rewrite]) |
|
42787 | 971 |
apply merge_act_box |
21624 | 972 |
apply (frule 4 [temp_use]) |
973 |
apply assumption+ |
|
974 |
apply (drule STL6 [temp_use]) |
|
975 |
apply assumption |
|
60588 | 976 |
apply (erule_tac V = "sigmaa \<Turnstile> \<diamond>\<box>P" in thin_rl) |
977 |
apply (erule_tac V = "sigmaa \<Turnstile> \<box>F" in thin_rl) |
|
21624 | 978 |
apply (drule BoxWFI [temp_use]) |
60591 | 979 |
apply (erule_tac F = "ACT N \<and> [\<not>B]_f" in dup_boxE) |
42787 | 980 |
apply merge_temp_box |
21624 | 981 |
apply (erule DmdImpldup) |
982 |
apply assumption |
|
983 |
apply (auto simp: split_box_conj [try_rewrite] STL3 [try_rewrite] |
|
984 |
WF_Box [try_rewrite] box_stp_act [try_rewrite]) |
|
985 |
apply (force elim!: TLA2E [where P = P, temp_use]) |
|
986 |
apply (rule STL2 [temp_use]) |
|
987 |
apply (force simp: WF_def split_box_conj [try_rewrite] |
|
988 |
elim!: mp intro!: InitDmd [temp_use] 3 [THEN STL4, temp_use]) |
|
989 |
done |
|
990 |
||
991 |
lemma SF2: |
|
60591 | 992 |
assumes 1: "\<turnstile> N \<and> <B>_f \<longrightarrow> <M>_g" |
993 |
and 2: "\<turnstile> $P \<and> P` \<and> <N \<and> A>_f \<longrightarrow> B" |
|
994 |
and 3: "\<turnstile> P \<and> Enabled(<M>_g) \<longrightarrow> Enabled(<A>_f)" |
|
995 |
and 4: "\<turnstile> \<box>(N \<and> [\<not>B]_f) \<and> SF(A)_f \<and> \<box>F \<and> \<box>\<diamond>Enabled(<M>_g) \<longrightarrow> \<diamond>\<box>P" |
|
996 |
shows "\<turnstile> \<box>N \<and> SF(A)_f \<and> \<box>F \<longrightarrow> SF(M)_g" |
|
21624 | 997 |
apply (clarsimp dest!: BoxSFI [temp_use] simp: 2 [try_rewrite] SF_def [where A = M]) |
998 |
apply (erule_tac F = F in dup_boxE) |
|
60587 | 999 |
apply (erule_tac F = "TEMP \<diamond>Enabled (<M>_g) " in dup_boxE) |
42787 | 1000 |
apply merge_temp_box |
21624 | 1001 |
apply (erule STL4Edup) |
1002 |
apply assumption |
|
1003 |
apply (clarsimp intro!: BoxDmd_simple [temp_use, THEN 1 [THEN DmdImpl, temp_use]]) |
|
1004 |
apply (rule classical) |
|
60591 | 1005 |
apply (subgoal_tac "sigmaa \<Turnstile> \<diamond> (($P \<and> P` \<and> N) \<and> <A>_f)") |
21624 | 1006 |
apply (force simp: angle_def intro!: 2 [temp_use] elim!: DmdImplE [temp_use]) |
1007 |
apply (rule BoxDmd_simple [THEN DmdImpl, unfolded DmdDmd [temp_rewrite], temp_use]) |
|
1008 |
apply (simp add: NotDmd [temp_use] not_angle [try_rewrite]) |
|
42787 | 1009 |
apply merge_act_box |
21624 | 1010 |
apply (frule 4 [temp_use]) |
1011 |
apply assumption+ |
|
60588 | 1012 |
apply (erule_tac V = "sigmaa \<Turnstile> \<box>F" in thin_rl) |
21624 | 1013 |
apply (drule BoxSFI [temp_use]) |
60587 | 1014 |
apply (erule_tac F = "TEMP \<diamond>Enabled (<M>_g)" in dup_boxE) |
60591 | 1015 |
apply (erule_tac F = "ACT N \<and> [\<not>B]_f" in dup_boxE) |
42787 | 1016 |
apply merge_temp_box |
21624 | 1017 |
apply (erule DmdImpldup) |
1018 |
apply assumption |
|
1019 |
apply (auto simp: split_box_conj [try_rewrite] STL3 [try_rewrite] |
|
1020 |
SF_Box [try_rewrite] box_stp_act [try_rewrite]) |
|
1021 |
apply (force elim!: TLA2E [where P = P, temp_use]) |
|
1022 |
apply (rule STL2 [temp_use]) |
|
1023 |
apply (force simp: SF_def split_box_conj [try_rewrite] |
|
1024 |
elim!: mp InfImpl [temp_use] intro!: 3 [temp_use]) |
|
1025 |
done |
|
1026 |
||
1027 |
(* ------------------------------------------------------------------------- *) |
|
1028 |
(*** Liveness proofs by well-founded orderings ***) |
|
1029 |
(* ------------------------------------------------------------------------- *) |
|
1030 |
section "Well-founded orderings" |
|
1031 |
||
1032 |
lemma wf_leadsto: |
|
1033 |
assumes 1: "wf r" |
|
60591 | 1034 |
and 2: "\<And>x. sigma \<Turnstile> F x \<leadsto> (G \<or> (\<exists>y. #((y,x)\<in>r) \<and> F y)) " |
60588 | 1035 |
shows "sigma \<Turnstile> F x \<leadsto> G" |
21624 | 1036 |
apply (rule 1 [THEN wf_induct]) |
1037 |
apply (rule LatticeTriangle [temp_use]) |
|
1038 |
apply (rule 2) |
|
1039 |
apply (auto simp: leadsto_exists [try_rewrite]) |
|
60591 | 1040 |
apply (case_tac "(y,x) \<in> r") |
21624 | 1041 |
apply force |
1042 |
apply (force simp: leadsto_def Init_simps intro!: necT [temp_use]) |
|
1043 |
done |
|
1044 |
||
1045 |
(* If r is well-founded, state function v cannot decrease forever *) |
|
60591 | 1046 |
lemma wf_not_box_decrease: "\<And>r. wf r \<Longrightarrow> \<turnstile> \<box>[ (v`, $v) \<in> #r ]_v \<longrightarrow> \<diamond>\<box>[#False]_v" |
21624 | 1047 |
apply clarsimp |
1048 |
apply (rule ccontr) |
|
60588 | 1049 |
apply (subgoal_tac "sigma \<Turnstile> (\<exists>x. v=#x) \<leadsto> #False") |
21624 | 1050 |
apply (drule leadsto_false [temp_use, THEN iffD1, THEN STL2_gen [temp_use]]) |
1051 |
apply (force simp: Init_defs) |
|
1052 |
apply (clarsimp simp: leadsto_exists [try_rewrite] not_square [try_rewrite] more_temp_simps) |
|
1053 |
apply (erule wf_leadsto) |
|
1054 |
apply (rule ensures_simple [temp_use]) |
|
1055 |
apply (auto simp: square_def angle_def) |
|
1056 |
done |
|
1057 |
||
60588 | 1058 |
(* "wf r \<Longrightarrow> \<turnstile> \<diamond>\<box>[ (v`, $v) : #r ]_v \<longrightarrow> \<diamond>\<box>[#False]_v" *) |
21624 | 1059 |
lemmas wf_not_dmd_box_decrease = |
45605 | 1060 |
wf_not_box_decrease [THEN DmdImpl, unfolded more_temp_simps] |
21624 | 1061 |
|
1062 |
(* If there are infinitely many steps where v decreases, then there |
|
1063 |
have to be infinitely many non-stuttering steps where v doesn't decrease. |
|
1064 |
*) |
|
1065 |
lemma wf_box_dmd_decrease: |
|
1066 |
assumes 1: "wf r" |
|
60591 | 1067 |
shows "\<turnstile> \<box>\<diamond>((v`, $v) \<in> #r) \<longrightarrow> \<box>\<diamond><(v`, $v) \<notin> #r>_v" |
21624 | 1068 |
apply clarsimp |
1069 |
apply (rule ccontr) |
|
1070 |
apply (simp add: not_angle [try_rewrite] more_temp_simps) |
|
1071 |
apply (drule 1 [THEN wf_not_dmd_box_decrease [temp_use]]) |
|
1072 |
apply (drule BoxDmdDmdBox [temp_use]) |
|
1073 |
apply assumption |
|
60588 | 1074 |
apply (subgoal_tac "sigma \<Turnstile> \<box>\<diamond> ((#False) ::action)") |
21624 | 1075 |
apply force |
1076 |
apply (erule STL4E) |
|
1077 |
apply (rule DmdImpl) |
|
1078 |
apply (force intro: 1 [THEN wf_irrefl, temp_use]) |
|
1079 |
done |
|
1080 |
||
1081 |
(* In particular, for natural numbers, if n decreases infinitely often |
|
1082 |
then it has to increase infinitely often. |
|
1083 |
*) |
|
60588 | 1084 |
lemma nat_box_dmd_decrease: "\<And>n::nat stfun. \<turnstile> \<box>\<diamond>(n` < $n) \<longrightarrow> \<box>\<diamond>($n < n`)" |
21624 | 1085 |
apply clarsimp |
60591 | 1086 |
apply (subgoal_tac "sigma \<Turnstile> \<box>\<diamond><\<not> ((n`,$n) \<in> #less_than)>_n") |
21624 | 1087 |
apply (erule thin_rl) |
1088 |
apply (erule STL4E) |
|
1089 |
apply (rule DmdImpl) |
|
1090 |
apply (clarsimp simp: angle_def [try_rewrite]) |
|
1091 |
apply (rule wf_box_dmd_decrease [temp_use]) |
|
1092 |
apply (auto elim!: STL4E [temp_use] DmdImplE [temp_use]) |
|
1093 |
done |
|
1094 |
||
1095 |
||
1096 |
(* ------------------------------------------------------------------------- *) |
|
1097 |
(*** Flexible quantification over state variables ***) |
|
1098 |
(* ------------------------------------------------------------------------- *) |
|
1099 |
section "Flexible quantification" |
|
1100 |
||
1101 |
lemma aallI: |
|
1102 |
assumes 1: "basevars vs" |
|
60588 | 1103 |
and 2: "(\<And>x. basevars (x,vs) \<Longrightarrow> sigma \<Turnstile> F x)" |
1104 |
shows "sigma \<Turnstile> (\<forall>\<forall>x. F x)" |
|
21624 | 1105 |
by (auto simp: aall_def elim!: eexE [temp_use] intro!: 1 dest!: 2 [temp_use]) |
1106 |
||
60588 | 1107 |
lemma aallE: "\<turnstile> (\<forall>\<forall>x. F x) \<longrightarrow> F x" |
21624 | 1108 |
apply (unfold aall_def) |
1109 |
apply clarsimp |
|
1110 |
apply (erule contrapos_np) |
|
1111 |
apply (force intro!: eexI [temp_use]) |
|
1112 |
done |
|
1113 |
||
1114 |
(* monotonicity of quantification *) |
|
1115 |
lemma eex_mono: |
|
60588 | 1116 |
assumes 1: "sigma \<Turnstile> \<exists>\<exists>x. F x" |
1117 |
and 2: "\<And>x. sigma \<Turnstile> F x \<longrightarrow> G x" |
|
1118 |
shows "sigma \<Turnstile> \<exists>\<exists>x. G x" |
|
21624 | 1119 |
apply (rule unit_base [THEN 1 [THEN eexE]]) |
1120 |
apply (rule eexI [temp_use]) |
|
1121 |
apply (erule 2 [unfolded intensional_rews, THEN mp]) |
|
1122 |
done |
|
1123 |
||
1124 |
lemma aall_mono: |
|
60588 | 1125 |
assumes 1: "sigma \<Turnstile> \<forall>\<forall>x. F(x)" |
1126 |
and 2: "\<And>x. sigma \<Turnstile> F(x) \<longrightarrow> G(x)" |
|
1127 |
shows "sigma \<Turnstile> \<forall>\<forall>x. G(x)" |
|
21624 | 1128 |
apply (rule unit_base [THEN aallI]) |
1129 |
apply (rule 2 [unfolded intensional_rews, THEN mp]) |
|
1130 |
apply (rule 1 [THEN aallE [temp_use]]) |
|
1131 |
done |
|
1132 |
||
1133 |
(* Derived history introduction rule *) |
|
1134 |
lemma historyI: |
|
60588 | 1135 |
assumes 1: "sigma \<Turnstile> Init I" |
1136 |
and 2: "sigma \<Turnstile> \<box>N" |
|
21624 | 1137 |
and 3: "basevars vs" |
60591 | 1138 |
and 4: "\<And>h. basevars(h,vs) \<Longrightarrow> \<turnstile> I \<and> h = ha \<longrightarrow> HI h" |
60588 | 1139 |
and 5: "\<And>h s t. \<lbrakk> basevars(h,vs); N (s,t); h t = hb (h s) (s,t) \<rbrakk> \<Longrightarrow> HN h (s,t)" |
60591 | 1140 |
shows "sigma \<Turnstile> \<exists>\<exists>h. Init (HI h) \<and> \<box>(HN h)" |
21624 | 1141 |
apply (rule history [temp_use, THEN eexE]) |
1142 |
apply (rule 3) |
|
1143 |
apply (rule eexI [temp_use]) |
|
1144 |
apply clarsimp |
|
1145 |
apply (rule conjI) |
|
1146 |
prefer 2 |
|
1147 |
apply (insert 2) |
|
42787 | 1148 |
apply merge_box |
21624 | 1149 |
apply (force elim!: STL4E [temp_use] 5 [temp_use]) |
1150 |
apply (insert 1) |
|
1151 |
apply (force simp: Init_defs elim!: 4 [temp_use]) |
|
1152 |
done |
|
1153 |
||
1154 |
(* ---------------------------------------------------------------------- |
|
1155 |
example of a history variable: existence of a clock |
|
1156 |
*) |
|
1157 |
||
60591 | 1158 |
lemma "\<turnstile> \<exists>\<exists>h. Init(h = #True) \<and> \<box>(h` = (\<not>$h))" |
21624 | 1159 |
apply (rule tempI) |
1160 |
apply (rule historyI) |
|
1161 |
apply (force simp: Init_defs intro!: unit_base [temp_use] necT [temp_use])+ |
|
1162 |
done |
|
1163 |
||
1164 |
end |