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