| author | wenzelm | 
| Fri, 02 Jan 2009 19:29:18 +0100 | |
| changeset 29316 | 0a7fcdd77f4b | 
| parent 28524 | 644b62cf678f | 
| child 35174 | e15040ae75d7 | 
| permissions | -rw-r--r-- | 
| 4559 | 1 | (* Title: HOLCF/IOA/meta_theory/TLS.thy | 
| 2 | ID: $Id$ | |
| 12218 | 3 | Author: Olaf Müller | 
| 17233 | 4 | *) | 
| 4559 | 5 | |
| 17233 | 6 | header {* A General Temporal Logic *}
 | 
| 4559 | 7 | |
| 17233 | 8 | theory TL | 
| 9 | imports Pred Sequence | |
| 10 | begin | |
| 11 | ||
| 12 | defaultsort type | |
| 4559 | 13 | |
| 14 | types | |
| 17233 | 15 | 'a temporal = "'a Seq predicate" | 
| 4559 | 16 | |
| 17 | ||
| 17233 | 18 | consts | 
| 4559 | 19 | suffix :: "'a Seq => 'a Seq => bool" | 
| 20 | tsuffix :: "'a Seq => 'a Seq => bool" | |
| 21 | ||
| 22 | validT :: "'a Seq predicate => bool" | |
| 23 | ||
| 24 | unlift :: "'a lift => 'a" | |
| 25 | ||
| 26 | Init         ::"'a predicate => 'a temporal"          ("<_>" [0] 1000)
 | |
| 27 | ||
| 28 | Box          ::"'a temporal => 'a temporal"   ("[] (_)" [80] 80)
 | |
| 29 | Diamond      ::"'a temporal => 'a temporal"   ("<> (_)" [80] 80)
 | |
| 17233 | 30 | Next ::"'a temporal => 'a temporal" | 
| 4559 | 31 | Leadsto ::"'a temporal => 'a temporal => 'a temporal" (infixr "~>" 22) | 
| 32 | ||
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19741diff
changeset | 33 | notation (xsymbols) | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19741diff
changeset | 34 |   Box  ("\<box> (_)" [80] 80) and
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19741diff
changeset | 35 |   Diamond  ("\<diamond> (_)" [80] 80) and
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19741diff
changeset | 36 | Leadsto (infixr "\<leadsto>" 22) | 
| 17233 | 37 | |
| 4559 | 38 | defs | 
| 39 | ||
| 17233 | 40 | unlift_def: | 
| 28524 | 41 | "unlift x == (case x of Def y => y)" | 
| 4559 | 42 | |
| 43 | (* this means that for nil and UU the effect is unpredictable *) | |
| 17233 | 44 | Init_def: | 
| 45 | "Init P s == (P (unlift (HD$s)))" | |
| 4559 | 46 | |
| 17233 | 47 | suffix_def: | 
| 48 | "suffix s2 s == ? s1. (Finite s1 & s = s1 @@ s2)" | |
| 4559 | 49 | |
| 17233 | 50 | tsuffix_def: | 
| 4559 | 51 | "tsuffix s2 s == s2 ~= nil & s2 ~= UU & suffix s2 s" | 
| 52 | ||
| 17233 | 53 | Box_def: | 
| 4559 | 54 | "([] P) s == ! s2. tsuffix s2 s --> P s2" | 
| 55 | ||
| 17233 | 56 | Next_def: | 
| 10835 | 57 | "(Next P) s == if (TL$s=UU | TL$s=nil) then (P s) else P (TL$s)" | 
| 4559 | 58 | |
| 17233 | 59 | Diamond_def: | 
| 4559 | 60 | "<> P == .~ ([] (.~ P))" | 
| 61 | ||
| 17233 | 62 | Leadsto_def: | 
| 63 | "P ~> Q == ([] (P .--> (<> Q)))" | |
| 4559 | 64 | |
| 17233 | 65 | validT_def: | 
| 4559 | 66 | "validT P == ! s. s~=UU & s~=nil --> (s |= P)" | 
| 67 | ||
| 19741 | 68 | |
| 69 | lemma simple: "[] <> (.~ P) = (.~ <> [] P)" | |
| 70 | apply (rule ext) | |
| 71 | apply (simp add: Diamond_def NOT_def Box_def) | |
| 72 | done | |
| 73 | ||
| 74 | lemma Boxnil: "nil |= [] P" | |
| 75 | apply (simp add: satisfies_def Box_def tsuffix_def suffix_def nil_is_Conc) | |
| 76 | done | |
| 77 | ||
| 78 | lemma Diamondnil: "~(nil |= <> P)" | |
| 79 | apply (simp add: Diamond_def satisfies_def NOT_def) | |
| 80 | apply (cut_tac Boxnil) | |
| 81 | apply (simp add: satisfies_def) | |
| 82 | done | |
| 83 | ||
| 84 | lemma Diamond_def2: "(<> F) s = (? s2. tsuffix s2 s & F s2)" | |
| 85 | apply (simp add: Diamond_def NOT_def Box_def) | |
| 86 | done | |
| 87 | ||
| 88 | ||
| 89 | ||
| 90 | subsection "TLA Axiomatization by Merz" | |
| 91 | ||
| 92 | lemma suffix_refl: "suffix s s" | |
| 93 | apply (simp add: suffix_def) | |
| 94 | apply (rule_tac x = "nil" in exI) | |
| 95 | apply auto | |
| 96 | done | |
| 97 | ||
| 98 | lemma reflT: "s~=UU & s~=nil --> (s |= [] F .--> F)" | |
| 99 | apply (simp add: satisfies_def IMPLIES_def Box_def) | |
| 100 | apply (rule impI)+ | |
| 101 | apply (erule_tac x = "s" in allE) | |
| 102 | apply (simp add: tsuffix_def suffix_refl) | |
| 103 | done | |
| 104 | ||
| 105 | ||
| 106 | lemma suffix_trans: "[| suffix y x ; suffix z y |] ==> suffix z x" | |
| 107 | apply (simp add: suffix_def) | |
| 108 | apply auto | |
| 109 | apply (rule_tac x = "s1 @@ s1a" in exI) | |
| 110 | apply auto | |
| 111 | apply (simp (no_asm) add: Conc_assoc) | |
| 112 | done | |
| 113 | ||
| 114 | lemma transT: "s |= [] F .--> [] [] F" | |
| 115 | apply (simp (no_asm) add: satisfies_def IMPLIES_def Box_def tsuffix_def) | |
| 116 | apply auto | |
| 117 | apply (drule suffix_trans) | |
| 118 | apply assumption | |
| 119 | apply (erule_tac x = "s2a" in allE) | |
| 120 | apply auto | |
| 121 | done | |
| 122 | ||
| 123 | ||
| 124 | lemma normalT: "s |= [] (F .--> G) .--> [] F .--> [] G" | |
| 125 | apply (simp (no_asm) add: satisfies_def IMPLIES_def Box_def) | |
| 126 | done | |
| 127 | ||
| 128 | ||
| 129 | subsection "TLA Rules by Lamport" | |
| 130 | ||
| 131 | lemma STL1a: "validT P ==> validT ([] P)" | |
| 132 | apply (simp add: validT_def satisfies_def Box_def tsuffix_def) | |
| 133 | done | |
| 134 | ||
| 135 | lemma STL1b: "valid P ==> validT (Init P)" | |
| 136 | apply (simp add: valid_def validT_def satisfies_def Init_def) | |
| 137 | done | |
| 138 | ||
| 139 | lemma STL1: "valid P ==> validT ([] (Init P))" | |
| 140 | apply (rule STL1a) | |
| 141 | apply (erule STL1b) | |
| 142 | done | |
| 143 | ||
| 144 | (* Note that unlift and HD is not at all used !!! *) | |
| 145 | lemma STL4: "valid (P .--> Q) ==> validT ([] (Init P) .--> [] (Init Q))" | |
| 146 | apply (simp add: valid_def validT_def satisfies_def IMPLIES_def Box_def Init_def) | |
| 147 | done | |
| 148 | ||
| 149 | ||
| 150 | subsection "LTL Axioms by Manna/Pnueli" | |
| 151 | ||
| 152 | lemma tsuffix_TL [rule_format (no_asm)]: | |
| 153 | "s~=UU & s~=nil --> tsuffix s2 (TL$s) --> tsuffix s2 s" | |
| 154 | apply (unfold tsuffix_def suffix_def) | |
| 155 | apply auto | |
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
25131diff
changeset | 156 | apply (tactic {* Seq_case_simp_tac @{context} "s" 1 *})
 | 
