PIPBasics.thy
author zhangx
Thu, 28 Jan 2016 21:14:17 +0800
changeset 90 ed938e2246b9
parent 84 cfd644dfc3b4
child 92 4763aa246dbd
permissions -rw-r--r--
Retrofiting of: CpsG.thy (the parallel copy of PIPBasics.thy), ExtGG.thy (The paralell copy of Implemenation.thy), PrioG.thy (The paralell copy of Correctness.thy) has completed. The next step is to overwite original copies with the paralell ones.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
64
b4bcd1edbb6d renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
     1
theory PIPBasics
b4bcd1edbb6d renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
     2
imports PIPDefs 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     3
begin
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
     5
locale valid_trace = 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
     6
  fixes s
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
     7
  assumes vt : "vt s"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
     8
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
     9
locale valid_trace_e = valid_trace +
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    10
  fixes e
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    11
  assumes vt_e: "vt (e#s)"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    12
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    13
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    14
lemma pip_e: "PIP s e"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    15
  using vt_e by (cases, simp)  
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    16
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    17
end
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    18
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
lemma runing_ready: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
  shows "runing s \<subseteq> readys s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
  unfolding runing_def readys_def
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
  by auto 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
lemma readys_threads:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
  shows "readys s \<subseteq> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
  unfolding readys_def
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
  by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    29
