author | nipkow |
Tue, 17 Jun 2025 14:11:40 +0200 | |
changeset 82733 | 8b537e1af2ec |
parent 81734 | 78d95ff11ade |
permissions | -rw-r--r-- |
62008 | 1 |
(* Title: HOL/HOLCF/IOA/TL.thy |
40945 | 2 |
Author: Olaf Müller |
17233 | 3 |
*) |
4559 | 4 |
|
62002 | 5 |
section \<open>A General Temporal Logic\<close> |
4559 | 6 |
|
17233 | 7 |
theory TL |
8 |
imports Pred Sequence |
|
9 |
begin |
|
10 |
||
36452 | 11 |
default_sort type |
4559 | 12 |
|
62005 | 13 |
type_synonym 'a temporal = "'a Seq predicate" |
4559 | 14 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
63648
diff
changeset
|
15 |
definition validT :: "'a Seq predicate \<Rightarrow> bool" (\<open>\<^bold>\<TTurnstile> _\<close> [9] 8) |
62194 | 16 |
where "(\<^bold>\<TTurnstile> P) \<longleftrightarrow> (\<forall>s. s \<noteq> UU \<and> s \<noteq> nil \<longrightarrow> (s \<Turnstile> P))" |
4559 | 17 |
|
62005 | 18 |
definition unlift :: "'a lift \<Rightarrow> 'a" |
19 |
where "unlift x = (case x of Def y \<Rightarrow> y)" |
|
4559 | 20 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
63648
diff
changeset
|
21 |
definition Init :: "'a predicate \<Rightarrow> 'a temporal" (\<open>\<langle>_\<rangle>\<close> [0] 1000) |
63549 | 22 |
where "Init P s = P (unlift (HD \<cdot> s))" |
62005 | 23 |
\<comment> \<open>this means that for \<open>nil\<close> and \<open>UU\<close> the effect is unpredictable\<close> |
17233 | 24 |
|
81095 | 25 |
definition Next :: "'a temporal \<Rightarrow> 'a temporal" (\<open>(\<open>indent=1 notation=\<open>prefix Next\<close>\<close>\<circle>_)\<close> [80] 80) |
63549 | 26 |
where "(\<circle>P) s \<longleftrightarrow> (if TL \<cdot> s = UU \<or> TL \<cdot> s = nil then P s else P (TL \<cdot> s))" |
4559 | 27 |
|
62005 | 28 |
definition suffix :: "'a Seq \<Rightarrow> 'a Seq \<Rightarrow> bool" |
29 |
where "suffix s2 s \<longleftrightarrow> (\<exists>s1. Finite s1 \<and> s = s1 @@ s2)" |
|
4559 | 30 |
|
62005 | 31 |
definition tsuffix :: "'a Seq \<Rightarrow> 'a Seq \<Rightarrow> bool" |
32 |
where "tsuffix s2 s \<longleftrightarrow> s2 \<noteq> nil \<and> s2 \<noteq> UU \<and> suffix s2 s" |
|
4559 | 33 |
|
81095 | 34 |
definition Box :: "'a temporal \<Rightarrow> 'a temporal" (\<open>(\<open>indent=1 notation=\<open>prefix Box\<close>\<close>\<box>_)\<close> [80] 80) |
62005 | 35 |
where "(\<box>P) s \<longleftrightarrow> (\<forall>s2. tsuffix s2 s \<longrightarrow> P s2)" |
4559 | 36 |
|
81095 | 37 |
definition Diamond :: "'a temporal \<Rightarrow> 'a temporal" (\<open>(\<open>indent=1 notation=\<open>prefix Diamond\<close>\<close>\<diamond>_)\<close> [80] 80) |
62005 | 38 |
where "\<diamond>P = (\<^bold>\<not> (\<box>(\<^bold>\<not> P)))" |
4559 | 39 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
63648
diff
changeset
|
40 |
definition Leadsto :: "'a temporal \<Rightarrow> 'a temporal \<Rightarrow> 'a temporal" (infixr \<open>\<leadsto>\<close> 22) |
62005 | 41 |
where "(P \<leadsto> Q) = (\<box>(P \<^bold>\<longrightarrow> (\<diamond>Q)))" |
4559 | 42 |
|
19741 | 43 |
|
62000 | 44 |
lemma simple: "\<box>\<diamond>(\<^bold>\<not> P) = (\<^bold>\<not> \<diamond>\<box>P)" |
62116 | 45 |
by (auto simp add: Diamond_def NOT_def Box_def) |
19741 | 46 |
|
62000 | 47 |
lemma Boxnil: "nil \<Turnstile> \<box>P" |
62116 | 48 |
by (simp add: satisfies_def Box_def tsuffix_def suffix_def nil_is_Conc) |
19741 | 49 |
|
62116 | 50 |
lemma Diamondnil: "\<not> (nil \<Turnstile> \<diamond>P)" |
51 |
using Boxnil by (simp add: Diamond_def satisfies_def NOT_def) |
|
19741 | 52 |
|
62116 | 53 |
lemma Diamond_def2: "(\<diamond>F) s \<longleftrightarrow> (\<exists>s2. tsuffix s2 s \<and> F s2)" |
54 |
by (simp add: Diamond_def NOT_def Box_def) |
|
19741 | 55 |
|
56 |
||
62116 | 57 |
subsection \<open>TLA Axiomatization by Merz\<close> |
19741 | 58 |
|
59 |
lemma suffix_refl: "suffix s s" |
|
81734 | 60 |
proof - |
61 |
have "Finite nil \<and> s = nil @@ s" by simp |
|
62 |
then show ?thesis unfolding suffix_def .. |
|
63 |
qed |
|
64 |
||
65 |
lemma suffix_trans: |
|
66 |
assumes "suffix y x" |
|
67 |
and "suffix z y" |
|
68 |
shows "suffix z x" |
|
69 |
proof - |
|
70 |
from assms obtain s1 s2 |
|
71 |
where "Finite s1 \<and> x = s1 @@ y" |
|
72 |
and "Finite s2 \<and> y = s2 @@ z" |
|
73 |
unfolding suffix_def by iprover |
|
74 |
then have "Finite (s1 @@ s2) \<and> x = (s1 @@ s2) @@ z" |
|
75 |
by (simp add: Conc_assoc) |
|
76 |
then show ?thesis unfolding suffix_def .. |
|
77 |
qed |
|
19741 | 78 |
|
62116 | 79 |
lemma reflT: "s \<noteq> UU \<and> s \<noteq> nil \<longrightarrow> (s \<Turnstile> \<box>F \<^bold>\<longrightarrow> F)" |
81734 | 80 |
proof |
81 |
assume s: "s \<noteq> UU \<and> s \<noteq> nil" |
|
82 |
have "F s" if box_F: "\<forall>s2. tsuffix s2 s \<longrightarrow> F s2" |
|
83 |
proof - |
|
84 |
from s and suffix_refl [of s] have "tsuffix s s" |
|
85 |
by (simp add: tsuffix_def) |
|
86 |
also from box_F have "?this \<longrightarrow> F s" .. |
|
87 |
finally show ?thesis . |
|
88 |
qed |
|
89 |
then show "s \<Turnstile> \<box>F \<^bold>\<longrightarrow> F" |
|
90 |
by (simp add: satisfies_def IMPLIES_def Box_def) |
|
91 |
qed |
|
19741 | 92 |
|
62000 | 93 |
lemma transT: "s \<Turnstile> \<box>F \<^bold>\<longrightarrow> \<box>\<box>F" |
81734 | 94 |
proof - |
95 |
have "F s2" if *: "tsuffix s1 s" "tsuffix s2 s1" |
|
96 |
and **: "\<forall>s'. tsuffix s' s \<longrightarrow> F s'" |
|
97 |
for s1 s2 |
|
98 |
proof - |
|
99 |
from * have "tsuffix s2 s" |
|
100 |
by (auto simp: tsuffix_def elim: suffix_trans) |
|
101 |
also from ** have "?this \<longrightarrow> F s2" .. |
|
102 |
finally show ?thesis . |
|
103 |
qed |
|
104 |
then show ?thesis |
|
105 |
by (simp add: satisfies_def IMPLIES_def Box_def) |
|
106 |
qed |
|
19741 | 107 |
|
62000 | 108 |
lemma normalT: "s \<Turnstile> \<box>(F \<^bold>\<longrightarrow> G) \<^bold>\<longrightarrow> \<box>F \<^bold>\<longrightarrow> \<box>G" |
62116 | 109 |
by (simp add: satisfies_def IMPLIES_def Box_def) |
19741 | 110 |
|
111 |
||
62116 | 112 |
subsection \<open>TLA Rules by Lamport\<close> |
19741 | 113 |
|
62194 | 114 |
lemma STL1a: "\<^bold>\<TTurnstile> P \<Longrightarrow> \<^bold>\<TTurnstile> (\<box>P)" |
62116 | 115 |
by (simp add: validT_def satisfies_def Box_def tsuffix_def) |
19741 | 116 |
|
62194 | 117 |
lemma STL1b: "\<TTurnstile> P \<Longrightarrow> \<^bold>\<TTurnstile> (Init P)" |
62116 | 118 |
by (simp add: valid_def validT_def satisfies_def Init_def) |
19741 | 119 |
|
81734 | 120 |
lemma STL1: assumes "\<TTurnstile> P" shows "\<^bold>\<TTurnstile> (\<box>(Init P))" |
121 |
proof - |
|
122 |
from assms have "\<^bold>\<TTurnstile> (Init P)" by (rule STL1b) |
|
123 |
then show ?thesis by (rule STL1a) |
|
124 |
qed |
|
19741 | 125 |
|
62116 | 126 |
(*Note that unlift and HD is not at all used!*) |
62194 | 127 |
lemma STL4: "\<TTurnstile> (P \<^bold>\<longrightarrow> Q) \<Longrightarrow> \<^bold>\<TTurnstile> (\<box>(Init P) \<^bold>\<longrightarrow> \<box>(Init Q))" |
62116 | 128 |
by (simp add: valid_def validT_def satisfies_def IMPLIES_def Box_def Init_def) |
19741 | 129 |
|
130 |
||
62116 | 131 |
subsection \<open>LTL Axioms by Manna/Pnueli\<close> |
19741 | 132 |
|
63549 | 133 |
lemma tsuffix_TL [rule_format]: "s \<noteq> UU \<and> s \<noteq> nil \<longrightarrow> tsuffix s2 (TL \<cdot> s) \<longrightarrow> tsuffix s2 s" |
62116 | 134 |
apply (unfold tsuffix_def suffix_def) |
135 |
apply auto |
|
62195 | 136 |
apply (Seq_case_simp s) |
62116 | 137 |
apply (rule_tac x = "a \<leadsto> s1" in exI) |
138 |
apply auto |
|
139 |
done |
|
19741 | 140 |
|
141 |
lemmas tsuffix_TL2 = conjI [THEN tsuffix_TL] |
|
142 |
||
62441 | 143 |
lemma LTL1: "s \<noteq> UU \<and> s \<noteq> nil \<longrightarrow> (s \<Turnstile> \<box>F \<^bold>\<longrightarrow> (F \<^bold>\<and> (\<circle>(\<box>F))))" |
62390 | 144 |
supply if_split [split del] |
62116 | 145 |
apply (unfold Next_def satisfies_def NOT_def IMPLIES_def AND_def Box_def) |
146 |
apply auto |
|
147 |
text \<open>\<open>\<box>F \<^bold>\<longrightarrow> F\<close>\<close> |
|
148 |
apply (erule_tac x = "s" in allE) |
|
149 |
apply (simp add: tsuffix_def suffix_refl) |
|
62441 | 150 |
text \<open>\<open>\<box>F \<^bold>\<longrightarrow> \<circle>\<box>F\<close>\<close> |
63648 | 151 |
apply (simp split: if_split) |
62116 | 152 |
apply auto |
153 |
apply (drule tsuffix_TL2) |
|
154 |
apply assumption+ |
|
155 |
apply auto |
|
156 |
done |
|
19741 | 157 |
|
62441 | 158 |
lemma LTL2a: "s \<Turnstile> \<^bold>\<not> (\<circle>F) \<^bold>\<longrightarrow> (\<circle>(\<^bold>\<not> F))" |
62116 | 159 |
by (simp add: Next_def satisfies_def NOT_def IMPLIES_def) |
160 |
||
62441 | 161 |
lemma LTL2b: "s \<Turnstile> (\<circle>(\<^bold>\<not> F)) \<^bold>\<longrightarrow> (\<^bold>\<not> (\<circle>F))" |
62116 | 162 |
by (simp add: Next_def satisfies_def NOT_def IMPLIES_def) |
19741 | 163 |
|
62441 | 164 |
lemma LTL3: "ex \<Turnstile> (\<circle>(F \<^bold>\<longrightarrow> G)) \<^bold>\<longrightarrow> (\<circle>F) \<^bold>\<longrightarrow> (\<circle>G)" |
62116 | 165 |
by (simp add: Next_def satisfies_def NOT_def IMPLIES_def) |
19741 | 166 |
|
62194 | 167 |
lemma ModusPonens: "\<^bold>\<TTurnstile> (P \<^bold>\<longrightarrow> Q) \<Longrightarrow> \<^bold>\<TTurnstile> P \<Longrightarrow> \<^bold>\<TTurnstile> Q" |
62116 | 168 |
by (simp add: validT_def satisfies_def IMPLIES_def) |
17233 | 169 |
|
170 |
end |