proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
clarified tool context in some boundary cases;
1 (* Title: HOL/TLA/TLA.thy
3 Copyright: 1998 University of Munich
6 header {* The temporal level of TLA *}
13 (** abstract syntax **)
14 Box :: "('w::world) form => temporal"
15 Dmd :: "('w::world) form => temporal"
16 leadsto :: "['w::world form, 'v::world form] => temporal"
17 Stable :: "stpred => temporal"
18 WF :: "[action, 'a stfun] => temporal"
19 SF :: "[action, 'a stfun] => temporal"
21 (* Quantification over (flexible) state variables *)
22 EEx :: "('a stfun => temporal) => temporal" (binder "Eex " 10)
23 AAll :: "('a stfun => temporal) => temporal" (binder "Aall " 10)
25 (** concrete syntax **)
27 "_Box" :: "lift => lift" ("([]_)" [40] 40)
28 "_Dmd" :: "lift => lift" ("(<>_)" [40] 40)
29 "_leadsto" :: "[lift,lift] => lift" ("(_ ~> _)" [23,22] 22)
30 "_stable" :: "lift => lift" ("(stable/ _)")
31 "_WF" :: "[lift,lift] => lift" ("(WF'(_')'_(_))" [0,60] 55)
32 "_SF" :: "[lift,lift] => lift" ("(SF'(_')'_(_))" [0,60] 55)
34 "_EEx" :: "[idts, lift] => lift" ("(3EEX _./ _)" [0,10] 10)
35 "_AAll" :: "[idts, lift] => lift" ("(3AALL _./ _)" [0,10] 10)
40 "_leadsto" == "CONST leadsto"
41 "_stable" == "CONST Stable"
44 "_EEx v A" == "Eex v. A"
45 "_AAll v A" == "Aall v. A"
47 "sigma |= []F" <= "_Box F sigma"
48 "sigma |= <>F" <= "_Dmd F sigma"
49 "sigma |= F ~> G" <= "_leadsto F G sigma"
50 "sigma |= stable P" <= "_stable P sigma"
51 "sigma |= WF(A)_v" <= "_WF A v sigma"
52 "sigma |= SF(A)_v" <= "_SF A v sigma"
53 "sigma |= EEX x. F" <= "_EEx x F sigma"
54 "sigma |= AALL x. F" <= "_AAll x F sigma"
57 "_Box" :: "lift => lift" ("(\<box>_)" [40] 40)
58 "_Dmd" :: "lift => lift" ("(\<diamond>_)" [40] 40)
59 "_leadsto" :: "[lift,lift] => lift" ("(_ \<leadsto> _)" [23,22] 22)
60 "_EEx" :: "[idts, lift] => lift" ("(3\<exists>\<exists> _./ _)" [0,10] 10)
61 "_AAll" :: "[idts, lift] => lift" ("(3\<forall>\<forall> _./ _)" [0,10] 10)
64 "_EEx" :: "[idts, lift] => lift" ("(3\<exists>\<exists> _./ _)" [0,10] 10)
65 "_AAll" :: "[idts, lift] => lift" ("(3\<forall>\<forall> _./ _)" [0,10] 10)
68 (* Definitions of derived operators *)
69 dmd_def: "\<And>F. TEMP <>F == TEMP ~[]~F"
72 boxInit: "\<And>F. TEMP []F == TEMP []Init F" and
73 leadsto_def: "\<And>F G. TEMP F ~> G == TEMP [](Init F --> <>G)" and
74 stable_def: "\<And>P. TEMP stable P == TEMP []($P --> P$)" and
75 WF_def: "TEMP WF(A)_v == TEMP <>[] Enabled(<A>_v) --> []<><A>_v" and
76 SF_def: "TEMP SF(A)_v == TEMP []<> Enabled(<A>_v) --> []<><A>_v" and
77 aall_def: "TEMP (AALL x. F x) == TEMP ~ (EEX x. ~ F x)"
80 (* Base axioms for raw TLA. *)
81 normalT: "\<And>F G. |- [](F --> G) --> ([]F --> []G)" and (* polymorphic *)
82 reflT: "\<And>F. |- []F --> F" and (* F::temporal *)
83 transT: "\<And>F. |- []F --> [][]F" and (* polymorphic *)
84 linT: "\<And>F G. |- <>F & <>G --> (<>(F & <>G)) | (<>(G & <>F))" and
85 discT: "\<And>F. |- [](F --> <>(~F & <>F)) --> (F --> []<>F)" and
86 primeI: "\<And>P. |- []P --> Init P`" and
87 primeE: "\<And>P F. |- [](Init P --> []F) --> Init P` --> (F --> []F)" and
88 indT: "\<And>P F. |- [](Init P & ~[]F --> Init P` & F) --> Init P --> []F" and
89 allT: "\<And>F. |- (ALL x. [](F x)) = ([](ALL x. F x))"
92 necT: "\<And>F. |- F ==> |- []F" (* polymorphic *)
95 (* Flexible quantification: refinement mappings, history variables *)
96 eexI: "|- F x --> (EEX x. F x)" and
97 eexE: "[| sigma |= (EEX x. F x); basevars vs;
98 (!!x. [| basevars (x, vs); sigma |= F x |] ==> (G sigma)::bool)
100 history: "|- EEX h. Init(h = ha) & [](!x. $h = #x --> h` = hb x)"
103 (* Specialize intensional introduction/elimination rules for temporal formulas *)
105 lemma tempI [intro!]: "(!!sigma. sigma |= (F::temporal)) ==> |- F"
107 apply (erule meta_spec)
110 lemma tempD [dest]: "|- (F::temporal) ==> sigma |= F"
114 (* ======== Functions to "unlift" temporal theorems ====== *)
117 (* The following functions are specialized versions of the corresponding
118 functions defined in theory Intensional in that they introduce a
119 "world" parameter of type "behavior".
121 fun temp_unlift ctxt th =
122 (rewrite_rule ctxt @{thms action_rews} (th RS @{thm tempD}))
123 handle THM _ => action_unlift ctxt th;
125 (* Turn |- F = G into meta-level rewrite rule F == G *)
126 val temp_rewrite = int_rewrite
128 fun temp_use ctxt th =
129 case (concl_of th) of
130 Const _ $ (Const (@{const_name Intensional.Valid}, _) $ _) =>
131 ((flatten (temp_unlift ctxt th)) handle THM _ => th)
134 fun try_rewrite ctxt th = temp_rewrite ctxt th handle THM _ => temp_use ctxt th;
137 attribute_setup temp_unlift =
138 {* Scan.succeed (Thm.rule_attribute (temp_unlift o Context.proof_of)) *}
139 attribute_setup temp_rewrite =
140 {* Scan.succeed (Thm.rule_attribute (temp_rewrite o Context.proof_of)) *}
141 attribute_setup temp_use =
142 {* Scan.succeed (Thm.rule_attribute (temp_use o Context.proof_of)) *}
143 attribute_setup try_rewrite =
144 {* Scan.succeed (Thm.rule_attribute (try_rewrite o Context.proof_of)) *}
147 (* ------------------------------------------------------------------------- *)
148 (*** "Simple temporal logic": only [] and <> ***)
149 (* ------------------------------------------------------------------------- *)
150 section "Simple temporal logic"
152 (* []~F == []~Init F *)
153 lemmas boxNotInit = boxInit [of "LIFT ~F", unfolded Init_simps] for F
155 lemma dmdInit: "TEMP <>F == TEMP <> Init F"
156 apply (unfold dmd_def)
157 apply (unfold boxInit [of "LIFT ~F"])
158 apply (simp (no_asm) add: Init_simps)
161 lemmas dmdNotInit = dmdInit [of "LIFT ~F", unfolded Init_simps] for F
163 (* boxInit and dmdInit cannot be used as rewrites, because they loop.
164 Non-looping instances for state predicates and actions are occasionally useful.
166 lemmas boxInit_stp = boxInit [where 'a = state]
167 lemmas boxInit_act = boxInit [where 'a = "state * state"]
168 lemmas dmdInit_stp = dmdInit [where 'a = state]
169 lemmas dmdInit_act = dmdInit [where 'a = "state * state"]
171 (* The symmetric equations can be used to get rid of Init *)
172 lemmas boxInitD = boxInit [symmetric]
173 lemmas dmdInitD = dmdInit [symmetric]
174 lemmas boxNotInitD = boxNotInit [symmetric]
175 lemmas dmdNotInitD = dmdNotInit [symmetric]
177 lemmas Init_simps = Init_simps boxInitD dmdInitD boxNotInitD dmdNotInitD
179 (* ------------------------ STL2 ------------------------------------------- *)
182 (* The "polymorphic" (generic) variant *)
183 lemma STL2_gen: "|- []F --> Init F"
184 apply (unfold boxInit [of F])
188 (* see also STL2_pr below: "|- []P --> Init P & Init (P`)" *)
191 (* Dual versions for <> *)
192 lemma InitDmd: "|- F --> <> F"
193 apply (unfold dmd_def)
194 apply (auto dest!: STL2 [temp_use])
197 lemma InitDmd_gen: "|- Init F --> <>F"
199 apply (drule InitDmd [temp_use])
200 apply (simp add: dmdInitD)
204 (* ------------------------ STL3 ------------------------------------------- *)
205 lemma STL3: "|- ([][]F) = ([]F)"
206 by (auto elim: transT [temp_use] STL2 [temp_use])
208 (* corresponding elimination rule introduces double boxes:
209 [| (sigma |= []F); (sigma |= [][]F) ==> PROP W |] ==> PROP W
211 lemmas dup_boxE = STL3 [temp_unlift, THEN iffD2, elim_format]
212 lemmas dup_boxD = STL3 [temp_unlift, THEN iffD1]
214 (* dual versions for <> *)
215 lemma DmdDmd: "|- (<><>F) = (<>F)"
216 by (auto simp add: dmd_def [try_rewrite] STL3 [try_rewrite])
218 lemmas dup_dmdE = DmdDmd [temp_unlift, THEN iffD2, elim_format]
219 lemmas dup_dmdD = DmdDmd [temp_unlift, THEN iffD1]
222 (* ------------------------ STL4 ------------------------------------------- *)
225 shows "|- []F --> []G"
227 apply (rule normalT [temp_use])
228 apply (rule assms [THEN necT, temp_use])
232 (* Unlifted version as an elimination rule *)
233 lemma STL4E: "[| sigma |= []F; |- F --> G |] ==> sigma |= []G"
234 by (erule (1) STL4 [temp_use])
236 lemma STL4_gen: "|- Init F --> Init G ==> |- []F --> []G"
238 apply (simp add: boxInitD)
241 lemma STL4E_gen: "[| sigma |= []F; |- Init F --> Init G |] ==> sigma |= []G"
242 by (erule (1) STL4_gen [temp_use])
244 (* see also STL4Edup below, which allows an auxiliary boxed formula:
250 (* The dual versions for <> *)
252 assumes prem: "|- F --> G"
253 shows "|- <>F --> <>G"
254 apply (unfold dmd_def)
255 apply (fastforce intro!: prem [temp_use] elim!: STL4E [temp_use])
258 lemma DmdImplE: "[| sigma |= <>F; |- F --> G |] ==> sigma |= <>G"
259 by (erule (1) DmdImpl [temp_use])
261 (* ------------------------ STL5 ------------------------------------------- *)
262 lemma STL5: "|- ([]F & []G) = ([](F & G))"
264 apply (subgoal_tac "sigma |= [] (G --> (F & G))")
265 apply (erule normalT [temp_use])
266 apply (fastforce elim!: STL4E [temp_use])+
269 (* rewrite rule to split conjunctions under boxes *)
270 lemmas split_box_conj = STL5 [temp_unlift, symmetric]
273 (* the corresponding elimination rule allows to combine boxes in the hypotheses
274 (NB: F and G must have the same type, i.e., both actions or temporals.)
275 Use "addSE2" etc. if you want to add this to a claset, otherwise it will loop!
278 assumes "sigma |= []F"
280 and "sigma |= [](F&G) ==> PROP R"
282 by (rule assms STL5 [temp_unlift, THEN iffD1] conjI)+
284 (* Instances of box_conjE for state predicates, actions, and temporals
285 in case the general rule is "too polymorphic".
287 lemmas box_conjE_temp = box_conjE [where 'a = behavior]
288 lemmas box_conjE_stp = box_conjE [where 'a = state]
289 lemmas box_conjE_act = box_conjE [where 'a = "state * state"]
291 (* Define a tactic that tries to merge all boxes in an antecedent. The definition is
292 a bit kludgy in order to simulate "double elim-resolution".
295 lemma box_thin: "[| sigma |= []F; PROP W |] ==> PROP W" .
298 fun merge_box_tac i =
299 REPEAT_DETERM (EVERY [etac @{thm box_conjE} i, atac i, etac @{thm box_thin} i])
301 fun merge_temp_box_tac ctxt i =
302 REPEAT_DETERM (EVERY [etac @{thm box_conjE_temp} i, atac i,
303 eres_inst_tac ctxt [(("'a", 0), "behavior")] @{thm box_thin} i])
305 fun merge_stp_box_tac ctxt i =
306 REPEAT_DETERM (EVERY [etac @{thm box_conjE_stp} i, atac i,
307 eres_inst_tac ctxt [(("'a", 0), "state")] @{thm box_thin} i])
309 fun merge_act_box_tac ctxt i =
310 REPEAT_DETERM (EVERY [etac @{thm box_conjE_act} i, atac i,
311 eres_inst_tac ctxt [(("'a", 0), "state * state")] @{thm box_thin} i])
314 method_setup merge_box = {* Scan.succeed (K (SIMPLE_METHOD' merge_box_tac)) *}
315 method_setup merge_temp_box = {* Scan.succeed (SIMPLE_METHOD' o merge_temp_box_tac) *}
316 method_setup merge_stp_box = {* Scan.succeed (SIMPLE_METHOD' o merge_stp_box_tac) *}
317 method_setup merge_act_box = {* Scan.succeed (SIMPLE_METHOD' o merge_act_box_tac) *}
319 (* rewrite rule to push universal quantification through box:
320 (sigma |= [](! x. F x)) = (! x. (sigma |= []F x))
322 lemmas all_box = allT [temp_unlift, symmetric]
324 lemma DmdOr: "|- (<>(F | G)) = (<>F | <>G)"
325 apply (auto simp add: dmd_def split_box_conj [try_rewrite])
326 apply (erule contrapos_np, merge_box, fastforce elim!: STL4E [temp_use])+
329 lemma exT: "|- (EX x. <>(F x)) = (<>(EX x. F x))"
330 by (auto simp: dmd_def Not_Rex [try_rewrite] all_box [try_rewrite])
332 lemmas ex_dmd = exT [temp_unlift, symmetric]
334 lemma STL4Edup: "!!sigma. [| sigma |= []A; sigma |= []F; |- F & []A --> G |] ==> sigma |= []G"
335 apply (erule dup_boxE)
342 "!!sigma. [| sigma |= <>F; sigma |= [](F --> G) |] ==> sigma |= <>G"
343 apply (unfold dmd_def)
347 apply (fastforce elim!: STL4E [temp_use])
351 assumes 1: "sigma |= []<>F"
352 and 2: "sigma |= []G"
353 and 3: "|- F & G --> H"
354 shows "sigma |= []<>H"
356 apply (erule_tac F = G in dup_boxE)
358 apply (fastforce elim!: STL4E [temp_use] DmdImpl2 [temp_use] intro!: 3 [temp_use])
361 (* ------------------------ STL6 ------------------------------------------- *)
362 (* Used in the proof of STL6, but useful in itself. *)
363 lemma BoxDmd: "|- []F & <>G --> <>([]F & G)"
364 apply (unfold dmd_def)
366 apply (erule dup_boxE)
368 apply (erule contrapos_np)
369 apply (fastforce elim!: STL4E [temp_use])
372 (* weaker than BoxDmd, but more polymorphic (and often just right) *)
373 lemma BoxDmd_simple: "|- []F & <>G --> <>(F & G)"
374 apply (unfold dmd_def)
377 apply (fastforce elim!: notE STL4E [temp_use])
380 lemma BoxDmd2_simple: "|- []F & <>G --> <>(G & F)"
381 apply (unfold dmd_def)
384 apply (fastforce elim!: notE STL4E [temp_use])
388 assumes 1: "sigma |= []A"
389 and 2: "sigma |= <>F"
390 and 3: "|- []A & F --> G"
392 apply (rule 2 [THEN 1 [THEN BoxDmd [temp_use]], THEN DmdImplE])
396 lemma STL6: "|- <>[]F & <>[]G --> <>[](F & G)"
397 apply (auto simp: STL5 [temp_rewrite, symmetric])
398 apply (drule linT [temp_use])
400 apply (erule thin_rl)
401 apply (rule DmdDmd [temp_unlift, THEN iffD1])
403 apply (erule DmdImplE)
405 apply (erule DmdImplE)
407 apply (drule BoxDmd [temp_use])
409 apply (erule thin_rl)
410 apply (fastforce elim!: DmdImplE [temp_use])
414 (* ------------------------ True / False ----------------------------------------- *)
415 section "Simplification of constants"
417 lemma BoxConst: "|- ([]#P) = #P"
420 apply (auto intro!: necT [temp_use] dest: STL2_gen [temp_use] simp: Init_simps)
423 lemma DmdConst: "|- (<>#P) = #P"
424 apply (unfold dmd_def)
426 apply (simp_all add: BoxConst [try_rewrite])
429 lemmas temp_simps [temp_rewrite, simp] = BoxConst DmdConst
432 (* ------------------------ Further rewrites ----------------------------------------- *)
433 section "Further rewrites"
435 lemma NotBox: "|- (~[]F) = (<>~F)"
436 by (simp add: dmd_def)
438 lemma NotDmd: "|- (~<>F) = ([]~F)"
439 by (simp add: dmd_def)
441 (* These are not declared by default, because they could be harmful,
442 e.g. []F & ~[]F becomes []F & <>~F !! *)
443 lemmas more_temp_simps1 =
444 STL3 [temp_rewrite] DmdDmd [temp_rewrite] NotBox [temp_rewrite] NotDmd [temp_rewrite]
445 NotBox [temp_unlift, THEN eq_reflection]
446 NotDmd [temp_unlift, THEN eq_reflection]
448 lemma BoxDmdBox: "|- ([]<>[]F) = (<>[]F)"
449 apply (auto dest!: STL2 [temp_use])
451 apply (subgoal_tac "sigma |= <>[][]F & <>[]~[]F")
452 apply (erule thin_rl)
454 apply (drule STL6 [temp_use])
457 apply (simp_all add: more_temp_simps1)
460 lemma DmdBoxDmd: "|- (<>[]<>F) = ([]<>F)"
461 apply (unfold dmd_def)
462 apply (auto simp: BoxDmdBox [unfolded dmd_def, try_rewrite])
465 lemmas more_temp_simps2 = more_temp_simps1 BoxDmdBox [temp_rewrite] DmdBoxDmd [temp_rewrite]
468 (* ------------------------ Miscellaneous ----------------------------------- *)
470 lemma BoxOr: "!!sigma. [| sigma |= []F | []G |] ==> sigma |= [](F | G)"
471 by (fastforce elim!: STL4E [temp_use])
473 (* "persistently implies infinitely often" *)
474 lemma DBImplBD: "|- <>[]F --> []<>F"
477 apply (simp add: more_temp_simps2)
478 apply (drule STL6 [temp_use])
483 lemma BoxDmdDmdBox: "|- []<>F & <>[]G --> []<>(F & G)"
486 apply (unfold more_temp_simps2)
487 apply (drule STL6 [temp_use])
489 apply (subgoal_tac "sigma |= <>[]~F")
490 apply (force simp: dmd_def)
491 apply (fastforce elim: DmdImplE [temp_use] STL4E [temp_use])
495 (* ------------------------------------------------------------------------- *)
496 (*** TLA-specific theorems: primed formulas ***)
497 (* ------------------------------------------------------------------------- *)
500 (* ------------------------ TLA2 ------------------------------------------- *)
501 lemma STL2_pr: "|- []P --> Init P & Init P`"
502 by (fastforce intro!: STL2_gen [temp_use] primeI [temp_use])
504 (* Auxiliary lemma allows priming of boxed actions *)
505 lemma BoxPrime: "|- []P --> []($P & P$)"
507 apply (erule dup_boxE)
508 apply (unfold boxInit_act)
510 apply (auto simp: Init_simps dest!: STL2_pr [temp_use])
514 assumes "|- $P & P$ --> A"
515 shows "|- []P --> []A"
517 apply (drule BoxPrime [temp_use])
518 apply (auto simp: Init_stp_act_rev [try_rewrite] intro!: assms [temp_use]
519 elim!: STL4E [temp_use])
522 lemma TLA2E: "[| sigma |= []P; |- $P & P$ --> A |] ==> sigma |= []A"
523 by (erule (1) TLA2 [temp_use])
525 lemma DmdPrime: "|- (<>P`) --> (<>P)"
526 apply (unfold dmd_def)
527 apply (fastforce elim!: TLA2E [temp_use])
530 lemmas PrimeDmd = InitDmd_gen [temp_use, THEN DmdPrime [temp_use]]
532 (* ------------------------ INV1, stable --------------------------------------- *)
533 section "stable, invariant"
536 "[| sigma |= []H; sigma |= Init P; |- H --> (Init P & ~[]F --> Init(P`) & F) |]
538 apply (rule indT [temp_use])
539 apply (erule (2) STL4E)
542 lemma box_stp_act: "|- ([]$P) = ([]P)"
543 by (simp add: boxInit_act Init_simps)
545 lemmas box_stp_actI = box_stp_act [temp_use, THEN iffD2]
546 lemmas box_stp_actD = box_stp_act [temp_use, THEN iffD1]
548 lemmas more_temp_simps3 = box_stp_act [temp_rewrite] more_temp_simps2
551 "|- (Init P) --> (stable P) --> []P"
552 apply (unfold stable_def boxInit_stp boxInit_act)
554 apply (erule ind_rule)
555 apply (auto simp: Init_simps elim: ind_rule)
559 "!!P. |- $P & A --> P` ==> |- []A --> stable P"
560 apply (unfold stable_def)
561 apply (fastforce elim!: STL4E [temp_use])
564 lemma Stable: "[| sigma |= []A; |- $P & A --> P` |] ==> sigma |= stable P"
565 by (erule (1) StableT [temp_use])
567 (* Generalization of INV1 *)
568 lemma StableBox: "|- (stable P) --> [](Init P --> []P)"
569 apply (unfold stable_def)
571 apply (erule dup_boxE)
572 apply (force simp: stable_def elim: STL4E [temp_use] INV1 [temp_use])
575 lemma DmdStable: "|- (stable P) & <>P --> <>[]P"
577 apply (rule DmdImpl2)
579 apply (erule StableBox [temp_use])
580 apply (simp add: dmdInitD)
583 (* ---------------- (Semi-)automatic invariant tactics ---------------------- *)
586 (* inv_tac reduces goals of the form ... ==> sigma |= []P *)
591 TRY (merge_box_tac 1),
592 rtac (temp_use ctxt @{thm INV1}) 1, (* fail if the goal is not a box *)
593 TRYALL (etac @{thm Stable})]);
595 (* auto_inv_tac applies inv_tac and then tries to attack the subgoals
596 in simple cases it may be able to handle goals like |- MyProg --> []Inv.
597 In these simple cases the simplifier seems to be more useful than the
598 auto-tactic, which applies too much propositional logic and simplifies
601 fun auto_inv_tac ctxt =
604 (TRYALL (action_simp_tac
605 (ctxt addsimps [@{thm Init_stp}, @{thm Init_act}]) [] [@{thm squareE}])));
608 method_setup invariant = {*
609 Method.sections Clasimp.clasimp_modifiers >> (K (SIMPLE_METHOD' o inv_tac))
612 method_setup auto_invariant = {*
613 Method.sections Clasimp.clasimp_modifiers >> (K (SIMPLE_METHOD' o auto_inv_tac))
616 lemma unless: "|- []($P --> P` | Q`) --> (stable P) | <>Q"
617 apply (unfold dmd_def)
618 apply (clarsimp dest!: BoxPrime [temp_use])
620 apply (erule contrapos_np)
621 apply (fastforce elim!: Stable [temp_use])
625 (* --------------------- Recursive expansions --------------------------------------- *)
626 section "recursive expansions"
628 (* Recursive expansions of [] and <> for state predicates *)
629 lemma BoxRec: "|- ([]P) = (Init P & []P`)"
630 apply (auto intro!: STL2_gen [temp_use])
631 apply (fastforce elim!: TLA2E [temp_use])
632 apply (auto simp: stable_def elim!: INV1 [temp_use] STL4E [temp_use])
635 lemma DmdRec: "|- (<>P) = (Init P | <>P`)"
636 apply (unfold dmd_def BoxRec [temp_rewrite])
637 apply (auto simp: Init_simps)
640 lemma DmdRec2: "!!sigma. [| sigma |= <>P; sigma |= []~P` |] ==> sigma |= Init P"
641 apply (force simp: DmdRec [temp_rewrite] dmd_def)
644 lemma InfinitePrime: "|- ([]<>P) = ([]<>P`)"
646 apply (rule classical)
647 apply (rule DBImplBD [temp_use])
648 apply (subgoal_tac "sigma |= <>[]P")
649 apply (fastforce elim!: DmdImplE [temp_use] TLA2E [temp_use])
650 apply (subgoal_tac "sigma |= <>[] (<>P & []~P`)")
651 apply (force simp: boxInit_stp [temp_use]
652 elim!: DmdImplE [temp_use] STL4E [temp_use] DmdRec2 [temp_use])
653 apply (force intro!: STL6 [temp_use] simp: more_temp_simps3)
654 apply (fastforce intro: DmdPrime [temp_use] elim!: STL4E [temp_use])
657 lemma InfiniteEnsures:
658 "[| sigma |= []N; sigma |= []<>A; |- A & N --> P` |] ==> sigma |= []<>P"
659 apply (unfold InfinitePrime [temp_rewrite])
664 (* ------------------------ fairness ------------------------------------------- *)
667 (* alternative definitions of fairness *)
668 lemma WF_alt: "|- WF(A)_v = ([]<>~Enabled(<A>_v) | []<><A>_v)"
669 apply (unfold WF_def dmd_def)
673 lemma SF_alt: "|- SF(A)_v = (<>[]~Enabled(<A>_v) | []<><A>_v)"
674 apply (unfold SF_def dmd_def)
678 (* theorems to "box" fairness conditions *)
679 lemma BoxWFI: "|- WF(A)_v --> []WF(A)_v"
680 by (auto simp: WF_alt [try_rewrite] more_temp_simps3 intro!: BoxOr [temp_use])
682 lemma WF_Box: "|- ([]WF(A)_v) = WF(A)_v"
683 by (fastforce intro!: BoxWFI [temp_use] dest!: STL2 [temp_use])
685 lemma BoxSFI: "|- SF(A)_v --> []SF(A)_v"
686 by (auto simp: SF_alt [try_rewrite] more_temp_simps3 intro!: BoxOr [temp_use])
688 lemma SF_Box: "|- ([]SF(A)_v) = SF(A)_v"
689 by (fastforce intro!: BoxSFI [temp_use] dest!: STL2 [temp_use])
691 lemmas more_temp_simps = more_temp_simps3 WF_Box [temp_rewrite] SF_Box [temp_rewrite]
693 lemma SFImplWF: "|- SF(A)_v --> WF(A)_v"
694 apply (unfold SF_def WF_def)
695 apply (fastforce dest!: DBImplBD [temp_use])
698 (* A tactic that "boxes" all fairness conditions. Apply more_temp_simps to "unbox". *)
700 val box_fair_tac = SELECT_GOAL (REPEAT (dresolve_tac [@{thm BoxWFI}, @{thm BoxSFI}] 1))
704 (* ------------------------------ leads-to ------------------------------ *)
708 lemma leadsto_init: "|- (Init F) & (F ~> G) --> <>G"
709 apply (unfold leadsto_def)
710 apply (auto dest!: STL2 [temp_use])
713 (* |- F & (F ~> G) --> <>G *)
714 lemmas leadsto_init_temp = leadsto_init [where 'a = behavior, unfolded Init_simps]
716 lemma streett_leadsto: "|- ([]<>Init F --> []<>G) = (<>(F ~> G))"
717 apply (unfold leadsto_def)
719 apply (simp add: more_temp_simps)
720 apply (fastforce elim!: DmdImplE [temp_use] STL4E [temp_use])
721 apply (fastforce intro!: InitDmd [temp_use] elim!: STL4E [temp_use])
722 apply (subgoal_tac "sigma |= []<><>G")
723 apply (simp add: more_temp_simps)
724 apply (drule BoxDmdDmdBox [temp_use])
726 apply (fastforce elim!: DmdImplE [temp_use] STL4E [temp_use])
729 lemma leadsto_infinite: "|- []<>F & (F ~> G) --> []<>G"
731 apply (erule InitDmd [temp_use, THEN streett_leadsto [temp_unlift, THEN iffD2, THEN mp]])
732 apply (simp add: dmdInitD)
735 (* In particular, strong fairness is a Streett condition. The following
736 rules are sometimes easier to use than WF2 or SF2 below.
738 lemma leadsto_SF: "|- (Enabled(<A>_v) ~> <A>_v) --> SF(A)_v"
739 apply (unfold SF_def)
740 apply (clarsimp elim!: leadsto_infinite [temp_use])
743 lemma leadsto_WF: "|- (Enabled(<A>_v) ~> <A>_v) --> WF(A)_v"
744 by (clarsimp intro!: SFImplWF [temp_use] leadsto_SF [temp_use])
746 (* introduce an invariant into the proof of a leadsto assertion.
747 []I --> ((P ~> Q) = (P /\ I ~> Q))
749 lemma INV_leadsto: "|- []I & (P & I ~> Q) --> (P ~> Q)"
750 apply (unfold leadsto_def)
752 apply (erule STL4Edup)
754 apply (auto simp: Init_simps dest!: STL2_gen [temp_use])
757 lemma leadsto_classical: "|- (Init F & []~G ~> G) --> (F ~> G)"
758 apply (unfold leadsto_def dmd_def)
759 apply (force simp: Init_simps elim!: STL4E [temp_use])
762 lemma leadsto_false: "|- (F ~> #False) = ([]~F)"
763 apply (unfold leadsto_def)
764 apply (simp add: boxNotInitD)
767 lemma leadsto_exists: "|- ((EX x. F x) ~> G) = (ALL x. (F x ~> G))"
768 apply (unfold leadsto_def)
769 apply (auto simp: allT [try_rewrite] Init_simps elim!: STL4E [temp_use])
772 (* basic leadsto properties, cf. Unity *)
774 lemma ImplLeadsto_gen: "|- [](Init F --> Init G) --> (F ~> G)"
775 apply (unfold leadsto_def)
776 apply (auto intro!: InitDmd_gen [temp_use]
777 elim!: STL4E_gen [temp_use] simp: Init_simps)
781 ImplLeadsto_gen [where 'a = behavior and 'b = behavior, unfolded Init_simps]
783 lemma ImplLeadsto_simple: "!!F G. |- F --> G ==> |- F ~> G"
784 by (auto simp: Init_def intro!: ImplLeadsto_gen [temp_use] necT [temp_use])
786 lemma EnsuresLeadsto:
787 assumes "|- A & $P --> Q`"
788 shows "|- []A --> (P ~> Q)"
789 apply (unfold leadsto_def)
790 apply (clarsimp elim!: INV_leadsto [temp_use])
791 apply (erule STL4E_gen)
792 apply (auto simp: Init_defs intro!: PrimeDmd [temp_use] assms [temp_use])
795 lemma EnsuresLeadsto2: "|- []($P --> Q`) --> (P ~> Q)"
796 apply (unfold leadsto_def)
798 apply (erule STL4E_gen)
799 apply (auto simp: Init_simps intro!: PrimeDmd [temp_use])
803 assumes 1: "|- $P & N --> P` | Q`"
804 and 2: "|- ($P & N) & A --> Q`"
805 shows "|- []N & []([]P --> <>A) --> (P ~> Q)"
806 apply (unfold leadsto_def)
808 apply (erule STL4Edup)
811 apply (subgoal_tac "sigmaa |= [] ($P --> P` | Q`) ")
812 apply (drule unless [temp_use])
813 apply (clarsimp dest!: INV1 [temp_use])
814 apply (rule 2 [THEN DmdImpl, temp_use, THEN DmdPrime [temp_use]])
815 apply (force intro!: BoxDmd_simple [temp_use]
816 simp: split_box_conj [try_rewrite] box_stp_act [try_rewrite])
817 apply (force elim: STL4E [temp_use] dest: 1 [temp_use])
820 lemma ensures_simple:
821 "[| |- $P & N --> P` | Q`;
822 |- ($P & N) & A --> Q`
823 |] ==> |- []N & []<>A --> (P ~> Q)"
825 apply (erule (2) ensures [temp_use])
826 apply (force elim!: STL4E [temp_use])
829 lemma EnsuresInfinite:
830 "[| sigma |= []<>P; sigma |= []A; |- A & $P --> Q` |] ==> sigma |= []<>Q"
831 apply (erule leadsto_infinite [temp_use])
832 apply (erule EnsuresLeadsto [temp_use])
837 (*** Gronning's lattice rules (taken from TLP) ***)
838 section "Lattice rules"
840 lemma LatticeReflexivity: "|- F ~> F"
841 apply (unfold leadsto_def)
842 apply (rule necT InitDmd_gen)+
845 lemma LatticeTransitivity: "|- (G ~> H) & (F ~> G) --> (F ~> H)"
846 apply (unfold leadsto_def)
848 apply (erule dup_boxE) (* [][] (Init G --> H) *)
850 apply (clarsimp elim!: STL4E [temp_use])
851 apply (rule dup_dmdD)
852 apply (subgoal_tac "sigmaa |= <>Init G")
853 apply (erule DmdImpl2)
855 apply (simp add: dmdInitD)
858 lemma LatticeDisjunctionElim1: "|- (F | G ~> H) --> (F ~> H)"
859 apply (unfold leadsto_def)
860 apply (auto simp: Init_simps elim!: STL4E [temp_use])
863 lemma LatticeDisjunctionElim2: "|- (F | G ~> H) --> (G ~> H)"
864 apply (unfold leadsto_def)
865 apply (auto simp: Init_simps elim!: STL4E [temp_use])
868 lemma LatticeDisjunctionIntro: "|- (F ~> H) & (G ~> H) --> (F | G ~> H)"
869 apply (unfold leadsto_def)
872 apply (auto simp: Init_simps elim!: STL4E [temp_use])
875 lemma LatticeDisjunction: "|- (F | G ~> H) = ((F ~> H) & (G ~> H))"
876 by (auto intro: LatticeDisjunctionIntro [temp_use]
877 LatticeDisjunctionElim1 [temp_use]
878 LatticeDisjunctionElim2 [temp_use])
880 lemma LatticeDiamond: "|- (A ~> B | C) & (B ~> D) & (C ~> D) --> (A ~> D)"
882 apply (subgoal_tac "sigma |= (B | C) ~> D")
883 apply (erule_tac G = "LIFT (B | C)" in LatticeTransitivity [temp_use])
884 apply (fastforce intro!: LatticeDisjunctionIntro [temp_use])+
887 lemma LatticeTriangle: "|- (A ~> D | B) & (B ~> D) --> (A ~> D)"
889 apply (subgoal_tac "sigma |= (D | B) ~> D")
890 apply (erule_tac G = "LIFT (D | B)" in LatticeTransitivity [temp_use])
892 apply (auto intro: LatticeDisjunctionIntro [temp_use] LatticeReflexivity [temp_use])
895 lemma LatticeTriangle2: "|- (A ~> B | D) & (B ~> D) --> (A ~> D)"
897 apply (subgoal_tac "sigma |= B | D ~> D")
898 apply (erule_tac G = "LIFT (B | D)" in LatticeTransitivity [temp_use])
900 apply (auto intro: LatticeDisjunctionIntro [temp_use] LatticeReflexivity [temp_use])
903 (*** Lamport's fairness rules ***)
904 section "Fairness rules"
907 "[| |- $P & N --> P` | Q`;
908 |- ($P & N) & <A>_v --> Q`;
909 |- $P & N --> $(Enabled(<A>_v)) |]
910 ==> |- []N & WF(A)_v --> (P ~> Q)"
911 apply (clarsimp dest!: BoxWFI [temp_use])
912 apply (erule (2) ensures [temp_use])
913 apply (erule (1) STL4Edup)
914 apply (clarsimp simp: WF_def)
915 apply (rule STL2 [temp_use])
916 apply (clarsimp elim!: mp intro!: InitDmd [temp_use])
917 apply (erule STL4 [temp_use, THEN box_stp_actD [temp_use]])
918 apply (simp add: split_box_conj box_stp_actI)
921 (* Sometimes easier to use; designed for action B rather than state predicate Q *)
923 assumes 1: "|- N & $P --> $Enabled (<A>_v)"
924 and 2: "|- N & <A>_v --> B"
925 and 3: "|- [](N & [~A]_v) --> stable P"
926 shows "|- []N & WF(A)_v --> (P ~> B)"
927 apply (unfold leadsto_def)
928 apply (clarsimp dest!: BoxWFI [temp_use])
929 apply (erule (1) STL4Edup)
931 apply (rule 2 [THEN DmdImpl, temp_use])
932 apply (rule BoxDmd_simple [temp_use])
934 apply (rule classical)
935 apply (rule STL2 [temp_use])
936 apply (clarsimp simp: WF_def elim!: mp intro!: InitDmd [temp_use])
937 apply (rule 1 [THEN STL4, temp_use, THEN box_stp_actD])
938 apply (simp (no_asm_simp) add: split_box_conj [try_rewrite] box_stp_act [try_rewrite])
939 apply (erule INV1 [temp_use])
940 apply (rule 3 [temp_use])
941 apply (simp add: split_box_conj [try_rewrite] NotDmd [temp_use] not_angle [try_rewrite])
945 "[| |- $P & N --> P` | Q`;
946 |- ($P & N) & <A>_v --> Q`;
947 |- []P & []N & []F --> <>Enabled(<A>_v) |]
948 ==> |- []N & SF(A)_v & []F --> (P ~> Q)"
949 apply (clarsimp dest!: BoxSFI [temp_use])
950 apply (erule (2) ensures [temp_use])
951 apply (erule_tac F = F in dup_boxE)
953 apply (erule STL4Edup)
955 apply (clarsimp simp: SF_def)
956 apply (rule STL2 [temp_use])
958 apply (erule STL4 [temp_use])
959 apply (simp add: split_box_conj [try_rewrite] STL3 [try_rewrite])
963 assumes 1: "|- N & <B>_f --> <M>_g"
964 and 2: "|- $P & P` & <N & A>_f --> B"
965 and 3: "|- P & Enabled(<M>_g) --> Enabled(<A>_f)"
966 and 4: "|- [](N & [~B]_f) & WF(A)_f & []F & <>[]Enabled(<M>_g) --> <>[]P"
967 shows "|- []N & WF(A)_f & []F --> WF(M)_g"
968 apply (clarsimp dest!: BoxWFI [temp_use] BoxDmdBox [temp_use, THEN iffD2]
969 simp: WF_def [where A = M])
970 apply (erule_tac F = F in dup_boxE)
972 apply (erule STL4Edup)
974 apply (clarsimp intro!: BoxDmd_simple [temp_use, THEN 1 [THEN DmdImpl, temp_use]])
975 apply (rule classical)
976 apply (subgoal_tac "sigmaa |= <> (($P & P` & N) & <A>_f)")
977 apply (force simp: angle_def intro!: 2 [temp_use] elim!: DmdImplE [temp_use])
978 apply (rule BoxDmd_simple [THEN DmdImpl, unfolded DmdDmd [temp_rewrite], temp_use])
979 apply (simp add: NotDmd [temp_use] not_angle [try_rewrite])
981 apply (frule 4 [temp_use])
983 apply (drule STL6 [temp_use])
985 apply (erule_tac V = "sigmaa |= <>[]P" in thin_rl)
986 apply (erule_tac V = "sigmaa |= []F" in thin_rl)
987 apply (drule BoxWFI [temp_use])
988 apply (erule_tac F = "ACT N & [~B]_f" in dup_boxE)
990 apply (erule DmdImpldup)
992 apply (auto simp: split_box_conj [try_rewrite] STL3 [try_rewrite]
993 WF_Box [try_rewrite] box_stp_act [try_rewrite])
994 apply (force elim!: TLA2E [where P = P, temp_use])
995 apply (rule STL2 [temp_use])
996 apply (force simp: WF_def split_box_conj [try_rewrite]
997 elim!: mp intro!: InitDmd [temp_use] 3 [THEN STL4, temp_use])
1001 assumes 1: "|- N & <B>_f --> <M>_g"
1002 and 2: "|- $P & P` & <N & A>_f --> B"
1003 and 3: "|- P & Enabled(<M>_g) --> Enabled(<A>_f)"
1004 and 4: "|- [](N & [~B]_f) & SF(A)_f & []F & []<>Enabled(<M>_g) --> <>[]P"
1005 shows "|- []N & SF(A)_f & []F --> SF(M)_g"
1006 apply (clarsimp dest!: BoxSFI [temp_use] simp: 2 [try_rewrite] SF_def [where A = M])
1007 apply (erule_tac F = F in dup_boxE)
1008 apply (erule_tac F = "TEMP <>Enabled (<M>_g) " in dup_boxE)
1009 apply merge_temp_box
1010 apply (erule STL4Edup)
1012 apply (clarsimp intro!: BoxDmd_simple [temp_use, THEN 1 [THEN DmdImpl, temp_use]])
1013 apply (rule classical)
1014 apply (subgoal_tac "sigmaa |= <> (($P & P` & N) & <A>_f)")
1015 apply (force simp: angle_def intro!: 2 [temp_use] elim!: DmdImplE [temp_use])
1016 apply (rule BoxDmd_simple [THEN DmdImpl, unfolded DmdDmd [temp_rewrite], temp_use])
1017 apply (simp add: NotDmd [temp_use] not_angle [try_rewrite])
1019 apply (frule 4 [temp_use])
1021 apply (erule_tac V = "sigmaa |= []F" in thin_rl)
1022 apply (drule BoxSFI [temp_use])
1023 apply (erule_tac F = "TEMP <>Enabled (<M>_g)" in dup_boxE)
1024 apply (erule_tac F = "ACT N & [~B]_f" in dup_boxE)
1025 apply merge_temp_box
1026 apply (erule DmdImpldup)
1028 apply (auto simp: split_box_conj [try_rewrite] STL3 [try_rewrite]
1029 SF_Box [try_rewrite] box_stp_act [try_rewrite])
1030 apply (force elim!: TLA2E [where P = P, temp_use])
1031 apply (rule STL2 [temp_use])
1032 apply (force simp: SF_def split_box_conj [try_rewrite]
1033 elim!: mp InfImpl [temp_use] intro!: 3 [temp_use])
1036 (* ------------------------------------------------------------------------- *)
1037 (*** Liveness proofs by well-founded orderings ***)
1038 (* ------------------------------------------------------------------------- *)
1039 section "Well-founded orderings"
1043 and 2: "!!x. sigma |= F x ~> (G | (EX y. #((y,x):r) & F y)) "
1044 shows "sigma |= F x ~> G"
1045 apply (rule 1 [THEN wf_induct])
1046 apply (rule LatticeTriangle [temp_use])
1048 apply (auto simp: leadsto_exists [try_rewrite])
1049 apply (case_tac "(y,x) :r")
1051 apply (force simp: leadsto_def Init_simps intro!: necT [temp_use])
1054 (* If r is well-founded, state function v cannot decrease forever *)
1055 lemma wf_not_box_decrease: "!!r. wf r ==> |- [][ (v`, $v) : #r ]_v --> <>[][#False]_v"
1058 apply (subgoal_tac "sigma |= (EX x. v=#x) ~> #False")
1059 apply (drule leadsto_false [temp_use, THEN iffD1, THEN STL2_gen [temp_use]])
1060 apply (force simp: Init_defs)
1061 apply (clarsimp simp: leadsto_exists [try_rewrite] not_square [try_rewrite] more_temp_simps)
1062 apply (erule wf_leadsto)
1063 apply (rule ensures_simple [temp_use])
1064 apply (auto simp: square_def angle_def)
1067 (* "wf r ==> |- <>[][ (v`, $v) : #r ]_v --> <>[][#False]_v" *)
1068 lemmas wf_not_dmd_box_decrease =
1069 wf_not_box_decrease [THEN DmdImpl, unfolded more_temp_simps]
1071 (* If there are infinitely many steps where v decreases, then there
1072 have to be infinitely many non-stuttering steps where v doesn't decrease.
1074 lemma wf_box_dmd_decrease:
1076 shows "|- []<>((v`, $v) : #r) --> []<><(v`, $v) ~: #r>_v"
1079 apply (simp add: not_angle [try_rewrite] more_temp_simps)
1080 apply (drule 1 [THEN wf_not_dmd_box_decrease [temp_use]])
1081 apply (drule BoxDmdDmdBox [temp_use])
1083 apply (subgoal_tac "sigma |= []<> ((#False) ::action)")
1086 apply (rule DmdImpl)
1087 apply (force intro: 1 [THEN wf_irrefl, temp_use])
1090 (* In particular, for natural numbers, if n decreases infinitely often
1091 then it has to increase infinitely often.
1093 lemma nat_box_dmd_decrease: "!!n::nat stfun. |- []<>(n` < $n) --> []<>($n < n`)"
1095 apply (subgoal_tac "sigma |= []<><~ ((n`,$n) : #less_than) >_n")
1096 apply (erule thin_rl)
1098 apply (rule DmdImpl)
1099 apply (clarsimp simp: angle_def [try_rewrite])
1100 apply (rule wf_box_dmd_decrease [temp_use])
1101 apply (auto elim!: STL4E [temp_use] DmdImplE [temp_use])
1105 (* ------------------------------------------------------------------------- *)
1106 (*** Flexible quantification over state variables ***)
1107 (* ------------------------------------------------------------------------- *)
1108 section "Flexible quantification"
1111 assumes 1: "basevars vs"
1112 and 2: "(!!x. basevars (x,vs) ==> sigma |= F x)"
1113 shows "sigma |= (AALL x. F x)"
1114 by (auto simp: aall_def elim!: eexE [temp_use] intro!: 1 dest!: 2 [temp_use])
1116 lemma aallE: "|- (AALL x. F x) --> F x"
1117 apply (unfold aall_def)
1119 apply (erule contrapos_np)
1120 apply (force intro!: eexI [temp_use])
1123 (* monotonicity of quantification *)
1125 assumes 1: "sigma |= EEX x. F x"
1126 and 2: "!!x. sigma |= F x --> G x"
1127 shows "sigma |= EEX x. G x"
1128 apply (rule unit_base [THEN 1 [THEN eexE]])
1129 apply (rule eexI [temp_use])
1130 apply (erule 2 [unfolded intensional_rews, THEN mp])
1134 assumes 1: "sigma |= AALL x. F(x)"
1135 and 2: "!!x. sigma |= F(x) --> G(x)"
1136 shows "sigma |= AALL x. G(x)"
1137 apply (rule unit_base [THEN aallI])
1138 apply (rule 2 [unfolded intensional_rews, THEN mp])
1139 apply (rule 1 [THEN aallE [temp_use]])
1142 (* Derived history introduction rule *)
1144 assumes 1: "sigma |= Init I"
1145 and 2: "sigma |= []N"
1146 and 3: "basevars vs"
1147 and 4: "!!h. basevars(h,vs) ==> |- I & h = ha --> HI h"
1148 and 5: "!!h s t. [| basevars(h,vs); N (s,t); h t = hb (h s) (s,t) |] ==> HN h (s,t)"
1149 shows "sigma |= EEX h. Init (HI h) & [](HN h)"
1150 apply (rule history [temp_use, THEN eexE])
1152 apply (rule eexI [temp_use])
1158 apply (force elim!: STL4E [temp_use] 5 [temp_use])
1160 apply (force simp: Init_defs elim!: 4 [temp_use])
1163 (* ----------------------------------------------------------------------
1164 example of a history variable: existence of a clock
1167 lemma "|- EEX h. Init(h = #True) & [](h` = (~$h))"
1169 apply (rule historyI)
1170 apply (force simp: Init_defs intro!: unit_base [temp_use] necT [temp_use])+