lemma wq_v_neq:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    30
   "cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    31
  by (auto simp:wq_def Let_def cp_def split:list.splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    33
lemma runing_head:
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    34
  assumes "th \<in> runing s"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    35
  and "th \<in> set (wq_fun (schs s) cs)"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    36
  shows "th = hd (wq_fun (schs s) cs)"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    37
  using assms
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    38
  by (simp add:runing_def readys_def s_waiting_def wq_def)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    39
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    40
context valid_trace
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    41
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    42
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    43
lemma actor_inv: 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    44
  assumes "PIP s e"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    45
  and "\<not> isCreate e"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    46
  shows "actor e \<in> runing s"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    47
  using assms
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    48
  by (induct, auto)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    49
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    50
lemma ind [consumes 0, case_names Nil Cons, induct type]:
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    51
  assumes "PP []"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    52
     and "(\<And>s e. valid_trace s \<Longrightarrow> valid_trace (e#s) \<Longrightarrow>
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    53
                   PP s \<Longrightarrow> PIP s e \<Longrightarrow> PP (e # s))"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    54
     shows "PP s"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    55
proof(rule vt.induct[OF vt])
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    56
  from assms(1) show "PP []" .
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    57
next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    58
  fix s e
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    59
  assume h: "vt s" "PP s" "PIP s e"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    60
  show "PP (e # s)"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    61
  proof(cases rule:assms(2))
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    62
    from h(1) show v1: "valid_trace s" by (unfold_locales, simp)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    63
  next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    64
    from h(1,3) have "vt (e#s)" by auto
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    65
    thus "valid_trace (e # s)" by (unfold_locales, simp)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    66
  qed (insert h, auto)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    67
qed
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    68
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    69
lemma wq_distinct: "distinct (wq s cs)"
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    70
proof(induct rule:ind)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    71
  case (Cons s e)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    72
  from Cons(4,3)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    73
  show ?case 
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    74
  proof(induct)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    75
    case (thread_P th s cs1)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    76
    show ?case 
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    77
    proof(cases "cs = cs1")
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    78
      case True
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    79
      thus ?thesis (is "distinct ?L")
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    80
      proof - 
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    81
        have "?L = wq_fun (schs s) cs1 @ [th]" using True
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    82
          by (simp add:wq_def wf_def Let_def split:list.splits)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    83
        moreover have "distinct ..."
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    84
        proof -
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    85
          have "th \<notin> set (wq_fun (schs s) cs1)"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    86
          proof
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    87
            assume otherwise: "th \<in> set (wq_fun (schs s) cs1)"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    88
            from runing_head[OF thread_P(1) this]
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    89
            have "th = hd (wq_fun (schs s) cs1)" .
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    90
            hence "(Cs cs1, Th th) \<in> (RAG s)" using otherwise
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    91
              by (simp add:s_RAG_def s_holding_def wq_def cs_holding_def)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    92
            with thread_P(2) show False by auto
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    93
          qed
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    94
          moreover have "distinct (wq_fun (schs s) cs1)"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    95
              using True thread_P wq_def by auto 
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    96
          ultimately show ?thesis by auto
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    97
        qed
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    98
        ultimately show ?thesis by simp
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    99
      qed
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   100
    next
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   101
      case False
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   102
      with thread_P(3)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   103
      show ?thesis
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   104
        by (auto simp:wq_def wf_def Let_def split:list.splits)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   105
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   106
  next
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   107
    case (thread_V th s cs1)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   108
    thus ?case
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   109
    proof(cases "cs = cs1")
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   110
      case True
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   111
      show ?thesis (is "distinct ?L")
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   112
      proof(cases "(wq s cs)")
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   113
        case Nil
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   114
        thus ?thesis
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   115
          by (auto simp:wq_def wf_def Let_def split:list.splits)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   116
      next
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   117
        case (Cons w_hd w_tl)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   118
        moreover have "distinct (SOME q. distinct q \<and> set q = set w_tl)"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   119
        proof(rule someI2)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   120
          from thread_V(3)[unfolded Cons]
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   121
          show  "distinct w_tl \<and> set w_tl = set w_tl" by auto
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   122
        qed auto
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   123
        ultimately show ?thesis
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   124
          by (auto simp:wq_def wf_def Let_def True split:list.splits)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   125
      qed 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   126
    next
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   127
      case False
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   128
      with thread_V(3)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   129
      show ?thesis
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   130
        by (auto simp:wq_def wf_def Let_def split:list.splits)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   131
    qed
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   132
  qed (insert Cons, auto simp: wq_def Let_def split:list.splits)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   133
qed (unfold wq_def Let_def, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   134
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   135
end
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   136
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   137
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   138
context valid_trace_e
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   139
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   140
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   141
text {*
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   142
  The following lemma shows that only the @{text "P"}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   143
  operation can add new thread into waiting queues. 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   144
  Such kind of lemmas are very obvious, but need to be checked formally.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   145
  This is a kind of confirmation that our modelling is correct.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   146
*}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   147
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
lemma block_pre: 
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   149
  assumes s_ni: "thread \<notin> set (wq s cs)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   150
  and s_i: "thread \<in> set (wq (e#s) cs)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   151
  shows "e = P thread cs"
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   152
proof(cases e)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   153
  -- {* This is the only non-trivial case: *}
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   154
  case (V th cs1)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   155
  have False
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   156
  proof(cases "cs1 = cs")
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   157
    case True
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   158
    show ?thesis
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   159
    proof(cases "(wq s cs1)")
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   160
      case (Cons w_hd w_tl)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   161
      have "set (wq (e#s) cs) \<subseteq> set (wq s cs)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   162
      proof -
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   163
        have "(wq (e#s) cs) = (SOME q. distinct q \<and> set q = set w_tl)"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   164
          using  Cons V by (auto simp:wq_def Let_def True split:if_splits)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   165
        moreover have "set ... \<subseteq> set (wq s cs)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   166
        proof(rule someI2)
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   167
          show "distinct w_tl \<and> set w_tl = set w_tl"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   168
            by (metis distinct.simps(2) local.Cons wq_distinct)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   169
        qed (insert Cons True, auto)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   170
        ultimately show ?thesis by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   171
      qed
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   172
      with assms show ?thesis by auto
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   173
    qed (insert assms V True, auto simp:wq_def Let_def split:if_splits)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   174
  qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   175
  thus ?thesis by auto
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   176
qed (insert assms, auto simp:wq_def Let_def split:if_splits)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   177
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   178
end
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   179
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   180
text {*
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   181
  The following lemmas is also obvious and shallow. It says
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   182
  that only running thread can request for a critical resource 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   183
  and that the requested resource must be one which is
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   184
  not current held by the thread.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   185
*}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   186
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   187
lemma p_pre: "\<lbrakk>vt ((P thread cs)#s)\<rbrakk> \<Longrightarrow> 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   188
  thread \<in> runing s \<and> (Cs cs, Th thread)  \<notin> (RAG s)^+"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   189
apply (ind_cases "vt ((P thread cs)#s)")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   190
apply (ind_cases "step s (P thread cs)")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   191
by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   192
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   193
lemma abs1:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   194
  assumes ein: "e \<in> set es"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   195
  and neq: "hd es \<noteq> hd (es @ [x])"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   196
  shows "False"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   197
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   198
  from ein have "es \<noteq> []" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   199
  then obtain e ess where "es = e # ess" by (cases es, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   200
  with neq show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   201
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   202
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   203
lemma q_head: "Q (hd es) \<Longrightarrow> hd es = hd [th\<leftarrow>es . Q th]"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   204
  by (cases es, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   205
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   206
inductive_cases evt_cons: "vt (a#s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   207
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   208
context valid_trace_e
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   209
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   210
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   211
lemma abs2:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   212
  assumes inq: "thread \<in> set (wq s cs)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   213
  and nh: "thread = hd (wq s cs)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   214
  and qt: "thread \<noteq> hd (wq (e#s) cs)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   215
  and inq': "thread \<in> set (wq (e#s) cs)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   216
  shows "False"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   217
proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   218
  from vt_e assms show "False"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   219
    apply (cases e)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   220
    apply ((simp split:if_splits add:Let_def wq_def)[1])+
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   221
    apply (insert abs1, fast)[1]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   222
    apply (auto simp:wq_def simp:Let_def split:if_splits list.splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   223
  proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   224
    fix th qs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   225
    assume vt: "vt (V th cs # s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   226
      and th_in: "thread \<in> set (SOME q. distinct q \<and> set q = set qs)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   227
      and eq_wq: "wq_fun (schs s) cs = thread # qs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   228
    show "False"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   229
    proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   230
      from wq_distinct[of cs]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   231
        and eq_wq[folded wq_def] have "distinct (thread#qs)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   232
      moreover have "thread \<in> set qs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   233
      proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   234
        have "set (SOME q. distinct q \<and> set q = set qs) = set qs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   235
        proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   236
          from wq_distinct [of cs]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   237
          and eq_wq [folded wq_def]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   238
          show "distinct qs \<and> set qs = set qs" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   239
        next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   240
          fix x assume "distinct x \<and> set x = set qs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   241
          thus "set x = set qs" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   242
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   243
        with th_in show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   244
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   245
      ultimately show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   246
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   247
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   248
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   249
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   250
end
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   251
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   252
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   253
context valid_trace
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   254
begin
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   255
lemma  vt_moment: "\<And> t. vt (moment t s)"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   256
proof(induct rule:ind)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   257
  case Nil
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   258
  thus ?case by (simp add:vt_nil)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   259
next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   260
  case (Cons s e t)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   261
  show ?case
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   262
  proof(cases "t \<ge> length (e#s)")
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   263
    case True
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   264
    from True have "moment t (e#s) = e#s" by simp
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   265
    thus ?thesis using Cons
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   266
      by (simp add:valid_trace_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   267
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   268
    case False
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   269
    from Cons have "vt (moment t s)" by simp
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   270
    moreover have "moment t (e#s) = moment t s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   271
    proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   272
      from False have "t \<le> length s" by simp
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   273
      from moment_app [OF this, of "[e]"] 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   274
      show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   275
    qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   276
    ultimately show ?thesis by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   277
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   278
qed
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   279
end
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   280
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   281
locale valid_moment = valid_trace + 
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   282
  fixes i :: nat
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   283
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   284
sublocale valid_moment < vat_moment: valid_trace "(moment i s)"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   285
  by (unfold_locales, insert vt_moment, auto)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   286
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   287
context valid_trace
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   288
begin
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   289
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   290
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   291
text {* (* ddd *)
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   292
  The nature of the work is like this: since it starts from a very simple and basic 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   293
  model, even intuitively very `basic` and `obvious` properties need to derived from scratch.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   294
  For instance, the fact 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   295
  that one thread can not be blocked by two critical resources at the same time
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   296
  is obvious, because only running threads can make new requests, if one is waiting for 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   297
  a critical resource and get blocked, it can not make another resource request and get 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   298
  blocked the second time (because it is not running). 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   299
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   300
  To derive this fact, one needs to prove by contraction and 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   301
  reason about time (or @{text "moement"}). The reasoning is based on a generic theorem
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   302
  named @{text "p_split"}, which is about status changing along the time axis. It says if 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   303
  a condition @{text "Q"} is @{text "True"} at a state @{text "s"},
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   304
  but it was @{text "False"} at the very beginning, then there must exits a moment @{text "t"} 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   305
  in the history of @{text "s"} (notice that @{text "s"} itself is essentially the history 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   306
  of events leading to it), such that @{text "Q"} switched 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   307
  from being @{text "False"} to @{text "True"} and kept being @{text "True"}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   308
  till the last moment of @{text "s"}.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   309
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   310
  Suppose a thread @{text "th"} is blocked
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   311
  on @{text "cs1"} and @{text "cs2"} in some state @{text "s"}, 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   312
  since no thread is blocked at the very beginning, by applying 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   313
  @{text "p_split"} to these two blocking facts, there exist 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   314
  two moments @{text "t1"} and @{text "t2"}  in @{text "s"}, such that 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   315
  @{text "th"} got blocked on @{text "cs1"} and @{text "cs2"} 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   316
  and kept on blocked on them respectively ever since.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   317
 
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   318
  Without lost of generality, we assume @{text "t1"} is earlier than @{text "t2"}.
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   319
  However, since @{text "th"} was blocked ever since memonent @{text "t1"}, so it was still
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   320
  in blocked state at moment @{text "t2"} and could not
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   321
  make any request and get blocked the second time: Contradiction.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   322
*}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   323
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   324
lemma waiting_unique_pre: (* ccc *)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   325
  assumes h11: "thread \<in> set (wq s cs1)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   326
  and h12: "thread \<noteq> hd (wq s cs1)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   327
  assumes h21: "thread \<in> set (wq s cs2)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   328
  and h22: "thread \<noteq> hd (wq s cs2)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   329
  and neq12: "cs1 \<noteq> cs2"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   330
  shows "False"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   331
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   332
  let "?Q cs s" = "thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   333
  from h11 and h12 have q1: "?Q cs1 s" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   334
  from h21 and h22 have q2: "?Q cs2 s" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   335
  have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   336
  have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   337
  from p_split [of "?Q cs1", OF q1 nq1]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   338
  obtain t1 where lt1: "t1 < length s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   339
    and np1: "\<not>(thread \<in> set (wq (moment t1 s) cs1) \<and>
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   340
        thread \<noteq> hd (wq (moment t1 s) cs1))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   341
    and nn1: "(\<forall>i'>t1. thread \<in> set (wq (moment i' s) cs1) \<and>
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   342
             thread \<noteq> hd (wq (moment i' s) cs1))" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   343
  from p_split [of "?Q cs2", OF q2 nq2]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   344
  obtain t2 where lt2: "t2 < length s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   345
    and np2: "\<not>(thread \<in> set (wq (moment t2 s) cs2) \<and>
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   346
        thread \<noteq> hd (wq (moment t2 s) cs2))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   347
    and nn2: "(\<forall>i'>t2. thread \<in> set (wq (moment i' s) cs2) \<and>
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   348
             thread \<noteq> hd (wq (moment i' s) cs2))" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   349
  show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   350
  proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   351
    { 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   352
      assume lt12: "t1 < t2"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   353
      let ?t3 = "Suc t2"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   354
      from lt2 have le_t3: "?t3 \<le> length s" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   355
      from moment_plus [OF this] 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   356
      obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   357
      have "t2 < ?t3" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   358
      from nn2 [rule_format, OF this] and eq_m
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   359
      have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   360
        h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   361
      have "vt (e#moment t2 s)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   362
      proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   363
        from vt_moment 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   364
        have "vt (moment ?t3 s)" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   365
        with eq_m show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   366
      qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   367
      then interpret vt_e: valid_trace_e "moment t2 s" "e"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   368
        by (unfold_locales, auto, cases, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   369
      have ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   370
      proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   371
        case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   372
        from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   373
          by auto 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   374
        from vt_e.abs2 [OF True eq_th h2 h1]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   375
        show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   376
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   377
        case False
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   378
        from vt_e.block_pre[OF False h1]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   379
        have "e = P thread cs2" .
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   380
        with vt_e.vt_e have "vt ((P thread cs2)# moment t2 s)" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   381
        from p_pre [OF this] have "thread \<in> runing (moment t2 s)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   382
        with runing_ready have "thread \<in> readys (moment t2 s)" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   383
        with nn1 [rule_format, OF lt12]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   384
        show ?thesis  by (simp add:readys_def wq_def s_waiting_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   385
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   386
    } moreover {
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   387
      assume lt12: "t2 < t1"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   388
      let ?t3 = "Suc t1"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   389
      from lt1 have le_t3: "?t3 \<le> length s" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   390
      from moment_plus [OF this] 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   391
      obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   392
      have lt_t3: "t1 < ?t3" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   393
      from nn1 [rule_format, OF this] and eq_m
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   394
      have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   395
        h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   396
      have "vt  (e#moment t1 s)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   397
      proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   398
        from vt_moment
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   399
        have "vt (moment ?t3 s)" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   400
        with eq_m show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   401
      qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   402
      then interpret vt_e: valid_trace_e "moment t1 s" e
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   403
        by (unfold_locales, auto, cases, auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   404
      have ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   405
      proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   406
        case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   407
        from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   408
          by auto
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   409
        from vt_e.abs2 True eq_th h2 h1
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   410
        show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   411
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   412
        case False
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   413
        from vt_e.block_pre [OF False h1]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   414
        have "e = P thread cs1" .
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   415
        with vt_e.vt_e have "vt ((P thread cs1)# moment t1 s)" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   416
        from p_pre [OF this] have "thread \<in> runing (moment t1 s)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   417
        with runing_ready have "thread \<in> readys (moment t1 s)" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   418
        with nn2 [rule_format, OF lt12]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   419
        show ?thesis  by (simp add:readys_def wq_def s_waiting_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   420
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   421
    } moreover {
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   422
      assume eqt12: "t1 = t2"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   423
      let ?t3 = "Suc t1"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   424
      from lt1 have le_t3: "?t3 \<le> length s" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   425
      from moment_plus [OF this] 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   426
      obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   427
      have lt_t3: "t1 < ?t3" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   428
      from nn1 [rule_format, OF this] and eq_m
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   429
      have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   430
        h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   431
      have vt_e: "vt (e#moment t1 s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   432
      proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   433
        from vt_moment
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   434
        have "vt (moment ?t3 s)" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   435
        with eq_m show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   436
      qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   437
      then interpret vt_e: valid_trace_e "moment t1 s" e
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   438
        by (unfold_locales, auto, cases, auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   439
      have ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   440
      proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   441
        case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   442
        from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   443
          by auto
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   444
        from vt_e.abs2 [OF True eq_th h2 h1]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   445
        show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   446
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   447
        case False
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   448
        from vt_e.block_pre [OF False h1]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   449
        have eq_e1: "e = P thread cs1" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   450
        have lt_t3: "t1 < ?t3" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   451
        with eqt12 have "t2 < ?t3" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   452
        from nn2 [rule_format, OF this] and eq_m and eqt12
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   453
        have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   454
          h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   455
        show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   456
        proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   457
          case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   458
          from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   459
            by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   460
          from vt_e and eqt12 have "vt (e#moment t2 s)" by simp 
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   461
          then interpret vt_e2: valid_trace_e "moment t2 s" e
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   462
            by (unfold_locales, auto, cases, auto)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   463
          from vt_e2.abs2 [OF True eq_th h2 h1]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   464
          show ?thesis .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   465
        next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   466
          case False
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   467
          have "vt (e#moment t2 s)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   468
          proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   469
            from vt_moment eqt12
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   470
            have "vt (moment (Suc t2) s)" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   471
            with eq_m eqt12 show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   472
          qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   473
          then interpret vt_e2: valid_trace_e "moment t2 s" e
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   474
            by (unfold_locales, auto, cases, auto)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   475
          from vt_e2.block_pre [OF False h1]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   476
          have "e = P thread cs2" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   477
          with eq_e1 neq12 show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   478
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   479
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   480
    } ultimately show ?thesis by arith
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   481
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   482
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   483
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   484
text {*
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   485
  This lemma is a simple corrolary of @{text "waiting_unique_pre"}.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   486
*}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   487
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   488
lemma waiting_unique:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   489
  assumes "waiting s th cs1"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   490
  and "waiting s th cs2"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   491
  shows "cs1 = cs2"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   492
using waiting_unique_pre assms
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   493
unfolding wq_def s_waiting_def
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   494
by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   495
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   496
end
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   497
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   498
(* not used *)
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   499
text {*
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   500
  Every thread can only be blocked on one critical resource, 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   501
  symmetrically, every critical resource can only be held by one thread. 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   502
  This fact is much more easier according to our definition. 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   503
*}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   504
lemma held_unique:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   505
  assumes "holding (s::event list) th1 cs"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   506
  and "holding s th2 cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   507
  shows "th1 = th2"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   508
 by (insert assms, unfold s_holding_def, auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   509
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   510
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
   511
lemma last_set_lt: "th \<in> threads s \<Longrightarrow> last_set th s < length s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   512
  apply (induct s, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   513
  by (case_tac a, auto split:if_splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   514
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
   515
lemma last_set_unique: 
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
   516
  "\<lbrakk>last_set th1 s = last_set th2 s; th1 \<in> threads s; th2 \<in> threads s\<rbrakk>
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   517
          \<Longrightarrow> th1 = th2"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   518
  apply (induct s, auto)
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
   519
  by (case_tac a, auto split:if_splits dest:last_set_lt)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   520
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   521
lemma preced_unique : 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   522
  assumes pcd_eq: "preced th1 s = preced th2 s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   523
  and th_in1: "th1 \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   524
  and th_in2: " th2 \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   525
  shows "th1 = th2"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   526
proof -
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
   527
  from pcd_eq have "last_set th1 s = last_set th2 s" by (simp add:preced_def)
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
   528
  from last_set_unique [OF this th_in1 th_in2]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   529
  show ?thesis .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   530
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   531
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   532
lemma preced_linorder: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   533
  assumes neq_12: "th1 \<noteq> th2"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   534
  and th_in1: "th1 \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   535
  and th_in2: " th2 \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   536
  shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   537
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   538
  from preced_unique [OF _ th_in1 th_in2] and neq_12 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   539
  have "preced th1 s \<noteq> preced th2 s" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   540
  thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   541
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   542
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   543
(* An aux lemma used later *)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   544
lemma unique_minus:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   545
  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   546
  and xy: "(x, y) \<in> r"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   547
  and xz: "(x, z) \<in> r^+"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   548
  and neq: "y \<noteq> z"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   549
  shows "(y, z) \<in> r^+"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   550
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   551
 from xz and neq show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   552
 proof(induct)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   553
   case (base ya)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   554
   have "(x, ya) \<in> r" by fact
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   555
   from unique [OF xy this] have "y = ya" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   556
   with base show ?case by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   557
 next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   558
   case (step ya z)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   559
   show ?case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   560
   proof(cases "y = ya")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   561
     case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   562
     from step True show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   563
   next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   564
     case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   565
     from step False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   566
     show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   567
   qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   568
 qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   569
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   570
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   571
lemma unique_base:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   572
  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   573
  and xy: "(x, y) \<in> r"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   574
  and xz: "(x, z) \<in> r^+"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   575
  and neq_yz: "y \<noteq> z"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   576
  shows "(y, z) \<in> r^+"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   577
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   578
  from xz neq_yz show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   579
  proof(induct)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   580
    case (base ya)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   581
    from xy unique base show ?case by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   582
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   583
    case (step ya z)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   584
    show ?case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   585
    proof(cases "y = ya")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   586
      case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   587
      from True step show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   588
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   589
      case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   590
      from False step 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   591
      have "(y, ya) \<in> r\<^sup>+" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   592
      with step show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   593
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   594
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   595
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   596
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   597
lemma unique_chain:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   598
  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   599
  and xy: "(x, y) \<in> r^+"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   600
  and xz: "(x, z) \<in> r^+"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   601
  and neq_yz: "y \<noteq> z"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   602
  shows "(y, z) \<in> r^+ \<or> (z, y) \<in> r^+"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   603
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   604
  from xy xz neq_yz show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   605
  proof(induct)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   606
    case (base y)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   607
    have h1: "(x, y) \<in> r" and h2: "(x, z) \<in> r\<^sup>+" and h3: "y \<noteq> z" using base by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   608
    from unique_base [OF _ h1 h2 h3] and unique show ?case by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   609
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   610
    case (step y za)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   611
    show ?case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   612
    proof(cases "y = z")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   613
      case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   614
      from True step show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   615
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   616
      case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   617
      from False step have "(y, z) \<in> r\<^sup>+ \<or> (z, y) \<in> r\<^sup>+" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   618
      thus ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   619
      proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   620
        assume "(z, y) \<in> r\<^sup>+"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   621
        with step have "(z, za) \<in> r\<^sup>+" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   622
        thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   623
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   624
        assume h: "(y, z) \<in> r\<^sup>+"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   625
        from step have yza: "(y, za) \<in> r" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   626
        from step have "za \<noteq> z" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   627
        from unique_minus [OF _ yza h this] and unique
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   628
        have "(za, z) \<in> r\<^sup>+" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   629
        thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   630
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   631
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   632
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   633
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   634
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   635
text {*
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   636
  The following three lemmas show that @{text "RAG"} does not change
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   637
  by the happening of @{text "Set"}, @{text "Create"} and @{text "Exit"}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   638
  events, respectively.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   639
*}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   640
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   641
lemma RAG_set_unchanged: "(RAG (Set th prio # s)) = RAG s"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   642
apply (unfold s_RAG_def s_waiting_def wq_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   643
by (simp add:Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   644
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   645
lemma RAG_create_unchanged: "(RAG (Create th prio # s)) = RAG s"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   646
apply (unfold s_RAG_def s_waiting_def wq_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   647
by (simp add:Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   648
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   649
lemma RAG_exit_unchanged: "(RAG (Exit th # s)) = RAG s"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   650
apply (unfold s_RAG_def s_waiting_def wq_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   651
by (simp add:Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   652
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   653
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   654
text {* 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   655
  The following lemmas are used in the proof of 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   656
  lemma @{text "step_RAG_v"}, which characterizes how the @{text "RAG"} is changed
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   657
  by @{text "V"}-events. 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   658
  However, since our model is very concise, such  seemingly obvious lemmas need to be derived from scratch,
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   659
  starting from the model definitions.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   660
*}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   661
lemma step_v_hold_inv[elim_format]:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   662
  "\<And>c t. \<lbrakk>vt (V th cs # s); 
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   663
          \<not> holding (wq s) t c; holding (wq (V th cs # s)) t c\<rbrakk> \<Longrightarrow> 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   664
            next_th s th cs t \<and> c = cs"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   665
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   666
  fix c t
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   667
  assume vt: "vt (V th cs # s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   668
    and nhd: "\<not> holding (wq s) t c"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   669
    and hd: "holding (wq (V th cs # s)) t c"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   670
  show "next_th s th cs t \<and> c = cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   671
  proof(cases "c = cs")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   672
    case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   673
    with nhd hd show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   674
      by (unfold cs_holding_def wq_def, auto simp:Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   675
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   676
    case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   677
    with step_back_step [OF vt] 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   678
    have "step s (V th c)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   679
    hence "next_th s th cs t"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   680
    proof(cases)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   681
      assume "holding s th c"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   682
      with nhd hd show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   683
        apply (unfold s_holding_def cs_holding_def wq_def next_th_def,
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   684
               auto simp:Let_def split:list.splits if_splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   685
        proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   686
          assume " hd (SOME q. distinct q \<and> q = []) \<in> set (SOME q. distinct q \<and> q = [])"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   687
          moreover have "\<dots> = set []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   688
          proof(rule someI2)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   689
            show "distinct [] \<and> [] = []" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   690
          next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   691
            fix x assume "distinct x \<and> x = []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   692
            thus "set x = set []" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   693
          qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   694
          ultimately show False by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   695
        next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   696
          assume " hd (SOME q. distinct q \<and> q = []) \<in> set (SOME q. distinct q \<and> q = [])"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   697
          moreover have "\<dots> = set []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   698
          proof(rule someI2)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   699
            show "distinct [] \<and> [] = []" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   700
          next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   701
            fix x assume "distinct x \<and> x = []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   702
            thus "set x = set []" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   703
          qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   704
          ultimately show False by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   705
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   706
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   707
    with True show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   708
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   709
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   710
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   711
text {* 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   712
  The following @{text "step_v_wait_inv"} is also an obvious lemma, which, however, needs to be
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   713
  derived from scratch, which confirms the correctness of the definition of @{text "next_th"}.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   714
*}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   715
lemma step_v_wait_inv[elim_format]:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   716
    "\<And>t c. \<lbrakk>vt (V th cs # s); \<not> waiting (wq (V th cs # s)) t c; waiting (wq s) t c
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   717
           \<rbrakk>
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   718
          \<Longrightarrow> (next_th s th cs t \<and> cs = c)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   719
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   720
  fix t c 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   721
  assume vt: "vt (V th cs # s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   722
    and nw: "\<not> waiting (wq (V th cs # s)) t c"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   723
    and wt: "waiting (wq s) t c"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   724
  from vt interpret vt_v: valid_trace_e s "V th cs" 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   725
    by  (cases, unfold_locales, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   726
  show "next_th s th cs t \<and> cs = c"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   727
  proof(cases "cs = c")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   728
    case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   729
    with nw wt show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   730
      by (auto simp:cs_waiting_def wq_def Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   731
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   732
    case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   733
    from nw[folded True] wt[folded True]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   734
    have "next_th s th cs t"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   735
      apply (unfold next_th_def, auto simp:cs_waiting_def wq_def Let_def split:list.splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   736
    proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   737
      fix a list
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   738
      assume t_in: "t \<in> set list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   739
        and t_ni: "t \<notin> set (SOME q. distinct q \<and> set q = set list)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   740
        and eq_wq: "wq_fun (schs s) cs = a # list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   741
      have " set (SOME q. distinct q \<and> set q = set list) = set list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   742
      proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   743
        from vt_v.wq_distinct[of cs] and eq_wq[folded wq_def]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   744
        show "distinct list \<and> set list = set list" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   745
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   746
        show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   747
          by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   748
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   749
      with t_ni and t_in show "a = th" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   750
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   751
      fix a list
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   752
      assume t_in: "t \<in> set list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   753
        and t_ni: "t \<notin> set (SOME q. distinct q \<and> set q = set list)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   754
        and eq_wq: "wq_fun (schs s) cs = a # list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   755
      have " set (SOME q. distinct q \<and> set q = set list) = set list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   756
      proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   757
        from vt_v.wq_distinct[of cs] and eq_wq[folded wq_def]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   758
        show "distinct list \<and> set list = set list" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   759
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   760
        show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   761
          by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   762
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   763
      with t_ni and t_in show "t = hd (SOME q. distinct q \<and> set q = set list)" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   764
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   765
      fix a list
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   766
      assume eq_wq: "wq_fun (schs s) cs = a # list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   767
      from step_back_step[OF vt]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   768
      show "a = th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   769
      proof(cases)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   770
        assume "holding s th cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   771
        with eq_wq show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   772
          by (unfold s_holding_def wq_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   773
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   774
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   775
    with True show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   776
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   777
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   778
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   779
lemma step_v_not_wait[consumes 3]:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   780
  "\<lbrakk>vt (V th cs # s); next_th s th cs t; waiting (wq (V th cs # s)) t cs\<rbrakk> \<Longrightarrow> False"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   781
  by (unfold next_th_def cs_waiting_def wq_def, auto simp:Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   782
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   783
lemma step_v_release:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   784
  "\<lbrakk>vt (V th cs # s); holding (wq (V th cs # s)) th cs\<rbrakk> \<Longrightarrow> False"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   785
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   786
  assume vt: "vt (V th cs # s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   787
    and hd: "holding (wq (V th cs # s)) th cs"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   788
  from vt interpret vt_v: valid_trace_e s "V th cs"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   789
    by (cases, unfold_locales, simp+)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   790
  from step_back_step [OF vt] and hd
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   791
  show "False"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   792
  proof(cases)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   793
    assume "holding (wq (V th cs # s)) th cs" and "holding s th cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   794
    thus ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   795
      apply (unfold s_holding_def wq_def cs_holding_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   796
      apply (auto simp:Let_def split:list.splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   797
    proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   798
      fix list
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   799
      assume eq_wq[folded wq_def]: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   800
        "wq_fun (schs s) cs = hd (SOME q. distinct q \<and> set q = set list) # list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   801
      and hd_in: "hd (SOME q. distinct q \<and> set q = set list)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   802
            \<in> set (SOME q. distinct q \<and> set q = set list)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   803
      have "set (SOME q. distinct q \<and> set q = set list) = set list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   804
      proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   805
        from vt_v.wq_distinct[of cs] and eq_wq
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   806
        show "distinct list \<and> set list = set list" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   807
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   808
        show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   809
          by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   810
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   811
      moreover have "distinct  (hd (SOME q. distinct q \<and> set q = set list) # list)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   812
      proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   813
        from vt_v.wq_distinct[of cs] and eq_wq
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   814
        show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   815
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   816
      moreover note eq_wq and hd_in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   817
      ultimately show "False" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   818
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   819
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   820
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   821
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   822
lemma step_v_get_hold:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   823
  "\<And>th'. \<lbrakk>vt (V th cs # s); \<not> holding (wq (V th cs # s)) th' cs; next_th s th cs th'\<rbrakk> \<Longrightarrow> False"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   824
  apply (unfold cs_holding_def next_th_def wq_def,
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   825
         auto simp:Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   826
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   827
  fix rest
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   828
  assume vt: "vt (V th cs # s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   829
    and eq_wq[folded wq_def]: " wq_fun (schs s) cs = th # rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   830
    and nrest: "rest \<noteq> []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   831
    and ni: "hd (SOME q. distinct q \<and> set q = set rest)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   832
            \<notin> set (SOME q. distinct q \<and> set q = set rest)"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   833
  from vt interpret vt_v: valid_trace_e s "V th cs"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   834
    by (cases, unfold_locales, simp+)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   835
  have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   836
  proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   837
    from vt_v.wq_distinct[of cs] and eq_wq
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   838
    show "distinct rest \<and> set rest = set rest" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   839
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   840
    fix x assume "distinct x \<and> set x = set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   841
    hence "set x = set rest" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   842
    with nrest
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   843
    show "x \<noteq> []" by (case_tac x, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   844
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   845
  with ni show "False" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   846
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   847
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   848
lemma step_v_release_inv[elim_format]:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   849
"\<And>c t. \<lbrakk>vt (V th cs # s); \<not> holding (wq (V th cs # s)) t c; holding (wq s) t c\<rbrakk> \<Longrightarrow> 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   850
  c = cs \<and> t = th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   851
  apply (unfold cs_holding_def wq_def, auto simp:Let_def split:if_splits list.splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   852
  proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   853
    fix a list
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   854
    assume vt: "vt (V th cs # s)" and eq_wq: "wq_fun (schs s) cs = a # list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   855
    from step_back_step [OF vt] show "a = th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   856
    proof(cases)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   857
      assume "holding s th cs" with eq_wq
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   858
      show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   859
        by (unfold s_holding_def wq_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   860
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   861
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   862
    fix a list
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   863
    assume vt: "vt (V th cs # s)" and eq_wq: "wq_fun (schs s) cs = a # list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   864
    from step_back_step [OF vt] show "a = th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   865
    proof(cases)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   866
      assume "holding s th cs" with eq_wq
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   867
      show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   868
        by (unfold s_holding_def wq_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   869
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   870
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   871
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   872
lemma step_v_waiting_mono:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   873
  "\<And>t c. \<lbrakk>vt (V th cs # s); waiting (wq (V th cs # s)) t c\<rbrakk> \<Longrightarrow> waiting (wq s) t c"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   874
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   875
  fix t c
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   876
  let ?s' = "(V th cs # s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   877
  assume vt: "vt ?s'" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   878
    and wt: "waiting (wq ?s') t c"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   879
  from vt interpret vt_v: valid_trace_e s "V th cs"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   880
    by (cases, unfold_locales, simp+)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   881
  show "waiting (wq s) t c"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   882
  proof(cases "c = cs")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   883
    case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   884
    assume neq_cs: "c \<noteq> cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   885
    hence "waiting (wq ?s') t c = waiting (wq s) t c"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   886
      by (unfold cs_waiting_def wq_def, auto simp:Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   887
    with wt show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   888
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   889
    case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   890
    with wt show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   891
      apply (unfold cs_waiting_def wq_def, auto simp:Let_def split:list.splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   892
    proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   893
      fix a list
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   894
      assume not_in: "t \<notin> set list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   895
        and is_in: "t \<in> set (SOME q. distinct q \<and> set q = set list)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   896
        and eq_wq: "wq_fun (schs s) cs = a # list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   897
      have "set (SOME q. distinct q \<and> set q = set list) = set list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   898
      proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   899
        from vt_v.wq_distinct [of cs]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   900
        and eq_wq[folded wq_def]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   901
        show "distinct list \<and> set list = set list" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   902
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   903
        fix x assume "distinct x \<and> set x = set list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   904
        thus "set x = set list" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   905
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   906
      with not_in is_in show "t = a" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   907
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   908
      fix list
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   909
      assume is_waiting: "waiting (wq (V th cs # s)) t cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   910
      and eq_wq: "wq_fun (schs s) cs = t # list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   911
      hence "t \<in> set list"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   912
        apply (unfold wq_def, auto simp:Let_def cs_waiting_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   913
      proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   914
        assume " t \<in> set (SOME q. distinct q \<and> set q = set list)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   915
        moreover have "\<dots> = set list" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   916
        proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   917
          from vt_v.wq_distinct [of cs]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   918
            and eq_wq[folded wq_def]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   919
          show "distinct list \<and> set list = set list" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   920
        next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   921
          fix x assume "distinct x \<and> set x = set list" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   922
          thus "set x = set list" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   923
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   924
        ultimately show "t \<in> set list" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   925
      qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   926
      with eq_wq and vt_v.wq_distinct [of cs, unfolded wq_def]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   927
      show False by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   928
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   929
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   930
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   931
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   932
text {* (* ddd *) 
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   933
  The following @{text "step_RAG_v"} lemma charaterizes how @{text "RAG"} is changed
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   934
  with the happening of @{text "V"}-events:
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   935
*}
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   936
lemma step_RAG_v:
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   937
assumes vt:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   938
  "vt (V th cs#s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   939
shows "
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   940
  RAG (V th cs # s) =
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   941
  RAG s - {(Cs cs, Th th)} -
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   942
  {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   943
  {(Cs cs, Th th') |th'.  next_th s th cs th'}"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   944
  apply (insert vt, unfold s_RAG_def) 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   945
  apply (auto split:if_splits list.splits simp:Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   946
  apply (auto elim: step_v_waiting_mono step_v_hold_inv 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   947
              step_v_release step_v_wait_inv
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   948
              step_v_get_hold step_v_release_inv)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   949
  apply (erule_tac step_v_not_wait, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   950
  done
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   951
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   952
text {* 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   953
  The following @{text "step_RAG_p"} lemma charaterizes how @{text "RAG"} is changed
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   954
  with the happening of @{text "P"}-events:
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   955
*}
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   956
lemma step_RAG_p:
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   957
  "vt (P th cs#s) \<Longrightarrow>
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   958
  RAG (P th cs # s) =  (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)}
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   959
                                             else RAG s \<union> {(Th th, Cs cs)})"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   960
  apply(simp only: s_RAG_def wq_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   961
  apply (auto split:list.splits prod.splits simp:Let_def wq_def cs_waiting_def cs_holding_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   962
  apply(case_tac "csa = cs", auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   963
  apply(fold wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   964
  apply(drule_tac step_back_step)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   965
  apply(ind_cases " step s (P (hd (wq s cs)) cs)")
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   966
  apply(simp add:s_RAG_def wq_def cs_holding_def)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   967
  apply(auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   968
  done
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   969
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   970
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   971
lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   972
  by (unfold s_RAG_def, auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   973
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   974
context valid_trace
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   975
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   976
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   977
text {*
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   978
  The following lemma shows that @{text "RAG"} is acyclic.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   979
  The overall structure is by induction on the formation of @{text "vt s"}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   980
  and then case analysis on event @{text "e"}, where the non-trivial cases 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   981
  for those for @{text "V"} and @{text "P"} events.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
   982
*}
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   983
lemma acyclic_RAG:
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   984
  shows "acyclic (RAG s)"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   985
using vt
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   986
proof(induct)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   987
  case (vt_cons s e)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   988
  interpret vt_s: valid_trace s using vt_cons(1)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   989
    by (unfold_locales, simp)
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   990
  assume ih: "acyclic (RAG s)"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   991
    and stp: "step s e"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   992
    and vt: "vt s"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   993
  show ?case
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   994
  proof(cases e)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   995
    case (Create th prio)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   996
    with ih
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   997
    show ?thesis by (simp add:RAG_create_unchanged)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   998
  next
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   999
    case (Exit th)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1000
    with ih show ?thesis by (simp add:RAG_exit_unchanged)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1001
  next
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1002
    case (V th cs)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1003
    from V vt stp have vtt: "vt (V th cs#s)" by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1004
    from step_RAG_v [OF this]
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1005
    have eq_de: 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1006
      "RAG (e # s) = 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1007
      RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1008
      {(Cs cs, Th th') |th'. next_th s th cs th'}"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1009
      (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1010
    from ih have ac: "acyclic (?A - ?B - ?C)" by (auto elim:acyclic_subset)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1011
    from step_back_step [OF vtt]
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1012
    have "step s (V th cs)" .
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1013
    thus ?thesis
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1014
    proof(cases)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1015
      assume "holding s th cs"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1016
      hence th_in: "th \<in> set (wq s cs)" and
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1017
        eq_hd: "th = hd (wq s cs)" unfolding s_holding_def wq_def by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1018
      then obtain rest where
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1019
        eq_wq: "wq s cs = th#rest"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1020
        by (cases "wq s cs", auto)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1021
      show ?thesis
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1022
      proof(cases "rest = []")
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1023
        case False
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1024
        let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1025
        from eq_wq False have eq_D: "?D = {(Cs cs, Th ?th')}"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1026
          by (unfold next_th_def, auto)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1027
        let ?E = "(?A - ?B - ?C)"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1028
        have "(Th ?th', Cs cs) \<notin> ?E\<^sup>*"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1029
        proof
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1030
          assume "(Th ?th', Cs cs) \<in> ?E\<^sup>*"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1031
          hence " (Th ?th', Cs cs) \<in> ?E\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1032
          from tranclD [OF this]
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1033
          obtain x where th'_e: "(Th ?th', x) \<in> ?E" by blast
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1034
          hence th_d: "(Th ?th', x) \<in> ?A" by simp
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1035
          from RAG_target_th [OF this]
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1036
          obtain cs' where eq_x: "x = Cs cs'" by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1037
          with th_d have "(Th ?th', Cs cs') \<in> ?A" by simp
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1038
          hence wt_th': "waiting s ?th' cs'"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1039
            unfolding s_RAG_def s_waiting_def cs_waiting_def wq_def by simp
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1040
          hence "cs' = cs"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1041
          proof(rule vt_s.waiting_unique)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1042
            from eq_wq vt_s.wq_distinct[of cs]
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1043
            show "waiting s ?th' cs" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1044
              apply (unfold s_waiting_def wq_def, auto)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1045
            proof -
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1046
              assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1047
                and eq_wq: "wq_fun (schs s) cs = th # rest"
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1048
              have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1049
              proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1050
                from vt_s.wq_distinct[of cs] and eq_wq
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1051
                show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1052
              next
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1053
                fix x assume "distinct x \<and> set x = set rest"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1054
                with False show "x \<noteq> []" by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1055
              qed
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1056
              hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1057
                set (SOME q. distinct q \<and> set q = set rest)" by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1058
              moreover have "\<dots> = set rest" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1059
              proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1060
                from vt_s.wq_distinct[of cs] and eq_wq
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1061
                show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1062
              next
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1063
                show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1064
              qed
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1065
              moreover note hd_in
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1066
              ultimately show "hd (SOME q. distinct q \<and> set q = set rest) = th" by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1067
            next
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1068
              assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1069
                and eq_wq: "wq s cs = hd (SOME q. distinct q \<and> set q = set rest) # rest"
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1070
              have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1071
              proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1072
                from vt_s.wq_distinct[of cs] and eq_wq
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1073
                show "distinct rest \<and> set rest = set rest" by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1074
              next
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1075
                fix x assume "distinct x \<and> set x = set rest"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1076
                with False show "x \<noteq> []" by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1077
              qed
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1078
              hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1079
                set (SOME q. distinct q \<and> set q = set rest)" by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1080
              moreover have "\<dots> = set rest" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1081
              proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1082
                from vt_s.wq_distinct[of cs] and eq_wq
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1083
                show "distinct rest \<and> set rest = set rest" by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1084
              next
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1085
                show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1086
              qed
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1087
              moreover note hd_in
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1088
              ultimately show False by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1089
            qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1090
          qed
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1091
          with th'_e eq_x have "(Th ?th', Cs cs) \<in> ?E" by simp
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1092
          with False
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1093
          show "False" by (auto simp: next_th_def eq_wq)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1094
        qed
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1095
        with acyclic_insert[symmetric] and ac
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1096
          and eq_de eq_D show ?thesis by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1097
      next
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1098
        case True
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1099
        with eq_wq
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1100
        have eq_D: "?D = {}"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1101
          by (unfold next_th_def, auto)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1102
        with eq_de ac
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1103
        show ?thesis by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1104
      qed 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1105
    qed
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1106
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1107
    case (P th cs)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1108
    from P vt stp have vtt: "vt (P th cs#s)" by auto
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1109
    from step_RAG_p [OF this] P
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1110
    have "RAG (e # s) = 
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1111
      (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1112
      RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1113
      by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1114
    moreover have "acyclic ?R"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1115
    proof(cases "wq s cs = []")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1116
      case True
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1117
      hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1118
      have "(Th th, Cs cs) \<notin> (RAG s)\<^sup>*"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1119
      proof
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1120
        assume "(Th th, Cs cs) \<in> (RAG s)\<^sup>*"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1121
        hence "(Th th, Cs cs) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1122
        from tranclD2 [OF this]
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1123
        obtain x where "(x, Cs cs) \<in> RAG s" by auto
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1124
        with True show False by (auto simp:s_RAG_def cs_waiting_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1125
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1126
      with acyclic_insert ih eq_r show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1127
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1128
      case False
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1129
      hence eq_r: "?R =  RAG s \<union> {(Th th, Cs cs)}" by simp
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1130
      have "(Cs cs, Th th) \<notin> (RAG s)\<^sup>*"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1131
      proof
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1132
        assume "(Cs cs, Th th) \<in> (RAG s)\<^sup>*"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1133
        hence "(Cs cs, Th th) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1134
        moreover from step_back_step [OF vtt] have "step s (P th cs)" .
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1135
        ultimately show False
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1136
        proof -
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1137
          show " \<lbrakk>(Cs cs, Th th) \<in> (RAG s)\<^sup>+; step s (P th cs)\<rbrakk> \<Longrightarrow> False"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1138
            by (ind_cases "step s (P th cs)", simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1139
        qed
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1140
      qed
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1141
      with acyclic_insert ih eq_r show ?thesis by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1142
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1143
      ultimately show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1144
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1145
      case (Set thread prio)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1146
      with ih
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1147
      thm RAG_set_unchanged
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1148
      show ?thesis by (simp add:RAG_set_unchanged)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1149
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1150
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1151
    case vt_nil
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1152
    show "acyclic (RAG ([]::state))"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1153
      by (auto simp: s_RAG_def cs_waiting_def 
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1154
        cs_holding_def wq_def acyclic_def)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1155
qed
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
  1156
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1157
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1158
lemma finite_RAG:
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1159
  shows "finite (RAG s)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1160
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1161
  from vt show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1162
  proof(induct)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1163
    case (vt_cons s e)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1164
    interpret vt_s: valid_trace s using vt_cons(1)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1165
      by (unfold_locales, simp)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1166
    assume ih: "finite (RAG s)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1167
      and stp: "step s e"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1168
      and vt: "vt s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1169
    show ?case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1170
    proof(cases e)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1171
      case (Create th prio)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1172
      with ih
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1173
      show ?thesis by (simp add:RAG_create_unchanged)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1174
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1175
      case (Exit th)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1176
      with ih show ?thesis by (simp add:RAG_exit_unchanged)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1177
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1178
      case (V th cs)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1179
      from V vt stp have vtt: "vt (V th cs#s)" by auto
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1180
      from step_RAG_v [OF this]
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1181
      have eq_de: "RAG (e # s) = 
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1182
                   RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1183
                      {(Cs cs, Th th') |th'. next_th s th cs th'}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1184
"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1185
        (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1186
      moreover from ih have ac: "finite (?A - ?B - ?C)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1187
      moreover have "finite ?D"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1188
      proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1189
        have "?D = {} \<or> (\<exists> a. ?D = {a})" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1190
          by (unfold next_th_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1191
        thus ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1192
        proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1193
          assume h: "?D = {}"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1194
          show ?thesis by (unfold h, simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1195
        next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1196
          assume "\<exists> a. ?D = {a}"
3
51019d035a79 made everything working
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 0
diff changeset
  1197
          thus ?thesis
51019d035a79 made everything working
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 0
diff changeset
  1198
            by (metis finite.simps)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1199
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1200
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1201
      ultimately show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1202
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1203
      case (P th cs)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1204
      from P vt stp have vtt: "vt (P th cs#s)" by auto
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1205
      from step_RAG_p [OF this] P
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1206
      have "RAG (e # s) = 
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1207
              (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1208
                                    RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1209
        by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1210
      moreover have "finite ?R"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1211
      proof(cases "wq s cs = []")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1212
        case True
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1213
        hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1214
        with True and ih show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1215
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1216
        case False
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1217
        hence "?R = RAG s \<union> {(Th th, Cs cs)}" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1218
        with False and ih show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1219
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1220
      ultimately show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1221
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1222
      case (Set thread prio)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1223
      with ih
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1224
      show ?thesis by (simp add:RAG_set_unchanged)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1225
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1226
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1227
    case vt_nil
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1228
    show "finite (RAG ([]::state))"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1229
      by (auto simp: s_RAG_def cs_waiting_def 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1230
                   cs_holding_def wq_def acyclic_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1231
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1232
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1233
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1234
text {* Several useful lemmas *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1235
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1236
lemma wf_dep_converse: 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1237
  shows "wf ((RAG s)^-1)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1238
proof(rule finite_acyclic_wf_converse)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1239
  from finite_RAG 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1240
  show "finite (RAG s)" .
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1241
next
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1242
  from acyclic_RAG
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1243
  show "acyclic (RAG s)" .
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1244
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1245
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1246
end
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1247
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1248
lemma hd_np_in: "x \<in> set l \<Longrightarrow> hd l \<in> set l"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1249
  by (induct l, auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1250
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1251
lemma th_chasing: "(Th th, Cs cs) \<in> RAG (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> RAG s"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1252
  by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1253
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1254
context valid_trace
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1255
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1256
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1257
lemma wq_threads: 
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1258
  assumes h: "th \<in> set (wq s cs)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1259
  shows "th \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1260
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1261
 from vt and h show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1262
  proof(induct arbitrary: th cs)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1263
    case (vt_cons s e)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1264
    interpret vt_s: valid_trace s
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1265
      using vt_cons(1) by (unfold_locales, auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1266
    assume ih: "\<And>th cs. th \<in> set (wq s cs) \<Longrightarrow> th \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1267
      and stp: "step s e"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1268
      and vt: "vt s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1269
      and h: "th \<in> set (wq (e # s) cs)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1270
    show ?case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1271
    proof(cases e)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1272
      case (Create th' prio)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1273
      with ih h show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1274
        by (auto simp:wq_def Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1275
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1276
      case (Exit th')
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1277
      with stp ih h show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1278
        apply (auto simp:wq_def Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1279
        apply (ind_cases "step s (Exit th')")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1280
        apply (auto simp:runing_def readys_def s_holding_def s_waiting_def holdents_def
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1281
               s_RAG_def s_holding_def cs_holding_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1282
        done
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1283
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1284
      case (V th' cs')
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1285
      show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1286
      proof(cases "cs' = cs")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1287
        case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1288
        with h
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1289
        show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1290
          apply(unfold wq_def V, auto simp:Let_def V split:prod.splits, fold wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1291
          by (drule_tac ih, simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1292
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1293
        case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1294
        from h
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1295
        show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1296
        proof(unfold V wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1297
          assume th_in: "th \<in> set (wq_fun (schs (V th' cs' # s)) cs)" (is "th \<in> set ?l")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1298
          show "th \<in> threads (V th' cs' # s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1299
          proof(cases "cs = cs'")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1300
            case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1301
            hence "?l = wq_fun (schs s) cs" by (simp add:Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1302
            with th_in have " th \<in> set (wq s cs)" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1303
              by (fold wq_def, simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1304
            from ih [OF this] show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1305
          next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1306
            case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1307
            show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1308
            proof(cases "wq_fun (schs s) cs'")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1309
              case Nil
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1310
              with h V show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1311
                apply (auto simp:wq_def Let_def split:if_splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1312
                by (fold wq_def, drule_tac ih, simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1313
            next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1314
              case (Cons a rest)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1315
              assume eq_wq: "wq_fun (schs s) cs' = a # rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1316
              with h V show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1317
                apply (auto simp:Let_def wq_def split:if_splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1318
              proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1319
                assume th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1320
                have "set (SOME q. distinct q \<and> set q = set rest) = set rest" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1321
                proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1322
                  from vt_s.wq_distinct[of cs'] and eq_wq[folded wq_def]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1323
                  show "distinct rest \<and> set rest = set rest" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1324
                next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1325
                  show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1326
                    by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1327
                qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1328
                with eq_wq th_in have "th \<in> set (wq_fun (schs s) cs')" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1329
                from ih[OF this[folded wq_def]] show "th \<in> threads s" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1330
              next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1331
                assume th_in: "th \<in> set (wq_fun (schs s) cs)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1332
                from ih[OF this[folded wq_def]]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1333
                show "th \<in> threads s" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1334
              qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1335
            qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1336
          qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1337
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1338
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1339
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1340
      case (P th' cs')
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1341
      from h stp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1342
      show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1343
        apply (unfold P wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1344
        apply (auto simp:Let_def split:if_splits, fold wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1345
        apply (auto intro:ih)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1346
        apply(ind_cases "step s (P th' cs')")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1347
        by (unfold runing_def readys_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1348
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1349
      case (Set thread prio)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1350
      with ih h show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1351
        by (auto simp:wq_def Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1352
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1353
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1354
    case vt_nil
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1355
    thus ?case by (auto simp:wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1356
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1357
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1358
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1359
lemma range_in: "\<lbrakk>(Th th) \<in> Range (RAG (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1360
  apply(unfold s_RAG_def cs_waiting_def cs_holding_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1361
  by (auto intro:wq_threads)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1362
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1363
lemma readys_v_eq:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1364
  assumes neq_th: "th \<noteq> thread"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1365
  and eq_wq: "wq s cs = thread#rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1366
  and not_in: "th \<notin>  set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1367
  shows "(th \<in> readys (V thread cs#s)) = (th \<in> readys s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1368
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1369
  from assms show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1370
    apply (auto simp:readys_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1371
    apply(simp add:s_waiting_def[folded wq_def])
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1372
    apply (erule_tac x = csa in allE)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1373
    apply (simp add:s_waiting_def wq_def Let_def split:if_splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1374
    apply (case_tac "csa = cs", simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1375
    apply (erule_tac x = cs in allE)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1376
    apply(auto simp add: s_waiting_def[folded wq_def] Let_def split: list.splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1377
    apply(auto simp add: wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1378
    apply (auto simp:s_waiting_def wq_def Let_def split:list.splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1379
    proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1380
       assume th_nin: "th \<notin> set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1381
        and th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1382
        and eq_wq: "wq_fun (schs s) cs = thread # rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1383
      have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1384
      proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1385
        from wq_distinct[of cs, unfolded wq_def] and eq_wq[unfolded wq_def]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1386
        show "distinct rest \<and> set rest = set rest" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1387
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1388
        show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1389
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1390
      with th_nin th_in show False by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1391
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1392
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1393
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  1394
text {* \noindent
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  1395
  The following lemmas shows that: starting from any node in @{text "RAG"}, 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  1396
  by chasing out-going edges, it is always possible to reach a node representing a ready
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  1397
  thread. In this lemma, it is the @{text "th'"}.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  1398
*}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  1399
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1400
lemma chain_building:
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1401
  shows "node \<in> Domain (RAG s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (RAG s)^+)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1402
proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1403
  from wf_dep_converse
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1404
  have h: "wf ((RAG s)\<inverse>)" .
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1405
  show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1406
  proof(induct rule:wf_induct [OF h])
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1407
    fix x
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1408
    assume ih [rule_format]: 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1409
      "\<forall>y. (y, x) \<in> (RAG s)\<inverse> \<longrightarrow> 
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1410
           y \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (RAG s)\<^sup>+)"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1411
    show "x \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1412
    proof
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1413
      assume x_d: "x \<in> Domain (RAG s)"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1414
      show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1415
      proof(cases x)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1416
        case (Th th)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1417
        from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> RAG s" by (auto simp:s_RAG_def)
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1418
        with Th have x_in_r: "(Cs cs, x) \<in> (RAG s)^-1" by simp
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1419
        from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> RAG s" by blast
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1420
        hence "Cs cs \<in> Domain (RAG s)" by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1421
        from ih [OF x_in_r this] obtain th'
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1422
          where th'_ready: " th' \<in> readys s" and cs_in: "(Cs cs, Th th') \<in> (RAG s)\<^sup>+" by auto
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1423
        have "(x, Th th') \<in> (RAG s)\<^sup>+" using Th x_in cs_in by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1424
        with th'_ready show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1425
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1426
        case (Cs cs)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1427
        from x_d Cs obtain th' where th'_d: "(Th th', x) \<in> (RAG s)^-1" by (auto simp:s_RAG_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1428
        show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1429
        proof(cases "th' \<in> readys s")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1430
          case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1431
          from True and th'_d show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1432
        next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1433
          case False
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1434
          from th'_d and range_in  have "th' \<in> threads s" by auto
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1435
          with False have "Th th' \<in> Domain (RAG s)" 
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1436
            by (auto simp:readys_def wq_def s_waiting_def s_RAG_def cs_waiting_def Domain_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1437
          from ih [OF th'_d this]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1438
          obtain th'' where 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1439
            th''_r: "th'' \<in> readys s" and 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1440
            th''_in: "(Th th', Th th'') \<in> (RAG s)\<^sup>+" by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1441
          from th'_d and th''_in 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1442
          have "(x, Th th'') \<in> (RAG s)\<^sup>+" by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1443
          with th''_r show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1444
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1445
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1446
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1447
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1448
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1449
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  1450
text {* \noindent
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  1451
  The following is just an instance of @{text "chain_building"}.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  1452
*}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1453
lemma th_chain_to_ready:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1454
  assumes th_in: "th \<in> threads s"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1455
  shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (RAG s)^+)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1456
proof(cases "th \<in> readys s")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1457
  case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1458
  thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1459
next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1460
  case False
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1461
  from False and th_in have "Th th \<in> Domain (RAG s)" 
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1462
    by (auto simp:readys_def s_waiting_def s_RAG_def wq_def cs_waiting_def Domain_def)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1463
  from chain_building [rule_format, OF this]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1464
  show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1465
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1466
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1467
end
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1468
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1469
lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1470
  by  (unfold s_waiting_def cs_waiting_def wq_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1471
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1472
lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1473
  by (unfold s_holding_def wq_def cs_holding_def, simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1474
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1475
lemma holding_unique: "\<lbrakk>holding (s::state) th1 cs; holding s th2 cs\<rbrakk> \<Longrightarrow> th1 = th2"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1476
  by (unfold s_holding_def cs_holding_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1477
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1478
context valid_trace
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1479
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1480
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1481
lemma unique_RAG: "\<lbrakk>(n, n1) \<in> RAG s; (n, n2) \<in> RAG s\<rbrakk> \<Longrightarrow> n1 = n2"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1482
  apply(unfold s_RAG_def, auto, fold waiting_eq holding_eq)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1483
  by(auto elim:waiting_unique holding_unique)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1484
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1485
end
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1486
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1487
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1488
lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1489
by (induct rule:trancl_induct, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1490
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1491
context valid_trace
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1492
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1493
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1494
lemma dchain_unique:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1495
  assumes th1_d: "(n, Th th1) \<in> (RAG s)^+"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1496
  and th1_r: "th1 \<in> readys s"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1497
  and th2_d: "(n, Th th2) \<in> (RAG s)^+"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1498
  and th2_r: "th2 \<in> readys s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1499
  shows "th1 = th2"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1500
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1501
  { assume neq: "th1 \<noteq> th2"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1502
    hence "Th th1 \<noteq> Th th2" by simp
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1503
    from unique_chain [OF _ th1_d th2_d this] and unique_RAG 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1504
    have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (RAG s)\<^sup>+" by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1505
    hence "False"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1506
    proof
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1507
      assume "(Th th1, Th th2) \<in> (RAG s)\<^sup>+"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1508
      from trancl_split [OF this]
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1509
      obtain n where dd: "(Th th1, n) \<in> RAG s" by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1510
      then obtain cs where eq_n: "n = Cs cs"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1511
        by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1512
      from dd eq_n have "th1 \<notin> readys s"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1513
        by (auto simp:readys_def s_RAG_def wq_def s_waiting_def cs_waiting_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1514
      with th1_r show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1515
    next
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1516
      assume "(Th th2, Th th1) \<in> (RAG s)\<^sup>+"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1517
      from trancl_split [OF this]
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1518
      obtain n where dd: "(Th th2, n) \<in> RAG s" by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1519
      then obtain cs where eq_n: "n = Cs cs"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1520
        by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1521
      from dd eq_n have "th2 \<notin> readys s"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1522
        by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1523
      with th2_r show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1524
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1525
  } thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1526
qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1527
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1528
end
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1529
             
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1530
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1531
lemma step_holdents_p_add:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1532
  assumes vt: "vt (P th cs#s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1533
  and "wq s cs = []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1534
  shows "holdents (P th cs#s) th = holdents s th \<union> {cs}"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1535
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1536
  from assms show ?thesis
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1537
  unfolding  holdents_test step_RAG_p[OF vt] by (auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1538
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1539
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1540
lemma step_holdents_p_eq:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1541
  assumes vt: "vt (P th cs#s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1542
  and "wq s cs \<noteq> []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1543
  shows "holdents (P th cs#s) th = holdents s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1544
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1545
  from assms show ?thesis
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1546
  unfolding  holdents_test step_RAG_p[OF vt] by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1547
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1548
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1549
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1550
lemma (in valid_trace) finite_holding :
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1551
  shows "finite (holdents s th)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1552
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1553
  let ?F = "\<lambda> (x, y). the_cs x"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1554
  from finite_RAG 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1555
  have "finite (RAG s)" .
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1556
  hence "finite (?F `(RAG s))" by simp
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1557
  moreover have "{cs . (Cs cs, Th th) \<in> RAG s} \<subseteq> \<dots>" 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1558
  proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1559
    { have h: "\<And> a A f. a \<in> A \<Longrightarrow> f a \<in> f ` A" by auto
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1560
      fix x assume "(Cs x, Th th) \<in> RAG s"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1561
      hence "?F (Cs x, Th th) \<in> ?F `(RAG s)" by (rule h)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1562
      moreover have "?F (Cs x, Th th) = x" by simp
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1563
      ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` RAG s" by simp 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1564
    } thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1565
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1566
  ultimately show ?thesis by (unfold holdents_test, auto intro:finite_subset)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1567
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1568
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1569
lemma cntCS_v_dec: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1570
  assumes vtv: "vt (V thread cs#s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1571
  shows "(cntCS (V thread cs#s) thread + 1) = cntCS s thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1572
proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1573
  from vtv interpret vt_s: valid_trace s
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1574
    by (cases, unfold_locales, simp)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1575
  from vtv interpret vt_v: valid_trace "V thread cs#s"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1576
     by (unfold_locales, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1577
  from step_back_step[OF vtv]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1578
  have cs_in: "cs \<in> holdents s thread" 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1579
    apply (cases, unfold holdents_test s_RAG_def, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1580
    by (unfold cs_holding_def s_holding_def wq_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1581
  moreover have cs_not_in: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1582
    "(holdents (V thread cs#s) thread) = holdents s thread - {cs}"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1583
    apply (insert vt_s.wq_distinct[of cs])
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1584
    apply (unfold holdents_test, unfold step_RAG_v[OF vtv],
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1585
            auto simp:next_th_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1586
  proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1587
    fix rest
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1588
    assume dst: "distinct (rest::thread list)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1589
      and ne: "rest \<noteq> []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1590
    and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1591
    moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1592
    proof(rule someI2)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1593
      from dst show "distinct rest \<and> set rest = set rest" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1594
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1595
      show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1596
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1597
    ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1598
                     set (SOME q. distinct q \<and> set q = set rest)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1599
    moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1600
    proof(rule someI2)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1601
      from dst show "distinct rest \<and> set rest = set rest" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1602
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1603
      fix x assume " distinct x \<and> set x = set rest" with ne
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1604
      show "x \<noteq> []" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1605
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1606
    ultimately 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1607
    show "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1608
      by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1609
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1610
    fix rest
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1611
    assume dst: "distinct (rest::thread list)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1612
      and ne: "rest \<noteq> []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1613
    and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1614
    moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1615
    proof(rule someI2)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1616
      from dst show "distinct rest \<and> set rest = set rest" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1617
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1618
      show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1619
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1620
    ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1621
                     set (SOME q. distinct q \<and> set q = set rest)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1622
    moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1623
    proof(rule someI2)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1624
      from dst show "distinct rest \<and> set rest = set rest" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1625
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1626
      fix x assume " distinct x \<and> set x = set rest" with ne
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1627
      show "x \<noteq> []" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1628
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1629
    ultimately show "False" by auto 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1630
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1631
  ultimately 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1632
  have "holdents s thread = insert cs (holdents (V thread cs#s) thread)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1633
    by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1634
  moreover have "card \<dots> = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1635
                    Suc (card ((holdents (V thread cs#s) thread) - {cs}))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1636
  proof(rule card_insert)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1637
    from vt_v.finite_holding
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1638
    show " finite (holdents (V thread cs # s) thread)" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1639
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1640
  moreover from cs_not_in 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1641
  have "cs \<notin> (holdents (V thread cs#s) thread)" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1642
  ultimately show ?thesis by (simp add:cntCS_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1643
qed 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1644
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1645
lemma count_rec1 [simp]: 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1646
  assumes "Q e"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1647
  shows "count Q (e#es) = Suc (count Q es)"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1648
  using assms
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1649
  by (unfold count_def, auto)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1650
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1651
lemma count_rec2 [simp]: 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1652
  assumes "\<not>Q e"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1653
  shows "count Q (e#es) = (count Q es)"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1654
  using assms
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1655
  by (unfold count_def, auto)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1656
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1657
lemma count_rec3 [simp]: 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1658
  shows "count Q [] =  0"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1659
  by (unfold count_def, auto)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1660
  
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1661
lemma cntP_diff_inv:
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1662
  assumes "cntP (e#s) th \<noteq> cntP s th"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1663
  shows "isP e \<and> actor e = th"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1664
proof(cases e)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1665
  case (P th' pty)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1666
  show ?thesis
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1667
  by (cases "(\<lambda>e. \<exists>cs. e = P th cs) (P th' pty)", 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1668
        insert assms P, auto simp:cntP_def)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1669
qed (insert assms, auto simp:cntP_def)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1670
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1671
lemma isP_E:
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1672
  assumes "isP e"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1673
  obtains cs where "e = P (actor e) cs"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1674
  using assms by (cases e, auto)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1675
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1676
lemma isV_E:
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1677
  assumes "isV e"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1678
  obtains cs where "e = V (actor e) cs"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1679
  using assms by (cases e, auto) (* ccc *)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1680
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1681
lemma cntV_diff_inv:
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1682
  assumes "cntV (e#s) th \<noteq> cntV s th"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1683
  shows "isV e \<and> actor e = th"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1684
proof(cases e)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1685
  case (V th' pty)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1686
  show ?thesis
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1687
  by (cases "(\<lambda>e. \<exists>cs. e = V th cs) (V th' pty)", 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1688
        insert assms V, auto simp:cntV_def)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1689
qed (insert assms, auto simp:cntV_def)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
  1690
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1691
context valid_trace
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1692
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1693
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
  1694
text {* (* ddd *) \noindent
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  1695
  The relationship between @{text "cntP"}, @{text "cntV"} and @{text "cntCS"} 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  1696
  of one particular thread. 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  1697
*} 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  1698
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1699
lemma cnp_cnv_cncs:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1700
  shows "cntP s th = cntV s th + (if (th \<in> readys s \<or> th \<notin> threads s) 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1701
                                       then cntCS s th else cntCS s th + 1)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1702
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1703
  from vt show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1704
  proof(induct arbitrary:th)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1705
    case (vt_cons s e)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1706
    interpret vt_s: valid_trace s using vt_cons(1) by (unfold_locales, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1707
    assume vt: "vt s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1708
    and ih: "\<And>th. cntP s th  = cntV s th +
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1709
               (if (th \<in> readys s \<or> th \<notin> threads s) then cntCS s th else cntCS s th + 1)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1710
    and stp: "step s e"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1711
    from stp show ?case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1712
    proof(cases)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1713
      case (thread_create thread prio)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1714
      assume eq_e: "e = Create thread prio"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1715
        and not_in: "thread \<notin> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1716
      show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1717
      proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1718
        { fix cs 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1719
          assume "thread \<in> set (wq s cs)"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1720
          from vt_s.wq_threads [OF this] have "thread \<in> threads s" .
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1721
          with not_in have "False" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1722
        } with eq_e have eq_readys: "readys (e#s) = readys s \<union> {thread}"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1723
          by (auto simp:readys_def threads.simps s_waiting_def 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1724
            wq_def cs_waiting_def Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1725
        from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1726
        from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1727
        have eq_cncs: "cntCS (e#s) th = cntCS s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1728
          unfolding cntCS_def holdents_test
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1729
          by (simp add:RAG_create_unchanged eq_e)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1730
        { assume "th \<noteq> thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1731
          with eq_readys eq_e
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1732
          have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1733
                      (th \<in> readys (s) \<or> th \<notin> threads (s))" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1734
            by (simp add:threads.simps)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1735
          with eq_cnp eq_cnv eq_cncs ih not_in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1736
          have ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1737
        } moreover {
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1738
          assume eq_th: "th = thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1739
          with not_in ih have " cntP s th  = cntV s th + cntCS s th" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1740
          moreover from eq_th and eq_readys have "th \<in> readys (e#s)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1741
          moreover note eq_cnp eq_cnv eq_cncs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1742
          ultimately have ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1743
        } ultimately show ?thesis by blast
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1744
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1745
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1746
      case (thread_exit thread)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1747
      assume eq_e: "e = Exit thread" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1748
      and is_runing: "thread \<in> runing s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1749
      and no_hold: "holdents s thread = {}"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1750
      from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1751
      from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1752
      have eq_cncs: "cntCS (e#s) th = cntCS s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1753
        unfolding cntCS_def holdents_test
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1754
        by (simp add:RAG_exit_unchanged eq_e)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1755
      { assume "th \<noteq> thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1756
        with eq_e
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1757
        have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1758
          (th \<in> readys (s) \<or> th \<notin> threads (s))" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1759
          apply (simp add:threads.simps readys_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1760
          apply (subst s_waiting_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1761
          apply (simp add:Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1762
          apply (subst s_waiting_def, simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1763
          done
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1764
        with eq_cnp eq_cnv eq_cncs ih
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1765
        have ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1766
      } moreover {
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1767
        assume eq_th: "th = thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1768
        with ih is_runing have " cntP s th = cntV s th + cntCS s th" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1769
          by (simp add:runing_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1770
        moreover from eq_th eq_e have "th \<notin> threads (e#s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1771
          by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1772
        moreover note eq_cnp eq_cnv eq_cncs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1773
        ultimately have ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1774
      } ultimately show ?thesis by blast
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1775
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1776
      case (thread_P thread cs)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1777
      assume eq_e: "e = P thread cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1778
        and is_runing: "thread \<in> runing s"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1779
        and no_dep: "(Cs cs, Th thread) \<notin> (RAG s)\<^sup>+"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1780
      from thread_P vt stp ih  have vtp: "vt (P thread cs#s)" by auto
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1781
      then interpret vt_p: valid_trace "(P thread cs#s)"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1782
        by (unfold_locales, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1783
      show ?thesis 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1784
      proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1785
        { have hh: "\<And> A B C. (B = C) \<Longrightarrow> (A \<and> B) = (A \<and> C)" by blast
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1786
          assume neq_th: "th \<noteq> thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1787
          with eq_e
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1788
          have eq_readys: "(th \<in> readys (e#s)) = (th \<in> readys (s))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1789
            apply (simp add:readys_def s_waiting_def wq_def Let_def)
44
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  1790
            apply (rule_tac hh)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  1791
             apply (intro iffI allI, clarify)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1792
            apply (erule_tac x = csa in allE, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1793
            apply (subgoal_tac "wq_fun (schs s) cs \<noteq> []", auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1794
            apply (erule_tac x = cs in allE, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1795
            by (case_tac "(wq_fun (schs s) cs)", auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1796
          moreover from neq_th eq_e have "cntCS (e # s) th = cntCS s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1797
            apply (simp add:cntCS_def holdents_test)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1798
            by (unfold  step_RAG_p [OF vtp], auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1799
          moreover from eq_e neq_th have "cntP (e # s) th = cntP s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1800
            by (simp add:cntP_def count_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1801
          moreover from eq_e neq_th have "cntV (e#s) th = cntV s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1802
            by (simp add:cntV_def count_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1803
          moreover from eq_e neq_th have "threads (e#s) = threads s" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1804
          moreover note ih [of th] 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1805
          ultimately have ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1806
        } moreover {
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1807
          assume eq_th: "th = thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1808
          have ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1809
          proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1810
            from eq_e eq_th have eq_cnp: "cntP (e # s) th  = 1 + (cntP s th)" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1811
              by (simp add:cntP_def count_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1812
            from eq_e eq_th have eq_cnv: "cntV (e#s) th = cntV s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1813
              by (simp add:cntV_def count_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1814
            show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1815
            proof (cases "wq s cs = []")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1816
              case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1817
              with is_runing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1818
              have "th \<in> readys (e#s)"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1819
                apply (unfold eq_e wq_def, unfold readys_def s_RAG_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1820
                apply (simp add: wq_def[symmetric] runing_def eq_th s_waiting_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1821
                by (auto simp:readys_def wq_def Let_def s_waiting_def wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1822
              moreover have "cntCS (e # s) th = 1 + cntCS s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1823
              proof -
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1824
                have "card {csa. csa = cs \<or> (Cs csa, Th thread) \<in> RAG s} =
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1825
                  Suc (card {cs. (Cs cs, Th thread) \<in> RAG s})" (is "card ?L = Suc (card ?R)")
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1826
                proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1827
                  have "?L = insert cs ?R" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1828
                  moreover have "card \<dots> = Suc (card (?R - {cs}))" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1829
                  proof(rule card_insert)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1830
                    from vt_s.finite_holding [of thread]
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1831
                    show " finite {cs. (Cs cs, Th thread) \<in> RAG s}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1832
                      by (unfold holdents_test, simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1833
                  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1834
                  moreover have "?R - {cs} = ?R"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1835
                  proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1836
                    have "cs \<notin> ?R"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1837
                    proof
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1838
                      assume "cs \<in> {cs. (Cs cs, Th thread) \<in> RAG s}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1839
                      with no_dep show False by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1840
                    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1841
                    thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1842
                  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1843
                  ultimately show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1844
                qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1845
                thus ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1846
                  apply (unfold eq_e eq_th cntCS_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1847
                  apply (simp add: holdents_test)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1848
                  by (unfold step_RAG_p [OF vtp], auto simp:True)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1849
              qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1850
              moreover from is_runing have "th \<in> readys s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1851
                by (simp add:runing_def eq_th)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1852
              moreover note eq_cnp eq_cnv ih [of th]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1853
              ultimately show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1854
            next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1855
              case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1856
              have eq_wq: "wq (e#s) cs = wq s cs @ [th]"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1857
                    by (unfold eq_th eq_e wq_def, auto simp:Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1858
              have "th \<notin> readys (e#s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1859
              proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1860
                assume "th \<in> readys (e#s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1861
                hence "\<forall>cs. \<not> waiting (e # s) th cs" by (simp add:readys_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1862
                from this[rule_format, of cs] have " \<not> waiting (e # s) th cs" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1863
                hence "th \<in> set (wq (e#s) cs) \<Longrightarrow> th = hd (wq (e#s) cs)" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1864
                  by (simp add:s_waiting_def wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1865
                moreover from eq_wq have "th \<in> set (wq (e#s) cs)" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1866
                ultimately have "th = hd (wq (e#s) cs)" by blast
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1867
                with eq_wq have "th = hd (wq s cs @ [th])" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1868
                hence "th = hd (wq s cs)" using False by auto
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1869
                with False eq_wq vt_p.wq_distinct [of cs]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1870
                show False by (fold eq_e, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1871
              qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1872
              moreover from is_runing have "th \<in> threads (e#s)" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1873
                by (unfold eq_e, auto simp:runing_def readys_def eq_th)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1874
              moreover have "cntCS (e # s) th = cntCS s th"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1875
                apply (unfold cntCS_def holdents_test eq_e step_RAG_p[OF vtp])
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1876
                by (auto simp:False)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1877
              moreover note eq_cnp eq_cnv ih[of th]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1878
              moreover from is_runing have "th \<in> readys s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1879
                by (simp add:runing_def eq_th)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1880
              ultimately show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1881
            qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1882
          qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1883
        } ultimately show ?thesis by blast
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1884
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1885
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1886
      case (thread_V thread cs)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1887
      from assms vt stp ih thread_V have vtv: "vt (V thread cs # s)" by auto
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1888
      then interpret vt_v: valid_trace "(V thread cs # s)" by (unfold_locales, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1889
      assume eq_e: "e = V thread cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1890
        and is_runing: "thread \<in> runing s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1891
        and hold: "holding s thread cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1892
      from hold obtain rest 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1893
        where eq_wq: "wq s cs = thread # rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1894
        by (case_tac "wq s cs", auto simp: wq_def s_holding_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1895
      have eq_threads: "threads (e#s) = threads s" by (simp add: eq_e)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1896
      have eq_set: "set (SOME q. distinct q \<and> set q = set rest) = set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1897
      proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1898
        from vt_v.wq_distinct[of cs] and eq_wq
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1899
        show "distinct rest \<and> set rest = set rest"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1900
          by (metis distinct.simps(2) vt_s.wq_distinct)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1901
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1902
        show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1903
          by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1904
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1905
      show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1906
      proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1907
        { assume eq_th: "th = thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1908
          from eq_th have eq_cnp: "cntP (e # s) th = cntP s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1909
            by (unfold eq_e, simp add:cntP_def count_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1910
          moreover from eq_th have eq_cnv: "cntV (e#s) th = 1 + cntV s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1911
            by (unfold eq_e, simp add:cntV_def count_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1912
          moreover from cntCS_v_dec [OF vtv] 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1913
          have "cntCS (e # s) thread + 1 = cntCS s thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1914
            by (simp add:eq_e)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1915
          moreover from is_runing have rd_before: "thread \<in> readys s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1916
            by (unfold runing_def, simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1917
          moreover have "thread \<in> readys (e # s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1918
          proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1919
            from is_runing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1920
            have "thread \<in> threads (e#s)" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1921
              by (unfold eq_e, auto simp:runing_def readys_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1922
            moreover have "\<forall> cs1. \<not> waiting (e#s) thread cs1"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1923
            proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1924
              fix cs1
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1925
              { assume eq_cs: "cs1 = cs" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1926
                have "\<not> waiting (e # s) thread cs1"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1927
                proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1928
                  from eq_wq
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1929
                  have "thread \<notin> set (wq (e#s) cs1)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1930
                    apply(unfold eq_e wq_def eq_cs s_holding_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1931
                    apply (auto simp:Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1932
                  proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1933
                    assume "thread \<in> set (SOME q. distinct q \<and> set q = set rest)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1934
                    with eq_set have "thread \<in> set rest" by simp
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1935
                    with vt_v.wq_distinct[of cs]
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1936
                    and eq_wq show False
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1937
                        by (metis distinct.simps(2) vt_s.wq_distinct)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1938
                  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1939
                  thus ?thesis by (simp add:wq_def s_waiting_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1940
                qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1941
              } moreover {
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1942
                assume neq_cs: "cs1 \<noteq> cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1943
                  have "\<not> waiting (e # s) thread cs1" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1944
                  proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1945
                    from wq_v_neq [OF neq_cs[symmetric]]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1946
                    have "wq (V thread cs # s) cs1 = wq s cs1" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1947
                    moreover have "\<not> waiting s thread cs1" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1948
                    proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1949
                      from runing_ready and is_runing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1950
                      have "thread \<in> readys s" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1951
                      thus ?thesis by (simp add:readys_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1952
                    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1953
                    ultimately show ?thesis 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1954
                      by (auto simp:wq_def s_waiting_def eq_e)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1955
                  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1956
              } ultimately show "\<not> waiting (e # s) thread cs1" by blast
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1957
            qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1958
            ultimately show ?thesis by (simp add:readys_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1959
          qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1960
          moreover note eq_th ih
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1961
          ultimately have ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1962
        } moreover {
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1963
          assume neq_th: "th \<noteq> thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1964
          from neq_th eq_e have eq_cnp: "cntP (e # s) th = cntP s th" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1965
            by (simp add:cntP_def count_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1966
          from neq_th eq_e have eq_cnv: "cntV (e # s) th = cntV s th" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1967
            by (simp add:cntV_def count_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1968
          have ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1969
          proof(cases "th \<in> set rest")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1970
            case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1971
            have "(th \<in> readys (e # s)) = (th \<in> readys s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1972
              apply (insert step_back_vt[OF vtv])
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1973
              by (simp add: False eq_e eq_wq neq_th vt_s.readys_v_eq)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1974
            moreover have "cntCS (e#s) th = cntCS s th"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1975
              apply (insert neq_th, unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1976
              proof -
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1977
                have "{csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1978
                      {cs. (Cs cs, Th th) \<in> RAG s}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1979
                proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1980
                  from False eq_wq
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1981
                  have " next_th s thread cs th \<Longrightarrow> (Cs cs, Th th) \<in> RAG s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1982
                    apply (unfold next_th_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1983
                  proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1984
                    assume ne: "rest \<noteq> []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1985
                      and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1986
                      and eq_wq: "wq s cs = thread # rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1987
                    from eq_set ni have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1988
                                  set (SOME q. distinct q \<and> set q = set rest)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1989
                                  " by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1990
                    moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1991
                    proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1992
                      from vt_s.wq_distinct[ of cs] and eq_wq
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1993
                      show "distinct rest \<and> set rest = set rest" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1994
                    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1995
                      fix x assume "distinct x \<and> set x = set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1996
                      with ne show "x \<noteq> []" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1997
                    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1998
                    ultimately show 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  1999
                      "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2000
                      by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2001
                  qed    
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2002
                  thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2003
                qed
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2004
                thus "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2005
                             card {cs. (Cs cs, Th th) \<in> RAG s}" by simp 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2006
              qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2007
            moreover note ih eq_cnp eq_cnv eq_threads
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2008
            ultimately show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2009
          next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2010
            case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2011
            assume th_in: "th \<in> set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2012
            show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2013
            proof(cases "next_th s thread cs th")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2014
              case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2015
              with eq_wq and th_in have 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2016
                neq_hd: "th \<noteq> hd (SOME q. distinct q \<and> set q = set rest)" (is "th \<noteq> hd ?rest")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2017
                by (auto simp:next_th_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2018
              have "(th \<in> readys (e # s)) = (th \<in> readys s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2019
              proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2020
                from eq_wq and th_in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2021
                have "\<not> th \<in> readys s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2022
                  apply (auto simp:readys_def s_waiting_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2023
                  apply (rule_tac x = cs in exI, auto)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2024
                  by (insert vt_s.wq_distinct[of cs], auto simp add: wq_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2025
                moreover 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2026
                from eq_wq and th_in and neq_hd
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2027
                have "\<not> (th \<in> readys (e # s))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2028
                  apply (auto simp:readys_def s_waiting_def eq_e wq_def Let_def split:list.splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2029
                  by (rule_tac x = cs in exI, auto simp:eq_set)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2030
                ultimately show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2031
              qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2032
              moreover have "cntCS (e#s) th = cntCS s th" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2033
              proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2034
                from eq_wq and  th_in and neq_hd
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2035
                have "(holdents (e # s) th) = (holdents s th)"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2036
                  apply (unfold eq_e step_RAG_v[OF vtv], 
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2037
                         auto simp:next_th_def eq_set s_RAG_def holdents_test wq_def
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2038
                                   Let_def cs_holding_def)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2039
                  by (insert vt_s.wq_distinct[of cs], auto simp:wq_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2040
                thus ?thesis by (simp add:cntCS_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2041
              qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2042
              moreover note ih eq_cnp eq_cnv eq_threads
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2043
              ultimately show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2044
            next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2045
              case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2046
              let ?rest = " (SOME q. distinct q \<and> set q = set rest)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2047
              let ?t = "hd ?rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2048
              from True eq_wq th_in neq_th
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2049
              have "th \<in> readys (e # s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2050
                apply (auto simp:eq_e readys_def s_waiting_def wq_def
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2051
                        Let_def next_th_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2052
              proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2053
                assume eq_wq: "wq_fun (schs s) cs = thread # rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2054
                  and t_in: "?t \<in> set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2055
                show "?t \<in> threads s"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2056
                proof(rule vt_s.wq_threads)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2057
                  from eq_wq and t_in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2058
                  show "?t \<in> set (wq s cs)" by (auto simp:wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2059
                qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2060
              next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2061
                fix csa
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2062
                assume eq_wq: "wq_fun (schs s) cs = thread # rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2063
                  and t_in: "?t \<in> set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2064
                  and neq_cs: "csa \<noteq> cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2065
                  and t_in': "?t \<in>  set (wq_fun (schs s) csa)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2066
                show "?t = hd (wq_fun (schs s) csa)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2067
                proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2068
                  { assume neq_hd': "?t \<noteq> hd (wq_fun (schs s) csa)"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2069
                    from vt_s.wq_distinct[of cs] and 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2070
                    eq_wq[folded wq_def] and t_in eq_wq
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2071
                    have "?t \<noteq> thread" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2072
                    with eq_wq and t_in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2073
                    have w1: "waiting s ?t cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2074
                      by (auto simp:s_waiting_def wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2075
                    from t_in' neq_hd'
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2076
                    have w2: "waiting s ?t csa"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2077
                      by (auto simp:s_waiting_def wq_def)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2078
                    from vt_s.waiting_unique[OF w1 w2]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2079
                    and neq_cs have "False" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2080
                  } thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2081
                qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2082
              qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2083
              moreover have "cntP s th = cntV s th + cntCS s th + 1"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2084
              proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2085
                have "th \<notin> readys s" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2086
                proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2087
                  from True eq_wq neq_th th_in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2088
                  show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2089
                    apply (unfold readys_def s_waiting_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2090
                    by (rule_tac x = cs in exI, auto simp add: wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2091
                qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2092
                moreover have "th \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2093
                proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2094
                  from th_in eq_wq
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2095
                  have "th \<in> set (wq s cs)" by simp
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2096
                  from vt_s.wq_threads [OF this] 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2097
                  show ?thesis .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2098
                qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2099
                ultimately show ?thesis using ih by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2100
              qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2101
              moreover from True neq_th have "cntCS (e # s) th = 1 + cntCS s th"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2102
                apply (unfold cntCS_def holdents_test eq_e step_RAG_v[OF vtv], auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2103
              proof -
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2104
                show "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs} =
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2105
                               Suc (card {cs. (Cs cs, Th th) \<in> RAG s})"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2106
                  (is "card ?A = Suc (card ?B)")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2107
                proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2108
                  have "?A = insert cs ?B" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2109
                  hence "card ?A = card (insert cs ?B)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2110
                  also have "\<dots> = Suc (card ?B)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2111
                  proof(rule card_insert_disjoint)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2112
                    have "?B \<subseteq> ((\<lambda> (x, y). the_cs x) ` RAG s)" 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2113
                      apply (auto simp:image_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2114
                      by (rule_tac x = "(Cs x, Th th)" in bexI, auto)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2115
                    with vt_s.finite_RAG
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2116
                    show "finite {cs. (Cs cs, Th th) \<in> RAG s}" by (auto intro:finite_subset)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2117
                  next
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2118
                    show "cs \<notin> {cs. (Cs cs, Th th) \<in> RAG s}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2119
                    proof
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2120
                      assume "cs \<in> {cs. (Cs cs, Th th) \<in> RAG s}"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2121
                      hence "(Cs cs, Th th) \<in> RAG s" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2122
                      with True neq_th eq_wq show False
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2123
                        by (auto simp:next_th_def s_RAG_def cs_holding_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2124
                    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2125
                  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2126
                  finally show ?thesis .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2127
                qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2128
              qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2129
              moreover note eq_cnp eq_cnv
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2130
              ultimately show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2131
            qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2132
          qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2133
        } ultimately show ?thesis by blast
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2134
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2135
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2136
      case (thread_set thread prio)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2137
      assume eq_e: "e = Set thread prio"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2138
        and is_runing: "thread \<in> runing s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2139
      show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2140
      proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2141
        from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2142
        from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2143
        have eq_cncs: "cntCS (e#s) th = cntCS s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2144
          unfolding cntCS_def holdents_test
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2145
          by (simp add:RAG_set_unchanged eq_e)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2146
        from eq_e have eq_readys: "readys (e#s) = readys s" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2147
          by (simp add:readys_def cs_waiting_def s_waiting_def wq_def,
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2148
                  auto simp:Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2149
        { assume "th \<noteq> thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2150
          with eq_readys eq_e
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2151
          have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2152
                      (th \<in> readys (s) \<or> th \<notin> threads (s))" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2153
            by (simp add:threads.simps)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2154
          with eq_cnp eq_cnv eq_cncs ih is_runing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2155
          have ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2156
        } moreover {
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2157
          assume eq_th: "th = thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2158
          with is_runing ih have " cntP s th  = cntV s th + cntCS s th" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2159
            by (unfold runing_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2160
          moreover from eq_th and eq_readys is_runing have "th \<in> readys (e#s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2161
            by (simp add:runing_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2162
          moreover note eq_cnp eq_cnv eq_cncs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2163
          ultimately have ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2164
        } ultimately show ?thesis by blast
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2165
      qed   
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2166
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2167
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2168
    case vt_nil
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2169
    show ?case 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2170
      by (unfold cntP_def cntV_def cntCS_def, 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2171
        auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2172
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2173
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2174
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2175
lemma not_thread_cncs:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2176
  assumes not_in: "th \<notin> threads s" 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2177
  shows "cntCS s th = 0"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2178
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2179
  from vt not_in show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2180
  proof(induct arbitrary:th)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2181
    case (vt_cons s e th)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2182
    interpret vt_s: valid_trace s using vt_cons(1)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2183
       by (unfold_locales, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2184
    assume vt: "vt s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2185
      and ih: "\<And>th. th \<notin> threads s \<Longrightarrow> cntCS s th = 0"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2186
      and stp: "step s e"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2187
      and not_in: "th \<notin> threads (e # s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2188
    from stp show ?case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2189
    proof(cases)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2190
      case (thread_create thread prio)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2191
      assume eq_e: "e = Create thread prio"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2192
        and not_in': "thread \<notin> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2193
      have "cntCS (e # s) th = cntCS s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2194
        apply (unfold eq_e cntCS_def holdents_test)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2195
        by (simp add:RAG_create_unchanged)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2196
      moreover have "th \<notin> threads s" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2197
      proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2198
        from not_in eq_e show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2199
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2200
      moreover note ih ultimately show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2201
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2202
      case (thread_exit thread)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2203
      assume eq_e: "e = Exit thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2204
      and nh: "holdents s thread = {}"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2205
      have eq_cns: "cntCS (e # s) th = cntCS s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2206
        apply (unfold eq_e cntCS_def holdents_test)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2207
        by (simp add:RAG_exit_unchanged)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2208
      show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2209
      proof(cases "th = thread")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2210
        case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2211
        have "cntCS s th = 0" by (unfold cntCS_def, auto simp:nh True)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2212
        with eq_cns show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2213
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2214
        case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2215
        with not_in and eq_e
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2216
        have "th \<notin> threads s" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2217
        from ih[OF this] and eq_cns show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2218
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2219
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2220
      case (thread_P thread cs)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2221
      assume eq_e: "e = P thread cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2222
      and is_runing: "thread \<in> runing s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2223
      from assms thread_P ih vt stp thread_P have vtp: "vt (P thread cs#s)" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2224
      have neq_th: "th \<noteq> thread" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2225
      proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2226
        from not_in eq_e have "th \<notin> threads s" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2227
        moreover from is_runing have "thread \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2228
          by (simp add:runing_def readys_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2229
        ultimately show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2230
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2231
      hence "cntCS (e # s) th  = cntCS s th "
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2232
        apply (unfold cntCS_def holdents_test eq_e)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2233
        by (unfold step_RAG_p[OF vtp], auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2234
      moreover have "cntCS s th = 0"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2235
      proof(rule ih)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2236
        from not_in eq_e show "th \<notin> threads s" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2237
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2238
      ultimately show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2239
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2240
      case (thread_V thread cs)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2241
      assume eq_e: "e = V thread cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2242
        and is_runing: "thread \<in> runing s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2243
        and hold: "holding s thread cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2244
      have neq_th: "th \<noteq> thread" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2245
      proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2246
        from not_in eq_e have "th \<notin> threads s" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2247
        moreover from is_runing have "thread \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2248
          by (simp add:runing_def readys_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2249
        ultimately show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2250
      qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2251
      from assms thread_V vt stp ih 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2252
      have vtv: "vt (V thread cs#s)" by auto
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2253
      then interpret vt_v: valid_trace "(V thread cs#s)"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2254
        by (unfold_locales, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2255
      from hold obtain rest 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2256
        where eq_wq: "wq s cs = thread # rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2257
        by (case_tac "wq s cs", auto simp: wq_def s_holding_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2258
      from not_in eq_e eq_wq
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2259
      have "\<not> next_th s thread cs th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2260
        apply (auto simp:next_th_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2261
      proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2262
        assume ne: "rest \<noteq> []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2263
          and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> threads s" (is "?t \<notin> threads s")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2264
        have "?t \<in> set rest"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2265
        proof(rule someI2)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2266
          from vt_v.wq_distinct[of cs] and eq_wq
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2267
          show "distinct rest \<and> set rest = set rest"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2268
            by (metis distinct.simps(2) vt_s.wq_distinct) 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2269
        next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2270
          fix x assume "distinct x \<and> set x = set rest" with ne
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2271
          show "hd x \<in> set rest" by (cases x, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2272
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2273
        with eq_wq have "?t \<in> set (wq s cs)" by simp
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2274
        from vt_s.wq_threads[OF this] and ni
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2275
        show False
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2276
          using `hd (SOME q. distinct q \<and> set q = set rest) \<in> set (wq s cs)` 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2277
            ni vt_s.wq_threads by blast 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2278
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2279
      moreover note neq_th eq_wq
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2280
      ultimately have "cntCS (e # s) th  = cntCS s th"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2281
        by (unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2282
      moreover have "cntCS s th = 0"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2283
      proof(rule ih)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2284
        from not_in eq_e show "th \<notin> threads s" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2285
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2286
      ultimately show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2287
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2288
      case (thread_set thread prio)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2289
      print_facts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2290
      assume eq_e: "e = Set thread prio"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2291
        and is_runing: "thread \<in> runing s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2292
      from not_in and eq_e have "th \<notin> threads s" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2293
      from ih [OF this] and eq_e
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2294
      show ?thesis 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2295
        apply (unfold eq_e cntCS_def holdents_test)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2296
        by (simp add:RAG_set_unchanged)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2297
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2298
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2299
      case vt_nil
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2300
      show ?case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2301
      by (unfold cntCS_def, 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2302
        auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2303
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2304
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2305
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2306
end
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2307
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2308
lemma eq_waiting: "waiting (wq (s::state)) th cs = waiting s th cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2309
  by (auto simp:s_waiting_def cs_waiting_def wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2310
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2311
context valid_trace
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2312
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2313
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2314
lemma dm_RAG_threads:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2315
  assumes in_dom: "(Th th) \<in> Domain (RAG s)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2316
  shows "th \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2317
proof -
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2318
  from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2319
  moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2320
  ultimately have "(Th th, Cs cs) \<in> RAG s" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2321
  hence "th \<in> set (wq s cs)"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2322
    by (unfold s_RAG_def, auto simp:cs_waiting_def)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2323
  from wq_threads [OF this] show ?thesis .
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2324
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2325
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2326
end
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2327
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2328
lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2329
unfolding cp_def wq_def
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2330
apply(induct s rule: schs.induct)
44
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  2331
thm cpreced_initial
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2332
apply(simp add: Let_def cpreced_initial)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2333
apply(simp add: Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2334
apply(simp add: Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2335
apply(simp add: Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2336
apply(subst (2) schs.simps)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2337
apply(simp add: Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2338
apply(subst (2) schs.simps)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2339
apply(simp add: Let_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2340
done
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2341
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2342
context valid_trace
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2343
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2344
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2345
lemma runing_unique:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2346
  assumes runing_1: "th1 \<in> runing s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2347
  and runing_2: "th2 \<in> runing s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2348
  shows "th1 = th2"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2349
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2350
  from runing_1 and runing_2 have "cp s th1 = cp s th2"
44
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  2351
    unfolding runing_def
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  2352
    apply(simp)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  2353
    done
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2354
  hence eq_max: "Max ((\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1)) =
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2355
                 Max ((\<lambda>th. preced th s) ` ({th2} \<union> dependants (wq s) th2))"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2356
    (is "Max (?f ` ?A) = Max (?f ` ?B)")
44
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  2357
    unfolding cp_eq_cpreced 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  2358
    unfolding cpreced_def .
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2359
  obtain th1' where th1_in: "th1' \<in> ?A" and eq_f_th1: "?f th1' = Max (?f ` ?A)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2360
  proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2361
    have h1: "finite (?f ` ?A)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2362
    proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2363
      have "finite ?A" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2364
      proof -
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2365
        have "finite (dependants (wq s) th1)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2366
        proof-
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2367
          have "finite {th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2368
          proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2369
            let ?F = "\<lambda> (x, y). the_th x"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2370
            have "{th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2371
              apply (auto simp:image_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2372
              by (rule_tac x = "(Th x, Th th1)" in bexI, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2373
            moreover have "finite \<dots>"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2374
            proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2375
              from finite_RAG have "finite (RAG s)" .
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2376
              hence "finite ((RAG (wq s))\<^sup>+)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2377
                apply (unfold finite_trancl)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2378
                by (auto simp: s_RAG_def cs_RAG_def wq_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2379
              thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2380
            qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2381
            ultimately show ?thesis by (auto intro:finite_subset)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2382
          qed
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2383
          thus ?thesis by (simp add:cs_dependants_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2384
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2385
        thus ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2386
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2387
      thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2388
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2389
    moreover have h2: "(?f ` ?A) \<noteq> {}"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2390
    proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2391
      have "?A \<noteq> {}" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2392
      thus ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2393
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2394
    from Max_in [OF h1 h2]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2395
    have "Max (?f ` ?A) \<in> (?f ` ?A)" .
44
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  2396
    thus ?thesis 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  2397
      thm cpreced_def
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  2398
      unfolding cpreced_def[symmetric] 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  2399
      unfolding cp_eq_cpreced[symmetric] 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  2400
      unfolding cpreced_def 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  2401
      using that[intro] by (auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2402
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2403
  obtain th2' where th2_in: "th2' \<in> ?B" and eq_f_th2: "?f th2' = Max (?f ` ?B)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2404
  proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2405
    have h1: "finite (?f ` ?B)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2406
    proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2407
      have "finite ?B" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2408
      proof -
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2409
        have "finite (dependants (wq s) th2)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2410
        proof-
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2411
          have "finite {th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2412
          proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2413
            let ?F = "\<lambda> (x, y). the_th x"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2414
            have "{th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2415
              apply (auto simp:image_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2416
              by (rule_tac x = "(Th x, Th th2)" in bexI, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2417
            moreover have "finite \<dots>"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2418
            proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2419
              from finite_RAG have "finite (RAG s)" .
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2420
              hence "finite ((RAG (wq s))\<^sup>+)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2421
                apply (unfold finite_trancl)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2422
                by (auto simp: s_RAG_def cs_RAG_def wq_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2423
              thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2424
            qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2425
            ultimately show ?thesis by (auto intro:finite_subset)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2426
          qed
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2427
          thus ?thesis by (simp add:cs_dependants_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2428
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2429
        thus ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2430
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2431
      thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2432
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2433
    moreover have h2: "(?f ` ?B) \<noteq> {}"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2434
    proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2435
      have "?B \<noteq> {}" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2436
      thus ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2437
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2438
    from Max_in [OF h1 h2]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2439
    have "Max (?f ` ?B) \<in> (?f ` ?B)" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2440
    thus ?thesis by (auto intro:that)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2441
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2442
  from eq_f_th1 eq_f_th2 eq_max 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2443
  have eq_preced: "preced th1' s = preced th2' s" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2444
  hence eq_th12: "th1' = th2'"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2445
  proof (rule preced_unique)
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2446
    from th1_in have "th1' = th1 \<or> (th1' \<in> dependants (wq s) th1)" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2447
    thus "th1' \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2448
    proof
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2449
      assume "th1' \<in> dependants (wq s) th1"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2450
      hence "(Th th1') \<in> Domain ((RAG s)^+)"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2451
        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2452
        by (auto simp:Domain_def)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2453
      hence "(Th th1') \<in> Domain (RAG s)" by (simp add:trancl_domain)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2454
      from dm_RAG_threads[OF this] show ?thesis .
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2455
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2456
      assume "th1' = th1"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2457
      with runing_1 show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2458
        by (unfold runing_def readys_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2459
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2460
  next
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2461
    from th2_in have "th2' = th2 \<or> (th2' \<in> dependants (wq s) th2)" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2462
    thus "th2' \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2463
    proof
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2464
      assume "th2' \<in> dependants (wq s) th2"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2465
      hence "(Th th2') \<in> Domain ((RAG s)^+)"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2466
        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2467
        by (auto simp:Domain_def)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2468
      hence "(Th th2') \<in> Domain (RAG s)" by (simp add:trancl_domain)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2469
      from dm_RAG_threads[OF this] show ?thesis .
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2470
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2471
      assume "th2' = th2"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2472
      with runing_2 show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2473
        by (unfold runing_def readys_def, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2474
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2475
  qed
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2476
  from th1_in have "th1' = th1 \<or> th1' \<in> dependants (wq s) th1" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2477
  thus ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2478
  proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2479
    assume eq_th': "th1' = th1"
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2480
    from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2481
    thus ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2482
    proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2483
      assume "th2' = th2" thus ?thesis using eq_th' eq_th12 by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2484
    next
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2485
      assume "th2' \<in> dependants (wq s) th2"
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2486
      with eq_th12 eq_th' have "th1 \<in> dependants (wq s) th2" by simp
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2487
      hence "(Th th1, Th th2) \<in> (RAG s)^+"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2488
        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2489
      hence "Th th1 \<in> Domain ((RAG s)^+)" 
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2490
        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2491
        by (auto simp:Domain_def)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2492
      hence "Th th1 \<in> Domain (RAG s)" by (simp add:trancl_domain)
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2493
      then obtain n where d: "(Th th1, n) \<in> RAG s" by (auto simp:Domain_def)
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2494
      from RAG_target_th [OF this]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2495
      obtain cs' where "n = Cs cs'" by auto
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2496
      with d have "(Th th1, Cs cs') \<in> RAG s" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2497
      with runing_1 have "False"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2498
        apply (unfold runing_def readys_def s_RAG_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2499
        by (auto simp:eq_waiting)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2500
      thus ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2501
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2502
  next
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2503
    assume th1'_in: "th1' \<in> dependants (wq s) th1"
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2504
    from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2505
    thus ?thesis 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2506
    proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2507
      assume "th2' = th2"
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2508
      with th1'_in eq_th12 have "th2 \<in> dependants (wq s) th1" by simp
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2509
      hence "(Th th2, Th th1) \<in> (RAG s)^+"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2510
        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2511
      hence "Th th2 \<in> Domain ((RAG s)^+)" 
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2512
        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2513
        by (auto simp:Domain_def)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2514
      hence "Th th2 \<in> Domain (RAG s)" by (simp add:trancl_domain)
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2515
      then obtain n where d: "(Th th2, n) \<in> RAG s" by (auto simp:Domain_def)
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2516
      from RAG_target_th [OF this]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2517
      obtain cs' where "n = Cs cs'" by auto
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2518
      with d have "(Th th2, Cs cs') \<in> RAG s" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2519
      with runing_2 have "False"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2520
        apply (unfold runing_def readys_def s_RAG_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2521
        by (auto simp:eq_waiting)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2522
      thus ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2523
    next
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2524
      assume "th2' \<in> dependants (wq s) th2"
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2525
      with eq_th12 have "th1' \<in> dependants (wq s) th2" by simp
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2526
      hence h1: "(Th th1', Th th2) \<in> (RAG s)^+"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2527
        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2528
      from th1'_in have h2: "(Th th1', Th th1) \<in> (RAG s)^+"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2529
        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2530
      show ?thesis
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2531
      proof(rule dchain_unique[OF h1 _ h2, symmetric])
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2532
        from runing_1 show "th1 \<in> readys s" by (simp add:runing_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2533
        from runing_2 show "th2 \<in> readys s" by (simp add:runing_def) 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2534
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2535
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2536
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2537
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2538
39
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
  2539
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2540
lemma "card (runing s) \<le> 1"
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
  2541
apply(subgoal_tac "finite (runing s)")
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
  2542
prefer 2
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
  2543
apply (metis finite_nat_set_iff_bounded lessI runing_unique)
44
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  2544
apply(rule ccontr)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
  2545
apply(simp)
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
  2546
apply(case_tac "Suc (Suc 0) \<le> card (runing s)")
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
  2547
apply(subst (asm) card_le_Suc_iff)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
  2548
apply(simp)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
  2549
apply(auto)[1]
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
  2550
apply (metis insertCI runing_unique)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
  2551
apply(auto) 
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
  2552
done
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
  2553
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2554
end
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2555
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2556
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2557
lemma create_pre:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2558
  assumes stp: "step s e"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2559
  and not_in: "th \<notin> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2560
  and is_in: "th \<in> threads (e#s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2561
  obtains prio where "e = Create th prio"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2562
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2563
  from assms  
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2564
  show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2565
  proof(cases)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2566
    case (thread_create thread prio)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2567
    with is_in not_in have "e = Create th prio" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2568
    from that[OF this] show ?thesis .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2569
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2570
    case (thread_exit thread)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2571
    with assms show ?thesis by (auto intro!:that)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2572
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2573
    case (thread_P thread)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2574
    with assms show ?thesis by (auto intro!:that)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2575
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2576
    case (thread_V thread)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2577
    with assms show ?thesis by (auto intro!:that)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2578
  next 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2579
    case (thread_set thread)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2580
    with assms show ?thesis by (auto intro!:that)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2581
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2582
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2583
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2584
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2585
context valid_trace
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2586
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2587
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2588
lemma cnp_cnv_eq:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2589
  assumes "th \<notin> threads s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2590
  shows "cntP s th = cntV s th"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2591
  using assms
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2592
  using cnp_cnv_cncs not_thread_cncs by auto
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2593
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2594
end
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2595
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2596
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2597
lemma eq_RAG: 
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2598
  "RAG (wq s) = RAG s"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2599
by (unfold cs_RAG_def s_RAG_def, auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2600
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2601
context valid_trace
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2602
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2603
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2604
lemma count_eq_dependants:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2605
  assumes eq_pv: "cntP s th = cntV s th"
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2606
  shows "dependants (wq s) th = {}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2607
proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2608
  from cnp_cnv_cncs and eq_pv
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2609
  have "cntCS s th = 0" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2610
    by (auto split:if_splits)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2611
  moreover have "finite {cs. (Cs cs, Th th) \<in> RAG s}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2612
  proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2613
    from finite_holding[of th] show ?thesis
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2614
      by (simp add:holdents_test)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2615
  qed
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2616
  ultimately have h: "{cs. (Cs cs, Th th) \<in> RAG s} = {}"
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2617
    by (unfold cntCS_def holdents_test cs_dependants_def, auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2618
  show ?thesis
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2619
  proof(unfold cs_dependants_def)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2620
    { assume "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2621
      then obtain th' where "(Th th', Th th) \<in> (RAG (wq s))\<^sup>+" by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2622
      hence "False"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2623
      proof(cases)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2624
        assume "(Th th', Th th) \<in> RAG (wq s)"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2625
        thus "False" by (auto simp:cs_RAG_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2626
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2627
        fix c
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2628
        assume "(c, Th th) \<in> RAG (wq s)"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2629
        with h and eq_RAG show "False"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2630
          by (cases c, auto simp:cs_RAG_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2631
      qed
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2632
    } thus "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} = {}" by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2633
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2634
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2635
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2636
lemma dependants_threads:
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2637
  shows "dependants (wq s) th \<subseteq> threads s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2638
proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2639
  { fix th th'
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2640
    assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (RAG (wq s))\<^sup>+}"
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2641
    have "Th th \<in> Domain (RAG s)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2642
    proof -
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2643
      from h obtain th' where "(Th th, Th th') \<in> (RAG (wq s))\<^sup>+" by auto
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2644
      hence "(Th th) \<in> Domain ( (RAG (wq s))\<^sup>+)" by (auto simp:Domain_def)
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2645
      with trancl_domain have "(Th th) \<in> Domain (RAG (wq s))" by simp
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2646
      thus ?thesis using eq_RAG by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2647
    qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2648
    from dm_RAG_threads[OF this]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2649
    have "th \<in> threads s" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2650
  } note hh = this
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2651
  fix th1 
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2652
  assume "th1 \<in> dependants (wq s) th"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2653
  hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (RAG (wq s))\<^sup>+}"
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2654
    by (unfold cs_dependants_def, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2655
  from hh [OF this] show "th1 \<in> threads s" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2656
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2657
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2658
lemma finite_threads:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2659
  shows "finite (threads s)"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2660
using vt by (induct) (auto elim: step.cases)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2661
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2662
end
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2663
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2664
lemma Max_f_mono:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2665
  assumes seq: "A \<subseteq> B"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2666
  and np: "A \<noteq> {}"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2667
  and fnt: "finite B"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2668
  shows "Max (f ` A) \<le> Max (f ` B)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2669
proof(rule Max_mono)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2670
  from seq show "f ` A \<subseteq> f ` B" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2671
next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2672
  from np show "f ` A \<noteq> {}" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2673
next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2674
  from fnt and seq show "finite (f ` B)" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2675
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2676
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2677
context valid_trace
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2678
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2679
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2680
lemma cp_le:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2681
  assumes th_in: "th \<in> threads s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2682
  shows "cp s th \<le> Max ((\<lambda> th. (preced th s)) ` threads s)"
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2683
proof(unfold cp_eq_cpreced cpreced_def cs_dependants_def)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2684
  show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}))
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2685
         \<le> Max ((\<lambda>th. preced th s) ` threads s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2686
    (is "Max (?f ` ?A) \<le> Max (?f ` ?B)")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2687
  proof(rule Max_f_mono)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2688
    show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2689
  next
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2690
    from finite_threads
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2691
    show "finite (threads s)" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2692
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2693
    from th_in
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2694
    show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> threads s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2695
      apply (auto simp:Domain_def)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2696
      apply (rule_tac dm_RAG_threads)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2697
      apply (unfold trancl_domain [of "RAG s", symmetric])
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2698
      by (unfold cs_RAG_def s_RAG_def, auto simp:Domain_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2699
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2700
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2701
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2702
lemma le_cp:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2703
  shows "preced th s \<le> cp s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2704
proof(unfold cp_eq_cpreced preced_def cpreced_def, simp)
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2705
  show "Prc (priority th s) (last_set th s)
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2706
    \<le> Max (insert (Prc (priority th s) (last_set th s))
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2707
            ((\<lambda>th. Prc (priority th s) (last_set th s)) ` dependants (wq s) th))"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2708
    (is "?l \<le> Max (insert ?l ?A)")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2709
  proof(cases "?A = {}")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2710
    case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2711
    have "finite ?A" (is "finite (?f ` ?B)")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2712
    proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2713
      have "finite ?B" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2714
      proof-
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2715
        have "finite {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2716
        proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2717
          let ?F = "\<lambda> (x, y). the_th x"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2718
          have "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2719
            apply (auto simp:image_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2720
            by (rule_tac x = "(Th x, Th th)" in bexI, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2721
          moreover have "finite \<dots>"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2722
          proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2723
            from finite_RAG have "finite (RAG s)" .
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2724
            hence "finite ((RAG (wq s))\<^sup>+)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2725
              apply (unfold finite_trancl)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2726
              by (auto simp: s_RAG_def cs_RAG_def wq_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2727
            thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2728
          qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2729
          ultimately show ?thesis by (auto intro:finite_subset)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2730
        qed
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2731
        thus ?thesis by (simp add:cs_dependants_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2732
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2733
      thus ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2734
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2735
    from Max_insert [OF this False, of ?l] show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2736
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2737
    case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2738
    thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2739
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2740
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2741
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2742
lemma max_cp_eq: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2743
  shows "Max ((cp s) ` threads s) = Max ((\<lambda> th. (preced th s)) ` threads s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2744
  (is "?l = ?r")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2745
proof(cases "threads s = {}")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2746
  case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2747
  thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2748
next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2749
  case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2750
  have "?l \<in> ((cp s) ` threads s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2751
  proof(rule Max_in)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2752
    from finite_threads
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2753
    show "finite (cp s ` threads s)" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2754
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2755
    from False show "cp s ` threads s \<noteq> {}" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2756
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2757
  then obtain th 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2758
    where th_in: "th \<in> threads s" and eq_l: "?l = cp s th" by auto
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2759
  have "\<dots> \<le> ?r" by (rule cp_le[OF th_in])
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2760
  moreover have "?r \<le> cp s th" (is "Max (?f ` ?A) \<le> cp s th")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2761
  proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2762
    have "?r \<in> (?f ` ?A)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2763
    proof(rule Max_in)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2764
      from finite_threads
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2765
      show " finite ((\<lambda>th. preced th s) ` threads s)" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2766
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2767
      from False show " (\<lambda>th. preced th s) ` threads s \<noteq> {}" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2768
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2769
    then obtain th' where 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2770
      th_in': "th' \<in> ?A " and eq_r: "?r = ?f th'" by auto
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2771
    from le_cp [of th']  eq_r
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2772
    have "?r \<le> cp s th'" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2773
    moreover have "\<dots> \<le> cp s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2774
    proof(fold eq_l)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2775
      show " cp s th' \<le> Max (cp s ` threads s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2776
      proof(rule Max_ge)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2777
        from th_in' show "cp s th' \<in> cp s ` threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2778
          by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2779
      next
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2780
        from finite_threads
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2781
        show "finite (cp s ` threads s)" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2782
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2783
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2784
    ultimately show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2785
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2786
  ultimately show ?thesis using eq_l by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2787
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2788
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2789
lemma max_cp_readys_threads_pre:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2790
  assumes np: "threads s \<noteq> {}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2791
  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2792
proof(unfold max_cp_eq)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2793
  show "Max (cp s ` readys s) = Max ((\<lambda>th. preced th s) ` threads s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2794
  proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2795
    let ?p = "Max ((\<lambda>th. preced th s) ` threads s)" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2796
    let ?f = "(\<lambda>th. preced th s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2797
    have "?p \<in> ((\<lambda>th. preced th s) ` threads s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2798
    proof(rule Max_in)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2799
      from finite_threads show "finite (?f ` threads s)" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2800
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2801
      from np show "?f ` threads s \<noteq> {}" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2802
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2803
    then obtain tm where tm_max: "?f tm = ?p" and tm_in: "tm \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2804
      by (auto simp:Image_def)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2805
    from th_chain_to_ready [OF tm_in]
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2806
    have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+)" .
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2807
    thus ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2808
    proof
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2809
      assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+ "
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2810
      then obtain th' where th'_in: "th' \<in> readys s" 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2811
        and tm_chain:"(Th tm, Th th') \<in> (RAG s)\<^sup>+" by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2812
      have "cp s th' = ?f tm"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2813
      proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2814
        from dependants_threads finite_threads
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2815
        show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th'))" 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2816
          by (auto intro:finite_subset)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2817
      next
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2818
        fix p assume p_in: "p \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2819
        from tm_max have " preced tm s = Max ((\<lambda>th. preced th s) ` threads s)" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2820
        moreover have "p \<le> \<dots>"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2821
        proof(rule Max_ge)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2822
          from finite_threads
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2823
          show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2824
        next
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2825
          from p_in and th'_in and dependants_threads[of th']
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2826
          show "p \<in> (\<lambda>th. preced th s) ` threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2827
            by (auto simp:readys_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2828
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2829
        ultimately show "p \<le> preced tm s" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2830
      next
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2831
        show "preced tm s \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2832
        proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2833
          from tm_chain
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2834
          have "tm \<in> dependants (wq s) th'"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2835
            by (unfold cs_dependants_def s_RAG_def cs_RAG_def, auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2836
          thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2837
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2838
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2839
      with tm_max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2840
      have h: "cp s th' = Max ((\<lambda>th. preced th s) ` threads s)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2841
      show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2842
      proof (fold h, rule Max_eqI)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2843
        fix q 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2844
        assume "q \<in> cp s ` readys s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2845
        then obtain th1 where th1_in: "th1 \<in> readys s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2846
          and eq_q: "q = cp s th1" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2847
        show "q \<le> cp s th'"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2848
          apply (unfold h eq_q)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2849
          apply (unfold cp_eq_cpreced cpreced_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2850
          apply (rule Max_mono)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2851
        proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2852
          from dependants_threads [of th1] th1_in
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2853
          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<subseteq> 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2854
                 (\<lambda>th. preced th s) ` threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2855
            by (auto simp:readys_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2856
        next
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2857
          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2858
        next
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2859
          from finite_threads 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2860
          show " finite ((\<lambda>th. preced th s) ` threads s)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2861
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2862
      next
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2863
        from finite_threads
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2864
        show "finite (cp s ` readys s)" by (auto simp:readys_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2865
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2866
        from th'_in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2867
        show "cp s th' \<in> cp s ` readys s" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2868
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2869
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2870
      assume tm_ready: "tm \<in> readys s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2871
      show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2872
      proof(fold tm_max)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2873
        have cp_eq_p: "cp s tm = preced tm s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2874
        proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2875
          fix y 
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2876
          assume hy: "y \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2877
          show "y \<le> preced tm s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2878
          proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2879
            { fix y'
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2880
              assume hy' : "y' \<in> ((\<lambda>th. preced th s) ` dependants (wq s) tm)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2881
              have "y' \<le> preced tm s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2882
              proof(unfold tm_max, rule Max_ge)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2883
                from hy' dependants_threads[of tm]
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2884
                show "y' \<in> (\<lambda>th. preced th s) ` threads s" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2885
              next
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2886
                from finite_threads
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2887
                show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2888
              qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2889
            } with hy show ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2890
          qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2891
        next
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2892
          from dependants_threads[of tm] finite_threads
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2893
          show "finite ((\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm))"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2894
            by (auto intro:finite_subset)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2895
        next
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2896
          show "preced tm s \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2897
            by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2898
        qed 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2899
        moreover have "Max (cp s ` readys s) = cp s tm"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2900
        proof(rule Max_eqI)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2901
          from tm_ready show "cp s tm \<in> cp s ` readys s" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2902
        next
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2903
          from finite_threads
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2904
          show "finite (cp s ` readys s)" by (auto simp:readys_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2905
        next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2906
          fix y assume "y \<in> cp s ` readys s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2907
          then obtain th1 where th1_readys: "th1 \<in> readys s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2908
            and h: "y = cp s th1" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2909
          show "y \<le> cp s tm"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2910
            apply(unfold cp_eq_p h)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2911
            apply(unfold cp_eq_cpreced cpreced_def tm_max, rule Max_mono)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2912
          proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2913
            from finite_threads
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2914
            show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2915
          next
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2916
            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2917
              by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2918
          next
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2919
            from dependants_threads[of th1] th1_readys
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3
diff changeset
  2920
            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2921
                    \<subseteq> (\<lambda>th. preced th s) ` threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2922
              by (auto simp:readys_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2923
          qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2924
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2925
        ultimately show " Max (cp s ` readys s) = preced tm s" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2926
      qed 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2927
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2928
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2929
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2930
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  2931
text {* (* ccc *) \noindent
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  2932
  Since the current precedence of the threads in ready queue will always be boosted,
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  2933
  there must be one inside it has the maximum precedence of the whole system. 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  2934
*}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2935
lemma max_cp_readys_threads:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2936
  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2937
proof(cases "threads s = {}")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2938
  case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2939
  thus ?thesis 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2940
    by (auto simp:readys_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2941
next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2942
  case False
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2943
  show ?thesis by (rule max_cp_readys_threads_pre[OF False])
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2944
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2945
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2946
end
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2947
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2948
lemma eq_holding: "holding (wq s) th cs = holding s th cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2949
  apply (unfold s_holding_def cs_holding_def wq_def, simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2950
  done
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2951
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2952
lemma f_image_eq:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2953
  assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2954
  shows "f ` A = g ` A"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2955
proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2956
  show "f ` A \<subseteq> g ` A"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2957
    by(rule image_subsetI, auto intro:h)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2958
next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2959
  show "g ` A \<subseteq> f ` A"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2960
   by (rule image_subsetI, auto intro:h[symmetric])
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2961
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2962
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2963
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2964
definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2965
  where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2966
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2967
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2968
lemma detached_test:
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2969
  shows "detached s th = (Th th \<notin> Field (RAG s))"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2970
apply(simp add: detached_def Field_def)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2971
apply(simp add: s_RAG_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2972
apply(simp add: s_holding_abv s_waiting_abv)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2973
apply(simp add: Domain_iff Range_iff)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2974
apply(simp add: wq_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2975
apply(auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2976
done
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2977
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2978
context valid_trace
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2979
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2980
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2981
lemma detached_intro:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2982
  assumes eq_pv: "cntP s th = cntV s th"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2983
  shows "detached s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2984
proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2985
 from cnp_cnv_cncs
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2986
  have eq_cnt: "cntP s th =
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2987
    cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2988
  hence cncs_zero: "cntCS s th = 0"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2989
    by (auto simp:eq_pv split:if_splits)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2990
  with eq_cnt
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2991
  have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:eq_pv)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2992
  thus ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2993
  proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2994
    assume "th \<notin> threads s"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  2995
    with range_in dm_RAG_threads
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2996
    show ?thesis
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  2997
      by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def Domain_iff Range_iff)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2998
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2999
    assume "th \<in> readys s"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  3000
    moreover have "Th th \<notin> Range (RAG s)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3001
    proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3002
      from card_0_eq [OF finite_holding] and cncs_zero
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3003
      have "holdents s th = {}"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3004
        by (simp add:cntCS_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3005
      thus ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3006
        apply(auto simp:holdents_test)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3007
        apply(case_tac a)
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  3008
        apply(auto simp:holdents_test s_RAG_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3009
        done
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3010
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3011
    ultimately show ?thesis
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  3012
      by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def readys_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3013
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3014
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3016
lemma detached_elim:
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3017
  assumes dtc: "detached s th"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3018
  shows "cntP s th = cntV s th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3019
proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3020
  from cnp_cnv_cncs
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3021
  have eq_pv: " cntP s th =
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3022
    cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3023
  have cncs_z: "cntCS s th = 0"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3024
  proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3025
    from dtc have "holdents s th = {}"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  3026
      unfolding detached_def holdents_test s_RAG_def
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3027
      by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3028
    thus ?thesis by (auto simp:cntCS_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3029
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3030
  show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3031
  proof(cases "th \<in> threads s")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3032
    case True
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3033
    with dtc 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3034
    have "th \<in> readys s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3035
      by (unfold readys_def detached_def Field_def Domain_def Range_def, 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  3036
           auto simp:eq_waiting s_RAG_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3037
    with cncs_z and eq_pv show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3038
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3039
    case False
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3040
    with cncs_z and eq_pv show ?thesis by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3041
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3042
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3043
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3044
lemma detached_eq:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3045
  shows "(detached s th) = (cntP s th = cntV s th)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3046
  by (insert vt, auto intro:detached_intro detached_elim)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3047
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3048
end
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3049
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  3050
text {* 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  3051
  The lemmas in this .thy file are all obvious lemmas, however, they still needs to be derived
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  3052
  from the concise and miniature model of PIP given in PrioGDef.thy.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  3053
*}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 44
diff changeset
  3054
61
f8194fd6214f CpsG.thy has been cleaned up.
zhangx
parents: 58
diff changeset
  3055
lemma eq_dependants: "dependants (wq s) = dependants s"
f8194fd6214f CpsG.thy has been cleaned up.
zhangx
parents: 58
diff changeset
  3056
  by (simp add: s_dependants_abv wq_def)
f8194fd6214f CpsG.thy has been cleaned up.
zhangx
parents: 58
diff changeset
  3057
f8194fd6214f CpsG.thy has been cleaned up.
zhangx
parents: 58
diff changeset
  3058
lemma next_th_unique: 
f8194fd6214f CpsG.thy has been cleaned up.
zhangx
parents: 58
diff changeset
  3059
  assumes nt1: "next_th s th cs th1"
f8194fd6214f CpsG.thy has been cleaned up.
zhangx
parents: 58
diff changeset
  3060
  and nt2: "next_th s th cs th2"
f8194fd6214f CpsG.thy has been cleaned up.
zhangx
parents: 58
diff changeset
  3061
  shows "th1 = th2"
f8194fd6214f CpsG.thy has been cleaned up.
zhangx
parents: 58
diff changeset
  3062
using assms by (unfold next_th_def, auto)
f8194fd6214f CpsG.thy has been cleaned up.
zhangx
parents: 58
diff changeset
  3063
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3064
lemma birth_time_lt:  "s \<noteq> [] \<Longrightarrow> last_set th s < length s"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3065
  apply (induct s, simp)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3066
proof -
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3067
  fix a s
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3068
  assume ih: "s \<noteq> [] \<Longrightarrow> last_set th s < length s"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3069
    and eq_as: "a # s \<noteq> []"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3070
  show "last_set th (a # s) < length (a # s)"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3071
  proof(cases "s \<noteq> []")
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3072
    case False
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3073
    from False show ?thesis
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3074
      by (cases a, auto simp:last_set.simps)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3075
  next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3076
    case True
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3077
    from ih [OF True] show ?thesis
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3078
      by (cases a, auto simp:last_set.simps)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3079
  qed
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3080
qed
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3081
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3082
lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3083
  by (induct s, auto simp:threads.simps)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3084
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3085
lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3086
  apply (drule_tac th_in_ne)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3087
  by (unfold preced_def, auto intro: birth_time_lt)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  3088
65
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3089
lemma inj_the_preced: 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3090
  "inj_on (the_preced s) (threads s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3091
  by (metis inj_onI preced_unique the_preced_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3092
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3093
lemma tRAG_alt_def: 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3094
  "tRAG s = {(Th th1, Th th2) | th1 th2. 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3095
                  \<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3096
 by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3097
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3098
lemma tRAG_Field:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3099
  "Field (tRAG s) \<subseteq> Field (RAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3100
  by (unfold tRAG_alt_def Field_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3101
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3102
lemma tRAG_ancestorsE:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3103
  assumes "x \<in> ancestors (tRAG s) u"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3104
  obtains th where "x = Th th"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3105
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3106
  from assms have "(u, x) \<in> (tRAG s)^+" 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3107
      by (unfold ancestors_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3108
  from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3109
  then obtain th where "x = Th th"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3110
    by (unfold tRAG_alt_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3111
  from that[OF this] show ?thesis .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3112
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3113
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3114
lemma tRAG_mono:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3115
  assumes "RAG s' \<subseteq> RAG s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3116
  shows "tRAG s' \<subseteq> tRAG s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3117
  using assms 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3118
  by (unfold tRAG_alt_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3119
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3120
lemma holding_next_thI:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3121
  assumes "holding s th cs"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3122
  and "length (wq s cs) > 1"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3123
  obtains th' where "next_th s th cs th'"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3124
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3125
  from assms(1)[folded eq_holding, unfolded cs_holding_def]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3126
  have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3127
  then obtain rest where h1: "wq s cs = th#rest" 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3128
    by (cases "wq s cs", auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3129
  with assms(2) have h2: "rest \<noteq> []" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3130
  let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3131
  have "next_th s th cs ?th'" using  h1(1) h2 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3132
    by (unfold next_th_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3133
  from that[OF this] show ?thesis .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3134
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3135
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3136
lemma RAG_tRAG_transfer:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3137
  assumes "vt s'"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3138
  assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3139
  and "(Cs cs, Th th'') \<in> RAG s'"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3140
  shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R")
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3141
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3142
  interpret vt_s': valid_trace "s'" using assms(1)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3143
    by (unfold_locales, simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3144
  interpret rtree: rtree "RAG s'"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3145
  proof
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3146
  show "single_valued (RAG s')"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3147
  apply (intro_locales)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3148
    by (unfold single_valued_def, 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3149
        auto intro:vt_s'.unique_RAG)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3150
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3151
  show "acyclic (RAG s')"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3152
     by (rule vt_s'.acyclic_RAG)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3153
  qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3154
  { fix n1 n2
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3155
    assume "(n1, n2) \<in> ?L"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3156
    from this[unfolded tRAG_alt_def]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3157
    obtain th1 th2 cs' where 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3158
      h: "n1 = Th th1" "n2 = Th th2" 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3159
         "(Th th1, Cs cs') \<in> RAG s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3160
         "(Cs cs', Th th2) \<in> RAG s" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3161
    from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \<in> RAG s'" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3162
    from h(3) and assms(2) 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3163
    have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3164
          (Th th1, Cs cs') \<in> RAG s'" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3165
    hence "(n1, n2) \<in> ?R"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3166
    proof
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3167
      assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3168
      hence eq_th1: "th1 = th" by simp
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3169
      moreover have "th2 = th''"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3170
      proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3171
        from h1 have "cs' = cs" by simp
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3172
        from assms(3) cs_in[unfolded this] rtree.sgv
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3173
        show ?thesis
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3174
          by (unfold single_valued_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3175
      qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3176
      ultimately show ?thesis using h(1,2) by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3177
    next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3178
      assume "(Th th1, Cs cs') \<in> RAG s'"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3179
      with cs_in have "(Th th1, Th th2) \<in> tRAG s'"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3180
        by (unfold tRAG_alt_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3181
      from this[folded h(1, 2)] show ?thesis by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3182
    qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3183
  } moreover {
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3184
    fix n1 n2
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3185
    assume "(n1, n2) \<in> ?R"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3186
    hence "(n1, n2) \<in>tRAG s' \<or> (n1, n2) = (Th th, Th th'')" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3187
    hence "(n1, n2) \<in> ?L" 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3188
    proof
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3189
      assume "(n1, n2) \<in> tRAG s'"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3190
      moreover have "... \<subseteq> ?L"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3191
      proof(rule tRAG_mono)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3192
        show "RAG s' \<subseteq> RAG s" by (unfold assms(2), auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3193
      qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3194
      ultimately show ?thesis by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3195
    next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3196
      assume eq_n: "(n1, n2) = (Th th, Th th'')"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3197
      from assms(2, 3) have "(Cs cs, Th th'') \<in> RAG s" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3198
      moreover have "(Th th, Cs cs) \<in> RAG s" using assms(2) by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3199
      ultimately show ?thesis 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3200
        by (unfold eq_n tRAG_alt_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3201
    qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3202
  } ultimately show ?thesis by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3203
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3204
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3205
context valid_trace
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3206
begin
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3207
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3208
lemmas RAG_tRAG_transfer = RAG_tRAG_transfer[OF vt]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3209
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
  3210
end
65
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3211
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3212
lemma cp_alt_def:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3213
  "cp s th =  
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3214
           Max ((the_preced s) ` {th'. Th th' \<in> (subtree (RAG s) (Th th))})"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3215
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3216
  have "Max (the_preced s ` ({th} \<union> dependants (wq s) th)) =
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3217
        Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})" 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3218
          (is "Max (_ ` ?L) = Max (_ ` ?R)")
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3219
  proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3220
    have "?L = ?R" 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3221
    by (auto dest:rtranclD simp:cs_dependants_def cs_RAG_def s_RAG_def subtree_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3222
    thus ?thesis by simp
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3223
  qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3224
  thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3225
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3226
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3227
lemma cp_gen_alt_def:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3228
  "cp_gen s = (Max \<circ> (\<lambda>x. (the_preced s \<circ> the_thread) ` subtree (tRAG s) x))"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3229
    by (auto simp:cp_gen_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3230
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3231
lemma tRAG_nodeE:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3232
  assumes "(n1, n2) \<in> tRAG s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3233
  obtains th1 th2 where "n1 = Th th1" "n2 = Th th2"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3234
  using assms
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3235
  by (auto simp: tRAG_def wRAG_def hRAG_def tRAG_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3236
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3237
lemma subtree_nodeE:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3238
  assumes "n \<in> subtree (tRAG s) (Th th)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3239
  obtains th1 where "n = Th th1"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3240
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3241
  show ?thesis
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3242
  proof(rule subtreeE[OF assms])
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3243
    assume "n = Th th"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3244
    from that[OF this] show ?thesis .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3245
  next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3246
    assume "Th th \<in> ancestors (tRAG s) n"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3247
    hence "(n, Th th) \<in> (tRAG s)^+" by (auto simp:ancestors_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3248
    hence "\<exists> th1. n = Th th1"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3249
    proof(induct)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3250
      case (base y)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3251
      from tRAG_nodeE[OF this] show ?case by metis
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3252
    next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3253
      case (step y z)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3254
      thus ?case by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3255
    qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3256
    with that show ?thesis by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3257
  qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3258
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3259
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3260
lemma tRAG_star_RAG: "(tRAG s)^* \<subseteq> (RAG s)^*"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3261
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3262
  have "(wRAG s O hRAG s)^* \<subseteq> (RAG s O RAG s)^*" 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3263
    by (rule rtrancl_mono, auto simp:RAG_split)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3264
  also have "... \<subseteq> ((RAG s)^*)^*"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3265
    by (rule rtrancl_mono, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3266
  also have "... = (RAG s)^*" by simp
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3267
  finally show ?thesis by (unfold tRAG_def, simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3268
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3269
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3270
lemma tRAG_subtree_RAG: "subtree (tRAG s) x \<subseteq> subtree (RAG s) x"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3271
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3272
  { fix a
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3273
    assume "a \<in> subtree (tRAG s) x"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3274
    hence "(a, x) \<in> (tRAG s)^*" by (auto simp:subtree_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3275
    with tRAG_star_RAG[of s]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3276
    have "(a, x) \<in> (RAG s)^*" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3277
    hence "a \<in> subtree (RAG s) x" by (auto simp:subtree_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3278
  } thus ?thesis by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3279
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3280
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3281
lemma tRAG_trancl_eq:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3282
   "{th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3283
    {th'. (Th th', Th th)  \<in> (RAG s)^+}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3284
   (is "?L = ?R")
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3285
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3286
  { fix th'
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3287
    assume "th' \<in> ?L"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3288
    hence "(Th th', Th th) \<in> (tRAG s)^+" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3289
    from tranclD[OF this]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3290
    obtain z where h: "(Th th', z) \<in> tRAG s" "(z, Th th) \<in> (tRAG s)\<^sup>*" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3291
    from tRAG_subtree_RAG[of s] and this(2)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3292
    have "(z, Th th) \<in> (RAG s)^*" by (meson subsetCE tRAG_star_RAG) 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3293
    moreover from h(1) have "(Th th', z) \<in> (RAG s)^+" using tRAG_alt_def by auto 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3294
    ultimately have "th' \<in> ?R"  by auto 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3295
  } moreover 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3296
  { fix th'
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3297
    assume "th' \<in> ?R"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3298
    hence "(Th th', Th th) \<in> (RAG s)^+" by (auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3299
    from plus_rpath[OF this]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3300
    obtain xs where rp: "rpath (RAG s) (Th th') xs (Th th)" "xs \<noteq> []" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3301
    hence "(Th th', Th th) \<in> (tRAG s)^+"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3302
    proof(induct xs arbitrary:th' th rule:length_induct)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3303
      case (1 xs th' th)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3304
      then obtain x1 xs1 where Cons1: "xs = x1#xs1" by (cases xs, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3305
      show ?case
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3306
      proof(cases "xs1")
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3307
        case Nil
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3308
        from 1(2)[unfolded Cons1 Nil]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3309
        have rp: "rpath (RAG s) (Th th') [x1] (Th th)" .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3310
        hence "(Th th', x1) \<in> (RAG s)" by (cases, simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3311
        then obtain cs where "x1 = Cs cs" 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3312
              by (unfold s_RAG_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3313
        from rpath_nnl_lastE[OF rp[unfolded this]]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3314
        show ?thesis by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3315
      next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3316
        case (Cons x2 xs2)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3317
        from 1(2)[unfolded Cons1[unfolded this]]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3318
        have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3319
        from rpath_edges_on[OF this]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3320
        have eds: "edges_on (Th th' # x1 # x2 # xs2) \<subseteq> RAG s" .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3321
        have "(Th th', x1) \<in> edges_on (Th th' # x1 # x2 # xs2)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3322
            by (simp add: edges_on_unfold)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3323
        with eds have rg1: "(Th th', x1) \<in> RAG s" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3324
        then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3325
        have "(x1, x2) \<in> edges_on (Th th' # x1 # x2 # xs2)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3326
            by (simp add: edges_on_unfold)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3327
        from this eds
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3328
        have rg2: "(x1, x2) \<in> RAG s" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3329
        from this[unfolded eq_x1] 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3330
        obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3331
        from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3332
        have rt1: "(Th th', Th th1) \<in> tRAG s" by (unfold tRAG_alt_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3333
        from rp have "rpath (RAG s) x2 xs2 (Th th)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3334
           by  (elim rpath_ConsE, simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3335
        from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3336
        show ?thesis
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3337
        proof(cases "xs2 = []")
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3338
          case True
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3339
          from rpath_nilE[OF rp'[unfolded this]]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3340
          have "th1 = th" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3341
          from rt1[unfolded this] show ?thesis by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3342
        next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3343
          case False
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3344
          from 1(1)[rule_format, OF _ rp' this, unfolded Cons1 Cons]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3345
          have "(Th th1, Th th) \<in> (tRAG s)\<^sup>+" by simp
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3346
          with rt1 show ?thesis by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3347
        qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3348
      qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3349
    qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3350
    hence "th' \<in> ?L" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3351
  } ultimately show ?thesis by blast
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3352
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3353
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3354
lemma tRAG_trancl_eq_Th:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3355
   "{Th th' | th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3356
    {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3357
    using tRAG_trancl_eq by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3358
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3359
lemma dependants_alt_def:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3360
  "dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3361
  by (metis eq_RAG s_dependants_def tRAG_trancl_eq)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3362
  
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3363
context valid_trace
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3364
begin
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3365
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3366
lemma count_eq_tRAG_plus:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3367
  assumes "cntP s th = cntV s th"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3368
  shows "{th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3369
  using assms count_eq_dependants dependants_alt_def eq_dependants by auto 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3370
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3371
lemma count_eq_RAG_plus:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3372
  assumes "cntP s th = cntV s th"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3373
  shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3374
  using assms count_eq_dependants cs_dependants_def eq_RAG by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3375
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3376
lemma count_eq_RAG_plus_Th:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3377
  assumes "cntP s th = cntV s th"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3378
  shows "{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3379
  using count_eq_RAG_plus[OF assms] by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3380
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3381
lemma count_eq_tRAG_plus_Th:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3382
  assumes "cntP s th = cntV s th"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3383
  shows "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3384
   using count_eq_tRAG_plus[OF assms] by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3385
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3386
end
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3387
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3388
lemma tRAG_subtree_eq: 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3389
   "(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th'  \<in> (subtree (RAG s) (Th th))}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3390
   (is "?L = ?R")
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3391
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3392
  { fix n 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3393
    assume h: "n \<in> ?L"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3394
    hence "n \<in> ?R"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3395
    by (smt mem_Collect_eq subsetCE subtree_def subtree_nodeE tRAG_subtree_RAG) 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3396
  } moreover {
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3397
    fix n
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3398
    assume "n \<in> ?R"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3399
    then obtain th' where h: "n = Th th'" "(Th th', Th th) \<in> (RAG s)^*"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3400
      by (auto simp:subtree_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3401
    from rtranclD[OF this(2)]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3402
    have "n \<in> ?L"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3403
    proof
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3404
      assume "Th th' \<noteq> Th th \<and> (Th th', Th th) \<in> (RAG s)\<^sup>+"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3405
      with h have "n \<in> {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3406
      thus ?thesis using subtree_def tRAG_trancl_eq by fastforce
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3407
    qed (insert h, auto simp:subtree_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3408
  } ultimately show ?thesis by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3409
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3410
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3411
lemma threads_set_eq: 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3412
   "the_thread ` (subtree (tRAG s) (Th th)) = 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3413
                  {th'. Th th' \<in> (subtree (RAG s) (Th th))}" (is "?L = ?R")
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3414
   by (auto intro:rev_image_eqI simp:tRAG_subtree_eq)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3415
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3416
lemma cp_alt_def1: 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3417
  "cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3418
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3419
  have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) =
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3420
       ((the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th))"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3421
       by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3422
  thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3423
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3424
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3425
lemma cp_gen_def_cond: 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3426
  assumes "x = Th th"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3427
  shows "cp s th = cp_gen s (Th th)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3428
by (unfold cp_alt_def1 cp_gen_def, simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3429
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3430
lemma cp_gen_over_set:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3431
  assumes "\<forall> x \<in> A. \<exists> th. x = Th th"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3432
  shows "cp_gen s ` A = (cp s \<circ> the_thread) ` A"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3433
proof(rule f_image_eq)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3434
  fix a
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3435
  assume "a \<in> A"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3436
  from assms[rule_format, OF this]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3437
  obtain th where eq_a: "a = Th th" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3438
  show "cp_gen s a = (cp s \<circ> the_thread) a"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3439
    by  (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3440
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3441
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3442
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3443
context valid_trace
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3444
begin
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3445
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3446
lemma RAG_threads:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3447
  assumes "(Th th) \<in> Field (RAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3448
  shows "th \<in> threads s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3449
  using assms
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3450
  by (metis Field_def UnE dm_RAG_threads range_in vt)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3451
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3452
lemma subtree_tRAG_thread:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3453
  assumes "th \<in> threads s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3454
  shows "subtree (tRAG s) (Th th) \<subseteq> Th ` threads s" (is "?L \<subseteq> ?R")
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3455
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3456
  have "?L = {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3457
    by (unfold tRAG_subtree_eq, simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3458
  also have "... \<subseteq> ?R"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3459
  proof
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3460
    fix x
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3461
    assume "x \<in> {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3462
    then obtain th' where h: "x = Th th'" "Th th' \<in> subtree (RAG s) (Th th)" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3463
    from this(2)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3464
    show "x \<in> ?R"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3465
    proof(cases rule:subtreeE)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3466
      case 1
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3467
      thus ?thesis by (simp add: assms h(1)) 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3468
    next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3469
      case 2
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3470
      thus ?thesis by (metis ancestors_Field dm_RAG_threads h(1) image_eqI) 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3471
    qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3472
  qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3473
  finally show ?thesis .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3474
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3475
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3476
lemma readys_root:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3477
  assumes "th \<in> readys s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3478
  shows "root (RAG s) (Th th)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3479
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3480
  { fix x
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3481
    assume "x \<in> ancestors (RAG s) (Th th)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3482
    hence h: "(Th th, x) \<in> (RAG s)^+" by (auto simp:ancestors_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3483
    from tranclD[OF this]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3484
    obtain z where "(Th th, z) \<in> RAG s" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3485
    with assms(1) have False
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3486
         apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3487
         by (fold wq_def, blast)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3488
  } thus ?thesis by (unfold root_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3489
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3490
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3491
lemma readys_in_no_subtree:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3492
  assumes "th \<in> readys s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3493
  and "th' \<noteq> th"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3494
  shows "Th th \<notin> subtree (RAG s) (Th th')" 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3495
proof
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3496
   assume "Th th \<in> subtree (RAG s) (Th th')"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3497
   thus False
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3498
   proof(cases rule:subtreeE)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3499
      case 1
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3500
      with assms show ?thesis by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3501
   next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3502
      case 2
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3503
      with readys_root[OF assms(1)]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3504
      show ?thesis by (auto simp:root_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3505
   qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3506
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3507
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3508
lemma not_in_thread_isolated:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3509
  assumes "th \<notin> threads s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3510
  shows "(Th th) \<notin> Field (RAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3511
proof
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3512
  assume "(Th th) \<in> Field (RAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3513
  with dm_RAG_threads and range_in assms
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3514
  show False by (unfold Field_def, blast)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3515
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3516
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3517
lemma wf_RAG: "wf (RAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3518
proof(rule finite_acyclic_wf)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3519
  from finite_RAG show "finite (RAG s)" .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3520
next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3521
  from acyclic_RAG show "acyclic (RAG s)" .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3522
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3523
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3524
lemma sgv_wRAG: "single_valued (wRAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3525
  using waiting_unique
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3526
  by (unfold single_valued_def wRAG_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3527
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3528
lemma sgv_hRAG: "single_valued (hRAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3529
  using holding_unique 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3530
  by (unfold single_valued_def hRAG_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3531
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3532
lemma sgv_tRAG: "single_valued (tRAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3533
  by (unfold tRAG_def, rule single_valued_relcomp, 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3534
              insert sgv_wRAG sgv_hRAG, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3535
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3536
lemma acyclic_tRAG: "acyclic (tRAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3537
proof(unfold tRAG_def, rule acyclic_compose)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3538
  show "acyclic (RAG s)" using acyclic_RAG .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3539
next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3540
  show "wRAG s \<subseteq> RAG s" unfolding RAG_split by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3541
next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3542
  show "hRAG s \<subseteq> RAG s" unfolding RAG_split by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3543
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3544
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3545
lemma sgv_RAG: "single_valued (RAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3546
  using unique_RAG by (auto simp:single_valued_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3547
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3548
lemma rtree_RAG: "rtree (RAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3549
  using sgv_RAG acyclic_RAG
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3550
  by (unfold rtree_def rtree_axioms_def sgv_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3551
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3552
end
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3553
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3554
sublocale valid_trace < rtree_RAG: rtree "RAG s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3555
proof
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3556
  show "single_valued (RAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3557
  apply (intro_locales)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3558
    by (unfold single_valued_def, 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3559
        auto intro:unique_RAG)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3560
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3561
  show "acyclic (RAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3562
     by (rule acyclic_RAG)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3563
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3564
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3565
sublocale valid_trace < rtree_s: rtree "tRAG s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3566
proof(unfold_locales)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3567
  from sgv_tRAG show "single_valued (tRAG s)" .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3568
next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3569
  from acyclic_tRAG show "acyclic (tRAG s)" .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3570
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3571
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3572
sublocale valid_trace < fsbtRAGs : fsubtree "RAG s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3573
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3574
  show "fsubtree (RAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3575
  proof(intro_locales)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3576
    show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG] .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3577
  next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3578
    show "fsubtree_axioms (RAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3579
    proof(unfold fsubtree_axioms_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3580
      from wf_RAG show "wf (RAG s)" .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3581
    qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3582
  qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3583
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3584
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3585
sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3586
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3587
  have "fsubtree (tRAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3588
  proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3589
    have "fbranch (tRAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3590
    proof(unfold tRAG_def, rule fbranch_compose)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3591
        show "fbranch (wRAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3592
        proof(rule finite_fbranchI)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3593
           from finite_RAG show "finite (wRAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3594
           by (unfold RAG_split, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3595
        qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3596
    next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3597
        show "fbranch (hRAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3598
        proof(rule finite_fbranchI)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3599
           from finite_RAG 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3600
           show "finite (hRAG s)" by (unfold RAG_split, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3601
        qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3602
    qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3603
    moreover have "wf (tRAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3604
    proof(rule wf_subset)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3605
      show "wf (RAG s O RAG s)" using wf_RAG
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3606
        by (fold wf_comp_self, simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3607
    next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3608
      show "tRAG s \<subseteq> (RAG s O RAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3609
        by (unfold tRAG_alt_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3610
    qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3611
    ultimately show ?thesis
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3612
      by (unfold fsubtree_def fsubtree_axioms_def,auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3613
  qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3614
  from this[folded tRAG_def] show "fsubtree (tRAG s)" .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3615
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3616
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3617
lemma Max_UNION: 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3618
  assumes "finite A"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3619
  and "A \<noteq> {}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3620
  and "\<forall> M \<in> f ` A. finite M"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3621
  and "\<forall> M \<in> f ` A. M \<noteq> {}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3622
  shows "Max (\<Union>x\<in> A. f x) = Max (Max ` f ` A)" (is "?L = ?R")
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3623
  using assms[simp]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3624
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3625
  have "?L = Max (\<Union>(f ` A))"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3626
    by (fold Union_image_eq, simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3627
  also have "... = ?R"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3628
    by (subst Max_Union, simp+)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3629
  finally show ?thesis .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3630
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3631
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3632
lemma max_Max_eq:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3633
  assumes "finite A"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3634
    and "A \<noteq> {}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3635
    and "x = y"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3636
  shows "max x (Max A) = Max ({y} \<union> A)" (is "?L = ?R")
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3637
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3638
  have "?R = Max (insert y A)" by simp
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3639
  also from assms have "... = ?L"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3640
      by (subst Max.insert, simp+)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3641
  finally show ?thesis by simp
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3642
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3643
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3644
context valid_trace
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3645
begin
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3646
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3647
(* ddd *)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3648
lemma cp_gen_rec:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3649
  assumes "x = Th th"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3650
  shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3651
proof(cases "children (tRAG s) x = {}")
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3652
  case True
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3653
  show ?thesis
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3654
    by (unfold True cp_gen_def subtree_children, simp add:assms)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3655
next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3656
  case False
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3657
  hence [simp]: "children (tRAG s) x \<noteq> {}" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3658
  note fsbttRAGs.finite_subtree[simp]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3659
  have [simp]: "finite (children (tRAG s) x)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3660
     by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3661
            rule children_subtree)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3662
  { fix r x
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3663
    have "subtree r x \<noteq> {}" by (auto simp:subtree_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3664
  } note this[simp]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3665
  have [simp]: "\<exists>x\<in>children (tRAG s) x. subtree (tRAG s) x \<noteq> {}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3666
  proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3667
    from False obtain q where "q \<in> children (tRAG s) x" by blast
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3668
    moreover have "subtree (tRAG s) q \<noteq> {}" by simp
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3669
    ultimately show ?thesis by blast
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3670
  qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3671
  have h: "Max ((the_preced s \<circ> the_thread) `
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3672
                ({x} \<union> \<Union>(subtree (tRAG s) ` children (tRAG s) x))) =
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3673
        Max ({the_preced s th} \<union> cp_gen s ` children (tRAG s) x)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3674
                     (is "?L = ?R")
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3675
  proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3676
    let "Max (?f ` (?A \<union> \<Union> (?g ` ?B)))" = ?L
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3677
    let "Max (_ \<union> (?h ` ?B))" = ?R
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3678
    let ?L1 = "?f ` \<Union>(?g ` ?B)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3679
    have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3680
    proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3681
      have "?L1 = ?f ` (\<Union> x \<in> ?B.(?g x))" by simp
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3682
      also have "... =  (\<Union> x \<in> ?B. ?f ` (?g x))" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3683
      finally have "Max ?L1 = Max ..." by simp
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3684
      also have "... = Max (Max ` (\<lambda>x. ?f ` subtree (tRAG s) x) ` ?B)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3685
        by (subst Max_UNION, simp+)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3686
      also have "... = Max (cp_gen s ` children (tRAG s) x)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3687
          by (unfold image_comp cp_gen_alt_def, simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3688
      finally show ?thesis .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3689
    qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3690
    show ?thesis
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3691
    proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3692
      have "?L = Max (?f ` ?A \<union> ?L1)" by simp
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3693
      also have "... = max (the_preced s (the_thread x)) (Max ?L1)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3694
            by (subst Max_Un, simp+)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3695
      also have "... = max (?f x) (Max (?h ` ?B))"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3696
        by (unfold eq_Max_L1, simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3697
      also have "... =?R"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3698
        by (rule max_Max_eq, (simp)+, unfold assms, simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3699
      finally show ?thesis .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3700
    qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3701
  qed  thus ?thesis 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3702
          by (fold h subtree_children, unfold cp_gen_def, simp) 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3703
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3704
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3705
lemma cp_rec:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3706
  "cp s th = Max ({the_preced s th} \<union> 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3707
                     (cp s o the_thread) ` children (tRAG s) (Th th))"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3708
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3709
  have "Th th = Th th" by simp
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3710
  note h =  cp_gen_def_cond[OF this] cp_gen_rec[OF this]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3711
  show ?thesis 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3712
  proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3713
    have "cp_gen s ` children (tRAG s) (Th th) = 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3714
                (cp s \<circ> the_thread) ` children (tRAG s) (Th th)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3715
    proof(rule cp_gen_over_set)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3716
      show " \<forall>x\<in>children (tRAG s) (Th th). \<exists>th. x = Th th"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3717
        by (unfold tRAG_alt_def, auto simp:children_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3718
    qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3719
    thus ?thesis by (subst (1) h(1), unfold h(2), simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3720
  qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3721
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3722
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3723
end
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3724
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3725
(* keep *)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3726
lemma next_th_holding:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3727
  assumes vt: "vt s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3728
  and nxt: "next_th s th cs th'"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3729
  shows "holding (wq s) th cs"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3730
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3731
  from nxt[unfolded next_th_def]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3732
  obtain rest where h: "wq s cs = th # rest"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3733
                       "rest \<noteq> []" 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3734
                       "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3735
  thus ?thesis
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3736
    by (unfold cs_holding_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3737
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3738
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3739
context valid_trace
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3740
begin
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3741
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3742
lemma next_th_waiting:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3743
  assumes nxt: "next_th s th cs th'"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3744
  shows "waiting (wq s) th' cs"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3745
proof -
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3746
  from nxt[unfolded next_th_def]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3747
  obtain rest where h: "wq s cs = th # rest"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3748
                       "rest \<noteq> []" 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3749
                       "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3750
  from wq_distinct[of cs, unfolded h]
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3751
  have dst: "distinct (th # rest)" .
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3752
  have in_rest: "th' \<in> set rest"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3753
  proof(unfold h, rule someI2)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3754
    show "distinct rest \<and> set rest = set rest" using dst by auto
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3755
  next
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3756
    fix x assume "distinct x \<and> set x = set rest"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3757
    with h(2)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3758
    show "hd x \<in> set (rest)" by (cases x, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3759
  qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3760
  hence "th' \<in> set (wq s cs)" by (unfold h(1), auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3761
  moreover have "th' \<noteq> hd (wq s cs)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3762
    by (unfold h(1), insert in_rest dst, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3763
  ultimately show ?thesis by (auto simp:cs_waiting_def)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3764
qed
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3765
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3766
lemma next_th_RAG:
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3767
  assumes nxt: "next_th (s::event list) th cs th'"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3768
  shows "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3769
  using vt assms next_th_holding next_th_waiting
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3770
  by (unfold s_RAG_def, simp)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3771
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3772
end
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3773
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3774
-- {* A useless definition *}
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3775
definition cps:: "state \<Rightarrow> (thread \<times> precedence) set"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3776
where "cps s = {(th, cp s th) | th . th \<in> threads s}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3777
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 77
diff changeset
  3778
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 77
diff changeset
  3779
find_theorems "waiting" holding
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 77
diff changeset
  3780
context valid_trace
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 77
diff changeset
  3781
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 77
diff changeset
  3782
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 77
diff changeset
  3783
find_theorems "waiting" holding
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 69
diff changeset
  3784
65
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
  3785
end
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 77
diff changeset
  3786
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 77
diff changeset
  3787
end