src/HOL/HOLCF/IOA/meta_theory/RefCorrectness.thy
author haftmann
Thu, 27 Aug 2015 21:19:48 +0200
changeset 61032 b57df8eecad6
parent 58880 0baae4311a9f
child 62001 1f2788fb0b8b
permissions -rw-r--r--
standardized some occurences of ancient "split" alias
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 40945
diff changeset
     1
(*  Title:      HOL/HOLCF/IOA/meta_theory/RefCorrectness.thy
40945
b8703f63bfb2 recoded latin1 as utf8;
wenzelm
parents: 40774
diff changeset
     2
    Author:     Olaf Müller
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     3
*)
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     4
58880
0baae4311a9f modernized header;
wenzelm
parents: 42151
diff changeset
     5
section {* Correctness of Refinement Mappings in HOLCF/IOA *}
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     6
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     7
theory RefCorrectness
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     8
imports RefMappings
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
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    11
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    12
  corresp_exC :: "('a,'s2)ioa => ('s1 => 's2) => ('a,'s1)pairs
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    13
                   -> ('s1 => ('a,'s2)pairs)" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    14
  "corresp_exC A f = (fix$(LAM h ex. (%s. case ex of
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    15
      nil =>  nil
3433
2de17c994071 added deadlock freedom, polished definitions and proofs
mueller
parents: 3275
diff changeset
    16
    | x##xs => (flift1 (%pr. (@cex. move A cex (f s) (fst pr) (f (snd pr)))
10835
nipkow
parents: 5976
diff changeset
    17
                              @@ ((h$xs) (snd pr)))
nipkow
parents: 5976
diff changeset
    18
                        $x) )))"
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    19
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    20
  corresp_ex :: "('a,'s2)ioa => ('s1 => 's2) =>
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    21
                  ('a,'s1)execution => ('a,'s2)execution" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    22
  "corresp_ex A f ex = (f (fst ex),(corresp_exC A f$(snd ex)) (fst ex))"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    23
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    24
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    25
  is_fair_ref_map :: "('s1 => 's2) => ('a,'s1)ioa => ('a,'s2)ioa => bool" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    26
  "is_fair_ref_map f C A =
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    27
      (is_ref_map f C A &
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    28
       (! ex : executions(C). fair_ex C ex --> fair_ex A (corresp_ex A f ex)))"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    29
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    30
(* Axioms for fair trace inclusion proof support, not for the correctness proof
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    31
   of refinement mappings!
5976
44290b71a85f tuning to assimiliate it with PhD;
mueller
parents: 4559
diff changeset
    32
   Note: Everything is superseded by LiveIOA.thy! *)
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    33
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    34
axiomatization where
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    35
corresp_laststate:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    36
  "Finite ex ==> laststate (corresp_ex A f (s,ex)) = f (laststate (s,ex))"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    37
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    38
axiomatization where
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    39
corresp_Finite:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    40
  "Finite (snd (corresp_ex A f (s,ex))) = Finite ex"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    41
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    42
axiomatization where
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    43
FromAtoC:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    44
  "fin_often (%x. P (snd x)) (snd (corresp_ex A f (s,ex))) ==> fin_often (%y. P (f (snd y))) ex"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    45
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    46
axiomatization where
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    47
FromCtoA:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    48
  "inf_often (%y. P (fst y)) ex ==> inf_often (%x. P (fst x)) (snd (corresp_ex A f (s,ex)))"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    49
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    50
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    51
(* Proof by case on inf W in ex: If so, ok. If not, only fin W in ex, ie there is
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    52
   an index i from which on no W in ex. But W inf enabled, ie at least once after i
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    53
   W is enabled. As W does not occur after i and W is enabling_persistent, W keeps
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    54
   enabled until infinity, ie. indefinitely *)
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    55
axiomatization where
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    56
persistent:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    57
  "[|inf_often (%x. Enabled A W (snd x)) ex; en_persistent A W|]
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    58
   ==> inf_often (%x. fst x :W) ex | fin_often (%x. ~Enabled A W (snd x)) ex"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    59
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    60
axiomatization where
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    61
infpostcond:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    62
  "[| is_exec_frag A (s,ex); inf_often (%x. fst x:W) ex|]
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    63
    ==> inf_often (% x. set_was_enabled A W (snd x)) ex"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    64
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    65
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    66
subsection "corresp_ex"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    67
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    68
lemma corresp_exC_unfold: "corresp_exC A f  = (LAM ex. (%s. case ex of
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    69
       nil =>  nil
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    70
     | x##xs => (flift1 (%pr. (@cex. move A cex (f s) (fst pr) (f (snd pr)))
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    71
                               @@ ((corresp_exC A f $xs) (snd pr)))
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    72
                         $x) ))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    73