| 19741 | 157 | apply (rule_tac x = "a>>s1" in exI) | 
| 158 | apply auto | |
| 159 | done | |
| 160 | ||
| 161 | lemmas tsuffix_TL2 = conjI [THEN tsuffix_TL] | |
| 162 | ||
| 163 | declare split_if [split del] | |
| 164 | lemma LTL1: | |
| 165 | "s~=UU & s~=nil --> (s |= [] F .--> (F .& (Next ([] F))))" | |
| 166 | apply (unfold Next_def satisfies_def NOT_def IMPLIES_def AND_def Box_def) | |
| 167 | apply auto | |
| 168 | (* []F .--> F *) | |
| 169 | apply (erule_tac x = "s" in allE) | |
| 170 | apply (simp add: tsuffix_def suffix_refl) | |
| 171 | (* []F .--> Next [] F *) | |
| 172 | apply (simp split add: split_if) | |
| 173 | apply auto | |
| 174 | apply (drule tsuffix_TL2) | |
| 175 | apply assumption+ | |
| 176 | apply auto | |
| 177 | done | |
| 178 | declare split_if [split] | |
| 179 | ||
| 180 | ||
| 181 | lemma LTL2a: | |
| 182 | "s |= .~ (Next F) .--> (Next (.~ F))" | |
| 183 | apply (unfold Next_def satisfies_def NOT_def IMPLIES_def) | |
| 184 | apply simp | |
| 185 | done | |
| 186 | ||
| 187 | lemma LTL2b: | |
| 188 | "s |= (Next (.~ F)) .--> (.~ (Next F))" | |
| 189 | apply (unfold Next_def satisfies_def NOT_def IMPLIES_def) | |
| 190 | apply simp | |
| 191 | done | |
| 192 | ||
| 193 | lemma LTL3: | |
| 194 | "ex |= (Next (F .--> G)) .--> (Next F) .--> (Next G)" | |
| 195 | apply (unfold Next_def satisfies_def NOT_def IMPLIES_def) | |
| 196 | apply simp | |
| 197 | done | |
| 198 | ||
| 199 | ||
| 200 | lemma ModusPonens: "[| validT (P .--> Q); validT P |] ==> validT Q" | |
| 201 | apply (simp add: validT_def satisfies_def IMPLIES_def) | |
| 202 | done | |
| 17233 | 203 | |
| 204 | end |