src/HOL/HOLCF/IOA/RefCorrectness.thy
author haftmann
Fri, 26 Oct 2018 08:20:45 +0000
changeset 69194 6d514e128a85
parent 67613 ce654b0e6d69
permissions -rw-r--r--
dedicated theory for sorting algorithms
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62008
cbedaddc9351 clarified directory structure;
wenzelm
parents: 62002
diff changeset
     1
(*  Title:      HOL/HOLCF/IOA/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
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
     5
section \<open>Correctness of Refinement Mappings in HOLCF/IOA\<close>
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
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    11
definition corresp_exC ::
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    12
    "('a, 's2) ioa \<Rightarrow> ('s1 \<Rightarrow> 's2) \<Rightarrow> ('a, 's1) pairs \<rightarrow> ('s1 \<Rightarrow> ('a, 's2) pairs)"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    13
  where "corresp_exC A f =
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    14
    (fix \<cdot>
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    15
      (LAM h ex.
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    16
        (\<lambda>s. case ex of
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    17
          nil \<Rightarrow> nil
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    18
        | x ## xs \<Rightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    19
            flift1 (\<lambda>pr.
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    20
              (SOME cex. move A cex (f s) (fst pr) (f (snd pr))) @@ ((h \<cdot> xs) (snd pr))) \<cdot> x)))"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    21
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    22
definition corresp_ex ::
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    23
    "('a, 's2) ioa \<Rightarrow> ('s1 \<Rightarrow> 's2) \<Rightarrow> ('a, 's1) execution \<Rightarrow> ('a, 's2) execution"
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    24
  where "corresp_ex A f ex = (f (fst ex), (corresp_exC A f \<cdot> (snd ex)) (fst ex))"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    25
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    26
definition is_fair_ref_map ::
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    27
    "('s1 \<Rightarrow> 's2) \<Rightarrow> ('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    28
  where "is_fair_ref_map f C A \<longleftrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    29
    is_ref_map f C A \<and> (\<forall>ex \<in> executions C. fair_ex C ex \<longrightarrow> fair_ex A (corresp_ex A f ex))"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    30
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    31
text \<open>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    32
  Axioms for fair trace inclusion proof support, not for the correctness proof
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    33
  of refinement mappings!
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    34
63680
6e1e8b5abbfa more symbols;
wenzelm
parents: 63648
diff changeset
    35
  Note: Everything is superseded by \<^file>\<open>LiveIOA.thy\<close>.
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    36
\<close>
4559
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
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    39
  corresp_laststate:
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    40
    "Finite ex \<Longrightarrow> laststate (corresp_ex A f (s, ex)) = f (laststate (s, ex))"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    41
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    42
axiomatization where
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    43
  corresp_Finite: "Finite (snd (corresp_ex A f (s, ex))) = Finite ex"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    44
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    45
axiomatization where
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    46
  FromAtoC:
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    47
    "fin_often (\<lambda>x. P (snd x)) (snd (corresp_ex A f (s, ex))) \<Longrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    48
      fin_often (\<lambda>y. P (f (snd y))) ex"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    49
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    50
axiomatization where
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    51
  FromCtoA:
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    52
    "inf_often (\<lambda>y. P (fst y)) ex \<Longrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    53
      inf_often (\<lambda>x. P (fst x)) (snd (corresp_ex A f (s,ex)))"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    54
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    55
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    56
text \<open>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    57
  Proof by case on \<open>inf W\<close> in ex: If so, ok. If not, only \<open>fin W\<close> in ex, ie.
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    58
  there is an index \<open>i\<close> from which on no \<open>W\<close> in ex. But \<open>W\<close> inf enabled, ie at
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    59
  least once after \<open>i\<close> \<open>W\<close> is enabled. As \<open>W\<close> does not occur after \<open>i\<close> and \<open>W\<close>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    60
  is \<open>enabling_persistent\<close>, \<open>W\<close> keeps enabled until infinity, ie. indefinitely
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    61
\<close>
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    62
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    63
axiomatization where
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    64
  persistent:
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    65
    "inf_often (\<lambda>x. Enabled A W (snd x)) ex \<Longrightarrow> en_persistent A W \<Longrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    66
      inf_often (\<lambda>x. fst x \<in> W) ex \<or> fin_often (\<lambda>x. \<not> Enabled A W (snd x)) ex"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    67
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    68
axiomatization where
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    69
  infpostcond:
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    70
    "is_exec_frag A (s,ex) \<Longrightarrow> inf_often (\<lambda>x. fst x \<in> W) ex \<Longrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    71
      inf_often (\<lambda>x. set_was_enabled A W (snd x)) ex"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
    72
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    73
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    74
subsection \<open>\<open>corresp_ex\<close>\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    75
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    76
lemma corresp_exC_unfold:
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    77
  "corresp_exC A f =
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    78
    (LAM ex.
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    79
      (\<lambda>s.
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    80
        case ex of
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    81
          nil \<Rightarrow> nil
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    82
        | x ## xs \<Rightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    83
            (flift1 (\<lambda>pr.
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    84
              (SOME cex. move A cex (f s) (fst pr) (f (snd pr))) @@
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    85
              ((corresp_exC A f \<cdot> xs) (snd pr))) \<cdot> x)))"
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    86
  apply (rule trans)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    87
  apply (rule fix_eq2)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    88
  apply (simp only: corresp_exC_def)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    89
  apply (rule beta_cfun)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    90
  apply (simp add: flift1_def)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    91
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    92
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    93
lemma corresp_exC_UU: "(corresp_exC A f \<cdot> UU) s = UU"
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    94
  apply (subst corresp_exC_unfold)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    95
  apply simp
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    96
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    97
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    98
lemma corresp_exC_nil: "(corresp_exC A f \<cdot> nil) s = nil"
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    99
  apply (subst corresp_exC_unfold)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   100
  apply simp
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   101
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   102
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   103
lemma corresp_exC_cons:
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
   104
  "(corresp_exC A f \<cdot> (at \<leadsto> xs)) s =
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   105
     (SOME cex. move A cex (f s) (fst at) (f (snd at))) @@
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
   106
     ((corresp_exC A f \<cdot> xs) (snd at))"
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   107
  apply (rule trans)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   108
  apply (subst corresp_exC_unfold)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   109
  apply (simp add: Consq_def flift1_def)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   110
  apply simp
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   111
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   112
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   113
declare corresp_exC_UU [simp] corresp_exC_nil [simp] corresp_exC_cons [simp]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   114
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   115
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   116
subsection \<open>Properties of move\<close>
19741
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_is_move:
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   119
  "is_ref_map f C A \<Longrightarrow> reachable C s \<Longrightarrow> (s, a, t) \<in> trans_of C \<Longrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   120
    move A (SOME x. move A x (f s) a (f t)) (f s) a (f t)"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   121
  apply (unfold is_ref_map_def)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   122
  apply (subgoal_tac "\<exists>ex. move A ex (f s) a (f t) ")
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   123
  prefer 2
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   124
  apply simp
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   125
  apply (erule exE)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   126
  apply (rule someI)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   127
  apply assumption
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   128
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   129
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   130
lemma move_subprop1:
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   131
  "is_ref_map f C A \<Longrightarrow> reachable C s \<Longrightarrow> (s, a, t) \<in> trans_of C \<Longrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   132
    is_exec_frag A (f s, SOME x. move A x (f s) a (f t))"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   133
  apply (cut_tac move_is_move)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   134
  defer
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   135
  apply assumption+
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   136
  apply (simp add: move_def)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   137
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   138
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   139
lemma move_subprop2:
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   140
  "is_ref_map f C A \<Longrightarrow> reachable C s \<Longrightarrow> (s, a, t) \<in> trans_of C \<Longrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   141
    Finite ((SOME x. move A x (f s) a (f t)))"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   142
  apply (cut_tac move_is_move)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   143
  defer
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   144
  apply assumption+
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   145
  apply (simp add: move_def)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   146
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   147
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   148
lemma move_subprop3:
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   149
  "is_ref_map f C A \<Longrightarrow> reachable C s \<Longrightarrow> (s, a, t) \<in> trans_of C \<Longrightarrow>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63680
diff changeset
   150
     laststate (f s, SOME x. move A x (f s) a (f t)) = (f t)"
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   151
  apply (cut_tac move_is_move)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   152
  defer
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   153
  apply assumption+
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   154
  apply (simp add: move_def)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   155
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   156
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   157
lemma move_subprop4:
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   158
  "is_ref_map f C A \<Longrightarrow> reachable C s \<Longrightarrow> (s, a, t) \<in> trans_of C \<Longrightarrow>
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
   159
    mk_trace A \<cdot> ((SOME x. move A x (f s) a (f t))) =
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   160
      (if a \<in> ext A then a \<leadsto> nil else nil)"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   161
  apply (cut_tac move_is_move)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   162
  defer
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   163
  apply assumption+
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   164
  apply (simp add: move_def)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   165
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   166
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   167
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   168
subsection \<open>TRACE INCLUSION Part 1: Traces coincide\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   169
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   170
subsubsection \<open>Lemmata for \<open>\<Longleftarrow>\<close>\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   171
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   172
text \<open>Lemma 1.1: Distribution of \<open>mk_trace\<close> and \<open>@@\<close>\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   173
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   174
lemma mk_traceConc:
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
   175
  "mk_trace C \<cdot> (ex1 @@ ex2) = (mk_trace C \<cdot> ex1) @@ (mk_trace C \<cdot> ex2)"
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   176
  by (simp add: mk_trace_def filter_act_def MapConc)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   177
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   178
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   179
text \<open>Lemma 1 : Traces coincide\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   180
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   181
lemma lemma_1:
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   182
  "is_ref_map f C A \<Longrightarrow> ext C = ext A \<Longrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   183
    \<forall>s. reachable C s \<and> is_exec_frag C (s, xs) \<longrightarrow>
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
   184
      mk_trace C \<cdot> xs = mk_trace A \<cdot> (snd (corresp_ex A f (s, xs)))"
62390
842917225d56 more canonical names
nipkow
parents: 62195
diff changeset
   185
  supply if_split [split del]
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   186
  apply (unfold corresp_ex_def)
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62156
diff changeset
   187
  apply (pair_induct xs simp: is_exec_frag_def)
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   188
  text \<open>cons case\<close>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   189
  apply (auto simp add: mk_traceConc)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   190
  apply (frule reachable.reachable_n)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   191
  apply assumption
63648
f9f3006a5579 "split add" -> "split"
nipkow
parents: 63549
diff changeset
   192
  apply (auto simp add: move_subprop4 split: if_split)
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   193
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   194
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   195
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   196
subsection \<open>TRACE INCLUSION Part 2: corresp_ex is execution\<close>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   197
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   198
subsubsection \<open>Lemmata for \<open>==>\<close>\<close>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   199
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   200
text \<open>Lemma 2.1\<close>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   201
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   202
lemma lemma_2_1 [rule_format]:
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   203
  "Finite xs \<longrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   204
    (\<forall>s. is_exec_frag A (s, xs) \<and> is_exec_frag A (t, ys) \<and>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   205
      t = laststate (s, xs) \<longrightarrow> is_exec_frag A (s, xs @@ ys))"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   206
  apply (rule impI)
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62156
diff changeset
   207
  apply Seq_Finite_induct
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   208
  text \<open>main case\<close>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   209
  apply (auto simp add: split_paired_all)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   210
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   211
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   212
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   213
text \<open>Lemma 2 : \<open>corresp_ex\<close> is execution\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   214
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   215
lemma lemma_2:
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   216
  "is_ref_map f C A \<Longrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   217
    \<forall>s. reachable C s \<and> is_exec_frag C (s, xs) \<longrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   218
      is_exec_frag A (corresp_ex A f (s, xs))"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   219
  apply (unfold corresp_ex_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   220
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   221
  apply simp
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62156
diff changeset
   222
  apply (pair_induct xs simp: is_exec_frag_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   223
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   224
  text \<open>main case\<close>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   225
  apply auto
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   226
  apply (rule_tac t = "f x2" in lemma_2_1)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   227
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   228
  text \<open>\<open>Finite\<close>\<close>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   229
  apply (erule move_subprop2)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   230
  apply assumption+
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   231
  apply (rule conjI)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   232
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   233
  text \<open>\<open>is_exec_frag\<close>\<close>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   234
  apply (erule move_subprop1)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   235
  apply assumption+
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   236
  apply (rule conjI)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   237
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   238
  text \<open>Induction hypothesis\<close>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   239
  text \<open>\<open>reachable_n\<close> looping, therefore apply it manually\<close>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   240
  apply (erule_tac x = "x2" in allE)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   241
  apply simp
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   242
  apply (frule reachable.reachable_n)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   243
  apply assumption
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   244
  apply simp
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   245
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   246
  text \<open>\<open>laststate\<close>\<close>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   247
  apply (erule move_subprop3 [symmetric])
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   248
  apply assumption+
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   249
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   250
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   251
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   252
subsection \<open>Main Theorem: TRACE -- INCLUSION\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   253
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   254
theorem trace_inclusion:
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   255
  "ext C = ext A \<Longrightarrow> is_ref_map f C A \<Longrightarrow> traces C \<subseteq> traces A"
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   256
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   257
  apply (unfold traces_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   258
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   259
  apply (simp add: has_trace_def2)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   260
  apply auto
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   261
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   262
  text \<open>give execution of abstract automata\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   263
  apply (rule_tac x = "corresp_ex A f ex" in bexI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   264
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   265
  text \<open>Traces coincide, Lemma 1\<close>
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62156
diff changeset
   266
  apply (pair ex)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   267
  apply (erule lemma_1 [THEN spec, THEN mp])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   268
  apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   269
  apply (simp add: executions_def reachable.reachable_0)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   270
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   271
  text \<open>\<open>corresp_ex\<close> is execution, Lemma 2\<close>
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62156
diff changeset
   272
  apply (pair ex)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   273
  apply (simp add: executions_def)
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   274
  text \<open>start state\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   275
  apply (rule conjI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   276
  apply (simp add: is_ref_map_def corresp_ex_def)
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   277
  text \<open>\<open>is_execution_fragment\<close>\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   278
  apply (erule lemma_2 [THEN spec, THEN mp])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   279
  apply (simp add: reachable.reachable_0)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   280
  done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   281
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   282
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   283
subsection \<open>Corollary:  FAIR TRACE -- INCLUSION\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   284
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   285
lemma fininf: "(~inf_often P s) = fin_often P s"
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   286
  by (auto simp: fin_often_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   287
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   288
lemma WF_alt: "is_wfair A W (s, ex) =
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   289
  (fin_often (\<lambda>x. \<not> Enabled A W (snd x)) ex \<longrightarrow> inf_often (\<lambda>x. fst x \<in> W) ex)"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   290
  by (auto simp add: is_wfair_def fin_often_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   291
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   292
lemma WF_persistent:
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   293
  "is_wfair A W (s, ex) \<Longrightarrow> inf_often (\<lambda>x. Enabled A W (snd x)) ex \<Longrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   294
    en_persistent A W \<Longrightarrow> inf_often (\<lambda>x. fst x \<in> W) ex"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   295
  apply (drule persistent)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   296
  apply assumption
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   297
  apply (simp add: WF_alt)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   298
  apply auto
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   299
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   300
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   301
lemma fair_trace_inclusion:
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   302
  assumes "is_ref_map f C A"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   303
    and "ext C = ext A"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   304
    and "\<And>ex. ex \<in> executions C \<Longrightarrow> fair_ex C ex \<Longrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   305
      fair_ex A (corresp_ex A f ex)"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   306
  shows "fairtraces C \<subseteq> fairtraces A"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   307
  apply (insert assms)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   308
  apply (simp add: fairtraces_def fairexecutions_def)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   309
  apply auto
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   310
  apply (rule_tac x = "corresp_ex A f ex" in exI)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   311
  apply auto
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   312
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   313
  text \<open>Traces coincide, Lemma 1\<close>
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62156
diff changeset
   314
  apply (pair ex)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   315
  apply (erule lemma_1 [THEN spec, THEN mp])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   316
  apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   317
  apply (simp add: executions_def reachable.reachable_0)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   318
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   319
  text \<open>\<open>corresp_ex\<close> is execution, Lemma 2\<close>
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62156
diff changeset
   320
  apply (pair ex)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   321
  apply (simp add: executions_def)
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   322
  text \<open>start state\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   323
  apply (rule conjI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   324
  apply (simp add: is_ref_map_def corresp_ex_def)
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   325
  text \<open>\<open>is_execution_fragment\<close>\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   326
  apply (erule lemma_2 [THEN spec, THEN mp])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   327
  apply (simp add: reachable.reachable_0)
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   328
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   329
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   330
lemma fair_trace_inclusion2:
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   331
  "inp C = inp A \<Longrightarrow> out C = out A \<Longrightarrow> is_fair_ref_map f C A \<Longrightarrow>
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   332
    fair_implements C A"
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   333
  apply (simp add: is_fair_ref_map_def fair_implements_def fairtraces_def fairexecutions_def)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   334
  apply auto
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   335
  apply (rule_tac x = "corresp_ex A f ex" in exI)
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   336
  apply auto
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   337
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   338
  text \<open>Traces coincide, Lemma 1\<close>
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62156
diff changeset
   339
  apply (pair ex)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   340
  apply (erule lemma_1 [THEN spec, THEN mp])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   341
  apply (simp (no_asm) add: externals_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   342
  apply (auto)[1]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   343
  apply (simp add: executions_def reachable.reachable_0)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   344
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   345
  text \<open>\<open>corresp_ex\<close> is execution, Lemma 2\<close>
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62156
diff changeset
   346
  apply (pair ex)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   347
  apply (simp add: executions_def)
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   348
  text \<open>start state\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   349
  apply (rule conjI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   350
  apply (simp add: is_ref_map_def corresp_ex_def)
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   351
  text \<open>\<open>is_execution_fragment\<close>\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   352
  apply (erule lemma_2 [THEN spec, THEN mp])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   353
  apply (simp add: reachable.reachable_0)
62156
7355fd313cf8 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   354
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   355
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   356
end