apply (rule trans)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    74
apply (rule fix_eq2)
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    75
apply (simp only: corresp_exC_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    76
apply (rule beta_cfun)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    77
apply (simp add: flift1_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    78
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    79
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    80
lemma corresp_exC_UU: "(corresp_exC A f$UU) s=UU"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    81
apply (subst corresp_exC_unfold)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    82
apply simp
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    83
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    84
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    85
lemma corresp_exC_nil: "(corresp_exC A f$nil) s = nil"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    86
apply (subst corresp_exC_unfold)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    87
apply simp
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    88
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    89
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    90
lemma corresp_exC_cons: "(corresp_exC A f$(at>>xs)) s =
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    91
           (@cex. move A cex (f s) (fst at) (f (snd at)))
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    92
           @@ ((corresp_exC A f$xs) (snd at))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    93
apply (rule trans)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    94
apply (subst corresp_exC_unfold)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    95
apply (simp add: Consq_def flift1_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    96
apply simp
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    97
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    98
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    99
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   100
declare corresp_exC_UU [simp] corresp_exC_nil [simp] corresp_exC_cons [simp]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   101
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   102
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   103
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   104
subsection "properties of move"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   105
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   106
lemma move_is_move:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   107
   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   108
      move A (@x. move A x (f s) a (f t)) (f s) a (f t)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   109
apply (unfold is_ref_map_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   110
apply (subgoal_tac "? ex. move A ex (f s) a (f t) ")
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   111
prefer 2
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   112
apply simp
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   113
apply (erule exE)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   114
apply (rule someI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   115
apply assumption
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   116
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   117
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   118
lemma move_subprop1:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   119
   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   120
     is_exec_frag A (f s,@x. move A x (f s) a (f t))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   121
apply (cut_tac move_is_move)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   122
defer
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   123
apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   124
apply (simp add: move_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   125
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   126
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   127
lemma move_subprop2:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   128
   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   129
     Finite ((@x. move A x (f s) a (f t)))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   130
apply (cut_tac move_is_move)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   131
defer
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   132
apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   133
apply (simp add: move_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   134
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   135
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   136
lemma move_subprop3:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   137
   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   138
     laststate (f s,@x. move A x (f s) a (f t)) = (f t)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   139
apply (cut_tac move_is_move)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   140
defer
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   141
apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   142
apply (simp add: move_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   143
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   144
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   145
lemma move_subprop4:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   146
   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   147
      mk_trace A$((@x. move A x (f s) a (f t))) =
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   148
        (if a:ext A then a>>nil else nil)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   149
apply (cut_tac move_is_move)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   150
defer
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   151
apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   152
apply (simp add: move_def)
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
(* ------------------------------------------------------------------ *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   157
(*                   The following lemmata contribute to              *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   158
(*                 TRACE INCLUSION Part 1: Traces coincide            *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   159
(* ------------------------------------------------------------------ *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   160
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   161
section "Lemmata for <=="
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   162
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   163
(* --------------------------------------------------- *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   164
(*   Lemma 1.1: Distribution of mk_trace and @@        *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   165
(* --------------------------------------------------- *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   166
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   167
lemma mk_traceConc: "mk_trace C$(ex1 @@ ex2)= (mk_trace C$ex1) @@ (mk_trace C$ex2)"
35215
a03462cbf86f get rid of warnings about duplicate simp rules in all HOLCF theories
huffman
parents: 35174
diff changeset
   168
apply (simp add: mk_trace_def filter_act_def MapConc)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   169
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   170
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   171
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   172
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   173
(* ------------------------------------------------------
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   174
                 Lemma 1 :Traces coincide
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   175
   ------------------------------------------------------- *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   176
declare split_if [split del]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   177
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   178
lemma lemma_1:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   179
  "[|is_ref_map f C A; ext C = ext A|] ==>
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   180
         !s. reachable C s & is_exec_frag C (s,xs) -->
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   181
             mk_trace C$xs = mk_trace A$(snd (corresp_ex A f (s,xs)))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   182
apply (unfold corresp_ex_def)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   183
apply (tactic {* pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}] 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   184
(* cons case *)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   185
apply (auto simp add: mk_traceConc)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   186
apply (frule reachable.reachable_n)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   187
apply assumption
61032
b57df8eecad6 standardized some occurences of ancient "split" alias
haftmann
parents: 58880
diff changeset
   188
apply (auto simp add: move_subprop4 split add: split_if) 
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   189
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   190
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   191
declare split_if [split]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   192
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   193
(* ------------------------------------------------------------------ *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   194
(*                   The following lemmata contribute to              *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   195
(*              TRACE INCLUSION Part 2: corresp_ex is execution       *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   196
(* ------------------------------------------------------------------ *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   197
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   198
section "Lemmata for ==>"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   199
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   200
(* -------------------------------------------------- *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   201
(*                   Lemma 2.1                        *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   202
(* -------------------------------------------------- *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   203
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   204
lemma lemma_2_1 [rule_format (no_asm)]:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   205
"Finite xs -->
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   206
 (!s .is_exec_frag A (s,xs) & is_exec_frag A (t,ys) &
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   207
      t = laststate (s,xs)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   208
  --> is_exec_frag A (s,xs @@ ys))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   209
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   210
apply (rule impI)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   211
apply (tactic {* Seq_Finite_induct_tac @{context} 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   212
(* main case *)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   213
apply (auto simp add: split_paired_all)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   214
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   215
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   216
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   217
(* ----------------------------------------------------------- *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   218
(*               Lemma 2 : corresp_ex is execution             *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   219
(* ----------------------------------------------------------- *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   220
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   221
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   222
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   223
lemma lemma_2:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   224
 "[| is_ref_map f C A |] ==>
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   225
  !s. reachable C s & is_exec_frag C (s,xs)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   226
  --> is_exec_frag A (corresp_ex A f (s,xs))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   227
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   228
apply (unfold corresp_ex_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   229
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   230
apply simp
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   231
apply (tactic {* pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}] 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   232
(* main case *)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   233
apply auto
61032
b57df8eecad6 standardized some occurences of ancient "split" alias
haftmann
parents: 58880
diff changeset
   234
apply (rule_tac t = "f x2" in lemma_2_1)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   235
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   236
(* Finite *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   237
apply (erule move_subprop2)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   238
apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   239
apply (rule conjI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   240
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   241
(* is_exec_frag *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   242
apply (erule move_subprop1)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   243
apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   244
apply (rule conjI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   245
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   246
(* Induction hypothesis  *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   247
(* reachable_n looping, therefore apply it manually *)
61032
b57df8eecad6 standardized some occurences of ancient "split" alias
haftmann
parents: 58880
diff changeset
   248
apply (erule_tac x = "x2" in allE)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   249
apply simp
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   250
apply (frule reachable.reachable_n)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   251
apply assumption
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   252
apply simp
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   253
(* laststate *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   254
apply (erule move_subprop3 [symmetric])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   255
apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   256
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   257
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   258
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   259
subsection "Main Theorem: TRACE - INCLUSION"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   260
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   261
lemma trace_inclusion:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   262
  "[| ext C = ext A; is_ref_map f C A |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   263
           ==> traces C <= traces A"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   264
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   265
  apply (unfold traces_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   266
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   267
  apply (simp (no_asm) add: has_trace_def2)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   268
  apply auto
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   269
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   270
  (* give execution of abstract automata *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   271
  apply (rule_tac x = "corresp_ex A f ex" in bexI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   272
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   273
  (* Traces coincide, Lemma 1 *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   274
  apply (tactic {* pair_tac @{context} "ex" 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   275
  apply (erule lemma_1 [THEN spec, THEN mp])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   276
  apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   277
  apply (simp add: executions_def reachable.reachable_0)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   278
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   279
  (* corresp_ex is execution, Lemma 2 *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   280
  apply (tactic {* pair_tac @{context} "ex" 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   281
  apply (simp add: executions_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   282
  (* start state *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   283
  apply (rule conjI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   284
  apply (simp add: is_ref_map_def corresp_ex_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   285
  (* is-execution-fragment *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   286
  apply (erule lemma_2 [THEN spec, THEN mp])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   287
  apply (simp add: reachable.reachable_0)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   288
  done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   289
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   290
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   291
subsection "Corollary:  FAIR TRACE - INCLUSION"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   292
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   293
lemma fininf: "(~inf_often P s) = fin_often P s"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   294
apply (unfold fin_often_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   295
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   296
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   297
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   298
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   299
lemma WF_alt: "is_wfair A W (s,ex) =
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   300
  (fin_often (%x. ~Enabled A W (snd x)) ex --> inf_often (%x. fst x :W) ex)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   301
apply (simp add: is_wfair_def fin_often_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   302
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   303
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   304
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   305
lemma WF_persistent: "[|is_wfair A W (s,ex); inf_often (%x. Enabled A W (snd x)) ex;
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   306
          en_persistent A W|]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   307
    ==> inf_often (%x. fst x :W) ex"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   308
apply (drule persistent)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   309
apply assumption
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   310
apply (simp add: WF_alt)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   311
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   312
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   313
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   314
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   315
lemma fair_trace_inclusion: "!! C A.
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   316
          [| is_ref_map f C A; ext C = ext A;
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   317
          !! ex. [| ex:executions C; fair_ex C ex|] ==> fair_ex A (corresp_ex A f ex) |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   318
          ==> fairtraces C <= fairtraces A"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   319
apply (simp (no_asm) add: fairtraces_def fairexecutions_def)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   320
apply auto
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   321
apply (rule_tac x = "corresp_ex A f ex" in exI)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   322
apply auto
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   323
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   324
  (* Traces coincide, Lemma 1 *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   325
  apply (tactic {* pair_tac @{context} "ex" 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   326
  apply (erule lemma_1 [THEN spec, THEN mp])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   327
  apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   328
  apply (simp add: executions_def reachable.reachable_0)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   329
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   330
  (* corresp_ex is execution, Lemma 2 *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   331
  apply (tactic {* pair_tac @{context} "ex" 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   332
  apply (simp add: executions_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   333
  (* start state *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   334
  apply (rule conjI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   335
  apply (simp add: is_ref_map_def corresp_ex_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   336
  (* is-execution-fragment *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   337
  apply (erule lemma_2 [THEN spec, THEN mp])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   338
  apply (simp add: reachable.reachable_0)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   339
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   340
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   341
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   342
lemma fair_trace_inclusion2: "!! C A.
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   343
          [| inp(C) = inp(A); out(C)=out(A);
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   344
             is_fair_ref_map f C A |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   345
          ==> fair_implements C A"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   346
apply (simp add: is_fair_ref_map_def fair_implements_def fairtraces_def fairexecutions_def)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   347
apply auto
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   348
apply (rule_tac x = "corresp_ex A f ex" in exI)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   349
apply auto
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   350
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   351
  (* Traces coincide, Lemma 1 *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   352
  apply (tactic {* pair_tac @{context} "ex" 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   353
  apply (erule lemma_1 [THEN spec, THEN mp])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   354
  apply (simp (no_asm) add: externals_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   355
  apply (auto)[1]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   356
  apply (simp add: executions_def reachable.reachable_0)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   357
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   358
  (* corresp_ex is execution, Lemma 2 *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26359
diff changeset
   359
  apply (tactic {* pair_tac @{context} "ex" 1 *})
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   360
  apply (simp add: executions_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   361
  (* start state *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   362
  apply (rule conjI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   363
  apply (simp add: is_ref_map_def corresp_ex_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   364
  (* is-execution-fragment *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   365
  apply (erule lemma_2 [THEN spec, THEN mp])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   366
  apply (simp add: reachable.reachable_0)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   367
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   368
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   369
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   370
end