Theory Rules
section "A Proof System for TLA* "
theory Rules
imports PreFormulas
begin
text‹
We prove soundness of the proof system of \tlastar{}, from which the system
verification rules from Lamport's original TLA paper will be derived.
This theory is still state-independent, thus state-dependent enableness proofs,
required for proofs based on fairness assumptions, and flexible quantification,
are not discussed here.
The \tlastar{} paper \<^cite>‹"Merz99"› suggest both a \emph{hetereogeneous} and a
\emph{homogenous} proof system for \tlastar{}.
The homogeneous version eliminates the auxiliary definitions from the
‹Preformula› theory, creating a single provability relation.
This axiomatisation is based on the fact that a pre-formula can only be used
via the ‹sq› rule. In a nutshell, ‹sq› is applied to
‹pax1› to ‹pax5›, and ‹nex›, ‹pre› and ‹pmp›
are changed to accommodate this. It is argued that while the hetereogenous version
is easier to understand, the homogenous system avoids the introduction of an
auxiliary provability relation. However, the price to pay is that reasoning about
pre-formulas (in particular, actions) has to be performed in the scope of
temporal operators such as ‹□[P]_v›, which is notationally quite heavy,
We prefer here the heterogeneous approach, which exposes the pre-formulas and
lets us use standard HOL rules more directly.
›
subsection "The Basic Axioms"
theorem fmp: assumes "⊢ F" and "⊢ F ⟶ G" shows "⊢ G"
using assms[unlifted] by auto
theorem pmp: assumes "|~ F" and "|~ F ⟶ G" shows "|~ G"
using assms[unlifted] by auto
theorem sq: assumes "|~ P" shows "⊢ □[P]_v"
using assms[unlifted] by (auto simp: action_def)
theorem pre: assumes "⊢ F" shows "|~ F"
using assms by auto
theorem nex: assumes h1: "⊢ F" shows "|~ ○F"
using assms by (auto simp: nexts_def)
theorem ax0: "⊢ # True"
by auto
theorem ax1: "⊢ □F ⟶ F"
proof (clarsimp simp: always_def)
fix w
assume "∀n. (w |⇩s n) ⊨ F"
hence "(w |⇩s 0) ⊨ F" ..
thus "w ⊨ F" by simp
qed
theorem ax2: "⊢ □F ⟶ □[□F]_v"
by (auto simp: always_def action_def suffix_plus)
theorem ax3:
assumes H: "|~ F ∧ Unchanged v ⟶ ○F"
shows "⊢ □[F ⟶ ○F]_v ⟶ (F ⟶ □F)"
proof (clarsimp simp: always_def)
fix w n
assume a1: "w ⊨ □[F ⟶ ○F]_v" and a2: "w ⊨ F"
show "(w |⇩s n) ⊨ F"
proof (induct n)
from a2 show "(w |⇩s 0) ⊨ F" by simp
next
fix m
assume a3: "(w |⇩s m) ⊨ F"
with a1 H[unlifted] show "(w |⇩s (Suc m)) ⊨ F"
by (auto simp: nexts_def action_def tail_suffix_suc)
qed
qed
theorem ax4: "⊢ □[P ⟶ Q]_v ⟶ (□[P]_v ⟶ □[Q]_v)"
by (force simp: action_def)
theorem ax5: "⊢ □[v` ≠ $v]_v"
by (auto simp: action_def unch_def)
theorem pax0: "|~ # True"
by auto
theorem pax1 [simp_unl]: "|~ (○¬F) = (¬○F)"
by (auto simp: nexts_def)
theorem pax2: "|~ ○(F ⟶ G) ⟶ (○F ⟶ ○G)"
by (auto simp: nexts_def)
theorem pax3: "|~ □F ⟶ ○□F"
by (auto simp: always_def nexts_def tail_def suffix_plus)
theorem pax4: "|~ □[P]_v = ([P]_v ∧ ○□[P]_v)"
proof (auto)
fix w
assume "w ⊨ □[P]_v"
from this[unfolded action_def] have "((w |⇩s 0) ⊨ P) ∨ ((w |⇩s 0) ⊨ Unchanged v)" ..
thus "w ⊨ [P]_v" by (simp add: actrans_def)
next
fix w
assume "w ⊨ □[P]_v"
thus "w ⊨ ○□[P]_v" by (auto simp: nexts_def action_def tail_def suffix_plus)
next
fix w
assume 1: "w ⊨ [P]_v" and 2: "w ⊨ ○□[P]_v"
show "w ⊨ □[P]_v"
proof (auto simp: action_def)
fix i
assume 3: "¬ ((w |⇩s i) ⊨ Unchanged v)"
show "(w |⇩s i) ⊨ P"
proof (cases i)
assume "i = 0"
with 1 3 show ?thesis by (simp add: actrans_def)
next
fix j
assume "i = Suc j"
with 2 3 show ?thesis by (auto simp: nexts_def action_def tail_def suffix_plus)
qed
qed
qed
theorem pax5: "|~ ○□F ⟶ □[○F]_v"
by (auto simp: nexts_def always_def action_def tail_def suffix_plus)
text ‹
Theorem to show that universal quantification distributes over the always
operator. Since the \tlastar{} paper only addresses the propositional fragment,
this theorem does not appear there.
›
theorem allT: "⊢ (∀x. □(F x)) = (□(∀x. F x))"
by (auto simp: always_def)
theorem allActT: "⊢ (∀x. □[F x]_v) = (□[(∀x. F x)]_v)"
by (force simp: action_def)
subsection "Derived Theorems"
text‹
This section includes some derived theorems based on the axioms, taken
from the \tlastar{} paper~\<^cite>‹"Merz99"›. We mimic the proofs given there
and avoid semantic reasoning whenever possible.
The ‹alw› theorem of~\<^cite>‹"Merz99"› states that if F holds
in all worlds then it always holds, i.e. $F \vDash \Box F$. However,
the derivation of this theorem (using the proof rules above)
relies on access of the set of free variables (FV), which is not
available in a shallow encoding.
However, we can prove a similar rule ‹alw2› using an additional
hypothesis @{term "|~ F ∧ Unchanged v ⟶ ○F"}.
›
theorem alw2:
assumes h1: "⊢ F" and h2: "|~ F ∧ Unchanged v ⟶ ○F"
shows "⊢ □F"
proof -
from h1 have g2: "|~ ○F" by (rule nex)
hence g3: "|~ F ⟶ ○F" by auto
hence g4:"⊢ □[(F ⟶ ○F)]_v" by (rule sq)
from h2 have "⊢ □[(F ⟶ ○F)]_v ⟶ F ⟶ □F" by (rule ax3)
with g4[unlifted] have g5: "⊢ F ⟶ □F" by auto
with h1[unlifted] show ?thesis by auto
qed
text‹
Similar theorem, assuming that @{term "F"} is stuttering invariant.
›
theorem alw3:
assumes h1: "⊢ F" and h2: "stutinv F"
shows "⊢ □F"
proof -
from h2 have "|~ F ∧ Unchanged id ⟶ ○F" by (rule pre_id_unch)
with h1 show ?thesis by (rule alw2)
qed
text‹
In a deep embedding, we could prove that all (proper) \tlastar{}
formulas are stuttering invariant and then get rid of the second
hypothesis of rule ‹alw3›. In fact, the rule is even true
for pre-formulas, as shown by the following rule, whose proof relies
on semantical reasoning.
›
theorem alw: assumes H1: "⊢ F" shows "⊢ □F"
using H1 by (auto simp: always_def)
theorem alw_valid_iff_valid: "(⊢ □F) = (⊢ F)"
proof
assume "⊢ □F"
from this ax1 show "⊢ F" by (rule fmp)
qed (rule alw)
text ‹
\<^cite>‹"Merz99"› proves the following theorem using the deduction theorem of
\tlastar{}: ‹(⊢ F ⟹ ⊢ G) ⟹ ⊢ []F ⟶ G›, which can only be
proved by induction on the formula structure, in a deep embedding.
›
theorem T1[simp_unl]: "⊢ □□F = []F"
proof (auto simp: always_def suffix_plus)
fix w n
assume "∀m k. (w |⇩s (k+m)) ⊨ F"
hence "(w |⇩s (n+0)) ⊨ F" by blast
thus "(w |⇩s n) ⊨ F" by simp
qed
theorem T2[simp_unl]: "⊢ □□[P]_v = □[P]_v"
proof -
have 1: "|~ □[P]_v ⟶ ○□[P]_v" using pax4 by force
hence "⊢ □[□[P]_v ⟶ ○□[P]_v]_v" by (rule sq)
moreover
have "⊢ □[ □[P]_v ⟶ ○□[P]_v ]_v ⟶ □[P]_v ⟶ □□[P]_v"
by (rule ax3) (auto elim: 1[unlift_rule])
moreover
have "⊢ □□[P]_v ⟶ □[P]_v" by (rule ax1)
ultimately show ?thesis by force
qed
theorem T3[simp_unl]: "⊢ □[[P]_v]_v = □[P]_v"
proof -
have "|~ P ⟶ [P]_v" by (auto simp: actrans_def)
hence "⊢ □[(P ⟶ [P]_v)]_v" by (rule sq)
with ax4 have "⊢ □[P]_v ⟶ □[[P]_v]_v" by force
moreover
have "|~ [P]_v ⟶ v`≠ $v ⟶ P" by (auto simp: unch_def actrans_def)
hence "⊢ □[[P]_v ⟶ v`≠ $v ⟶ P]_v" by (rule sq)
with ax5 have "⊢ □[[P]_v]_v ⟶ □[P]_v" by (force intro: ax4[unlift_rule])
ultimately show ?thesis by force
qed
theorem M2:
assumes h: "|~ F ⟶ G"
shows "⊢ □[F]_v ⟶ □[G]_v"
using sq[OF h] ax4 by force
theorem N1:
assumes h: "⊢ F ⟶ G"
shows "|~ ○F ⟶ ○G"
by (rule pmp[OF nex[OF h] pax2])
theorem T4: "⊢ □[P]_v ⟶ □[[P]_v]_w"
proof -
have "⊢ □□[P]_v ⟶ □[□□[P]_v]_w" by (rule ax2)
moreover
from pax4 have "|~ □□[P]_v ⟶ [P]_v" unfolding T2[int_rewrite] by force
hence "⊢ □[□□[P]_v]_w ⟶ □[[P]_v]_w" by (rule M2)
ultimately show ?thesis unfolding T2[int_rewrite] by (rule lift_imp_trans)
qed
theorem T5: "⊢ □[[P]_w]_v ⟶ □[[P]_v]_w"
proof -
have "|~ [[P]_w]_v ⟶ [[P]_v]_w" by (auto simp: actrans_def)
hence "⊢ □[[[P]_w]_v]_w ⟶ □[[[P]_v]_w]_w" by (rule M2)
with T4 show ?thesis unfolding T3[int_rewrite] by (rule lift_imp_trans)
qed
theorem T6: "⊢ □F ⟶ □[○F]_v"
proof -
from ax1 have "|~ ○(□F ⟶ F)" by (rule nex)
with pax2 have "|~ ○□F ⟶ ○F" by force
with pax3 have "|~ □F ⟶ ○F" by (rule pref_imp_trans)
hence "⊢ □[□F]_v ⟶ □[○F]_v" by (rule M2)
with ax2 show ?thesis by (rule lift_imp_trans)
qed
theorem T7:
assumes h: "|~ F ∧ Unchanged v ⟶ ○F"
shows "|~ (F ∧ ○□F) = □F"
proof -
have "⊢ □[○F ⟶ F ⟶ ○F]_v" by (rule sq) auto
with ax4 have "⊢ □[○F]_v ⟶ □[(F ⟶ ○F)]_v" by force
with ax3[OF h, unlifted] have "⊢ □[○F]_v ⟶ (F ⟶ □F)" by force
with pax5 have "|~ F ∧ ○□F ⟶ □F" by force
with ax1[of "TEMP F",unlifted] pax3[of "TEMP F",unlifted] show ?thesis by force
qed
theorem T8: "|~ ○(F ∧ G) = (○F ∧ ○G)"
proof -
have "|~ ○(F ∧ G) ⟶ ○F" by (rule N1) auto
moreover
have "|~ ○(F ∧ G) ⟶ ○G" by (rule N1) auto
moreover
have "⊢ F ⟶ G ⟶ F ∧ G" by auto
from nex[OF this] have "|~ ○F ⟶ ○G ⟶ ○(F ∧ G)"
by (force intro: pax2[unlift_rule])
ultimately show ?thesis by force
qed
lemma T9: "|~ □[A]_v ⟶ [A]_v"
using pax4 by force
theorem H1:
assumes h1: "⊢ □[P]_v" and h2: "⊢ □[P ⟶ Q]_v"
shows "⊢ □[Q]_v"
using assms ax4[unlifted] by force
theorem H2: assumes h1: "⊢ F" shows "⊢ □[F]_v"
using h1 by (blast dest: pre sq)
theorem H3:
assumes h1: "⊢ □[P ⟶ Q]_v" and h2: "⊢ □[Q ⟶ R]_v"
shows "⊢ □[P ⟶ R]_v"
proof -
have "|~ (P ⟶ Q) ⟶ (Q ⟶ R) ⟶ (P ⟶ R)" by auto
hence "⊢ □[(P ⟶ Q) ⟶ (Q ⟶ R) ⟶ (P ⟶ R)]_v" by (rule sq)
with h1 have "⊢ □[(Q ⟶ R) ⟶ (P ⟶ R)]_v" by (rule H1)
with h2 show ?thesis by (rule H1)
qed
theorem H4: "⊢ □[[P]_v ⟶ P]_v"
proof -
have "|~ v` ≠ $v ⟶ ([P]_v ⟶ P)" by (auto simp: unch_def actrans_def)
hence "⊢ □[v` ≠ $v ⟶ ([P]_v ⟶ P)]_v" by (rule sq)
with ax5 show ?thesis by (rule H1)
qed
theorem H5: "⊢ □[□F ⟶ ○□F]_v"
by (rule pax3[THEN sq])
subsection "Some other useful derived theorems"
theorem P1: "|~ □F ⟶ ○F"
proof -
have "|~ ○□F ⟶ ○F" by (rule N1[OF ax1])
with pax3 show ?thesis by (rule pref_imp_trans)
qed
theorem P2: "|~ □F ⟶ F ∧ ○F"
using ax1[of F] P1[of F] by force
theorem P4: "⊢ □F ⟶ □[F]_v"
proof -
have "⊢ □[□F]_v ⟶ □[F]_v" by (rule M2[OF pre[OF ax1]])
with ax2 show ?thesis by (rule lift_imp_trans)
qed
theorem P5: "⊢ □[P]_v ⟶ □[□[P]_v]_w"
proof -
have "⊢ □□[P]_v ⟶ □[□[P]_v]_w" by (rule P4)
thus ?thesis by (unfold T2[int_rewrite])
qed
theorem M0: "⊢ □F ⟶ □[F ⟶ ○F]_v"
proof -
from P1 have "|~ □F ⟶ F ⟶ ○F" by force
hence "⊢ □[□F]_v ⟶ □[F ⟶ ○F]_v" by (rule M2)
with ax2 show ?thesis by (rule lift_imp_trans)
qed
theorem M1: "⊢ □F ⟶ □[F ∧ ○F]_v"
proof -
have "|~ □F ⟶ F ∧ ○F" by (rule P2)
hence "⊢ □[□F]_v ⟶ □[F ∧ ○F]_v" by (rule M2)
with ax2 show ?thesis by (rule lift_imp_trans)
qed
theorem M3: assumes h: "⊢ F" shows "⊢ □[○F]_v"
using alw[OF h] T6 by (rule fmp)
lemma M4: "⊢ □[○(F ∧ G) = (○F ∧ ○G)]_v"
by (rule sq[OF T8])
theorem M5: "⊢ □[ □[P]_v ⟶ ○□[P]_v ]_w"
proof (rule sq)
show "|~ □[P]_v ⟶ ○□[P]_v" by (auto simp: pax4[unlifted])
qed
theorem M6: "⊢ □[F ∧ G]_v ⟶ □[F]_v ∧ □[G]_v"
proof -
have "⊢ □[F ∧ G]_v ⟶ □[F]_v" by (rule M2) auto
moreover
have "⊢ □[F ∧ G]_v ⟶ □[G]_v" by (rule M2) auto
ultimately show ?thesis by force
qed
theorem M7: "⊢ □[F]_v ∧ □[G]_v ⟶ □[F ∧ G]_v"
proof -
have "|~ F ⟶ G ⟶ F ∧ G" by auto
hence "⊢ □[F]_v ⟶ □[G ⟶ F ∧ G]_v" by (rule M2)
with ax4 show ?thesis by force
qed
theorem M8: "⊢ □[F ∧ G]_v = (□[F]_v ∧ □[G]_v)"
by (rule int_iffI[OF M6 M7])
theorem M9: "|~ □F ⟶ F ∧ ○□F"
using pre[OF ax1[of "F"]] pax3[of "F"] by force
theorem M10:
assumes h: "|~ F ∧ Unchanged v ⟶ ○F"
shows "|~ F ∧ ○□F ⟶ □F"
using T7[OF h] by auto
theorem M11:
assumes h: "|~ [A]_f ⟶ [B]_g"
shows "⊢ □[A]_f ⟶ □[B]_g"
proof -
from h have "⊢ □[[A]_f]_g ⟶ □[[B]_g]_g" by (rule M2)
with T4 show ?thesis by force
qed
theorem M12: "⊢ (□[A]_f ∧ □[B]_g) = □[[A]_f ∧ [B]_g]_(f,g)"
proof -
have "⊢ □[A]_f ∧ □[B]_g ⟶ □[[A]_f ∧ [B]_g]_(f,g)"
by (auto simp: M8[int_rewrite] elim: T4[unlift_rule])
moreover
have "|~ [[A]_f ∧ [B]_g]_(f,g) ⟶ [A]_f"
by (auto simp: actrans_def unch_def all_before_eq all_after_eq)
hence "⊢ □[[A]_f ∧ [B]_g]_(f,g) ⟶ □[A]_f" by (rule M11)
moreover
have "|~ [[A]_f ∧ [B]_g]_(f,g) ⟶ [B]_g"
by (auto simp: actrans_def unch_def all_before_eq all_after_eq)
hence "⊢ □[[A]_f ∧ [B]_g]_(f,g) ⟶ □[B]_g"
by (rule M11)
ultimately show ?thesis by force
qed
text ‹
We now derive Lamport's 6 simple temporal logic rules (STL1)-(STL6) \<^cite>‹"Lamport94"›.
Firstly, STL1 is the same as @{thm alw} derived above.
›
lemmas STL1 = alw
text ‹
STL2 and STL3 have also already been derived.
›
lemmas STL2 = ax1
lemmas STL3 = T1
text ‹
As with the derivation of @{thm alw}, a purely syntactic derivation of
(STL4) relies on an additional argument -- either using ‹Unchanged›
or ‹STUTINV›.
›
theorem STL4_2:
assumes h1: "⊢ F ⟶ G" and h2: "|~ G ∧ Unchanged v ⟶ ○G"
shows "⊢ □F ⟶ □G"
proof -
from ax1[of F] h1 have "⊢ □F ⟶ G" by (rule lift_imp_trans)
moreover
from h1 have "|~ ○F ⟶ ○G" by (rule N1)
hence "|~ ○F ⟶ G ⟶ ○G" by auto
hence "⊢ □[○F]_v ⟶ □[G ⟶ ○G]_v" by (rule M2)
with T6 have "⊢ □F ⟶ □[G ⟶ ○G]_v" by (rule lift_imp_trans)
moreover
from h2 have "⊢ □[G ⟶ ○G]_v ⟶ G ⟶ □G" by (rule ax3)
ultimately
show ?thesis by force
qed
lemma STL4_3:
assumes h1: "⊢ F ⟶ G" and h2: "STUTINV G"
shows "⊢ □F ⟶ □G"
using h1 h2[THEN pre_id_unch] by (rule STL4_2)
text ‹Of course, the original rule can be derived semantically›
lemma STL4: assumes h: "⊢ F ⟶ G" shows "⊢ □F ⟶ □G"
using h by (force simp: always_def)
text ‹Dual rule for ‹◇››
lemma STL4_eve: assumes h: "⊢ F ⟶ G" shows "⊢ ◇F ⟶ ◇G"
using h by (force simp: eventually_defs)
text‹
Similarly, a purely syntactic derivation of (STL5) requires extra hypotheses.
›
theorem STL5_2:
assumes h1: "|~ F ∧ Unchanged f ⟶ ○F"
and h2: "|~ G ∧ Unchanged g ⟶ ○G"
shows "⊢ □(F ∧ G) = (□F ∧ □G)"
proof (rule int_iffI)
have "⊢ F ∧ G ⟶ F" by auto
from this h1 have "⊢ □(F ∧ G) ⟶ □F" by (rule STL4_2)
moreover
have "⊢ F ∧ G ⟶ G" by auto
from this h2 have "⊢ □(F ∧ G) ⟶ □G" by (rule STL4_2)
ultimately show "⊢ □(F ∧ G) ⟶ □F ∧ □G" by force
next
have "|~ Unchanged (f,g) ⟶ Unchanged f ∧ Unchanged g" by (auto simp: unch_defs)
with h1[unlifted] h2[unlifted] T8[of F G, unlifted]
have h3: "|~ (F ∧ G) ∧ Unchanged (f,g) ⟶ ○(F ∧ G)" by force
from ax1[of F] ax1[of G] have "⊢ □F ∧ □G ⟶ F ∧ G" by force
moreover
from ax2[of F] ax2[of G] have "⊢ □F ∧ □G ⟶ □[□F]_(f,g) ∧ □[□G]_(f,g)" by force
with M8 have "⊢ □F ∧ □G ⟶ □[□F ∧ □G]_(f,g)" by force
moreover
from P1[of F] P1[of G] have "|~ □F ∧ □G ⟶ F ∧ G ⟶ ○(F ∧ G)"
unfolding T8[int_rewrite] by force
hence "⊢ □[ □F ∧ □G ]_(f,g) ⟶ □[F ∧ G ⟶ ○(F ∧ G)]_(f,g)" by (rule M2)
from this ax3[OF h3] have "⊢ □[ □F ∧ □G ]_(f,g) ⟶ F ∧ G ⟶ □(F ∧ G)"
by (rule lift_imp_trans)
ultimately show "⊢ □F ∧ □G ⟶ □(F ∧ G)" by force
qed
theorem STL5_21:
assumes h1: "stutinv F" and h2: "stutinv G"
shows "⊢ □(F ∧ G) = (□F ∧ □G)"
using h1[THEN pre_id_unch] h2[THEN pre_id_unch] by (rule STL5_2)
text ‹We also derive STL5 semantically.›
lemma STL5: "⊢ □(F ∧ G) = (□F ∧ □G)"
by (auto simp: always_def)
text ‹Elimination rule corresponding to ‹STL5› in unlifted form.›
lemma box_conjE:
assumes "s ⊨ □F" and "s ⊨ □G" and "s ⊨ □(F ∧ G) ⟹ P"
shows "P"
using assms by (auto simp: STL5[unlifted])
lemma box_thin:
assumes h1: "s ⊨ □F" and h2: "PROP W"
shows "PROP W"
using h2 .
text ‹Finally, we derive STL6 (only semantically)›
lemma STL6: "⊢ ◇□(F ∧ G) = (◇□F ∧ ◇□G)"
proof auto
fix w
assume a1: "w ⊨ ◇□F" and a2: "w ⊨ ◇□G"
from a1 obtain nf where nf: "(w |⇩s nf) ⊨ □F" by (auto simp: eventually_defs)
from a2 obtain ng where ng: "(w |⇩s ng) ⊨ □G" by (auto simp: eventually_defs)
let ?n = "max nf ng"
have "nf ≤ ?n" by simp
from this nf have "(w |⇩s ?n) ⊨ □F" by (rule linalw)
moreover
have "ng ≤ ?n" by simp
from this ng have "(w |⇩s ?n) ⊨ □G" by (rule linalw)
ultimately
have "(w |⇩s ?n) ⊨ □(F ∧ G)" by (rule box_conjE)
thus "w ⊨ ◇□(F ∧ G)" by (auto simp: eventually_defs)
next
fix w
assume h: "w ⊨ ◇□(F ∧ G)"
have "⊢ F ∧ G ⟶ F" by auto
hence "⊢ ◇□(F ∧ G) ⟶ ◇□F" by (rule STL4_eve[OF STL4])
with h show "w ⊨ ◇□F" by auto
next
fix w
assume h: "w ⊨ ◇□(F ∧ G)"
have "⊢ F ∧ G ⟶ G" by auto
hence "⊢ ◇□(F ∧ G) ⟶ ◇□G" by (rule STL4_eve[OF STL4])
with h show "w ⊨ ◇□G" by auto
qed
lemma MM0: "⊢ □(F ⟶ G) ⟶ □F ⟶ □G"
proof -
have "⊢ □(F ∧ (F ⟶ G)) ⟶ □G" by (rule STL4) auto
thus ?thesis by (auto simp: STL5[int_rewrite])
qed
lemma MM1: assumes h: "⊢ F = G" shows "⊢ □F = □G"
by (auto simp: h[int_rewrite])
theorem MM2: "⊢ □A ∧ □(B ⟶ C) ⟶ □(A ∧ B ⟶ C)"
proof -
have "⊢ □(A ∧ (B ⟶ C)) ⟶ □(A ∧ B ⟶ C)" by (rule STL4) auto
thus ?thesis by (auto simp: STL5[int_rewrite])
qed
theorem MM3: "⊢ □¬A ⟶ □(A ∧ B ⟶ C)"
by (rule STL4) auto
theorem MM4[simp_unl]: "⊢ □#F = #F"
proof (cases "F")
assume "F"
hence 1: "⊢ #F" by auto
hence "⊢ □#F" by (rule alw)
with 1 show ?thesis by force
next
assume "¬F"
hence 1: "⊢ ¬ #F" by auto
from ax1 have "⊢ ¬ #F ⟶ ¬ □#F" by (rule lift_imp_neg)
with 1 show ?thesis by force
qed
theorem MM4b[simp_unl]: "⊢ □¬#F = ¬#F"
proof -
have "⊢ ¬#F = #(¬F)" by auto
hence "⊢ □¬#F = □#(¬F)" by (rule MM1)
thus ?thesis by auto
qed
theorem MM5: "⊢ □F ∨ □G ⟶ □(F ∨ G)"
proof -
have "⊢ □F ⟶ □(F ∨ G)" by (rule STL4) auto
moreover
have "⊢ □G ⟶ □(F ∨ G)" by (rule STL4) auto
ultimately show ?thesis by force
qed
theorem MM6: "⊢ □F ∨ □G ⟶ □(□F ∨ □G)"
proof -
have "⊢ □□F ∨ □□G ⟶ □(□F ∨ □G)" by (rule MM5)
thus ?thesis by simp
qed
lemma MM10:
assumes h: "|~ F = G" shows "⊢ □[F]_v = □[G]_v"
by (auto simp: h[int_rewrite])
lemma MM9:
assumes h: "⊢ F = G" shows "⊢ □[F]_v = □[G]_v"
by (rule MM10[OF pre[OF h]])
theorem MM11: "⊢ □[¬(P ∧ Q)]_v ⟶ □[P]_v ⟶ □[P ∧ ¬Q]_v"
proof -
have "⊢ □[¬(P ∧ Q)]_v ⟶ □[P ⟶ P ∧ ¬Q]_v" by (rule M2) auto
from this ax4 show ?thesis by (rule lift_imp_trans)
qed
theorem MM12[simp_unl]: "⊢ □[□[P]_v]_v = □[P]_v"
proof (rule int_iffI)
have "|~ □[P]_v ⟶ [P]_v" by (auto simp: pax4[unlifted])
hence "⊢ □[□[P]_v]_v ⟶ □[[P]_v]_v" by (rule M2)
thus "⊢ □[□[P]_v]_v ⟶ □[P]_v" by (unfold T3[int_rewrite])
next
have "⊢ □□[P]_v ⟶ □[□□[P]_v]_v" by (rule ax2)
thus "⊢ □[P]_v ⟶ □[□[P]_v]_v" by auto
qed
subsection "Theorems about the eventually operator"
theorem dualization:
"⊢ ¬□F = ◇¬F"
"⊢ ¬◇F = □¬F"
"⊢ ¬□[A]_v = ◇⟨¬A⟩_v"
"⊢ ¬◇⟨A⟩_v = □[¬A]_v"
unfolding eventually_def angle_action_def by simp_all
lemmas dualization_rew = dualization[int_rewrite]
lemmas dualization_unl = dualization[unlifted]
theorem E1: "⊢ ◇(F ∨ G) = (◇F ∨ ◇G)"
proof -
have "⊢ □¬(F ∨ G) = □(¬F ∧ ¬G)" by (rule MM1) auto
thus ?thesis unfolding eventually_def STL5[int_rewrite] by force
qed
theorem E3: "⊢ F ⟶ ◇F"
unfolding eventually_def by (force dest: ax1[unlift_rule])
theorem E4: "⊢ □F ⟶ ◇F"
by (rule lift_imp_trans[OF ax1 E3])
theorem E5: "⊢ □F ⟶ □◇F"
proof -
have "⊢ □□F ⟶ □◇F" by (rule STL4[OF E4])
thus ?thesis by simp
qed
theorem E6: "⊢ □F ⟶ ◇□F"
using E4[of "TEMP □F"] by simp
theorem E7:
assumes h: "|~ ¬F ∧ Unchanged v ⟶ ○¬F"
shows "|~ ◇F ⟶ F ∨ ○◇F"
proof -
from h have "|~ ¬F ∧ ○□¬F ⟶ □¬F" by (rule M10)
thus ?thesis by (auto simp: eventually_def)
qed
theorem E8: "⊢ ◇(F ⟶ G) ⟶ □F ⟶ ◇G"
proof -
have "⊢ □(F ∧ ¬G) ⟶ □¬(F ⟶ G)" by (rule STL4) auto
thus ?thesis unfolding eventually_def STL5[int_rewrite] by auto
qed
theorem E9: "⊢ □(F ⟶ G) ⟶ ◇F ⟶ ◇G"
proof -
have "⊢ □(F ⟶ G) ⟶ □(¬G ⟶ ¬F)" by (rule STL4) auto
with MM0[of "TEMP ¬G" "TEMP ¬F"] show ?thesis unfolding eventually_def by force
qed
theorem E10[simp_unl]: "⊢ ◇◇F = ◇F"
by (simp add: eventually_def)
theorem E22:
assumes h: "⊢ F = G"
shows "⊢ ◇F = ◇G"
by (auto simp: h[int_rewrite])
theorem E15[simp_unl]: "⊢ ◇#F = #F"
by (simp add: eventually_def)
theorem E15b[simp_unl]: "⊢ ◇¬#F = ¬#F"
by (simp add: eventually_def)
theorem E16: "⊢ ◇□F ⟶ ◇F"
by (rule STL4_eve[OF ax1])
text ‹An action version of STL6›
lemma STL6_act: "⊢ ◇(□[F]_v ∧ □[G]_w) = (◇□[F]_v ∧ ◇□[G]_w)"
proof -
have "⊢ (◇□(□[F]_v ∧ □[G]_w)) = ◇(□□[F]_v ∧ □□[G]_w)" by (rule E22[OF STL5])
thus ?thesis by (auto simp: STL6[int_rewrite])
qed
lemma SE1: "⊢ □F ∧ ◇G ⟶ ◇(□F ∧ G)"
proof -
have "⊢ □¬(□F ∧ G) ⟶ □(□F ⟶ ¬G)" by (rule STL4) auto
with MM0 show ?thesis by (force simp: eventually_def)
qed
lemma SE2: "⊢ □F ∧ ◇G ⟶ ◇(F ∧ G)"
proof -
have "⊢ □F ∧ G ⟶ F ∧ G" by (auto elim: ax1[unlift_rule])
hence "⊢ ◇(□F ∧ G) ⟶ ◇(F ∧ G)" by (rule STL4_eve)
with SE1 show ?thesis by (rule lift_imp_trans)
qed
lemma SE3: "⊢ □F ∧ ◇G ⟶ ◇(G ∧ F)"
proof -
have "⊢ ◇(F ∧ G) ⟶ ◇(G ∧ F)" by (rule STL4_eve) auto
with SE2 show ?thesis by (rule lift_imp_trans)
qed
lemma SE4:
assumes h1: "s ⊨ □F" and h2: "s ⊨ ◇G" and h3: "⊢ □F ∧ G ⟶ H"
shows "s ⊨ ◇H"
using h1 h2 h3[THEN STL4_eve] SE1 by force
theorem E17: "⊢ □◇□F ⟶ □◇F"
by (rule STL4[OF STL4_eve[OF ax1]])
theorem E18: "⊢ □◇□F ⟶ ◇□F"
by (rule ax1)
theorem E19: "⊢ ◇□F ⟶ □◇□F"
proof -
have "⊢ (□F ∧ ¬□F) = #False" by auto
hence "⊢ ◇□(□F ∧ ¬□F) = ◇□#False" by (rule E22[OF MM1])
thus ?thesis unfolding STL6[int_rewrite] by (auto simp: eventually_def)
qed
theorem E20: "⊢ ◇□F ⟶ □◇F"
by (rule lift_imp_trans[OF E19 E17])
theorem E21[simp_unl]: "⊢ □◇□F = ◇□F"
by (rule int_iffI[OF E18 E19])
theorem E27[simp_unl]: "⊢ ◇□◇F = □◇F"
using E21 unfolding eventually_def by force
lemma E28: "⊢ ◇□F ∧ □◇G ⟶ □◇(F ∧ G)"
proof -
have "⊢ ◇□(□F ∧ ◇G) ⟶ ◇□◇(F ∧ G)" by (rule STL4_eve[OF STL4[OF SE2]])
thus ?thesis by (simp add: STL6[int_rewrite])
qed
lemma E23: "|~ ○F ⟶ ◇F"
using P1 by (force simp: eventually_def)
lemma E24: "⊢ ◇□Q ⟶ □[◇Q]_v"
by (rule lift_imp_trans[OF E20 P4])
lemma E25: "⊢ ◇⟨A⟩_v ⟶ ◇A"
using P4 by (force simp: eventually_def angle_action_def)
lemma E26: "⊢ □◇⟨A⟩_v ⟶ □◇A"
by (rule STL4[OF E25])
lemma allBox: "(s ⊨ □(∀x. F x)) = (∀x. s ⊨ □(F x))"
unfolding allT[unlifted] ..
lemma E29: "|~ ○◇F ⟶ ◇F"
unfolding eventually_def using pax3 by force
lemma E30:
assumes h1: "⊢ F ⟶ □F" and h2: "⊢ ◇F"
shows "⊢ ◇□F"
using h2 h1[THEN STL4_eve] by (rule fmp)
lemma E31: "⊢ □(F ⟶ □F) ∧ ◇F ⟶ ◇□F"
proof -
have "⊢ □(F ⟶ □F) ∧ ◇F ⟶ ◇(□(F ⟶ □F) ∧ F)" by (rule SE1)
moreover
have "⊢ □(F ⟶ □F) ∧ F ⟶ □F" using ax1[of "TEMP F ⟶ □F"] by auto
hence "⊢ ◇(□(F ⟶ □F) ∧ F) ⟶ ◇□F" by (rule STL4_eve)
ultimately show ?thesis by (rule lift_imp_trans)
qed
lemma allActBox: "(s ⊨ □[(∀x. F x)]_v) = (∀x. s ⊨ □[(F x)]_v)"
unfolding allActT[unlifted] ..
theorem exEE: "⊢ (∃x. ◇(F x)) = ◇(∃x. F x)"
proof -
have "⊢ ¬(∃ x. ◇(F x)) = ¬◇(∃ x. F x)"
by (auto simp: eventually_def Not_Rex[int_rewrite] allBox)
thus ?thesis by force
qed
theorem exActE: "⊢ (∃x. ◇⟨F x⟩_v) = ◇⟨(∃x. F x)⟩_v"
proof -
have "⊢ ¬(∃x. ◇⟨F x⟩_v) = ¬◇⟨(∃x. F x)⟩_v"
by (auto simp: angle_action_def Not_Rex[int_rewrite] allActBox)
thus ?thesis by force
qed
subsection "Theorems about the leadsto operator"
theorem LT1: "⊢ F ↝ F"
unfolding leadsto_def by (rule alw[OF E3])
theorem LT2: assumes h: "⊢ F ⟶ G" shows "⊢ F ⟶ ◇G"
by (rule lift_imp_trans[OF h E3])
theorem LT3: assumes h: "⊢ F ⟶ G" shows "⊢ F ↝ G"
unfolding leadsto_def by (rule alw[OF LT2[OF h]])
theorem LT4: "⊢ F ⟶ (F ↝ G) ⟶ ◇G"
unfolding leadsto_def using ax1[of "TEMP F ⟶ ◇G"] by auto
theorem LT5: "⊢ □(F ⟶ ◇G) ⟶ ◇F ⟶ ◇G"
using E9[of "F" "TEMP ◇G"] by simp
theorem LT6: "⊢ ◇F ⟶ (F ↝ G) ⟶ ◇G"
unfolding leadsto_def using LT5[of "F" "G"] by auto
theorem LT9[simp_unl]: "⊢ □(F ↝ G) = (F ↝ G)"
by (auto simp: leadsto_def)
theorem LT7: "⊢ □◇F ⟶ (F ↝ G) ⟶ □◇G"
proof -
have "⊢ □◇F ⟶ □((F ↝ G) ⟶ ◇G)" by (rule STL4[OF LT6])
from lift_imp_trans[OF this MM0] show ?thesis by simp
qed
theorem LT8: "⊢ □◇G ⟶ (F ↝ G)"
unfolding leadsto_def by (rule STL4) auto
theorem LT13: "⊢ (F ↝ G) ⟶ (G ↝ H) ⟶ (F ↝ H)"
proof -
have "⊢ ◇G ⟶ (G ↝ H) ⟶ ◇H" by (rule LT6)
hence "⊢ □(F ⟶ ◇G) ⟶ □((G ↝ H) ⟶ (F ⟶ ◇H))" by (intro STL4) auto
from lift_imp_trans[OF this MM0] show ?thesis by (simp add: leadsto_def)
qed
theorem LT11: "⊢ (F ↝ G) ⟶ (F ↝ (G ∨ H))"
proof -
have "⊢ G ↝ (G ∨ H)" by (rule LT3) auto
with LT13[of "F" "G" "TEMP (G ∨ H)"] show ?thesis by force
qed
theorem LT12: "⊢ (F ↝ H) ⟶ (F ↝ (G ∨ H))"
proof -
have "⊢ H ↝ (G ∨ H)" by (rule LT3) auto
with LT13[of "F" "H" "TEMP (G ∨ H)"] show ?thesis by force
qed
theorem LT14: "⊢ ((F ∨ G) ↝ H) ⟶ (F ↝ H)"
unfolding leadsto_def by (rule STL4) auto
theorem LT15: "⊢ ((F ∨ G) ↝ H) ⟶ (G ↝ H)"
unfolding leadsto_def by (rule STL4) auto
theorem LT16: "⊢ (F ↝ H) ⟶ (G ↝ H) ⟶ ((F ∨ G) ↝ H)"
proof -
have "⊢ □(F ⟶ ◇H) ⟶ □((G ⟶ ◇H) ⟶ (F ∨ G ⟶ ◇H))" by (rule STL4) auto
from lift_imp_trans[OF this MM0] show ?thesis by (unfold leadsto_def)
qed
theorem LT17: "⊢ ((F ∨ G) ↝ H) = ((F ↝ H) ∧ (G ↝ H))"
by (auto elim: LT14[unlift_rule] LT15[unlift_rule]
LT16[unlift_rule])
theorem LT10:
assumes h: "⊢ (F ∧ ¬G) ↝ G"
shows "⊢ F ↝ G"
proof -
from h have "⊢ ((F ∧ ¬G) ∨ G) ↝ G"
by (auto simp: LT17[int_rewrite] LT1[int_rewrite])
moreover
have "⊢ F ↝ ((F ∧ ¬G) ∨ G)" by (rule LT3, auto)
ultimately
show ?thesis by (force elim: LT13[unlift_rule])
qed
theorem LT18: "⊢ (A ↝ (B ∨ C)) ⟶ (B ↝ D) ⟶ (C ↝ D) ⟶ (A ↝ D)"
proof -
have "⊢ (B ↝ D) ⟶ (C ↝ D) ⟶ ((B ∨ C) ↝ D)" by (rule LT16)
thus ?thesis by (force elim: LT13[unlift_rule])
qed
theorem LT19: "⊢ (A ↝ (D ∨ B)) ⟶ (B ↝ D) ⟶ (A ↝ D)"
using LT18[of "A" "D" "B" "D"] LT1[of "D"] by force
theorem LT20: "⊢ (A ↝ (B ∨ D)) ⟶ (B ↝ D) ⟶ (A ↝ D)"
using LT18[of "A" "B" "D" "D"] LT1[of "D"] by force
theorem LT21: "⊢ ((∃x. F x) ↝ G) = (∀x. (F x ↝ G))"
proof -
have "⊢ □((∃x. F x) ⟶ ◇G) = □(∀x. (F x ⟶ ◇G))" by (rule MM1) auto
thus ?thesis by (unfold leadsto_def allT[int_rewrite])
qed
theorem LT22: "⊢ (F ↝ (G ∨ H)) ⟶ □¬G ⟶ (F ↝ H)"
proof -
have "⊢ □¬G ⟶ (G ↝ H)" unfolding leadsto_def by (rule STL4) auto
thus ?thesis by (force elim: LT20[unlift_rule])
qed
lemma LT23: "|~ (P ⟶ ○Q) ⟶ (P ⟶ ◇Q)"
by (auto dest: E23[unlift_rule])
theorem LT24: "⊢ □I ⟶ ((P ∧ I) ↝ Q) ⟶ P ↝ Q"
proof -
have "⊢ □I ⟶ □((P ∧ I ⟶ ◇Q) ⟶ (P ⟶ ◇Q))" by (rule STL4) auto
from lift_imp_trans[OF this MM0] show ?thesis by (unfold leadsto_def)
qed
theorem LT25[simp_unl]: "⊢ (F ↝ #False) = □¬F"
unfolding leadsto_def proof (rule MM1)
show "⊢ (F ⟶ ◇#False) = ¬F" by simp
qed
lemma LT28:
assumes h: "|~ P ⟶ ○P ∨ ○Q"
shows "|~ (P ⟶ ○P) ∨ ◇Q"
using h E23[of Q] by force
lemma LT29:
assumes h1: "|~ P ⟶ ○P ∨ ○Q" and h2: "|~ P ∧ Unchanged v ⟶ ○P"
shows "⊢ P ⟶ □P ∨ ◇Q"
proof -
from h1[THEN LT28] have "|~ □¬Q ⟶ (P ⟶ ○P)" unfolding eventually_def by auto
hence "⊢ □[□¬Q]_v ⟶ □[P ⟶ ○P]_v" by (rule M2)
moreover
have "⊢ ¬◇Q ⟶ □[□¬Q]_v" unfolding dualization_rew by (rule ax2)
moreover
note ax3[OF h2]
ultimately
show ?thesis by force
qed
lemma LT30:
assumes h: "|~ P ∧ N ⟶ ○P ∨ ○Q"
shows "|~ N ⟶ (P ⟶ ○P) ∨ ◇Q"
using h E23 by force
lemma LT31:
assumes h1: "|~ P ∧ N ⟶ ○P ∨ ○Q" and h2: "|~ P ∧ Unchanged v ⟶ ○P"
shows"⊢ □N ⟶ P ⟶ □P ∨ ◇Q"
proof -
from h1[THEN LT30] have "|~ N ⟶ □¬Q ⟶ P ⟶ ○P" unfolding eventually_def by auto
hence "⊢ □[N ⟶ □¬Q ⟶ P ⟶ ○P]_v" by (rule sq)
hence "⊢ □[N]_v ⟶ □[□¬Q]_v ⟶ □[P ⟶ ○P]_v"
by (force intro: ax4[unlift_rule])
with P4 have "⊢ □N ⟶ □[□¬Q]_v ⟶ □[P ⟶ ○P]_v" by (rule lift_imp_trans)
moreover
have "⊢ ¬◇Q ⟶ □[□¬Q]_v" unfolding dualization_rew by (rule ax2)
moreover
note ax3[OF h2]
ultimately
show ?thesis by force
qed
lemma LT33: "⊢ ((#P ∧ F) ↝ G) = (#P ⟶ (F ↝ G))"
by (cases "P", auto simp: leadsto_def)
lemma AA1: "⊢ □[#False]_v ⟶ ¬◇⟨Q⟩_v"
unfolding dualization_rew by (rule M2) auto
lemma AA2: "⊢ □[P]_v ∧ ◇⟨Q⟩_v ⟶ ◇⟨P ∧ Q⟩_v"
proof -
have "⊢ □[P ⟶ ~(P ∧ Q) ⟶ ¬Q]_v" by (rule sq) (auto simp: actrans_def)
hence "⊢ □[P]_v ⟶ □[~(P ∧ Q)]_v ⟶ □[¬Q]_v"
by (force intro: ax4[unlift_rule])
thus ?thesis by (auto simp: angle_action_def)
qed
lemma AA3: "⊢ □P ∧ □[P ⟶ Q]_v ∧ ◇⟨A⟩_v ⟶ ◇Q"
proof -
have "⊢ □P ∧ □[P ⟶ Q]_v ⟶ □[P ∧ (P ⟶ Q)]_v"
by (auto dest: P4[unlift_rule] simp: M8[int_rewrite])
moreover
have "⊢ □[P ∧ (P ⟶ Q)]_v ⟶ □[Q]_v" by (rule M2) auto
ultimately have "⊢ □P ∧ □[P ⟶ Q]_v ⟶ □[Q]_v" by (rule lift_imp_trans)
moreover
have "⊢ ◇(Q ∧ A) ⟶ ◇Q" by (rule STL4_eve) auto
hence "⊢ ◇⟨Q ∧ A⟩_v ⟶ ◇Q" by (force dest: E25[unlift_rule])
with AA2 have "⊢ □[Q]_v ∧ ◇⟨A⟩_v ⟶ ◇Q" by (rule lift_imp_trans)
ultimately show ?thesis by force
qed
lemma AA4: "⊢ ◇⟨⟨A⟩_v⟩_w ⟶ ◇⟨⟨A⟩_w⟩_v"
unfolding angle_action_def angle_actrans_def using T5 by force
lemma AA7: assumes h: "|~ F ⟶ G" shows "⊢ ◇⟨F⟩_v ⟶ ◇⟨G⟩_v"
proof -
from h have "⊢ □[¬G]_v ⟶ □[¬F]_v" by (intro M2) auto
thus ?thesis unfolding angle_action_def by force
qed
lemma AA6: "⊢ □[P ⟶ Q]_v ∧ ◇⟨P⟩_v ⟶ ◇⟨Q⟩_v"
proof -
have "⊢ ◇⟨(P ⟶ Q) ∧ P⟩_v ⟶ ◇⟨Q⟩_v" by (rule AA7) auto
with AA2 show ?thesis by (rule lift_imp_trans)
qed
lemma AA8: "⊢ □[P]_v ∧ ◇⟨A⟩_v ⟶ ◇⟨□[P]_v ∧ A⟩_v"
proof -
have "⊢ □[□[P]_v]_v ∧ ◇⟨A⟩_v ⟶ ◇⟨□[P]_v ∧ A⟩_v" by (rule AA2)
with P5 show ?thesis by force
qed
lemma AA9: "⊢ □[P]_v ∧ ◇⟨A⟩_v ⟶ ◇⟨[P]_v ∧ A⟩_v"
proof -
have "⊢ □[[P]_v]_v ∧ ◇⟨A⟩_v ⟶ ◇⟨[P]_v ∧ A⟩_v" by (rule AA2)
thus ?thesis by simp
qed
lemma AA10: "⊢ ¬(□[P]_v ∧ ◇⟨¬P⟩_v)"
unfolding angle_action_def by auto
lemma AA11: "⊢ ¬◇⟨v$ = $v⟩_v"
unfolding dualization_rew by (rule ax5)
lemma AA15: "⊢ ◇⟨P ∧ Q⟩_v ⟶ ◇⟨P⟩_v"
by (rule AA7) auto
lemma AA16: "⊢ ◇⟨P ∧ Q⟩_v ⟶ ◇⟨Q⟩_v"
by (rule AA7) auto
lemma AA13: "⊢ ◇⟨P⟩_v ⟶ ◇⟨v$ ≠ $v⟩_v"
proof -
have "⊢ □[v$ ≠ $v]_v ∧ ◇⟨P⟩_v ⟶ ◇⟨v$ ≠ $v ∧ P⟩_v" by (rule AA2)
hence "⊢ ◇⟨P⟩_v ⟶ ◇⟨v$ ≠ $v ∧ P⟩_v" by (simp add: ax5[int_rewrite])
from this AA15 show ?thesis by (rule lift_imp_trans)
qed
lemma AA14: "⊢ ◇⟨P ∨ Q⟩_v = (◇⟨P⟩_v ∨ ◇⟨Q⟩_v)"
proof -
have "⊢ □[¬(P ∨ Q)]_v = □[¬P ∧ ¬Q]_v" by (rule MM10) auto
hence "⊢ □[¬(P ∨ Q)]_v = (□[¬P]_v ∧ □[¬Q]_v)" by (unfold M8[int_rewrite])
thus ?thesis unfolding angle_action_def by auto
qed
lemma AA17: "⊢ ◇⟨[P]_v ∧ A⟩_v ⟶ ◇⟨P ∧ A⟩_v"
proof -
have "⊢ □[v$ ≠ $v ∧ ¬(P ∧ A)]_v ⟶ □[¬([P]_v ∧ A)]_v"
by (rule M2) (auto simp: actrans_def unch_def)
with ax5[of "v"] show ?thesis
unfolding angle_action_def M8[int_rewrite] by force
qed
lemma AA19: "⊢ □P ∧ ◇⟨A⟩_v ⟶ ◇⟨P ∧ A⟩_v"
using P4 by (force intro: AA2[unlift_rule])
lemma AA20:
assumes h1: "|~ P ⟶ ○P ∨ ○Q"
and h2: "|~ P ∧ A ⟶ ○Q"
and h3: "|~ P ∧ Unchanged w ⟶ ○P"
shows "⊢ □(□P ⟶ ◇⟨A⟩_v) ⟶ (P ↝ Q)"
proof -
from h2 E23 have "|~ P ∧ A ⟶ ◇Q" by force
hence "⊢ ◇⟨P ∧ A⟩_v ⟶ ◇⟨◇Q⟩_v" by (rule AA7)
with E25[of "TEMP ◇Q" "v"] have "⊢ ◇⟨P ∧ A⟩_v ⟶ ◇Q" by force
with AA19 have "⊢ □P ∧ ◇⟨A⟩_v ⟶ ◇Q" by (rule lift_imp_trans)
with LT29[OF h1 h3] have "⊢ (□P ⟶ ◇⟨A⟩_v) ⟶ (P ⟶ ◇Q)" by force
thus ?thesis unfolding leadsto_def by (rule STL4)
qed
lemma AA21: "|~ ◇⟨○F⟩_v ⟶ ○◇F"
using pax5[of "TEMP ¬F" "v"] unfolding angle_action_def eventually_def by auto
theorem AA24[simp_unl]: "⊢ ◇⟨⟨P⟩_f⟩_f = ◇⟨P⟩_f"
unfolding angle_action_def angle_actrans_def by simp
lemma AA22:
assumes h1: "|~ P ∧ N ⟶ ○P ∨ ○Q"
and h2: "|~ P ∧ N ∧ ⟨A⟩_v ⟶ ○Q"
and h3: "|~ P ∧ Unchanged w ⟶ ○P"
shows "⊢ □N ∧ □(□P ⟶ ◇⟨A⟩_v) ⟶ (P ↝ Q)"
proof -
from h2 have "|~ ⟨(N ∧ P) ∧ A⟩_v ⟶ ○Q" by (auto simp: angle_actrans_sem[int_rewrite])
from pref_imp_trans[OF this E23] have "⊢ ◇⟨⟨(N ∧ P) ∧ A⟩_v⟩_v ⟶ ◇⟨◇Q⟩_v" by (rule AA7)
hence "⊢ ◇⟨(N ∧ P) ∧ A⟩_v ⟶ ◇Q" by (force dest: E25[unlift_rule])
with AA19 have "⊢ □(N ∧ P) ∧ ◇⟨A⟩_v ⟶ ◇Q" by (rule lift_imp_trans)
hence "⊢ □N ∧ □P ∧ ◇⟨A⟩_v ⟶ ◇Q" by (auto simp: STL5[int_rewrite])
with LT31[OF h1 h3] have "⊢ □N ∧ (□P ⟶ ◇⟨A⟩_v) ⟶ (P ⟶ ◇Q)" by force
hence "⊢ □(□N ∧ (□P ⟶ ◇⟨A⟩_v)) ⟶ □(P ⟶ ◇Q)" by (rule STL4)
thus ?thesis by (simp add: leadsto_def STL5[int_rewrite])
qed
lemma AA23:
assumes "|~ P ∧ N ⟶ ○P ∨ ○Q"
and "|~ P ∧ N ∧ ⟨A⟩_v ⟶ ○Q"
and "|~ P ∧ Unchanged w ⟶ ○P"
shows "⊢ □N ∧ □◇⟨A⟩_v ⟶ (P ↝ Q)"
proof -
have "⊢ □◇⟨A⟩_v ⟶ □(□P ⟶ ◇⟨A⟩_v)" by (rule STL4) auto
with AA22[OF assms] show ?thesis by force
qed
lemma AA25:
assumes h: "|~ ⟨P⟩_v ⟶ ⟨Q⟩_w"
shows "⊢ ◇⟨P⟩_v ⟶ ◇⟨Q⟩_w"
proof -
from h have "⊢ ◇⟨⟨P⟩_v⟩_v ⟶ ◇⟨⟨P⟩_w⟩_v"
by (intro AA7) (auto simp: angle_actrans_def actrans_def)
with AA4 have "⊢ ◇⟨P⟩_v ⟶ ◇⟨⟨P⟩_v⟩_w" by force
from this AA7[OF h] have "⊢ ◇⟨P⟩_v ⟶ ◇⟨⟨Q⟩_w⟩_w" by (rule lift_imp_trans)
thus ?thesis by simp
qed
lemma AA26:
assumes h: "|~ ⟨A⟩_v = ⟨B⟩_w"
shows "⊢ ◇⟨A⟩_v = ◇⟨B⟩_w"
proof -
from h have "|~ ⟨A⟩_v ⟶ ⟨B⟩_w" by auto
hence "⊢ ◇⟨A⟩_v ⟶ ◇⟨B⟩_w" by (rule AA25)
moreover
from h have "|~ ⟨B⟩_w ⟶ ⟨A⟩_v" by auto
hence "⊢ ◇⟨B⟩_w ⟶ ◇⟨A⟩_v" by (rule AA25)
ultimately
show ?thesis by force
qed
theorem AA28[simp_unl]: "⊢ ◇◇⟨A⟩_v = ◇⟨A⟩_v"
unfolding eventually_def angle_action_def by simp
theorem AA29: "⊢ □[N]_v ∧ □◇⟨A⟩_v ⟶ □◇⟨N ∧ A⟩_v"
proof -
have "⊢ □(□[N]_v ∧ ◇⟨A⟩_v) ⟶ □◇⟨N ∧ A⟩_v" by (rule STL4[OF AA2])
thus ?thesis by (simp add: STL5[int_rewrite])
qed
theorem AA30[simp_unl]: "⊢ ◇⟨◇⟨P⟩_f⟩_f = ◇⟨P⟩_f"
unfolding angle_action_def by simp
theorem AA31: "⊢ ◇⟨○F⟩_v ⟶ ◇F"
using pref_imp_trans[OF AA21 E29] by auto
lemma AA32[simp_unl]: "⊢ □◇□[A]_v = ◇□[A]_v"
using E21[of "TEMP □[A]_v"] by simp
lemma AA33[simp_unl]: "⊢ ◇□◇⟨A⟩_v = □◇⟨A⟩_v"
using E27[of "TEMP ◇⟨A⟩_v"] by simp
subsection "Lemmas about the next operator"
lemma N2: assumes h: "⊢ F = G" shows "|~ ○F = ○G"
by (simp add: h[int_rewrite])
lemmas next_and = T8
lemma next_or: "|~ ○(F ∨ G) = (○F ∨ ○G)"
proof (rule pref_iffI)
have "|~ ○((F ∨ G) ∧ ¬F) ⟶ ○G" by (rule N1) auto
thus "|~ ○(F ∨ G) ⟶ ○F ∨ ○G" by (auto simp: T8[int_rewrite])
next
have "|~ ○F ⟶ ○(F ∨ G)" by (rule N1) auto
moreover have "|~ ○G ⟶ ○(F ∨ G)" by (rule N1) auto
ultimately show "|~ ○F ∨ ○G ⟶ ○(F ∨ G)" by force
qed
lemma next_imp: "|~ ○(F ⟶ G) = (○F ⟶ ○G)"
proof (rule pref_iffI)
have "|~ ○G ⟶ ○(F ⟶ G)" by (rule N1) auto
moreover have "|~ ○¬F ⟶ ○(F ⟶ G)" by (rule N1) auto
ultimately show "|~ (○F ⟶ ○G) ⟶ ○(F ⟶ G)" by force
qed (rule pax2)
lemmas next_not = pax1
lemma next_eq: "|~ ○(F = G) = (○F = ○G)"
proof -
have "|~ ○(F = G) = ○((F ⟶ G) ∧ (G ⟶ F))" by (rule N2) auto
from this[int_rewrite] show ?thesis
by (auto simp: next_and[int_rewrite] next_imp[int_rewrite])
qed
lemma next_noteq: "|~ ○(F ≠ G) = (○F ≠ ○G)"
by (simp add: next_eq[int_rewrite])
lemma next_const[simp_unl]: "|~ ○#P = #P"
proof (cases "P")
assume "P"
hence 1: "⊢ #P" by auto
hence "|~ ○#P" by (rule nex)
with 1 show ?thesis by force
next
assume "¬P"
hence 1: "⊢ ¬#P" by auto
hence "|~ ○¬#P" by (rule nex)
with 1 show ?thesis by force
qed
text ‹
The following are proved semantically because they are essentially
first-order theorems.
›
lemma next_fun1: "|~ ○f<x> = f<○x>"
by (auto simp: nexts_def)
lemma next_fun2: "|~ ○f<x,y> = f<○x,○y>"
by (auto simp: nexts_def)
lemma next_fun3: "|~ ○f<x,y,z> = f<○x,○y,○z>"
by (auto simp: nexts_def)
lemma next_fun4: "|~ ○f<x,y,z,zz> = f<○x,○y,○z,○zz>"
by (auto simp: nexts_def)
lemma next_forall: "|~ ○(∀ x. P x) = (∀ x. ○ P x)"
by (auto simp: nexts_def)
lemma next_exists: "|~ ○(∃ x. P x) = (∃ x. ○ P x)"
by (auto simp: nexts_def)
lemma next_exists1: "|~ ○(∃! x. P x) = (∃! x. ○ P x)"
by (auto simp: nexts_def)
text ‹
Rewrite rules to push the ``next'' operator inward over connectives.
(Note that axiom ‹pax1› and theorem ‹next_const› are anyway active
as rewrite rules.)
›
lemmas next_commutes[int_rewrite] =
next_and next_or next_imp next_eq
next_fun1 next_fun2 next_fun3 next_fun4
next_forall next_exists next_exists1
lemmas ifs_eq[int_rewrite] = after_fun3 next_fun3 before_fun3
lemmas next_always = pax3
lemma t1: "|~ ○$x = x$"
by (simp add: before_def after_def nexts_def first_tail_second)
text ‹
Theorem ‹next_eventually› should not be used "blindly".
›
lemma next_eventually:
assumes h: "stutinv F"
shows "|~ ◇F ⟶ ¬F ⟶ ○◇F"
proof -
from h have 1: "stutinv (TEMP ¬F)" by (rule stut_not)
have "|~ □¬F = (¬F ∧ ○□¬F)" unfolding T7[OF pre_id_unch[OF 1], int_rewrite] by simp
thus ?thesis by (auto simp: eventually_def)
qed
lemma next_action: "|~ □[P]_v ⟶ ○□[P]_v"
using pax4[of P v] by auto
subsection "Higher Level Derived Rules"
text ‹
In most verification tasks the low-level rules discussed above are not used directly.
Here, we derive some higher-level rules more suitable for verification. In particular,
variants of Lamport's rules ‹TLA1›, ‹TLA2›, ‹INV1› and ‹INV2›
are derived, where ‹|~› is used where appropriate.
›
theorem TLA1:
assumes H: "|~ P ∧ Unchanged f ⟶ ○P"
shows "⊢ □P = (P ∧ □[P ⟶ ○P]_f)"
proof (rule int_iffI)
from ax1[of P] M0[of P f] show "⊢ □P ⟶ P ∧ □[P ⟶ ○P]_f" by force
next
from ax3[OF H] show "⊢ P ∧ □[P ⟶ ○P]_f ⟶ □P" by auto
qed
theorem TLA2:
assumes h1: "⊢ P ⟶ Q"
and h2: "|~ P ∧ ○P ∧ [A]_f ⟶ [B]_g"
shows "⊢ □P ∧ □[A]_f ⟶ □Q ∧ □[B]_g"
proof -
from h2 have "⊢ □[P ∧ ○P ∧ [A]_f]_g ⟶ □[[B]_g]_g" by (rule M2)
hence "⊢ □[P ∧ ○P]_g ∧ □[[A]_f]_g ⟶ □[B]_g" by (auto simp add: M8[int_rewrite])
with M1[of P g] T4[of A f g] have "⊢ □P ∧ □[A]_f ⟶ □[B]_g" by force
with h1[THEN STL4] show ?thesis by force
qed
theorem INV1:
assumes H: "|~ I ∧ [N]_f ⟶ ○I"
shows "⊢ I ∧ □[N]_f ⟶ □I"
proof -
from H have "|~ [N]_f ⟶ I ⟶ ○I" by auto
hence "⊢ □[[N]_f]_f ⟶ □[I ⟶ ○I]_f" by (rule M2)
moreover
from H have "|~ I ∧ Unchanged f ⟶ ○I" by (auto simp: actrans_def)
hence "⊢ □[I ⟶ ○I]_f ⟶ I ⟶ □I" by (rule ax3)
ultimately show ?thesis by force
qed
theorem INV2: "⊢ □I ⟶ □[N]_f = □[N ∧ I ∧ ○I]_f"
proof -
from M1[of I f] have "⊢ □I ⟶ (□[N]_f = □[N]_f ∧ □[I ∧ ○I]_f)" by auto
thus ?thesis by (auto simp: M8[int_rewrite])
qed
lemma R1:
assumes H: "|~ Unchanged w ⟶ Unchanged v"
shows "⊢ □[F]_w ⟶ □[F]_v"
proof -
from H have "|~ [F]_w ⟶ [F]_v" by (auto simp: actrans_def)
thus ?thesis by (rule M11)
qed
theorem invmono:
assumes h1: "⊢ I ⟶ P"
and h2: "|~ P ∧ [N]_f ⟶ ○P"
shows "⊢ I ∧ □[N]_f ⟶ □P"
using h1 INV1[OF h2] by force
theorem preimpsplit:
assumes "|~ I ∧ N ⟶ Q"
and "|~ I ∧ Unchanged v ⟶ Q"
shows "|~ I ∧ [N]_v ⟶ Q"
using assms[unlift_rule] by (auto simp: actrans_def)
theorem refinement1:
assumes h1: "⊢ P ⟶ Q"
and h2: "|~ I ∧ ○I ∧ [A]_f ⟶ [B]_g"
shows "⊢ P ∧ □I ∧ □[A]_f ⟶ Q ∧ □[B]_g"
proof -
have "⊢ I ⟶ #True" by simp
from this h2 have "⊢ □I ∧ □[A]_f ⟶ □#True ∧ □[B]_g" by (rule TLA2)
with h1 show ?thesis by force
qed
theorem inv_join:
assumes "⊢ P ⟶ □Q" and "⊢ P ⟶ □R"
shows "⊢ P ⟶ □(Q ∧ R)"
using assms[unlift_rule] unfolding STL5[int_rewrite] by force
lemma inv_cases: "⊢ □(A ⟶ B) ∧ □(¬A ⟶ B) ⟶ □B"
proof -
have "⊢ □((A ⟶ B) ∧ (¬A ⟶ B)) ⟶ □B" by (rule STL4) auto
thus ?thesis by (simp add: STL5[int_rewrite])
qed
end