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