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