src/HOL/HOLCF/IOA/TL.thy
author Manuel Eberl <eberlm@in.tum.de>
Tue, 12 Dec 2017 10:01:14 +0100
changeset 67167 88d1c9d86f48
parent 63648 f9f3006a5579
child 80914 d97fdabd9e2b
permissions -rw-r--r--
Moved analysis material from AFP
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62008
cbedaddc9351 clarified directory structure;
wenzelm
parents: 62005
diff changeset
     1
(*  Title:      HOL/HOLCF/IOA/TL.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
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
     5
section \<open>A General Temporal Logic\<close>
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
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    13
type_synonym 'a temporal = "'a Seq predicate"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    14
62194
0aabc5931361 tuned syntax;
wenzelm
parents: 62116
diff changeset
    15
definition validT :: "'a Seq predicate \<Rightarrow> bool"    ("\<^bold>\<TTurnstile> _" [9] 8)
0aabc5931361 tuned syntax;
wenzelm
parents: 62116
diff changeset
    16
  where "(\<^bold>\<TTurnstile> P) \<longleftrightarrow> (\<forall>s. s \<noteq> UU \<and> s \<noteq> nil \<longrightarrow> (s \<Turnstile> P))"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    17
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    18
definition unlift :: "'a lift \<Rightarrow> 'a"
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    19
  where "unlift x = (case x of Def y \<Rightarrow> y)"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    20
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    21
definition Init :: "'a predicate \<Rightarrow> 'a temporal"  ("\<langle>_\<rangle>" [0] 1000)
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62441
diff changeset
    22
  where "Init P s = P (unlift (HD \<cdot> s))"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    23
    \<comment> \<open>this means that for \<open>nil\<close> and \<open>UU\<close> the effect is unpredictable\<close>
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    24
62441
e5e38e1f2dd4 more symbols;
wenzelm
parents: 62390
diff changeset
    25
definition Next :: "'a temporal \<Rightarrow> 'a temporal"  ("\<circle>(_)" [80] 80)
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62441
diff changeset
    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
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    27
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    28
definition suffix :: "'a Seq \<Rightarrow> 'a Seq \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    29
  where "suffix s2 s \<longleftrightarrow> (\<exists>s1. Finite s1 \<and> s = s1 @@ s2)"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    30
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    31
definition tsuffix :: "'a Seq \<Rightarrow> 'a Seq \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    32
  where "tsuffix s2 s \<longleftrightarrow> s2 \<noteq> nil \<and> s2 \<noteq> UU \<and> suffix s2 s"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    33
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    34
definition Box :: "'a temporal \<Rightarrow> 'a temporal"  ("\<box>(_)" [80] 80)
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    35
  where "(\<box>P) s \<longleftrightarrow> (\<forall>s2. tsuffix s2 s \<longrightarrow> P s2)"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    36
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    37
definition Diamond :: "'a temporal \<Rightarrow> 'a temporal"  ("\<diamond>(_)" [80] 80)
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    38
  where "\<diamond>P = (\<^bold>\<not> (\<box>(\<^bold>\<not> P)))"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    39
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    40
definition Leadsto :: "'a temporal \<Rightarrow> 'a temporal \<Rightarrow> 'a temporal"  (infixr "\<leadsto>" 22)
68db98c2cd97 modernized defs;
wenzelm
parents: 62004
diff changeset
    41
  where "(P \<leadsto> Q) = (\<box>(P \<^bold>\<longrightarrow> (\<diamond>Q)))"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    42
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    43
62000
8cba509ace9c more symbols;
wenzelm
parents: 61999
diff changeset
    44
lemma simple: "\<box>\<diamond>(\<^bold>\<not> P) = (\<^bold>\<not> \<diamond>\<box>P)"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    45
  by (auto simp add: Diamond_def NOT_def Box_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    46
62000
8cba509ace9c more symbols;
wenzelm
parents: 61999
diff changeset
    47
lemma Boxnil: "nil \<Turnstile> \<box>P"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    48
  by (simp add: satisfies_def Box_def tsuffix_def suffix_def nil_is_Conc)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    49
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    50
lemma Diamondnil: "\<not> (nil \<Turnstile> \<diamond>P)"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    51
  using Boxnil by (simp add: Diamond_def satisfies_def NOT_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    52
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    53
lemma Diamond_def2: "(\<diamond>F) s \<longleftrightarrow> (\<exists>s2. tsuffix s2 s \<and> F s2)"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    54
  by (simp add: Diamond_def NOT_def Box_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    55
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    56
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    57
subsection \<open>TLA Axiomatization by Merz\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    58
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    59
lemma suffix_refl: "suffix s s"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    60
  apply (simp add: suffix_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    61
  apply (rule_tac x = "nil" in exI)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    62
  apply auto
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    63
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    64
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    65
lemma reflT: "s \<noteq> UU \<and> s \<noteq> nil \<longrightarrow> (s \<Turnstile> \<box>F \<^bold>\<longrightarrow> F)"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    66
  apply (simp add: satisfies_def IMPLIES_def Box_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    67
  apply (rule impI)+
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    68
  apply (erule_tac x = "s" in allE)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    69
  apply (simp add: tsuffix_def suffix_refl)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    70
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    71
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    72
lemma suffix_trans: "suffix y x \<Longrightarrow> suffix z y \<Longrightarrow> suffix z x"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    73
  apply (simp add: suffix_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    74
  apply auto
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    75
  apply (rule_tac x = "s1 @@ s1a" in exI)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    76
  apply auto
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    77
  apply (simp add: Conc_assoc)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    78
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    79
62000
8cba509ace9c more symbols;
wenzelm
parents: 61999
diff changeset
    80
lemma transT: "s \<Turnstile> \<box>F \<^bold>\<longrightarrow> \<box>\<box>F"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    81
  apply (simp add: satisfies_def IMPLIES_def Box_def tsuffix_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    82
  apply auto
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    83
  apply (drule suffix_trans)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    84
  apply assumption
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    85
  apply (erule_tac x = "s2a" in allE)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    86
  apply auto
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    87
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    88
62000
8cba509ace9c more symbols;
wenzelm
parents: 61999
diff changeset
    89
lemma normalT: "s \<Turnstile> \<box>(F \<^bold>\<longrightarrow> G) \<^bold>\<longrightarrow> \<box>F \<^bold>\<longrightarrow> \<box>G"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    90
  by (simp add: satisfies_def IMPLIES_def Box_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    91
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    92
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    93
subsection \<open>TLA Rules by Lamport\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    94
62194
0aabc5931361 tuned syntax;
wenzelm
parents: 62116
diff changeset
    95
lemma STL1a: "\<^bold>\<TTurnstile> P \<Longrightarrow> \<^bold>\<TTurnstile> (\<box>P)"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    96
  by (simp add: validT_def satisfies_def Box_def tsuffix_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    97
62194
0aabc5931361 tuned syntax;
wenzelm
parents: 62116
diff changeset
    98
lemma STL1b: "\<TTurnstile> P \<Longrightarrow> \<^bold>\<TTurnstile> (Init P)"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    99
  by (simp add: valid_def validT_def satisfies_def Init_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   100
62194
0aabc5931361 tuned syntax;
wenzelm
parents: 62116
diff changeset
   101
lemma STL1: "\<TTurnstile> P \<Longrightarrow> \<^bold>\<TTurnstile> (\<box>(Init P))"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   102
  apply (rule STL1a)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   103
  apply (erule STL1b)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   104
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   105
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   106
(*Note that unlift and HD is not at all used!*)
62194
0aabc5931361 tuned syntax;
wenzelm
parents: 62116
diff changeset
   107
lemma STL4: "\<TTurnstile> (P \<^bold>\<longrightarrow> Q) \<Longrightarrow> \<^bold>\<TTurnstile> (\<box>(Init P) \<^bold>\<longrightarrow> \<box>(Init Q))"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   108
  by (simp add: valid_def validT_def satisfies_def IMPLIES_def Box_def Init_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   109
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   110
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   111
subsection \<open>LTL Axioms by Manna/Pnueli\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   112
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62441
diff changeset
   113
lemma tsuffix_TL [rule_format]: "s \<noteq> UU \<and> s \<noteq> nil \<longrightarrow> tsuffix s2 (TL \<cdot> s) \<longrightarrow> tsuffix s2 s"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   114
  apply (unfold tsuffix_def suffix_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   115
  apply auto
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62194
diff changeset
   116
  apply (Seq_case_simp s)
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   117
  apply (rule_tac x = "a \<leadsto> s1" in exI)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   118
  apply auto
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   119
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   120
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   121
lemmas tsuffix_TL2 = conjI [THEN tsuffix_TL]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   122
62441
e5e38e1f2dd4 more symbols;
wenzelm
parents: 62390
diff changeset
   123
lemma LTL1: "s \<noteq> UU \<and> s \<noteq> nil \<longrightarrow> (s \<Turnstile> \<box>F \<^bold>\<longrightarrow> (F \<^bold>\<and> (\<circle>(\<box>F))))"
62390
842917225d56 more canonical names
nipkow
parents: 62195
diff changeset
   124
  supply if_split [split del] 
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   125
  apply (unfold Next_def satisfies_def NOT_def IMPLIES_def AND_def Box_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   126
  apply auto
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   127
  text \<open>\<open>\<box>F \<^bold>\<longrightarrow> F\<close>\<close>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   128
  apply (erule_tac x = "s" in allE)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   129
  apply (simp add: tsuffix_def suffix_refl)
62441
e5e38e1f2dd4 more symbols;
wenzelm
parents: 62390
diff changeset
   130
  text \<open>\<open>\<box>F \<^bold>\<longrightarrow> \<circle>\<box>F\<close>\<close>
63648
f9f3006a5579 "split add" -> "split"
nipkow
parents: 63549
diff changeset
   131
  apply (simp split: if_split)
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   132
  apply auto
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   133
  apply (drule tsuffix_TL2)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   134
  apply assumption+
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   135
  apply auto
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   136
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   137
62441
e5e38e1f2dd4 more symbols;
wenzelm
parents: 62390
diff changeset
   138
lemma LTL2a: "s \<Turnstile> \<^bold>\<not> (\<circle>F) \<^bold>\<longrightarrow> (\<circle>(\<^bold>\<not> F))"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   139
  by (simp add: Next_def satisfies_def NOT_def IMPLIES_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   140
62441
e5e38e1f2dd4 more symbols;
wenzelm
parents: 62390
diff changeset
   141
lemma LTL2b: "s \<Turnstile> (\<circle>(\<^bold>\<not> F)) \<^bold>\<longrightarrow> (\<^bold>\<not> (\<circle>F))"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   142
  by (simp add: Next_def satisfies_def NOT_def IMPLIES_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   143
62441
e5e38e1f2dd4 more symbols;
wenzelm
parents: 62390
diff changeset
   144
lemma LTL3: "ex \<Turnstile> (\<circle>(F \<^bold>\<longrightarrow> G)) \<^bold>\<longrightarrow> (\<circle>F) \<^bold>\<longrightarrow> (\<circle>G)"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   145
  by (simp add: Next_def satisfies_def NOT_def IMPLIES_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   146
62194
0aabc5931361 tuned syntax;
wenzelm
parents: 62116
diff changeset
   147
lemma ModusPonens: "\<^bold>\<TTurnstile> (P \<^bold>\<longrightarrow> Q) \<Longrightarrow> \<^bold>\<TTurnstile> P \<Longrightarrow> \<^bold>\<TTurnstile> Q"
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   148
  by (simp add: validT_def satisfies_def IMPLIES_def)
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   149
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   150
end