src/HOL/HOLCF/IOA/TLS.thy
author wenzelm
Wed, 21 Mar 2018 18:30:17 +0100
changeset 67913 d58fa3ed236f
parent 63648 f9f3006a5579
child 71989 bad75618fb82
permissions -rw-r--r--
proper order of matches: Server.Error is an instance of Exn.ERROR;
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/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
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
     5
section \<open>Temporal Logic of Steps -- tailored for I/O automata\<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 TLS
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     8
imports IOA TL
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     9
begin
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    10
36452
d37c6eed8117 renamed command 'defaultsort' to 'default_sort';
wenzelm
parents: 35215
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: 62003
diff changeset
    13
type_synonym ('a, 's) ioa_temp  = "('a option, 's) transition temporal"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    14
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    15
type_synonym ('a, 's) step_pred = "('a option, 's) transition predicate"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    16
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    17
type_synonym 's state_pred = "'s predicate"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    18
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    19
definition mkfin :: "'a Seq \<Rightarrow> 'a Seq"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    20
  where "mkfin s = (if Partial s then SOME t. Finite t \<and> s = t @@ UU else s)"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    21
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    22
definition option_lift :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a option \<Rightarrow> 'b"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    23
  where "option_lift f s y = (case y of None \<Rightarrow> s | Some x \<Rightarrow> f x)"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    24
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    25
definition plift :: "('a \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> bool"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    26
(* plift is used to determine that None action is always false in
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    27
   transition predicates *)
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    28
  where "plift P = option_lift P False"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    29
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    30
definition xt1 :: "'s predicate \<Rightarrow> ('a, 's) step_pred"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    31
  where "xt1 P tr = P (fst tr)"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    32
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    33
definition xt2 :: "'a option predicate \<Rightarrow> ('a, 's) step_pred"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    34
  where "xt2 P tr = P (fst (snd tr))"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    35
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    36
definition ex2seqC :: "('a, 's) pairs \<rightarrow> ('s \<Rightarrow> ('a option, 's) transition Seq)"
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    37
  where "ex2seqC =
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62441
diff changeset
    38
    (fix \<cdot> (LAM h ex. (\<lambda>s. case ex of
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    39
      nil \<Rightarrow> (s, None, s) \<leadsto> nil
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62441
diff changeset
    40
    | x ## xs \<Rightarrow> (flift1 (\<lambda>pr. (s, Some (fst pr), snd pr) \<leadsto> (h \<cdot> xs) (snd pr)) \<cdot> x))))"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    41
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    42
definition ex2seq :: "('a, 's) execution \<Rightarrow> ('a option, 's) transition Seq"
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62441
diff changeset
    43
  where "ex2seq ex = (ex2seqC \<cdot> (mkfin (snd ex))) (fst ex)"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    44
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    45
definition temp_sat :: "('a, 's) execution \<Rightarrow> ('a, 's) ioa_temp \<Rightarrow> bool"  (infixr "\<TTurnstile>" 22)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    46
  where "(ex \<TTurnstile> P) \<longleftrightarrow> ((ex2seq ex) \<Turnstile> P)"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    47
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    48
definition validTE :: "('a, 's) ioa_temp \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    49
  where "validTE P \<longleftrightarrow> (\<forall>ex. (ex \<TTurnstile> P))"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    50
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    51
definition validIOA :: "('a, 's) ioa \<Rightarrow> ('a, 's) ioa_temp \<Rightarrow> bool"
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    52
  where "validIOA A P \<longleftrightarrow> (\<forall>ex \<in> executions A. (ex \<TTurnstile> P))"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    53
5976
44290b71a85f tuning to assimiliate it with PhD;
mueller
parents: 5677
diff changeset
    54
62193
wenzelm
parents: 62192
diff changeset
    55
lemma IMPLIES_temp_sat [simp]: "(ex \<TTurnstile> P \<^bold>\<longrightarrow> Q) \<longleftrightarrow> ((ex \<TTurnstile> P) \<longrightarrow> (ex \<TTurnstile> Q))"
wenzelm
parents: 62192
diff changeset
    56
  by (simp add: IMPLIES_def temp_sat_def satisfies_def)
wenzelm
parents: 62192
diff changeset
    57
wenzelm
parents: 62192
diff changeset
    58
lemma AND_temp_sat [simp]: "(ex \<TTurnstile> P \<^bold>\<and> Q) \<longleftrightarrow> ((ex \<TTurnstile> P) \<and> (ex \<TTurnstile> Q))"
wenzelm
parents: 62192
diff changeset
    59
  by (simp add: AND_def temp_sat_def satisfies_def)
wenzelm
parents: 62192
diff changeset
    60
wenzelm
parents: 62192
diff changeset
    61
lemma OR_temp_sat [simp]: "(ex \<TTurnstile> P \<^bold>\<or> Q) \<longleftrightarrow> ((ex \<TTurnstile> P) \<or> (ex \<TTurnstile> Q))"
wenzelm
parents: 62192
diff changeset
    62
  by (simp add: OR_def temp_sat_def satisfies_def)
wenzelm
parents: 62192
diff changeset
    63
wenzelm
parents: 62192
diff changeset
    64
lemma NOT_temp_sat [simp]: "(ex \<TTurnstile> \<^bold>\<not> P) \<longleftrightarrow> (\<not> (ex \<TTurnstile> P))"
wenzelm
parents: 62192
diff changeset
    65
  by (simp add: NOT_def temp_sat_def satisfies_def)
wenzelm
parents: 62192
diff changeset
    66
wenzelm
parents: 62192
diff changeset
    67
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    68
axiomatization
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    69
where mkfin_UU [simp]: "mkfin UU = nil"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    70
  and mkfin_nil [simp]: "mkfin nil = nil"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    71
  and mkfin_cons [simp]: "mkfin (a \<leadsto> s) = a \<leadsto> mkfin s"
5677
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    72
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    73
37598
893dcabf0c04 explicit is better than implicit
haftmann
parents: 36452
diff changeset
    74
lemmas [simp del] = HOL.ex_simps HOL.all_simps split_paired_Ex
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    75
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
    76
setup \<open>map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    77
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    78
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
    79
subsection \<open>ex2seqC\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    80
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    81
lemma ex2seqC_unfold:
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    82
  "ex2seqC =
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    83
    (LAM ex. (\<lambda>s. case ex of
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    84
      nil \<Rightarrow> (s, None, s) \<leadsto> nil
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    85
    | x ## xs \<Rightarrow>
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62441
diff changeset
    86
        (flift1 (\<lambda>pr. (s, Some (fst pr), snd pr) \<leadsto> (ex2seqC \<cdot> xs) (snd pr)) \<cdot> x)))"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    87
  apply (rule trans)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    88
  apply (rule fix_eq4)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    89
  apply (rule ex2seqC_def)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    90
  apply (rule beta_cfun)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    91
  apply (simp add: flift1_def)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    92
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    93
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62441
diff changeset
    94
lemma ex2seqC_UU [simp]: "(ex2seqC \<cdot> UU) s = UU"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    95
  apply (subst ex2seqC_unfold)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    96
  apply simp
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
    97
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    98
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62441
diff changeset
    99
lemma ex2seqC_nil [simp]: "(ex2seqC \<cdot> nil) s = (s, None, s) \<leadsto> nil"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   100
  apply (subst ex2seqC_unfold)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   101
  apply simp
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   102
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   103
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62441
diff changeset
   104
lemma ex2seqC_cons [simp]: "(ex2seqC \<cdot> ((a, t) \<leadsto> xs)) s = (s, Some a,t ) \<leadsto> (ex2seqC \<cdot> xs) t"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   105
  apply (rule trans)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   106
  apply (subst ex2seqC_unfold)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   107
  apply (simp add: Consq_def flift1_def)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   108
  apply (simp add: Consq_def flift1_def)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   109
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   110
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   111
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   112
lemma ex2seq_UU: "ex2seq (s, UU) = (s, None, s) \<leadsto> nil"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   113
  by (simp add: ex2seq_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   114
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   115
lemma ex2seq_nil: "ex2seq (s, nil) = (s, None, s) \<leadsto> nil"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   116
  by (simp add: ex2seq_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   117
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   118
lemma ex2seq_cons: "ex2seq (s, (a, t) \<leadsto> ex) = (s, Some a, t) \<leadsto> ex2seq (t, ex)"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   119
  by (simp add: ex2seq_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   120
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   121
declare ex2seqC_UU [simp del] ex2seqC_nil [simp del] ex2seqC_cons [simp del]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   122
declare ex2seq_UU [simp] ex2seq_nil [simp] ex2seq_cons [simp]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   123
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   124
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   125
lemma ex2seq_nUUnnil: "ex2seq exec \<noteq> UU \<and> ex2seq exec \<noteq> nil"
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   126
  apply (pair exec)
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   127
  apply (Seq_case_simp x2)
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   128
  apply (pair a)
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   129
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   130
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   131
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
   132
subsection \<open>Interface TL -- TLS\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   133
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   134
(* uses the fact that in executions states overlap, which is lost in
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   135
   after the translation via ex2seq !! *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   136
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   137
lemma TL_TLS:
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   138
  "\<forall>s a t. (P s) \<and> s \<midarrow>a\<midarrow>A\<rightarrow> t \<longrightarrow> (Q t)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   139
    \<Longrightarrow> ex \<TTurnstile> (Init (\<lambda>(s, a, t). P s) \<^bold>\<and> Init (\<lambda>(s, a, t). s \<midarrow>a\<midarrow>A\<rightarrow> t)
62441
e5e38e1f2dd4 more symbols;
wenzelm
parents: 62390
diff changeset
   140
              \<^bold>\<longrightarrow> (\<circle>(Init (\<lambda>(s, a, t). Q s))))"
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   141
  apply (unfold Init_def Next_def temp_sat_def satisfies_def IMPLIES_def AND_def)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   142
  apply clarify
63648
f9f3006a5579 "split add" -> "split"
nipkow
parents: 63549
diff changeset
   143
  apply (simp split: if_split)
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   144
  text \<open>\<open>TL = UU\<close>\<close>
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   145
  apply (rule conjI)
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   146
  apply (pair ex)
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   147
  apply (Seq_case_simp x2)
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   148
  apply (pair a)
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   149
  apply (Seq_case_simp s)
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   150
  apply (pair a)
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   151
  text \<open>\<open>TL = nil\<close>\<close>
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   152
  apply (rule conjI)
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   153
  apply (pair ex)
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   154
  apply (Seq_case x2)
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   155
  apply (simp add: unlift_def)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   156
  apply fast
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   157
  apply (simp add: unlift_def)
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   158
  apply fast
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   159
  apply (simp add: unlift_def)
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   160
  apply (pair a)
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   161
  apply (Seq_case_simp s)
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   162
  apply (pair a)
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   163
  text \<open>\<open>TL = cons\<close>\<close>
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   164
  apply (simp add: unlift_def)
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   165
  apply (pair ex)
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   166
  apply (Seq_case_simp x2)
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   167
  apply (pair a)
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   168
  apply (Seq_case_simp s)
799a5306e2ed more method definitions;
wenzelm
parents: 62193
diff changeset
   169
  apply (pair a)
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62003
diff changeset
   170
  done
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
   171
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
   172
end