src/HOLCF/IOA/meta_theory/TLS.thy
author haftmann
Wed, 22 Sep 2010 18:40:23 +0200
changeset 39641 bcee72a58a12
parent 37598 893dcabf0c04
permissions -rw-r--r--
merged
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
12218
wenzelm
parents: 10835
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 {* Temporal Logic of Steps -- tailored for I/O automata *}
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
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    13
types
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    14
  ('a, 's) ioa_temp  = "('a option,'s)transition temporal"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    15
  ('a, 's) step_pred = "('a option,'s)transition predicate"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    16
  's state_pred      = "'s predicate"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    17
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    18
consts
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    19
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    20
option_lift :: "('a => 'b) => 'b => ('a option => 'b)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    21
plift       :: "('a => bool) => ('a option => bool)"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    22
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    23
temp_sat   :: "('a,'s)execution => ('a,'s)ioa_temp => bool"    (infixr "|==" 22)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    24
xt1        :: "'s predicate => ('a,'s)step_pred"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    25
xt2        :: "'a option predicate => ('a,'s)step_pred"
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
validTE    :: "('a,'s)ioa_temp => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    28
validIOA   :: "('a,'s)ioa => ('a,'s)ioa_temp => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    29
5677
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    30
mkfin      :: "'a Seq => 'a Seq"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    31
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    32
ex2seq     :: "('a,'s)execution => ('a option,'s)transition Seq"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    33
ex2seqC    :: "('a,'s)pairs -> ('s => ('a option,'s)transition Seq)"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    34
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    35
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    36
defs
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    37
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    38
mkfin_def:
5677
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    39
  "mkfin s == if Partial s then @t. Finite t & s = t @@ UU
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    40
                           else s"
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    41
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    42
option_lift_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    43
  "option_lift f s y == case y of None => s | Some x => (f x)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    44
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    45
(* 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
    46
   transition predicates *)
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    47
plift_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    48
  "plift P == option_lift P False"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    49
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    50
temp_sat_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    51
  "ex |== P == ((ex2seq ex) |= P)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    52
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    53
xt1_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    54
  "xt1 P tr == P (fst tr)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    55
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    56
xt2_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    57
  "xt2 P tr == P (fst (snd tr))"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    58
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    59
ex2seq_def:
10835
nipkow
parents: 5976
diff changeset
    60
  "ex2seq ex == ((ex2seqC $(mkfin (snd ex))) (fst ex))"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    61
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    62
ex2seqC_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    63
  "ex2seqC == (fix$(LAM h ex. (%s. case ex of
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    64
      nil =>  (s,None,s)>>nil
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    65
    | x##xs => (flift1 (%pr.
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    66
                (s,Some (fst pr), snd pr)>> (h$xs) (snd pr))
10835
nipkow
parents: 5976
diff changeset
    67
                $x)
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    68
      )))"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    69
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    70
validTE_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    71
  "validTE P == ! ex. (ex |== P)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    72
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    73
validIOA_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    74
  "validIOA A P == ! ex : executions A . (ex |== P)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    75
5976
44290b71a85f tuning to assimiliate it with PhD;
mueller
parents: 5677
diff changeset
    76
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    77
axioms
5976
44290b71a85f tuning to assimiliate it with PhD;
mueller
parents: 5677
diff changeset
    78
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    79
mkfin_UU:
5677
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    80
  "mkfin UU = nil"
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    81
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    82
mkfin_nil:
5677
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    83
  "mkfin nil =nil"
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    84
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    85
mkfin_cons:
5677
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    86
  "(mkfin (a>>s)) = (a>>(mkfin s))"
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    87
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    88
37598
893dcabf0c04 explicit is better than implicit
haftmann
parents: 36452
diff changeset
    89
lemmas [simp del] = HOL.ex_simps HOL.all_simps split_paired_Ex
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    90
26339
7825c83c9eff eliminated change_claset/simpset;
wenzelm
parents: 19741
diff changeset
    91
declaration {* fn _ => Classical.map_cs (fn cs => cs delSWrapper "split_all_tac") *}
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    92
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    93
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    94
subsection {* ex2seqC *}
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    95
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    96
lemma ex2seqC_unfold: "ex2seqC  = (LAM ex. (%s. case ex of  
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    97
       nil =>  (s,None,s)>>nil    
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    98
     | x##xs => (flift1 (%pr.  
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    99
                 (s,Some (fst pr), snd pr)>> (ex2seqC$xs) (snd pr))   
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   100
                 $x)   
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   101
       ))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   102
