src/HOL/HOLCF/IOA/SimCorrectness.thy
author wenzelm
Fri, 04 Nov 2016 13:27:31 +0100
changeset 64460 e84fb8b4245c
parent 63648 f9f3006a5579
child 67613 ce654b0e6d69
permissions -rw-r--r--
back to stable scala-2.11.8;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62008
cbedaddc9351 clarified directory structure;
wenzelm
parents: 62003
diff changeset
     1
(*  Title:      HOL/HOLCF/IOA/SimCorrectness.thy
40945
b8703f63bfb2 recoded latin1 as utf8;
wenzelm
parents: 40774
diff changeset
     2
    Author:     Olaf Müller
4565
ea467ce15040 added forward simulation correectness;
mueller
parents:
diff changeset
     3
*)
ea467ce15040 added forward simulation correectness;
mueller
parents:
diff changeset
     4
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
     5
section \<open>Correctness of Simulations in HOLCF/IOA\<close>
4565
ea467ce15040 added forward simulation correectness;
mueller
parents:
diff changeset
     6
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     7
theory SimCorrectness
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     8
imports Simulations
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
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    11
(*Note: s2 instead of s1 in last argument type!*)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    12
definition corresp_ex_simC ::
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    13
    "('a, 's2) ioa \<Rightarrow> ('s1 \<times> 's2) set \<Rightarrow> ('a, 's1) pairs \<rightarrow> ('s2 \<Rightarrow> ('a, 's2) pairs)"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    14
  where "corresp_ex_simC A R =
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    15
    (fix \<cdot> (LAM h ex. (\<lambda>s. case ex of
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    16
      nil \<Rightarrow> nil
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    17
    | x ## xs \<Rightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    18
        (flift1
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    19
          (\<lambda>pr.
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    20
            let
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    21
              a = fst pr;
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    22
              t = snd pr;
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    23
              T' = SOME t'. \<exists>ex1. (t, t') \<in> R \<and> move A ex1 s a t'
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    24
            in (SOME cex. move A cex s a T') @@ ((h \<cdot> xs) T')) \<cdot> x))))"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    25
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    26
definition corresp_ex_sim ::
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    27
    "('a, 's2) ioa \<Rightarrow> ('s1 \<times> 's2) set \<Rightarrow> ('a, 's1) execution \<Rightarrow> ('a, 's2) execution"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    28
  where "corresp_ex_sim A R ex \<equiv>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    29
    let S' = SOME s'. (fst ex, s') \<in> R \<and> s' \<in> starts_of A
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    30
    in (S', (corresp_ex_simC A R \<cdot> (snd ex)) S')"
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    31
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    32
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    33
subsection \<open>\<open>corresp_ex_sim\<close>\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    34
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    35
lemma corresp_ex_simC_unfold:
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    36
  "corresp_ex_simC A R =
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    37
    (LAM ex. (\<lambda>s. case ex of
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    38
      nil \<Rightarrow> nil
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    39
    | x ## xs \<Rightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    40
        (flift1
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    41
          (\<lambda>pr.
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    42
            let
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    43
              a = fst pr;
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    44
              t = snd pr;
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    45
              T' = SOME t'. \<exists>ex1. (t, t') \<in> R \<and> move A ex1 s a t'
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    46
            in (SOME cex. move A cex s a T') @@ ((corresp_ex_simC A R \<cdot> xs) T')) \<cdot> x)))"
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    47
  apply (rule trans)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    48
  apply (rule fix_eq2)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    49
  apply (simp only: corresp_ex_simC_def)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    50
  apply (rule beta_cfun)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    51
  apply (simp add: flift1_def)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    52
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    53
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    54
lemma corresp_ex_simC_UU [simp]: "(corresp_ex_simC A R \<cdot> UU) s = UU"
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    55
  apply (subst corresp_ex_simC_unfold)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    56
  apply simp
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    57
  done
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    58
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    59
lemma corresp_ex_simC_nil [simp]: "(corresp_ex_simC A R \<cdot> nil) s = nil"
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    60
  apply (subst corresp_ex_simC_unfold)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    61
  apply simp
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    62
  done
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    63
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    64
lemma corresp_ex_simC_cons [simp]:
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    65
  "(corresp_ex_simC A R \<cdot> ((a, t) \<leadsto> xs)) s =
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    66
    (let T' = SOME t'. \<exists>ex1. (t, t') \<in> R \<and> move A ex1 s a t'
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    67
     in (SOME cex. move A cex s a T') @@ ((corresp_ex_simC A R \<cdot> xs) T'))"
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    68
  apply (rule trans)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    69
  apply (subst corresp_ex_simC_unfold)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    70
  apply (simp add: Consq_def flift1_def)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    71
  apply simp
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    72
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    73
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    74
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    75
subsection \<open>Properties of move\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    76
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    77
lemma move_is_move_sim:
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    78
   "is_simulation R C A \<Longrightarrow> reachable C s \<Longrightarrow> s \<midarrow>a\<midarrow>C\<rightarrow> t \<Longrightarrow> (s, s') \<in> R \<Longrightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    79
     let T' = SOME t'. \<exists>ex1. (t, t') \<in> R \<and> move A ex1 s' a t'
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    80
     in (t, T') \<in> R \<and> move A (SOME ex2. move A ex2 s' a T') s' a T'"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    81
  supply Let_def [simp del]
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    82
  apply (unfold is_simulation_def)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    83
  (* Does not perform conditional rewriting on assumptions automatically as
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    84
     usual. Instantiate all variables per hand. Ask Tobias?? *)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    85
  apply (subgoal_tac "\<exists>t' ex. (t, t') \<in> R \<and> move A ex s' a t'")
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    86
  prefer 2
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    87
  apply simp
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    88
  apply (erule conjE)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    89
  apply (erule_tac x = "s" in allE)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    90
  apply (erule_tac x = "s'" in allE)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    91
  apply (erule_tac x = "t" in allE)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    92
  apply (erule_tac x = "a" in allE)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    93
  apply simp
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    94
  (* Go on as usual *)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    95
  apply (erule exE)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    96
  apply (drule_tac x = "t'" and P = "\<lambda>t'. \<exists>ex. (t, t') \<in> R \<and> move A ex s' a t'" in someI)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    97
  apply (erule exE)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    98
  apply (erule conjE)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    99
  apply (simp add: Let_def)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   100
  apply (rule_tac x = "ex" in someI)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   101
  apply assumption
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   102
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   103
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   104
lemma move_subprop1_sim:
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   105
  "is_simulation R C A \<Longrightarrow> reachable C s \<Longrightarrow> s \<midarrow>a\<midarrow>C\<rightarrow> t \<Longrightarrow> (s, s') \<in> R \<Longrightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   106
    let T' = SOME t'. \<exists>ex1. (t, t') \<in> R \<and> move A ex1 s' a t'
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   107
    in is_exec_frag A (s', SOME x. move A x s' a T')"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   108
  apply (cut_tac move_is_move_sim)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   109
  defer
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   110
  apply assumption+
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   111
  apply (simp add: move_def)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   112
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   113
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   114
lemma move_subprop2_sim:
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   115
  "is_simulation R C A \<Longrightarrow> reachable C s \<Longrightarrow> s \<midarrow>a\<midarrow>C\<rightarrow> t \<Longrightarrow> (s, s') \<in> R \<Longrightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   116
    let T' = SOME t'. \<exists>ex1. (t, t') \<in> R \<and> move A ex1 s' a t'
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   117
    in Finite (SOME x. move A x s' a T')"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   118
  apply (cut_tac move_is_move_sim)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   119
  defer
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   120
  apply assumption+
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   121
  apply (simp add: move_def)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   122
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   123
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   124
lemma move_subprop3_sim:
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   125
  "is_simulation R C A \<Longrightarrow> reachable C s \<Longrightarrow> s \<midarrow>a\<midarrow>C\<rightarrow> t \<Longrightarrow> (s, s') \<in> R \<Longrightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   126
    let T' = SOME t'. \<exists>ex1. (t, t') \<in> R \<and> move A ex1 s' a t'
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   127
    in laststate (s', SOME x. move A x s' a T') = T'"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   128
  apply (cut_tac move_is_move_sim)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   129
  defer
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   130
  apply assumption+
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   131
  apply (simp add: move_def)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   132
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   133
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   134
lemma move_subprop4_sim:
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   135
  "is_simulation R C A \<Longrightarrow> reachable C s \<Longrightarrow> s \<midarrow>a\<midarrow>C\<rightarrow> t \<Longrightarrow> (s, s') \<in> R \<Longrightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   136
    let T' = SOME t'. \<exists>ex1. (t, t') \<in> R \<and> move A ex1 s' a t'
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
   137
    in mk_trace A \<cdot> (SOME x. move A x s' a T') = (if a \<in> ext A then a \<leadsto> nil else nil)"
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   138
  apply (cut_tac move_is_move_sim)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   139
  defer
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   140
  apply assumption+
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   141
  apply (simp add: move_def)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   142
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   143
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   144
lemma move_subprop5_sim:
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   145
  "is_simulation R C A \<Longrightarrow> reachable C s \<Longrightarrow> s \<midarrow>a\<midarrow>C\<rightarrow> t \<Longrightarrow> (s, s') \<in> R \<Longrightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   146
    let T' = SOME t'. \<exists>ex1. (t, t') \<in> R \<and> move A ex1 s' a t'
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   147
    in (t, T') \<in> R"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   148
  apply (cut_tac move_is_move_sim)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   149
  defer
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   150
  apply assumption+
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   151
  apply (simp add: move_def)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   152
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   153
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   154
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
   155
subsection \<open>TRACE INCLUSION Part 1: Traces coincide\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   156
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   157
subsubsection "Lemmata for <=="
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   158
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   159
text \<open>Lemma 1: Traces coincide\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   160
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   161
lemma traces_coincide_sim [rule_format (no_asm)]:
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   162
  "is_simulation R C A \<Longrightarrow> ext C = ext A \<Longrightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   163
    \<forall>s s'. reachable C s \<and> is_exec_frag C (s, ex) \<and> (s, s') \<in> R \<longrightarrow>
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
   164
      mk_trace C \<cdot> ex = mk_trace A \<cdot> ((corresp_ex_simC A R \<cdot> ex) s')"
62390
842917225d56 more canonical names
nipkow
parents: 62195
diff changeset
   165
  supply if_split [split del]
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62192
diff changeset
   166
  apply (pair_induct ex simp: is_exec_frag_def)
799a5306e2ed more method definitions;
wenzelm
parents: 62192
diff changeset
   167
  text \<open>cons case\<close>
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   168
  apply auto
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   169
  apply (rename_tac ex a t s s')
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   170
  apply (simp add: mk_traceConc)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   171
  apply (frule reachable.reachable_n)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   172
  apply assumption
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   173
  apply (erule_tac x = "t" in allE)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   174
  apply (erule_tac x = "SOME t'. \<exists>ex1. (t, t') \<in> R \<and> move A ex1 s' a t'" in allE)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   175
  apply (simp add: move_subprop5_sim [unfolded Let_def]
63648
f9f3006a5579 "split add" -> "split"
nipkow
parents: 63549
diff changeset
   176
    move_subprop4_sim [unfolded Let_def] split: if_split)
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   177
  done
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   178
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   179
text \<open>Lemma 2: \<open>corresp_ex_sim\<close> is execution\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   180
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   181
lemma correspsim_is_execution [rule_format]:
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   182
  "is_simulation R C A \<Longrightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   183
    \<forall>s s'. reachable C s \<and> is_exec_frag C (s, ex) \<and> (s, s') \<in> R
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
   184
      \<longrightarrow> is_exec_frag A (s', (corresp_ex_simC A R \<cdot> ex) s')"
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62192
diff changeset
   185
  apply (pair_induct ex simp: is_exec_frag_def)
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   186
  text \<open>main case\<close>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   187
  apply auto
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   188
  apply (rename_tac ex a t s s')
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   189
  apply (rule_tac t = "SOME t'. \<exists>ex1. (t, t') \<in> R \<and> move A ex1 s' a t'" in lemma_2_1)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   190
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   191
  text \<open>Finite\<close>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   192
  apply (erule move_subprop2_sim [unfolded Let_def])
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   193
  apply assumption+
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   194
  apply (rule conjI)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   195
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   196
  text \<open>\<open>is_exec_frag\<close>\<close>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   197
  apply (erule move_subprop1_sim [unfolded Let_def])
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   198
  apply assumption+
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   199
  apply (rule conjI)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   200
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   201
  text \<open>Induction hypothesis\<close>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   202
  text \<open>\<open>reachable_n\<close> looping, therefore apply it manually\<close>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   203
  apply (erule_tac x = "t" in allE)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   204
  apply (erule_tac x = "SOME t'. \<exists>ex1. (t, t') \<in> R \<and> move A ex1 s' a t'" in allE)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   205
  apply simp
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   206
  apply (frule reachable.reachable_n)
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   207
  apply assumption
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   208
  apply (simp add: move_subprop5_sim [unfolded Let_def])
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   209
  text \<open>laststate\<close>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   210
  apply (erule move_subprop3_sim [unfolded Let_def, symmetric])
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   211
  apply assumption+
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   212
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   213
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   214
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   215
subsection \<open>Main Theorem: TRACE - INCLUSION\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   216
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   217
text \<open>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   218
  Generate condition \<open>(s, S') \<in> R \<and> S' \<in> starts_of A\<close>, the first being
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   219
  interesting for the induction cases concerning the two lemmas
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   220
  \<open>correpsim_is_execution\<close> and \<open>traces_coincide_sim\<close>, the second for the start
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   221
  state case.
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   222
  \<open>S' := SOME s'. (s, s') \<in> R \<and> s' \<in> starts_of A\<close>, where \<open>s \<in> starts_of C\<close>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   223
\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   224
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   225
lemma simulation_starts:
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   226
  "is_simulation R C A \<Longrightarrow> s:starts_of C \<Longrightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   227
    let S' = SOME s'. (s, s') \<in> R \<and> s' \<in> starts_of A
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   228
    in (s, S') \<in> R \<and> S' \<in> starts_of A"
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   229
  apply (simp add: is_simulation_def corresp_ex_sim_def Int_non_empty Image_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   230
  apply (erule conjE)+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   231
  apply (erule ballE)
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   232
  prefer 2 apply blast
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   233
  apply (erule exE)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   234
  apply (rule someI2)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   235
  apply assumption
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   236
  apply blast
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   237
  done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   238
45606
b1e1508643b1 eliminated obsolete "standard";
wenzelm
parents: 42151
diff changeset
   239
lemmas sim_starts1 = simulation_starts [unfolded Let_def, THEN conjunct1]
b1e1508643b1 eliminated obsolete "standard";
wenzelm
parents: 42151
diff changeset
   240
lemmas sim_starts2 = simulation_starts [unfolded Let_def, THEN conjunct2]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   241
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   242
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
   243
lemma trace_inclusion_for_simulations:
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   244
  "ext C = ext A \<Longrightarrow> is_simulation R C A \<Longrightarrow> traces C \<subseteq> traces A"
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   245
  apply (unfold traces_def)
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   246
  apply (simp add: has_trace_def2)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   247
  apply auto
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   248
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   249
  text \<open>give execution of abstract automata\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   250
  apply (rule_tac x = "corresp_ex_sim A R ex" in bexI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   251
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   252
  text \<open>Traces coincide, Lemma 1\<close>
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62192
diff changeset
   253
  apply (pair ex)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   254
  apply (rename_tac s ex)
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   255
  apply (simp add: corresp_ex_sim_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   256
  apply (rule_tac s = "s" in traces_coincide_sim)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   257
  apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   258
  apply (simp add: executions_def reachable.reachable_0 sim_starts1)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   259
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   260
  text \<open>\<open>corresp_ex_sim\<close> is execution, Lemma 2\<close>
62195
799a5306e2ed more method definitions;
wenzelm
parents: 62192
diff changeset
   261
  apply (pair ex)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   262
  apply (simp add: executions_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   263
  apply (rename_tac s ex)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   264
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   265
  text \<open>start state\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   266
  apply (rule conjI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   267
  apply (simp add: sim_starts2 corresp_ex_sim_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   268
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   269
  text \<open>\<open>is_execution-fragment\<close>\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   270
  apply (simp add: corresp_ex_sim_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   271
  apply (rule_tac s = s in correspsim_is_execution)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   272
  apply assumption
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   273
  apply (simp add: reachable.reachable_0 sim_starts1)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   274
  done
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   275
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   276
end