| 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 | 
 | 
| 62194 |     15 | definition validT :: "'a Seq predicate \<Rightarrow> bool"    ("\<^bold>\<TTurnstile> _" [9] 8)
 | 
|  |     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 | 
 | 
| 62005 |     21 | definition Init :: "'a predicate \<Rightarrow> 'a temporal"  ("\<langle>_\<rangle>" [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 | 
 | 
| 62441 |     25 | definition Next :: "'a temporal \<Rightarrow> 'a temporal"  ("\<circle>(_)" [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 | 
 | 
| 62005 |     34 | definition Box :: "'a temporal \<Rightarrow> 'a temporal"  ("\<box>(_)" [80] 80)
 | 
|  |     35 |   where "(\<box>P) s \<longleftrightarrow> (\<forall>s2. tsuffix s2 s \<longrightarrow> P s2)"
 | 
| 4559 |     36 | 
 | 
| 62005 |     37 | definition Diamond :: "'a temporal \<Rightarrow> 'a temporal"  ("\<diamond>(_)" [80] 80)
 | 
|  |     38 |   where "\<diamond>P = (\<^bold>\<not> (\<box>(\<^bold>\<not> P)))"
 | 
| 4559 |     39 | 
 | 
| 62005 |     40 | definition Leadsto :: "'a temporal \<Rightarrow> 'a temporal \<Rightarrow> 'a temporal"  (infixr "\<leadsto>" 22)
 | 
|  |     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"
 | 
| 62116 |     60 |   apply (simp add: suffix_def)
 | 
|  |     61 |   apply (rule_tac x = "nil" in exI)
 | 
|  |     62 |   apply auto
 | 
|  |     63 |   done
 | 
| 19741 |     64 | 
 | 
| 62116 |     65 | lemma reflT: "s \<noteq> UU \<and> s \<noteq> nil \<longrightarrow> (s \<Turnstile> \<box>F \<^bold>\<longrightarrow> F)"
 | 
|  |     66 |   apply (simp add: satisfies_def IMPLIES_def Box_def)
 | 
|  |     67 |   apply (rule impI)+
 | 
|  |     68 |   apply (erule_tac x = "s" in allE)
 | 
|  |     69 |   apply (simp add: tsuffix_def suffix_refl)
 | 
|  |     70 |   done
 | 
| 19741 |     71 | 
 | 
| 62116 |     72 | lemma suffix_trans: "suffix y x \<Longrightarrow> suffix z y \<Longrightarrow> suffix z x"
 | 
|  |     73 |   apply (simp add: suffix_def)
 | 
|  |     74 |   apply auto
 | 
|  |     75 |   apply (rule_tac x = "s1 @@ s1a" in exI)
 | 
|  |     76 |   apply auto
 | 
|  |     77 |   apply (simp add: Conc_assoc)
 | 
|  |     78 |   done
 | 
| 19741 |     79 | 
 | 
| 62000 |     80 | lemma transT: "s \<Turnstile> \<box>F \<^bold>\<longrightarrow> \<box>\<box>F"
 | 
| 62116 |     81 |   apply (simp add: satisfies_def IMPLIES_def Box_def tsuffix_def)
 | 
|  |     82 |   apply auto
 | 
|  |     83 |   apply (drule suffix_trans)
 | 
|  |     84 |   apply assumption
 | 
|  |     85 |   apply (erule_tac x = "s2a" in allE)
 | 
|  |     86 |   apply auto
 | 
|  |     87 |   done
 | 
| 19741 |     88 | 
 | 
| 62000 |     89 | lemma normalT: "s \<Turnstile> \<box>(F \<^bold>\<longrightarrow> G) \<^bold>\<longrightarrow> \<box>F \<^bold>\<longrightarrow> \<box>G"
 | 
| 62116 |     90 |   by (simp add: satisfies_def IMPLIES_def Box_def)
 | 
| 19741 |     91 | 
 | 
|  |     92 | 
 | 
| 62116 |     93 | subsection \<open>TLA Rules by Lamport\<close>
 | 
| 19741 |     94 | 
 | 
| 62194 |     95 | lemma STL1a: "\<^bold>\<TTurnstile> P \<Longrightarrow> \<^bold>\<TTurnstile> (\<box>P)"
 | 
| 62116 |     96 |   by (simp add: validT_def satisfies_def Box_def tsuffix_def)
 | 
| 19741 |     97 | 
 | 
| 62194 |     98 | lemma STL1b: "\<TTurnstile> P \<Longrightarrow> \<^bold>\<TTurnstile> (Init P)"
 | 
| 62116 |     99 |   by (simp add: valid_def validT_def satisfies_def Init_def)
 | 
| 19741 |    100 | 
 | 
| 62194 |    101 | lemma STL1: "\<TTurnstile> P \<Longrightarrow> \<^bold>\<TTurnstile> (\<box>(Init P))"
 | 
| 62116 |    102 |   apply (rule STL1a)
 | 
|  |    103 |   apply (erule STL1b)
 | 
|  |    104 |   done
 | 
| 19741 |    105 | 
 | 
| 62116 |    106 | (*Note that unlift and HD is not at all used!*)
 | 
| 62194 |    107 | lemma STL4: "\<TTurnstile> (P \<^bold>\<longrightarrow> Q) \<Longrightarrow> \<^bold>\<TTurnstile> (\<box>(Init P) \<^bold>\<longrightarrow> \<box>(Init Q))"
 | 
| 62116 |    108 |   by (simp add: valid_def validT_def satisfies_def IMPLIES_def Box_def Init_def)
 | 
| 19741 |    109 | 
 | 
|  |    110 | 
 | 
| 62116 |    111 | subsection \<open>LTL Axioms by Manna/Pnueli\<close>
 | 
| 19741 |    112 | 
 | 
| 63549 |    113 | lemma tsuffix_TL [rule_format]: "s \<noteq> UU \<and> s \<noteq> nil \<longrightarrow> tsuffix s2 (TL \<cdot> s) \<longrightarrow> tsuffix s2 s"
 | 
| 62116 |    114 |   apply (unfold tsuffix_def suffix_def)
 | 
|  |    115 |   apply auto
 | 
| 62195 |    116 |   apply (Seq_case_simp s)
 | 
| 62116 |    117 |   apply (rule_tac x = "a \<leadsto> s1" in exI)
 | 
|  |    118 |   apply auto
 | 
|  |    119 |   done
 | 
| 19741 |    120 | 
 | 
|  |    121 | lemmas tsuffix_TL2 = conjI [THEN tsuffix_TL]
 | 
|  |    122 | 
 | 
| 62441 |    123 | lemma LTL1: "s \<noteq> UU \<and> s \<noteq> nil \<longrightarrow> (s \<Turnstile> \<box>F \<^bold>\<longrightarrow> (F \<^bold>\<and> (\<circle>(\<box>F))))"
 | 
| 62390 |    124 |   supply if_split [split del] 
 | 
| 62116 |    125 |   apply (unfold Next_def satisfies_def NOT_def IMPLIES_def AND_def Box_def)
 | 
|  |    126 |   apply auto
 | 
|  |    127 |   text \<open>\<open>\<box>F \<^bold>\<longrightarrow> F\<close>\<close>
 | 
|  |    128 |   apply (erule_tac x = "s" in allE)
 | 
|  |    129 |   apply (simp add: tsuffix_def suffix_refl)
 | 
| 62441 |    130 |   text \<open>\<open>\<box>F \<^bold>\<longrightarrow> \<circle>\<box>F\<close>\<close>
 | 
| 63648 |    131 |   apply (simp split: if_split)
 | 
| 62116 |    132 |   apply auto
 | 
|  |    133 |   apply (drule tsuffix_TL2)
 | 
|  |    134 |   apply assumption+
 | 
|  |    135 |   apply auto
 | 
|  |    136 |   done
 | 
| 19741 |    137 | 
 | 
| 62441 |    138 | lemma LTL2a: "s \<Turnstile> \<^bold>\<not> (\<circle>F) \<^bold>\<longrightarrow> (\<circle>(\<^bold>\<not> F))"
 | 
| 62116 |    139 |   by (simp add: Next_def satisfies_def NOT_def IMPLIES_def)
 | 
|  |    140 | 
 | 
| 62441 |    141 | lemma LTL2b: "s \<Turnstile> (\<circle>(\<^bold>\<not> F)) \<^bold>\<longrightarrow> (\<^bold>\<not> (\<circle>F))"
 | 
| 62116 |    142 |   by (simp add: Next_def satisfies_def NOT_def IMPLIES_def)
 | 
| 19741 |    143 | 
 | 
| 62441 |    144 | lemma LTL3: "ex \<Turnstile> (\<circle>(F \<^bold>\<longrightarrow> G)) \<^bold>\<longrightarrow> (\<circle>F) \<^bold>\<longrightarrow> (\<circle>G)"
 | 
| 62116 |    145 |   by (simp add: Next_def satisfies_def NOT_def IMPLIES_def)
 | 
| 19741 |    146 | 
 | 
| 62194 |    147 | lemma ModusPonens: "\<^bold>\<TTurnstile> (P \<^bold>\<longrightarrow> Q) \<Longrightarrow> \<^bold>\<TTurnstile> P \<Longrightarrow> \<^bold>\<TTurnstile> Q"
 | 
| 62116 |    148 |   by (simp add: validT_def satisfies_def IMPLIES_def)
 | 
| 17233 |    149 | 
 | 
|  |    150 | end
 |