Correctness.thy
author zhangx
Thu, 28 Jan 2016 21:14:17 +0800
changeset 90 ed938e2246b9
parent 82 c0a4e840aefe
child 91 0525670d8e6a
child 94 44df6ac30bd2
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 Correctness
65
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
     2
imports PIPBasics
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
69
1dc801552dfd simplified Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 68
diff changeset
     5
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
     6
text {* 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
     7
  The following two auxiliary lemmas are used to reason about @{term Max}.
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
lemma image_Max_eqI: 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    10
  assumes "finite B"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    11
  and "b \<in> B"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    12
  and "\<forall> x \<in> B. f x \<le> f b"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    13
  shows "Max (f ` B) = f b"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    14
  using assms
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    15
  using Max_eqI by blast 
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
lemma image_Max_subset:
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    18
  assumes "finite A"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    19
  and "B \<subseteq> A"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    20
  and "a \<in> B"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    21
  and "Max (f ` A) = f a"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    22
  shows "Max (f ` B) = f a"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    23
proof(rule image_Max_eqI)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    24
  show "finite B"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    25
    using assms(1) assms(2) finite_subset by auto 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    26
next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    27
  show "a \<in> B" using assms by simp
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    28
next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    29
  show "\<forall>x\<in>B. f x \<le> f a"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    30
    by (metis Max_ge assms(1) assms(2) assms(4) 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    31
            finite_imageI image_eqI subsetCE) 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    34
text {*
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    35
  The following locale @{text "highest_gen"} sets the basic context for our
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    36
  investigation: supposing thread @{text th} holds the highest @{term cp}-value
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    37
  in state @{text s}, which means the task for @{text th} is the 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    38
  most urgent. We want to show that  
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    39
  @{text th} is treated correctly by PIP, which means
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    40
  @{text th} will not be blocked unreasonably by other less urgent
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    41
  threads. 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    42
*}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    43
locale highest_gen =
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    44
  fixes s th prio tm
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    45
  assumes vt_s: "vt s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    46
  and threads_s: "th \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    47
  and highest: "preced th s = Max ((cp s)`threads s)"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    48
  -- {* The internal structure of @{term th}'s precedence is exposed:*}
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    49
  and preced_th: "preced th s = Prc prio tm" 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    50
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    51
-- {* @{term s} is a valid trace, so it will inherit all results derived for
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    52
      a valid trace: *}
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
    53
sublocale highest_gen < vat_s: valid_trace "s"
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
    54
  by (unfold_locales, insert vt_s, simp)
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
    55
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    56
context highest_gen
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    57
begin
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    58
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    59
text {*
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    60
  @{term tm} is the time when the precedence of @{term th} is set, so 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    61
  @{term tm} must be a valid moment index into @{term s}.
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    62
*}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    63
lemma lt_tm: "tm < length s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    64
  by (insert preced_tm_lt[OF threads_s preced_th], simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    65
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    66
text {*
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    67
  Since @{term th} holds the highest precedence and @{text "cp"}
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    68
  is the highest precedence of all threads in the sub-tree of 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    69
  @{text "th"} and @{text th} is among these threads, 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    70
  its @{term cp} must equal to its precedence:
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    71
*}
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
    72
lemma eq_cp_s_th: "cp s th = preced th s" (is "?L = ?R")
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    73
proof -
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
    74
  have "?L \<le> ?R"
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
    75
  by (unfold highest, rule Max_ge, 
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
    76
        auto simp:threads_s finite_threads)
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
    77
  moreover have "?R \<le> ?L"
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
    78
    by (unfold vat_s.cp_rec, rule Max_ge, 
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
    79
        auto simp:the_preced_def vat_s.fsbttRAGs.finite_children)
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
    80
  ultimately show ?thesis by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    81
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    82
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    83
lemma highest_cp_preced: "cp s th = Max (the_preced s ` threads s)"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    84
  using eq_cp_s_th highest max_cp_eq the_preced_def by presburger
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    85
  
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    87
lemma highest_preced_thread: "preced th s = Max (the_preced s ` threads s)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    88
  by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    90
lemma highest': "cp s th = Max (cp s ` threads s)"
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
    91
  by (simp add: eq_cp_s_th highest)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    92
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    93
end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    94
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    95
locale extend_highest_gen = highest_gen + 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    96
  fixes t 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    97
  assumes vt_t: "vt (t@s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    98
  and create_low: "Create th' prio' \<in> set t \<Longrightarrow> prio' \<le> prio"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    99
  and set_diff_low: "Set th' prio' \<in> set t \<Longrightarrow> th' \<noteq> th \<and> prio' \<le> prio"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   100
  and exit_diff: "Exit th' \<in> set t \<Longrightarrow> th' \<noteq> th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   101
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   102
sublocale extend_highest_gen < vat_t: valid_trace "t@s"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   103
  by (unfold_locales, insert vt_t, simp)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   104
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   105
lemma step_back_vt_app: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   106
  assumes vt_ts: "vt (t@s)" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   107
  shows "vt s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   108
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   109
  from vt_ts show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   110
  proof(induct t)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   111
    case Nil
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   112
    from Nil show ?case by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   113
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   114
    case (Cons e t)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   115
    assume ih: " vt (t @ s) \<Longrightarrow> vt s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   116
      and vt_et: "vt ((e # t) @ s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   117
    show ?case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   118
    proof(rule ih)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   119
      show "vt (t @ s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   120
      proof(rule step_back_vt)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   121
        from vt_et show "vt (e # t @ s)" by simp
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   122
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   123
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   124
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   125
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   126
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 69
diff changeset
   127
(* locale red_extend_highest_gen = extend_highest_gen +
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   128
   fixes i::nat
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 69
diff changeset
   129
*)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   130
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 69
diff changeset
   131
(*
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   132
sublocale red_extend_highest_gen <   red_moment: extend_highest_gen "s" "th" "prio" "tm" "(moment i t)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   133
  apply (insert extend_highest_gen_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   134
  apply (unfold extend_highest_gen_def extend_highest_gen_axioms_def, clarsimp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   135
  by (unfold highest_gen_def, auto dest:step_back_vt_app)
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 69
diff changeset
   136
*)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   137
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   138
context extend_highest_gen
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   139
begin
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   140
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   141
 lemma ind [consumes 0, case_names Nil Cons, induct type]:
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   142
  assumes 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   143
    h0: "R []"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   144
  and h2: "\<And> e t. \<lbrakk>vt (t@s); step (t@s) e; 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   145
                    extend_highest_gen s th prio tm t; 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   146
                    extend_highest_gen s th prio tm (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   147
  shows "R t"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   149
  from vt_t extend_highest_gen_axioms show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   150
  proof(induct t)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   151
    from h0 show "R []" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   152
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   153
    case (Cons e t')
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   154
    assume ih: "\<lbrakk>vt (t' @ s); extend_highest_gen s th prio tm t'\<rbrakk> \<Longrightarrow> R t'"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   155
      and vt_e: "vt ((e # t') @ s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   156
      and et: "extend_highest_gen s th prio tm (e # t')"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   157
    from vt_e and step_back_step have stp: "step (t'@s) e" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   158
    from vt_e and step_back_vt have vt_ts: "vt (t'@s)" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   159
    show ?case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   160
    proof(rule h2 [OF vt_ts stp _ _ _ ])
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   161
      show "R t'"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   162
      proof(rule ih)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   163
        from et show ext': "extend_highest_gen s th prio tm t'"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   164
          by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   165
      next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   166
        from vt_ts show "vt (t' @ s)" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   167
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   168
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   169
      from et show "extend_highest_gen s th prio tm (e # t')" .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   171
      from et show ext': "extend_highest_gen s th prio tm t'"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   172
          by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   173
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   174
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   175
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   176
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   177
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   178
lemma th_kept: "th \<in> threads (t @ s) \<and> 
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   179
                 preced th (t@s) = preced th s" (is "?Q t") 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   180
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   181
  show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   182
  proof(induct rule:ind)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   183
    case Nil
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   184
    from threads_s
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   185
    show ?case
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   186
      by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   187
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   188
    case (Cons e t)
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   189
    interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   190
    interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   191
    show ?case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   192
    proof(cases e)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   193
      case (Create thread prio)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   194
      show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   195
      proof -
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   196
        from Cons and Create have "step (t@s) (Create thread prio)" by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   197
        hence "th \<noteq> thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   198
        proof(cases)
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   199
          case thread_create
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   200
          with Cons 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
        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   203
          by (unfold Create, auto simp:preced_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   204
        moreover note Cons
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   205
        ultimately show ?thesis
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   206
          by (auto simp:Create)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   207
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   208
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   209
      case (Exit thread)
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   210
      from h_e.exit_diff and Exit
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   211
      have neq_th: "thread \<noteq> th" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   212
      with Cons
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   213
      show ?thesis
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   214
        by (unfold Exit, auto simp:preced_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   215
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   216
      case (P thread cs)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   217
      with Cons
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   218
      show ?thesis 
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   219
        by (auto simp:P preced_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   220
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   221
      case (V thread cs)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   222
      with Cons
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   223
      show ?thesis 
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   224
        by (auto simp:V preced_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   225
    next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   226
      case (Set thread prio')
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   227
      show ?thesis
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   228
      proof -
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   229
        from h_e.set_diff_low and Set
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   230
        have "th \<noteq> thread" by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   231
        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   232
          by (unfold Set, auto simp:preced_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   233
        moreover note Cons
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   234
        ultimately show ?thesis
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   235
          by (auto simp:Set)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   236
      qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   237
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   238
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   239
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   240
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   241
text {*
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   242
  According to @{thm th_kept}, thread @{text "th"} has its living status
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   243
  and precedence kept along the way of @{text "t"}. The following lemma
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   244
  shows that this preserved precedence of @{text "th"} remains as the highest
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   245
  along the way of @{text "t"}.
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   246
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   247
  The proof goes by induction over @{text "t"} using the specialized
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   248
  induction rule @{thm ind}, followed by case analysis of each possible 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   249
  operations of PIP. All cases follow the same pattern rendered by the 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   250
  generalized introduction rule @{thm "image_Max_eqI"}. 
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
  The very essence is to show that precedences, no matter whether they 
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   253
  are newly introduced or modified, are always lower than the one held 
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   254
  by @{term "th"}, which by @{thm th_kept} is preserved along the way.
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   255
*}
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   256
lemma max_kept: "Max (the_preced (t @ s) ` (threads (t@s))) = preced th s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   257
proof(induct rule:ind)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   258
  case Nil
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   259
  from highest_preced_thread
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   260
  show ?case by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   261
next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   262
  case (Cons e t)
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   263
    interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   264
    interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   265
  show ?case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   266
  proof(cases e)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   267
    case (Create thread prio')
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   268
    show ?thesis (is "Max (?f ` ?A) = ?t")
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   269
    proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   270
      -- {* The following is the common pattern of each branch of the case analysis. *}
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   271
      -- {* The major part is to show that @{text "th"} holds the highest precedence: *}
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   272
      have "Max (?f ` ?A) = ?f th"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   273
      proof(rule image_Max_eqI)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   274
        show "finite ?A" using h_e.finite_threads by auto 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   275
      next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   276
        show "th \<in> ?A" using h_e.th_kept by auto 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   277
      next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   278
        show "\<forall>x\<in>?A. ?f x \<le> ?f th"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   279
        proof 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   280
          fix x
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   281
          assume "x \<in> ?A"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   282
          hence "x = thread \<or> x \<in> threads (t@s)" by (auto simp:Create)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   283
          thus "?f x \<le> ?f th"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   284
          proof
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   285
            assume "x = thread"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   286
            thus ?thesis 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   287
              apply (simp add:Create the_preced_def preced_def, fold preced_def)
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   288
              using Create h_e.create_low h_t.th_kept lt_tm preced_leI2 
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   289
              preced_th by force
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   290
          next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   291
            assume h: "x \<in> threads (t @ s)"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   292
            from Cons(2)[unfolded Create] 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   293
            have "x \<noteq> thread" using h by (cases, auto)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   294
            hence "?f x = the_preced (t@s) x" 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   295
              by (simp add:Create the_preced_def preced_def)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   296
            hence "?f x \<le> Max (the_preced (t@s) ` threads (t@s))"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   297
              by (simp add: h_t.finite_threads h)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   298
            also have "... = ?f th"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   299
              by (metis Cons.hyps(5) h_e.th_kept the_preced_def) 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   300
            finally show ?thesis .
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   301
          qed
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   302
        qed
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   303
      qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   304
     -- {* The minor part is to show that the precedence of @{text "th"} 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   305
           equals to preserved one, given by the foregoing lemma @{thm th_kept} *}
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   306
      also have "... = ?t" using h_e.th_kept the_preced_def by auto
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   307
      -- {* Then it follows trivially that the precedence preserved
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   308
            for @{term "th"} remains the maximum of all living threads along the way. *}
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   309
      finally show ?thesis .
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   310
    qed 
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   311
  next 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   312
    case (Exit thread)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   313
    show ?thesis (is "Max (?f ` ?A) = ?t")
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   314
    proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   315
      have "Max (?f ` ?A) = ?f th"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   316
      proof(rule image_Max_eqI)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   317
        show "finite ?A" using h_e.finite_threads by auto 
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   318
      next
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   319
        show "th \<in> ?A" using h_e.th_kept by auto 
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   320
      next
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   321
        show "\<forall>x\<in>?A. ?f x \<le> ?f th"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   322
        proof 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   323
          fix x
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   324
          assume "x \<in> ?A"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   325
          hence "x \<in> threads (t@s)" by (simp add: Exit) 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   326
          hence "?f x \<le> Max (?f ` threads (t@s))" 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   327
            by (simp add: h_t.finite_threads) 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   328
          also have "... \<le> ?f th" 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   329
            apply (simp add:Exit the_preced_def preced_def, fold preced_def)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   330
            using Cons.hyps(5) h_t.th_kept the_preced_def by auto
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   331
          finally show "?f x \<le> ?f th" .
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   332
        qed
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   333
      qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   334
      also have "... = ?t" using h_e.th_kept the_preced_def by auto
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   335
      finally show ?thesis .
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   336
    qed 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   337
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   338
    case (P thread cs)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   339
    with Cons
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   340
    show ?thesis by (auto simp:preced_def the_preced_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   341
  next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   342
    case (V thread cs)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   343
    with Cons
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 35
diff changeset
   344
    show ?thesis by (auto simp:preced_def the_preced_def)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   345
  next 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   346
    case (Set thread prio')
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   347
    show ?thesis (is "Max (?f ` ?A) = ?t")
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   348
    proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   349
      have "Max (?f ` ?A) = ?f th"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   350
      proof(rule image_Max_eqI)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   351
        show "finite ?A" using h_e.finite_threads by auto 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   352
      next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   353
        show "th \<in> ?A" using h_e.th_kept by auto 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   354
      next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   355
        show "\<forall>x\<in>?A. ?f x \<le> ?f th"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   356
        proof 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   357
          fix x
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   358
          assume h: "x \<in> ?A"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   359
          show "?f x \<le> ?f th"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   360
          proof(cases "x = thread")
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   361
            case True
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   362
            moreover have "the_preced (Set thread prio' # t @ s) thread \<le> the_preced (t @ s) th"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   363
            proof -
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   364
              have "the_preced (t @ s) th = Prc prio tm"  
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   365
                using h_t.th_kept preced_th by (simp add:the_preced_def)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   366
              moreover have "prio' \<le> prio" using Set h_e.set_diff_low by auto
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   367
              ultimately show ?thesis by (insert lt_tm, auto simp:the_preced_def preced_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   368
            qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   369
            ultimately show ?thesis
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   370
              by (unfold Set, simp add:the_preced_def preced_def)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   371
          next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   372
            case False
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   373
            then have "?f x  = the_preced (t@s) x"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   374
              by (simp add:the_preced_def preced_def Set)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   375
            also have "... \<le> Max (the_preced (t@s) ` threads (t@s))"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   376
              using Set h h_t.finite_threads by auto 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   377
            also have "... = ?f th" by (metis Cons.hyps(5) h_e.th_kept the_preced_def) 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   378
            finally show ?thesis .
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   379
          qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   380
        qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   381
      qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   382
      also have "... = ?t" using h_e.th_kept the_preced_def by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   383
      finally show ?thesis .
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   384
    qed 
0
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
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   387
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   388
lemma max_preced: "preced th (t@s) = Max (the_preced (t@s) ` (threads (t@s)))"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   389
  by (insert th_kept max_kept, auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   390
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   391
text {*
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   392
  The reason behind the following lemma is that:
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   393
  Since @{term "cp"} is defined as the maximum precedence 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   394
  of those threads contained in the sub-tree of node @{term "Th th"} 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   395
  in @{term "RAG (t@s)"}, and all these threads are living threads, and 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   396
  @{term "th"} is also among them, the maximum precedence of 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   397
  them all must be the one for @{text "th"}.
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   398
*}
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   399
lemma th_cp_max_preced: 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   400
  "cp (t@s) th = Max (the_preced (t@s) ` (threads (t@s)))" (is "?L = ?R") 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   401
proof -
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   402
  let ?f = "the_preced (t@s)"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   403
  have "?L = ?f th"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   404
  proof(unfold cp_alt_def, rule image_Max_eqI)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   405
    show "finite {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   406
    proof -
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   407
      have "{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)} = 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   408
            the_thread ` {n . n \<in> subtree (RAG (t @ s)) (Th th) \<and>
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   409
                            (\<exists> th'. n = Th th')}"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   410
      by (smt Collect_cong Setcompr_eq_image mem_Collect_eq the_thread.simps)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   411
      moreover have "finite ..." by (simp add: vat_t.fsbtRAGs.finite_subtree) 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   412
      ultimately show ?thesis by simp
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   413
    qed
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   414
  next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   415
    show "th \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   416
      by (auto simp:subtree_def)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   417
  next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   418
    show "\<forall>x\<in>{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}.
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   419
               the_preced (t @ s) x \<le> the_preced (t @ s) th"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   420
    proof
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   421
      fix th'
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   422
      assume "th' \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   423
      hence "Th th' \<in> subtree (RAG (t @ s)) (Th th)" by auto
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   424
      moreover have "... \<subseteq> Field (RAG (t @ s)) \<union> {Th th}"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   425
        by (meson subtree_Field)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   426
      ultimately have "Th th' \<in> ..." by auto
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   427
      hence "th' \<in> threads (t@s)" 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   428
      proof
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   429
        assume "Th th' \<in> {Th th}"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   430
        thus ?thesis using th_kept by auto 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   431
      next
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   432
        assume "Th th' \<in> Field (RAG (t @ s))"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   433
        thus ?thesis using vat_t.not_in_thread_isolated by blast 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   434
      qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   435
      thus "the_preced (t @ s) th' \<le> the_preced (t @ s) th"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   436
        by (metis Max_ge finite_imageI finite_threads image_eqI 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   437
               max_kept th_kept the_preced_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   438
    qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   439
  qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   440
  also have "... = ?R" by (simp add: max_preced the_preced_def) 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   441
  finally show ?thesis .
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   442
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   443
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   444
lemma th_cp_max[simp]: "Max (cp (t@s) ` threads (t@s)) = cp (t@s) th"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   445
  using max_cp_eq th_cp_max_preced the_preced_def vt_t by presburger
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   446
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   447
lemma [simp]: "Max (cp (t@s) ` threads (t@s)) = Max (the_preced (t@s) ` threads (t@s))"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   448
  by (simp add: th_cp_max_preced)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   449
  
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   450
lemma [simp]: "Max (the_preced (t@s) ` threads (t@s)) = the_preced (t@s) th"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   451
  using max_kept th_kept the_preced_def by auto
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   452
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   453
lemma [simp]: "the_preced (t@s) th = preced th (t@s)"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   454
  using the_preced_def by auto
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   455
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   456
lemma [simp]: "preced th (t@s) = preced th s"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   457
  by (simp add: th_kept)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   458
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   459
lemma [simp]: "cp s th = preced th s"
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   460
  by (simp add: eq_cp_s_th)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   461
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   462
lemma th_cp_preced [simp]: "cp (t@s) th = preced th s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   463
  by (fold max_kept, unfold th_cp_max_preced, simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   464
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   465
lemma preced_less:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   466
  assumes th'_in: "th' \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   467
  and neq_th': "th' \<noteq> th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   468
  shows "preced th' s < preced th s"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   469
  using assms
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   470
by (metis Max.coboundedI finite_imageI highest not_le order.trans 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   471
    preced_linorder rev_image_eqI threads_s vat_s.finite_threads 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   472
    vat_s.le_cp)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   473
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   474
section {* The `blocking thread` *}
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   475
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   476
text {* 
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   477
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   478
  The purpose of PIP is to ensure that the most urgent thread @{term
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   479
  th} is not blocked unreasonably. Therefore, below, we will derive
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   480
  properties of the blocking thread. By blocking thread, we mean a
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   481
  thread in running state t @ s, but is different from thread @{term
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   482
  th}.
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   483
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   484
  The first lemmas shows that the @{term cp}-value of the blocking
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   485
  thread @{text th'} equals to the highest precedence in the whole
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   486
  system.
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   487
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   488
*}
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   489
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   490
lemma runing_preced_inversion:
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   491
  assumes runing': "th' \<in> runing (t @ s)"
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   492
  shows "cp (t @ s) th' = preced th s" 
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   493
proof -
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   494
  have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))" 
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   495
    using assms by (unfold runing_def, auto)
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   496
  also have "\<dots> = preced th s"
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   497
    by (metis th_cp_max th_cp_preced vat_t.max_cp_readys_threads) 
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   498
  finally show ?thesis .
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   499
qed
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   500
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   501
text {*
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   502
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   503
  The next lemma shows how the counters for @{term "P"} and @{term
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   504
  "V"} operations relate to the running threads in the states @{term
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   505
  s} and @{term "t @ s"}: if a thread's @{term "P"}-count equals its
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   506
  @{term "V"}-count (which means it no longer has any resource in its
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   507
  possession), it cannot be a running thread.
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   508
76
b6ea51cd2e88 some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   509
  The proof is by contraction with the assumption @{text "th' \<noteq> th"}.
b6ea51cd2e88 some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   510
  The key is the use of @{thm count_eq_dependants} to derive the
b6ea51cd2e88 some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   511
  emptiness of @{text th'}s @{term dependants}-set from the balance of
b6ea51cd2e88 some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   512
  its @{term P} and @{term V} counts.  From this, it can be shown
b6ea51cd2e88 some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   513
  @{text th'}s @{term cp}-value equals to its own precedence.
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   514
76
b6ea51cd2e88 some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   515
  On the other hand, since @{text th'} is running, by @{thm
b6ea51cd2e88 some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   516
  runing_preced_inversion}, its @{term cp}-value equals to the
b6ea51cd2e88 some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   517
  precedence of @{term th}.
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   518
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   519
  Combining the above two results we have that @{text th'} and @{term
76
b6ea51cd2e88 some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   520
  th} have the same precedence. By uniqueness of precedences, we have
b6ea51cd2e88 some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   521
  @{text "th' = th"}, which is in contradiction with the assumption
b6ea51cd2e88 some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   522
  @{text "th' \<noteq> th"}.
b6ea51cd2e88 some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   523
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   524
*} 
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   525
                      
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   526
lemma eq_pv_blocked: (* ddd *)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   527
  assumes neq_th': "th' \<noteq> th"
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   528
  and eq_pv: "cntP (t @ s) th' = cntV (t @ s) th'"
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   529
  shows "th' \<notin> runing (t @ s)"
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   530
proof
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   531
  assume otherwise: "th' \<in> runing (t @ s)"
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   532
  show False
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   533
  proof -
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   534
    have th'_in: "th' \<in> threads (t @ s)"
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   535
        using otherwise readys_threads runing_def by auto 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   536
    have "th' = th"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   537
    proof(rule preced_unique)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   538
      -- {* The proof goes like this: 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   539
            it is first shown that the @{term preced}-value of @{term th'} 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   540
            equals to that of @{term th}, then by uniqueness 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   541
            of @{term preced}-values (given by lemma @{thm preced_unique}), 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   542
            @{term th'} equals to @{term th}: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   543
      show "preced th' (t @ s) = preced th (t @ s)" (is "?L = ?R")
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   544
      proof -
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   545
        -- {* Since the counts of @{term th'} are balanced, the subtree
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   546
              of it contains only itself, so, its @{term cp}-value
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   547
              equals its @{term preced}-value: *}
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   548
        have "?L = cp (t @ s) th'"
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   549
          by (unfold cp_eq_cpreced cpreced_def count_eq_dependants[OF eq_pv], simp)
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   550
        -- {* Since @{term "th'"} is running, by @{thm runing_preced_inversion},
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   551
              its @{term cp}-value equals @{term "preced th s"}, 
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   552
              which equals to @{term "?R"} by simplification: *}
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   553
        also have "... = ?R" 
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   554
            using runing_preced_inversion[OF otherwise] by simp
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   555
        finally show ?thesis .
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   556
      qed
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   557
    qed (auto simp: th'_in th_kept)
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 67
diff changeset
   558
    with `th' \<noteq> th` show ?thesis by simp
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   559
 qed
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   560
qed
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   561
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   562
text {*
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   563
  The following lemma is the extrapolation of @{thm eq_pv_blocked}.
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   564
  It says if a thread, different from @{term th}, 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   565
  does not hold any resource at the very beginning,
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   566
  it will keep hand-emptied in the future @{term "t@s"}.
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   567
*}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   568
lemma eq_pv_persist: (* ddd *)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   569
  assumes neq_th': "th' \<noteq> th"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   570
  and eq_pv: "cntP s th' = cntV s th'"
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   571
  shows "cntP (t @ s) th' = cntV (t @ s) th'"
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   572
proof(induction rule: ind) 
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   573
  -- {* The nontrivial case is for the @{term Cons}: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   574
  case (Cons e t)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   575
  -- {* All results derived so far hold for both @{term s} and @{term "t@s"}: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   576
  interpret vat_t: extend_highest_gen s th prio tm t using Cons by simp
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   577
  interpret vat_e: extend_highest_gen s th prio tm "(e # t)" using Cons by simp
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   578
  show ?case
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   579
  proof -
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   580
    -- {* It can be proved that @{term cntP}-value of @{term th'} does not change
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   581
          by the happening of event @{term e}: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   582
    have "cntP ((e#t)@s) th' = cntP (t@s) th'"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   583
    proof(rule ccontr) -- {* Proof by contradiction. *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   584
      -- {* Suppose @{term cntP}-value of @{term th'} is changed by @{term e}: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   585
      assume otherwise: "cntP ((e # t) @ s) th' \<noteq> cntP (t @ s) th'"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   586
      -- {* Then the actor of @{term e} must be @{term th'} and @{term e}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   587
            must be a @{term P}-event: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   588
      hence "isP e" "actor e = th'" by (auto simp:cntP_diff_inv) 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   589
      with vat_t.actor_inv[OF Cons(2)]
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   590
      -- {* According to @{thm actor_inv}, @{term th'} must be running at 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   591
            the moment @{term "t@s"}: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   592
      have "th' \<in> runing (t@s)" by (cases e, auto)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   593
      -- {* However, an application of @{thm eq_pv_blocked} to induction hypothesis
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   594
            shows @{term th'} can not be running at moment  @{term "t@s"}: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   595
      moreover have "th' \<notin> runing (t@s)" 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   596
               using vat_t.eq_pv_blocked[OF neq_th' Cons(5)] .
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   597
      -- {* Contradiction is finally derived: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   598
      ultimately show False by simp
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   599
    qed
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   600
    -- {* It can also be proved that @{term cntV}-value of @{term th'} does not change
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   601
          by the happening of event @{term e}: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   602
    -- {* The proof follows exactly the same pattern as the case for @{term cntP}-value: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   603
    moreover have "cntV ((e#t)@s) th' = cntV (t@s) th'"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   604
    proof(rule ccontr) -- {* Proof by contradiction. *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   605
      assume otherwise: "cntV ((e # t) @ s) th' \<noteq> cntV (t @ s) th'"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   606
      hence "isV e" "actor e = th'" by (auto simp:cntV_diff_inv) 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   607
      with vat_t.actor_inv[OF Cons(2)]
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   608
      have "th' \<in> runing (t@s)" by (cases e, auto)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   609
      moreover have "th' \<notin> runing (t@s)"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   610
          using vat_t.eq_pv_blocked[OF neq_th' Cons(5)] .
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   611
      ultimately show False by simp
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   612
    qed
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   613
    -- {* Finally, it can be shown that the @{term cntP} and @{term cntV} 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   614
          value for @{term th'} are still in balance, so @{term th'} 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   615
          is still hand-emptied after the execution of event @{term e}: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   616
    ultimately show ?thesis using Cons(5) by metis
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   617
  qed
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   618
qed (auto simp:eq_pv)
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   619
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   620
text {*
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   621
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   622
  By combining @{thm eq_pv_blocked} and @{thm eq_pv_persist}, it can
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   623
  be derived easily that @{term th'} can not be running in the future:
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   624
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   625
*}
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   626
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   627
lemma eq_pv_blocked_persist:
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   628
  assumes neq_th': "th' \<noteq> th"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   629
  and eq_pv: "cntP s th' = cntV s th'"
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   630
  shows "th' \<notin> runing (t @ s)"
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   631
  using assms
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   632
  by (simp add: eq_pv_blocked eq_pv_persist) 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   633
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   634
text {*
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   635
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   636
  The following lemma shows the blocking thread @{term th'} must hold
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   637
  some resource in the very beginning.
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   638
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   639
*}
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   640
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   641
lemma runing_cntP_cntV_inv: (* ddd *)
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   642
  assumes is_runing: "th' \<in> runing (t @ s)"
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   643
  and neq_th': "th' \<noteq> th"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   644
  shows "cntP s th' > cntV s th'"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   645
  using assms
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   646
proof -
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   647
  -- {* First, it can be shown that the number of @{term P} and
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   648
        @{term V} operations can not be equal for thred @{term th'} *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   649
  have "cntP s th' \<noteq> cntV s th'"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   650
  proof
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   651
     -- {* The proof goes by contradiction, suppose otherwise: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   652
    assume otherwise: "cntP s th' = cntV s th'"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   653
    -- {* By applying @{thm  eq_pv_blocked_persist} to this: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   654
    from eq_pv_blocked_persist[OF neq_th' otherwise] 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   655
    -- {* we have that @{term th'} can not be running at moment @{term "t@s"}: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   656
    have "th' \<notin> runing (t@s)" .
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   657
    -- {* This is obvious in contradiction with assumption @{thm is_runing}  *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   658
    thus False using is_runing by simp
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   659
  qed
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   660
  -- {* However, the number of @{term V} is always less or equal to @{term P}: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   661
  moreover have "cntV s th' \<le> cntP s th'" using vat_s.cnp_cnv_cncs by auto
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   662
  -- {* Thesis is finally derived by combining the these two results: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   663
  ultimately show ?thesis by auto
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   664
qed
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   665
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   666
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   667
text {*
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   668
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   669
  The following lemmas shows the blocking thread @{text th'} must be
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   670
  live at the very beginning, i.e. the moment (or state) @{term s}.
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   671
  The proof is a  simple combination of the results above:
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   672
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   673
*}
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   674
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   675
lemma runing_threads_inv: 
66
2af87bb52fca Some small improvements in Correctness.thy.
zhangx
parents: 65
diff changeset
   676
  assumes runing': "th' \<in> runing (t@s)"
2af87bb52fca Some small improvements in Correctness.thy.
zhangx
parents: 65
diff changeset
   677
  and neq_th': "th' \<noteq> th"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   678
  shows "th' \<in> threads s"
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   679
proof(rule ccontr) -- {* Proof by contradiction: *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   680
  assume otherwise: "th' \<notin> threads s" 
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   681
  have "th' \<notin> runing (t @ s)"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   682
  proof -
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   683
    from vat_s.cnp_cnv_eq[OF otherwise]
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   684
    have "cntP s th' = cntV s th'" .
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   685
    from eq_pv_blocked_persist[OF neq_th' this]
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   686
    show ?thesis .
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   687
  qed
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   688
  with runing' show False by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   689
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   690
66
2af87bb52fca Some small improvements in Correctness.thy.
zhangx
parents: 65
diff changeset
   691
text {*
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   692
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   693
  The following lemma summarises the above lemmas to give an overall
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   694
  characterisationof the blocking thread @{text "th'"}:
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   695
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   696
*}
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   697
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   698
lemma runing_inversion: (* ddd, one of the main lemmas to present *)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   699
  assumes runing': "th' \<in> runing (t@s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   700
  and neq_th: "th' \<noteq> th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   701
  shows "th' \<in> threads s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   702
  and    "\<not>detached s th'"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   703
  and    "cp (t@s) th' = preced th s"
66
2af87bb52fca Some small improvements in Correctness.thy.
zhangx
parents: 65
diff changeset
   704
proof -
2af87bb52fca Some small improvements in Correctness.thy.
zhangx
parents: 65
diff changeset
   705
  from runing_threads_inv[OF assms]
2af87bb52fca Some small improvements in Correctness.thy.
zhangx
parents: 65
diff changeset
   706
  show "th' \<in> threads s" .
2af87bb52fca Some small improvements in Correctness.thy.
zhangx
parents: 65
diff changeset
   707
next
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   708
  from runing_cntP_cntV_inv[OF runing' neq_th]
66
2af87bb52fca Some small improvements in Correctness.thy.
zhangx
parents: 65
diff changeset
   709
  show "\<not>detached s th'" using vat_s.detached_eq by simp
2af87bb52fca Some small improvements in Correctness.thy.
zhangx
parents: 65
diff changeset
   710
next
2af87bb52fca Some small improvements in Correctness.thy.
zhangx
parents: 65
diff changeset
   711
  from runing_preced_inversion[OF runing']
2af87bb52fca Some small improvements in Correctness.thy.
zhangx
parents: 65
diff changeset
   712
  show "cp (t@s) th' = preced th s" .
2af87bb52fca Some small improvements in Correctness.thy.
zhangx
parents: 65
diff changeset
   713
qed
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   714
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   715
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   716
section {* The existence of `blocking thread` *}
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   717
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   718
text {* 
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   719
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   720
  Suppose @{term th} is not running, it is first shown that there is a
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   721
  path in RAG leading from node @{term th} to another thread @{text
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   722
  "th'"} in the @{term readys}-set (So @{text "th'"} is an ancestor of
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   723
  @{term th}}).
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   724
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   725
  Now, since @{term readys}-set is non-empty, there must be one in it
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   726
  which holds the highest @{term cp}-value, which, by definition, is
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   727
  the @{term runing}-thread. However, we are going to show more: this
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   728
  running thread is exactly @{term "th'"}.
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   729
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   730
*}
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   731
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 66
diff changeset
   732
lemma th_blockedE: (* ddd, the other main lemma to be presented: *)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   733
  assumes "th \<notin> runing (t@s)"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   734
  obtains th' where "Th th' \<in> ancestors (RAG (t @ s)) (Th th)"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   735
                    "th' \<in> runing (t@s)"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   736
proof -
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   737
  -- {* According to @{thm vat_t.th_chain_to_ready}, either 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   738
        @{term "th"} is in @{term "readys"} or there is path leading from it to 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   739
        one thread in @{term "readys"}. *}
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   740
  have "th \<in> readys (t @ s) \<or> (\<exists>th'. th' \<in> readys (t @ s) \<and> (Th th, Th th') \<in> (RAG (t @ s))\<^sup>+)" 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   741
    using th_kept vat_t.th_chain_to_ready by auto
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   742
  -- {* However, @{term th} can not be in @{term readys}, because otherwise, since 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   743
       @{term th} holds the highest @{term cp}-value, it must be @{term "runing"}. *}
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   744
  moreover have "th \<notin> readys (t@s)" 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   745
    using assms runing_def th_cp_max vat_t.max_cp_readys_threads by auto 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   746
  -- {* So, there must be a path from @{term th} to another thread @{text "th'"} in 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   747
        term @{term readys}: *}
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   748
  ultimately obtain th' where th'_in: "th' \<in> readys (t@s)"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   749
                          and dp: "(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+" by auto
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   750
  -- {* We are going to show that this @{term th'} is running. *}
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   751
  have "th' \<in> runing (t@s)"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   752
  proof -
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   753
    -- {* We only need to show that this @{term th'} holds the highest @{term cp}-value: *}
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   754
    have "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" (is "?L = ?R")
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   755
    proof -
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   756
      have "?L =  Max ((the_preced (t @ s) \<circ> the_thread) ` subtree (tRAG (t @ s)) (Th th'))"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   757
        by (unfold cp_alt_def1, simp)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   758
      also have "... = (the_preced (t @ s) \<circ> the_thread) (Th th)"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   759
      proof(rule image_Max_subset)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   760
        show "finite (Th ` (threads (t@s)))" by (simp add: vat_t.finite_threads)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   761
      next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   762
        show "subtree (tRAG (t @ s)) (Th th') \<subseteq> Th ` threads (t @ s)"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   763
          by (metis Range.intros dp trancl_range vat_t.range_in vat_t.subtree_tRAG_thread) 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   764
      next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   765
        show "Th th \<in> subtree (tRAG (t @ s)) (Th th')" using dp
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   766
                    by (unfold tRAG_subtree_eq, auto simp:subtree_def)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   767
      next
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   768
        show "Max ((the_preced (t @ s) \<circ> the_thread) ` Th ` threads (t @ s)) =
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   769
                      (the_preced (t @ s) \<circ> the_thread) (Th th)" (is "Max ?L = _")
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   770
        proof -
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   771
          have "?L = the_preced (t @ s) `  threads (t @ s)" 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   772
                     by (unfold image_comp, rule image_cong, auto)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   773
          thus ?thesis using max_preced the_preced_def by auto
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   774
        qed
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   775
      qed
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   776
      also have "... = ?R"
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   777
        using th_cp_max th_cp_preced th_kept 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   778
              the_preced_def vat_t.max_cp_readys_threads by auto
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   779
      finally show ?thesis .
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   780
    qed 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   781
    -- {* Now, since @{term th'} holds the highest @{term cp} 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   782
          and we have already show it is in @{term readys},
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   783
          it is @{term runing} by definition. *}
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   784
    with `th' \<in> readys (t@s)` show ?thesis by (simp add: runing_def) 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   785
  qed
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   786
  -- {* It is easy to show @{term th'} is an ancestor of @{term th}: *}
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   787
  moreover have "Th th' \<in> ancestors (RAG (t @ s)) (Th th)" 
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   788
    using `(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+` by (auto simp:ancestors_def)
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   789
  ultimately show ?thesis using that by metis
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   790
qed
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   791
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   792
text {*
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   793
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   794
  Now it is easy to see there is always a thread to run by case
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   795
  analysis on whether thread @{term th} is running: if the answer is
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   796
  yes, the the running thread is obviously @{term th} itself;
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   797
  otherwise, the running thread is the @{text th'} given by lemma
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   798
  @{thm th_blockedE}.
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   799
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   800
*}
82
c0a4e840aefe some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   801
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   802
lemma live: "runing (t@s) \<noteq> {}"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   803
proof(cases "th \<in> runing (t@s)") 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   804
  case True thus ?thesis by auto
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   805
next
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   806
  case False
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
   807
  thus ?thesis using th_blockedE by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   808
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   809
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 69
diff changeset
   810
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   811
end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   812
end