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