apply (rule trans)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   103
apply (rule fix_eq2)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   104
apply (rule ex2seqC_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   105
apply (rule beta_cfun)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   106
apply (simp add: flift1_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   107
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   108
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   109
lemma ex2seqC_UU: "(ex2seqC $UU) s=UU"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   110
apply (subst ex2seqC_unfold)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   111
apply simp
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   112
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   113
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   114
lemma ex2seqC_nil: "(ex2seqC $nil) s = (s,None,s)>>nil"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   115
apply (subst ex2seqC_unfold)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   116
apply simp
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   117
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   118
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   119
lemma ex2seqC_cons: "(ex2seqC $((a,t)>>xs)) s =  
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   120
           (s,Some a,t)>> ((ex2seqC$xs) t)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   121
apply (rule trans)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   122
apply (subst ex2seqC_unfold)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   123
apply (simp add: Consq_def flift1_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   124
apply (simp add: Consq_def flift1_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
declare ex2seqC_UU [simp] ex2seqC_nil [simp] ex2seqC_cons [simp]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   128
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
declare mkfin_UU [simp] mkfin_nil [simp] mkfin_cons [simp]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   132
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   133
lemma ex2seq_UU: "ex2seq (s, UU) = (s,None,s)>>nil"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   134
apply (simp add: ex2seq_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 ex2seq_nil: "ex2seq (s, nil) = (s,None,s)>>nil"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   138
apply (simp add: ex2seq_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 ex2seq_cons: "ex2seq (s, (a,t)>>ex) = (s,Some a,t) >> ex2seq (t, ex)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   142
apply (simp add: ex2seq_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   143
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   144
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   145
declare ex2seqC_UU [simp del] ex2seqC_nil [simp del] ex2seqC_cons [simp del]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   146
declare ex2seq_UU [simp] ex2seq_nil [simp] ex2seq_cons [simp]
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
lemma ex2seq_nUUnnil: "ex2seq exec ~= UU & ex2seq exec ~= nil"
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   150
apply (tactic {* pair_tac @{context} "exec" 1 *})
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   151
apply (tactic {* Seq_case_simp_tac @{context} "y" 1 *})
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   152
apply (tactic {* pair_tac @{context} "a" 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   153
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   154
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   155
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   156
subsection {* Interface TL -- TLS *}
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   157
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   158
(* uses the fact that in executions states overlap, which is lost in 
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   159
   after the translation via ex2seq !! *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   160
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   161
lemma TL_TLS: 
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   162
 "[| ! s a t. (P s) & s-a--A-> t --> (Q t) |] 
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   163
   ==> ex |== (Init (%(s,a,t). P s) .& Init (%(s,a,t). s -a--A-> t)  
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   164
              .--> (Next (Init (%(s,a,t).Q s))))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   165
apply (unfold Init_def Next_def temp_sat_def satisfies_def IMPLIES_def AND_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   166
26359
6d437bde2f1d more antiquotations
haftmann
parents: 26339
diff changeset
   167
apply clarify
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   168
apply (simp split add: split_if)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   169
(* TL = UU *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   170
apply (rule conjI)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   171
apply (tactic {* pair_tac @{context} "ex" 1 *})
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   172
apply (tactic {* Seq_case_simp_tac @{context} "y" 1 *})
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   173
apply (tactic {* pair_tac @{context} "a" 1 *})
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   174
apply (tactic {* Seq_case_simp_tac @{context} "s" 1 *})
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   175
apply (tactic {* pair_tac @{context} "a" 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   176
(* TL = nil *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   177
apply (rule conjI)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   178
apply (tactic {* pair_tac @{context} "ex" 1 *})
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   179
apply (tactic {* Seq_case_tac @{context} "y" 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   180
apply (simp add: unlift_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   181
apply fast
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   182
apply (simp add: unlift_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   183
apply fast
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   184
apply (simp add: unlift_def)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   185
apply (tactic {* pair_tac @{context} "a" 1 *})
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   186
apply (tactic {* Seq_case_simp_tac @{context} "s" 1 *})
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   187
apply (tactic {* pair_tac @{context} "a" 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   188
(* TL =cons *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   189
apply (simp add: unlift_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   190
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   191
apply (tactic {* pair_tac @{context} "ex" 1 *})
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   192
apply (tactic {* Seq_case_simp_tac @{context} "y" 1 *})
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   193
apply (tactic {* pair_tac @{context} "a" 1 *})
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   194
apply (tactic {* Seq_case_simp_tac @{context} "s" 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   195
apply blast
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   196
apply fastsimp
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   197
apply (tactic {* pair_tac @{context} "a" 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   198
 apply fastsimp
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   199
done
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
   200
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
   201
end