Attic/ExtGG_1.thy
author zhangx
Thu, 28 Jan 2016 21:14:17 +0800
changeset 90 ed938e2246b9
parent 1 c4783e4ef43f
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.

theory ExtGG
imports PrioG
begin

lemma birth_time_lt:  "s \<noteq> [] \<Longrightarrow> birthtime th s < length s"
  apply (induct s, simp)
proof -
  fix a s
  assume ih: "s \<noteq> [] \<Longrightarrow> birthtime th s < length s"
    and eq_as: "a # s \<noteq> []"
  show "birthtime th (a # s) < length (a # s)"
  proof(cases "s \<noteq> []")
    case False
    from False show ?thesis
      by (cases a, auto simp:birthtime.simps)
  next
    case True
    from ih [OF True] show ?thesis
      by (cases a, auto simp:birthtime.simps)
  qed
qed

lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []"
  by (induct s, auto simp:threads.simps)

lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s"
  apply (drule_tac th_in_ne)
  by (unfold preced_def, auto intro: birth_time_lt)

locale highest_gen =
  fixes s' th s e' prio tm
  defines s_def : "s \<equiv> (e'#s')"
  assumes vt_s: "vt step s"
  and threads_s: "th \<in> threads s"
  and highest: "preced th s = Max ((cp s)`threads s)"
  and nh: "preced th s' \<noteq> Max ((cp s)`threads s')"
  and preced_th: "preced th s = Prc prio tm"

context highest_gen
begin

lemma lt_tm: "tm < length s"
  by (insert preced_tm_lt[OF threads_s preced_th], simp)

lemma vt_s': "vt step s'"
  by (insert vt_s, unfold s_def, drule_tac step_back_vt, simp)

lemma eq_cp_s_th: "cp s th = preced th s"
proof -
  from highest and max_cp_eq[OF vt_s]
  have is_max: "preced th s = Max ((\<lambda>th. preced th s) ` threads s)" by simp
  have sbs: "({th} \<union> dependents (wq s) th) \<subseteq> threads s"
  proof -
    from threads_s and dependents_threads[OF vt_s, of th]
    show ?thesis by auto
  qed
  show ?thesis
  proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
    show "preced th s \<in> (\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)" by simp
  next
    fix y 
    assume "y \<in> (\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)"
    then obtain th1 where th1_in: "th1 \<in> ({th} \<union> dependents (wq s) th)"
      and eq_y: "y = preced th1 s" by auto
    show "y \<le> preced th s"
    proof(unfold is_max, rule Max_ge)
      from finite_threads[OF vt_s] 
      show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
    next
      from sbs th1_in and eq_y 
      show "y \<in> (\<lambda>th. preced th s) ` threads s" by auto
    qed
  next
    from sbs and finite_threads[OF vt_s]
    show "finite ((\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th))"
      by (auto intro:finite_subset)
  qed
qed

lemma highest_cp_preced: "cp s th = Max ((\<lambda> th'. preced th' s) ` threads s)"
  by (fold max_cp_eq[OF vt_s], unfold eq_cp_s_th, insert highest, simp)

lemma highest_preced_thread: "preced th s = Max ((\<lambda> th'. preced th' s) ` threads s)"
  by (fold eq_cp_s_th, unfold highest_cp_preced, simp)

lemma highest': "cp s th = Max (cp s ` threads s)"
proof -
  from highest_cp_preced max_cp_eq[OF vt_s, symmetric]
  show ?thesis by simp
qed

end

locale extend_highest_gen = highest_gen + 
  fixes t 
  assumes vt_t: "vt step (t@s)"
  and create_low: "Create th' prio' \<in> set t \<Longrightarrow> prio' \<le> prio"
  and set_diff_low: "Set th' prio' \<in> set t \<Longrightarrow> th' \<noteq> th \<and> prio' \<le> prio"
  and exit_diff: "Exit th' \<in> set t \<Longrightarrow> th' \<noteq> th"

lemma step_back_vt_app: 
  assumes vt_ts: "vt cs (t@s)" 
  shows "vt cs s"
proof -
  from vt_ts show ?thesis
  proof(induct t)
    case Nil
    from Nil show ?case by auto
  next
    case (Cons e t)
    assume ih: " vt cs (t @ s) \<Longrightarrow> vt cs s"
      and vt_et: "vt cs ((e # t) @ s)"
    show ?case
    proof(rule ih)
      show "vt cs (t @ s)"
      proof(rule step_back_vt)
        from vt_et show "vt cs (e # t @ s)" by simp
      qed
    qed
  qed
qed

context extend_highest_gen
begin

lemma red_moment:
  "extend_highest_gen s' th e' prio tm (moment i t)"
  apply (insert extend_highest_gen_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
  apply (unfold extend_highest_gen_def extend_highest_gen_axioms_def, clarsimp)
  by (unfold highest_gen_def, auto dest:step_back_vt_app)

lemma ind [consumes 0, case_names Nil Cons, induct type]:
  assumes 
    h0: "R []"
  and h2: "\<And> e t. \<lbrakk>vt step (t@s); step (t@s) e; 
                    extend_highest_gen s' th e' prio tm t; 
                    extend_highest_gen s' th e' prio tm (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)"
  shows "R t"
proof -
  from vt_t extend_highest_gen_axioms show ?thesis
  proof(induct t)
    from h0 show "R []" .
  next
    case (Cons e t')
    assume ih: "\<lbrakk>vt step (t' @ s); extend_highest_gen s' th e' prio tm t'\<rbrakk> \<Longrightarrow> R t'"
      and vt_e: "vt step ((e # t') @ s)"
      and et: "extend_highest_gen s' th e' prio tm (e # t')"
    from vt_e and step_back_step have stp: "step (t'@s) e" by auto
    from vt_e and step_back_vt have vt_ts: "vt step (t'@s)" by auto
    show ?case
    proof(rule h2 [OF vt_ts stp _ _ _ ])
      show "R t'"
      proof(rule ih)
        from et show ext': "extend_highest_gen s' th e' prio tm t'"
          by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
      next
        from vt_ts show "vt step (t' @ s)" .
      qed
    next
      from et show "extend_highest_gen s' th e' prio tm (e # t')" .
    next
      from et show ext': "extend_highest_gen s' th e' prio tm t'"
          by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
    qed
  qed
qed

lemma th_kept: "th \<in> threads (t @ s) \<and> 
        preced th (t@s) = preced th s" (is "?Q t")
proof -
  show ?thesis
  proof(induct rule:ind)
    case Nil
    from threads_s
    show "th \<in> threads ([] @ s) \<and> preced th ([] @ s) = preced th s"
      by auto
  next
    case (Cons e t)
    show ?case
    proof(cases e)
      case (Create thread prio)
      assume eq_e: " e = Create thread prio"
      show ?thesis
      proof -
        from Cons and eq_e have "step (t@s) (Create thread prio)" by auto
        hence "th \<noteq> thread"
        proof(cases)
          assume "thread \<notin> threads (t @ s)"
          with Cons show ?thesis by auto
        qed
        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
          by (unfold eq_e, auto simp:preced_def)
        moreover note Cons
        ultimately show ?thesis
          by (auto simp:eq_e)
      qed
    next
      case (Exit thread)
      assume eq_e: "e = Exit thread"
      from Cons have "extend_highest_gen s' th e' prio tm (e # t)" by auto
      from extend_highest_gen.exit_diff [OF this] and eq_e
      have neq_th: "thread \<noteq> th" by auto
      with Cons
      show ?thesis
        by (unfold eq_e, auto simp:preced_def)
    next
      case (P thread cs)
      assume eq_e: "e = P thread cs"
      with Cons
      show ?thesis 
        by (auto simp:eq_e preced_def)
    next
      case (V thread cs)
      assume eq_e: "e = V thread cs"
      with Cons
      show ?thesis 
        by (auto simp:eq_e preced_def)
    next
      case (Set thread prio')
      assume eq_e: " e = Set thread prio'"
      show ?thesis
      proof -
        from Cons have "extend_highest_gen s' th e' prio tm (e # t)" by auto
        from extend_highest_gen.set_diff_low[OF this] and eq_e
        have "th \<noteq> thread" by auto
        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
          by (unfold eq_e, auto simp:preced_def)
        moreover note Cons
        ultimately show ?thesis
          by (auto simp:eq_e)
      qed
    qed
  qed
qed

lemma max_kept: "Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s))) = preced th s"
proof(induct rule:ind)
  case Nil
  from highest_preced_thread
  show "Max ((\<lambda>th'. preced th' ([] @ s)) ` threads ([] @ s)) = preced th s"
    by simp
next
  case (Cons e t)
  show ?case
  proof(cases e)
    case (Create thread prio')
    assume eq_e: " e = Create thread prio'"
    from Cons and eq_e have stp: "step (t@s) (Create thread prio')" by auto
    hence neq_thread: "thread \<noteq> th"
    proof(cases)
      assume "thread \<notin> threads (t @ s)"
      moreover have "th \<in> threads (t@s)"
      proof -
        from Cons have "extend_highest_gen s' th e' prio tm t" by auto
        from extend_highest_gen.th_kept[OF this] show ?thesis by (simp add:s_def)
      qed
      ultimately show ?thesis by auto
    qed
    from Cons have "extend_highest_gen s' th e' prio tm t" by auto
    from extend_highest_gen.th_kept[OF this]
    have h': " th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" 
      by (auto simp:s_def)
    from stp
    have thread_ts: "thread \<notin> threads (t @ s)"
      by (cases, auto)
    show ?thesis (is "Max (?f ` ?A) = ?t")
    proof -
      have "Max (?f ` ?A) = Max(insert (?f thread) (?f ` (threads (t@s))))"
        by (unfold eq_e, simp)
      moreover have "\<dots> = max (?f thread) (Max (?f ` (threads (t@s))))"
      proof(rule Max_insert)
        from Cons have "vt step (t @ s)" by auto
        from finite_threads[OF this]
        show "finite (?f ` (threads (t@s)))" by simp
      next
        from h' show "(?f ` (threads (t@s))) \<noteq> {}" by auto
      qed
      moreover have "(Max (?f ` (threads (t@s)))) = ?t"
      proof -
        have "(\<lambda>th'. preced th' ((e # t) @ s)) ` threads (t @ s) = 
          (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" (is "?f1 ` ?B = ?f2 ` ?B")
        proof -
          { fix th' 
            assume "th' \<in> ?B"
            with thread_ts eq_e
            have "?f1 th' = ?f2 th'" by (auto simp:preced_def)
          } thus ?thesis 
            apply (auto simp:Image_def)
          proof -
            fix th'
            assume h: "\<And>th'. th' \<in> threads (t @ s) \<Longrightarrow> 
              preced th' (e # t @ s) = preced th' (t @ s)"
              and h1: "th' \<in> threads (t @ s)"
            show "preced th' (t @ s) \<in> (\<lambda>th'. preced th' (e # t @ s)) ` threads (t @ s)"
            proof -
              from h1 have "?f1 th' \<in> ?f1 ` ?B" by auto
              moreover from h[OF h1] have "?f1 th' = ?f2 th'" by simp
              ultimately show ?thesis by simp
            qed
          qed
        qed
        with Cons show ?thesis by auto
      qed
      moreover have "?f thread < ?t"
      proof -
        from Cons have "extend_highest_gen s' th e' prio tm (e # t)" by auto
        from extend_highest_gen.create_low[OF this] and eq_e
        have "prio' \<le> prio" by auto
        thus ?thesis
        by (unfold preced_th, unfold eq_e, insert lt_tm, 
          auto simp:preced_def s_def precedence_less_def preced_th)
    qed
    ultimately show ?thesis by (auto simp:max_def)
  qed
next
    case (Exit thread)
    assume eq_e: "e = Exit thread"
    from Cons have vt_e: "vt step (e#(t @ s))" by auto
    from Cons and eq_e have stp: "step (t@s) (Exit thread)" by auto
    from stp have thread_ts: "thread \<in> threads (t @ s)"
      by(cases, unfold runing_def readys_def, auto)
    from Cons have "extend_highest_gen s' th e' prio tm (e # t)" by auto
    from extend_highest_gen.exit_diff[OF this] and eq_e
    have neq_thread: "thread \<noteq> th" by auto
    from Cons have "extend_highest_gen s' th e' prio tm t" by auto
    from extend_highest_gen.th_kept[OF this, folded s_def]
    have h': "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
    show ?thesis (is "Max (?f ` ?A) = ?t")
    proof -
      have "threads (t@s) = insert thread ?A"
        by (insert stp thread_ts, unfold eq_e, auto)
      hence "Max (?f ` (threads (t@s))) = Max (?f ` \<dots>)" by simp
      also from this have "\<dots> = Max (insert (?f thread) (?f ` ?A))" by simp
      also have "\<dots> = max (?f thread) (Max (?f ` ?A))"
      proof(rule Max_insert)
        from finite_threads [OF vt_e]
        show "finite (?f ` ?A)" by simp
      next
        from Cons have "extend_highest_gen s' th e' prio tm (e # t)" by auto
        from extend_highest_gen.th_kept[OF this]
        show "?f ` ?A \<noteq> {}" by  (auto simp:s_def)
      qed
      finally have "Max (?f ` (threads (t@s))) = max (?f thread) (Max (?f ` ?A))" .
      moreover have "Max (?f ` (threads (t@s))) = ?t"
      proof -
        from Cons show ?thesis
          by (unfold eq_e, auto simp:preced_def)
      qed
      ultimately have "max (?f thread) (Max (?f ` ?A)) = ?t" by simp
      moreover have "?f thread < ?t"
      proof(unfold eq_e, simp add:preced_def, fold preced_def)
        show "preced thread (t @ s) < ?t"
        proof -
          have "preced thread (t @ s) \<le> ?t" 
          proof -
            from Cons
            have "?t = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" 
              (is "?t = Max (?g ` ?B)") by simp
            moreover have "?g thread \<le> \<dots>"
            proof(rule Max_ge)
              have "vt step (t@s)" by fact
              from finite_threads [OF this]
              show "finite (?g ` ?B)" by simp
            next
              from thread_ts
              show "?g thread \<in> (?g ` ?B)" by auto
            qed
            ultimately show ?thesis by auto
          qed
          moreover have "preced thread (t @ s) \<noteq> ?t"
          proof
            assume "preced thread (t @ s) = preced th s"
            with h' have "preced thread (t @ s) = preced th (t@s)" by simp
            from preced_unique [OF this] have "thread = th"
            proof
              from h' show "th \<in> threads (t @ s)" by simp
            next
              from thread_ts show "thread \<in> threads (t @ s)" .
            qed(simp)
            with neq_thread show "False" by simp
          qed
          ultimately show ?thesis by auto
        qed
      qed
      ultimately show ?thesis 
        by (auto simp:max_def split:if_splits)
    qed
  next
    case (P thread cs)
    with Cons
    show ?thesis by (auto simp:preced_def)
  next
    case (V thread cs)
    with Cons
    show ?thesis by (auto simp:preced_def)
  next
    case (Set thread prio')
    show ?thesis (is "Max (?f ` ?A) = ?t")
    proof -
      let ?B = "threads (t@s)"
      from Cons have "extend_highest_gen s' th e' prio tm (e # t)" by auto
      from extend_highest_gen.set_diff_low[OF this] and Set
      have neq_thread: "thread \<noteq> th" and le_p: "prio' \<le> prio" by auto
      from Set have "Max (?f ` ?A) = Max (?f ` ?B)" by simp
      also have "\<dots> = ?t"
      proof(rule Max_eqI)
        fix y
        assume y_in: "y \<in> ?f ` ?B"
        then obtain th1 where 
          th1_in: "th1 \<in> ?B" and eq_y: "y = ?f th1" by auto
        show "y \<le> ?t"
        proof(cases "th1 = thread")
          case True
          with neq_thread le_p eq_y s_def Set
          show ?thesis
            apply (subst preced_th, insert lt_tm)
            by (auto simp:preced_def precedence_le_def)
        next
          case False
          with Set eq_y
          have "y  = preced th1 (t@s)"
            by (simp add:preced_def)
          moreover have "\<dots> \<le> ?t"
          proof -
            from Cons
            have "?t = Max ((\<lambda> th'. preced th' (t@s)) ` (threads (t@s)))"
              by auto
            moreover have "preced th1 (t@s) \<le> \<dots>"
            proof(rule Max_ge)
              from th1_in 
              show "preced th1 (t @ s) \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)"
                by simp
            next
              show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
              proof -
                from Cons have "vt step (t @ s)" by auto
                from finite_threads[OF this] show ?thesis by auto
              qed
            qed
            ultimately show ?thesis by auto
          qed
          ultimately show ?thesis by auto
        qed
      next
        from Cons and finite_threads
        show "finite (?f ` ?B)" by auto
      next
        from Cons have "extend_highest_gen s' th e' prio tm t" by auto
        from extend_highest_gen.th_kept [OF this, folded s_def]
        have h: "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
        show "?t \<in> (?f ` ?B)" 
        proof -
          from neq_thread Set h
          have "?t = ?f th" by (auto simp:preced_def)
          with h show ?thesis by auto
        qed
      qed
      finally show ?thesis .
    qed
  qed
qed

lemma max_preced: "preced th (t@s) = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))"
  by (insert th_kept max_kept, auto)

lemma th_cp_max_preced: "cp (t@s) th = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))" 
  (is "?L = ?R")
proof -
  have "?L = cpreced (t@s) (wq (t@s)) th" 
    by (unfold cp_eq_cpreced, simp)
  also have "\<dots> = ?R"
  proof(unfold cpreced_def)
    show "Max ((\<lambda>th. preced th (t @ s)) ` ({th} \<union> dependents (wq (t @ s)) th)) =
          Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
      (is "Max (?f ` ({th} \<union> ?A)) = Max (?f ` ?B)")
    proof(cases "?A = {}")
      case False
      have "Max (?f ` ({th} \<union> ?A)) = Max (insert (?f th) (?f ` ?A))" by simp
      moreover have "\<dots> = max (?f th) (Max (?f ` ?A))"
      proof(rule Max_insert)
        show "finite (?f ` ?A)"
        proof -
          from dependents_threads[OF vt_t]
          have "?A \<subseteq> threads (t@s)" .
          moreover from finite_threads[OF vt_t] have "finite \<dots>" .
          ultimately show ?thesis 
            by (auto simp:finite_subset)
        qed
      next
        from False show "(?f ` ?A) \<noteq> {}" by simp
      qed
      moreover have "\<dots> = Max (?f ` ?B)"
      proof -
        from max_preced have "?f th = Max (?f ` ?B)" .
        moreover have "Max (?f ` ?A) \<le> \<dots>" 
        proof(rule Max_mono)
          from False show "(?f ` ?A) \<noteq> {}" by simp
        next
          show "?f ` ?A \<subseteq> ?f ` ?B" 
          proof -
            have "?A \<subseteq> ?B" by (rule dependents_threads[OF vt_t])
            thus ?thesis by auto
          qed
        next
          from finite_threads[OF vt_t] 
          show "finite (?f ` ?B)" by simp
        qed
        ultimately show ?thesis
          by (auto simp:max_def)
      qed
      ultimately show ?thesis by auto
    next
      case True
      with max_preced show ?thesis by auto
    qed
  qed
  finally show ?thesis .
qed

lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))"
  by (unfold max_cp_eq[OF vt_t] th_cp_max_preced, simp)

lemma th_cp_preced: "cp (t@s) th = preced th s"
  by (fold max_kept, unfold th_cp_max_preced, simp)

lemma preced_less':
  fixes th'
  assumes th'_in: "th' \<in> threads s"
  and neq_th': "th' \<noteq> th"
  shows "preced th' s < preced th s"
proof -
  have "preced th' s \<le> Max ((\<lambda>th'. preced th' s) ` threads s)"
  proof(rule Max_ge)
    from finite_threads [OF vt_s]
    show "finite ((\<lambda>th'. preced th' s) ` threads s)" by simp
  next
    from th'_in show "preced th' s \<in> (\<lambda>th'. preced th' s) ` threads s"
      by simp
  qed
  moreover have "preced th' s \<noteq> preced th s"
  proof
    assume "preced th' s = preced th s"
    from preced_unique[OF this th'_in] neq_th' threads_s
    show "False" by  (auto simp:readys_def)
  qed
  ultimately show ?thesis using highest_preced_thread
    by auto
qed

lemma pv_blocked:
  fixes th'
  assumes th'_in: "th' \<in> threads (t@s)"
  and neq_th': "th' \<noteq> th"
  and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
  shows "th' \<notin> runing (t@s)"
proof
  assume "th' \<in> runing (t@s)"
  hence "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" 
    by (auto simp:runing_def)
  with max_cp_readys_threads [OF vt_t]
  have "cp (t @ s) th' = Max (cp (t@s) ` threads (t@s))"
    by auto
  moreover from th_cp_max have "cp (t @ s) th = \<dots>" by simp
  ultimately have "cp (t @ s) th' = cp (t @ s) th" by simp
  moreover from th_cp_preced and th_kept have "\<dots> = preced th (t @ s)"
    by simp
  finally have h: "cp (t @ s) th' = preced th (t @ s)" .
  show False
  proof -
    have "dependents (wq (t @ s)) th' = {}" 
      by (rule count_eq_dependents [OF vt_t eq_pv])
    moreover have "preced th' (t @ s) \<noteq> preced th (t @ s)"
    proof
      assume "preced th' (t @ s) = preced th (t @ s)"
      hence "th' = th"
      proof(rule preced_unique)
        from th_kept show "th \<in> threads (t @ s)" by simp
      next
        from th'_in show "th' \<in> threads (t @ s)" by simp
      qed
      with assms show False by simp
    qed
    ultimately show ?thesis
      by (insert h, unfold cp_eq_cpreced cpreced_def, simp)
  qed
qed

lemma runing_precond_pre:
  fixes th'
  assumes th'_in: "th' \<in> threads s"
  and eq_pv: "cntP s th' = cntV s th'"
  and neq_th': "th' \<noteq> th"
  shows "th' \<in> threads (t@s) \<and>
         cntP (t@s) th' = cntV (t@s) th'"
proof -
  show ?thesis
  proof(induct rule:ind)
    case (Cons e t)
    from Cons
    have in_thread: "th' \<in> threads (t @ s)"
      and not_holding: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
    from Cons have "extend_highest_gen s' th e' prio tm t" by auto
    from extend_highest_gen.pv_blocked 
    [OF this, folded s_def, OF in_thread neq_th' not_holding]
    have not_runing: "th' \<notin> runing (t @ s)" .
    show ?case
    proof(cases e)
      case (V thread cs)
      from Cons and V have vt_v: "vt step (V thread cs#(t@s))" by auto

      show ?thesis
      proof -
        from Cons and V have "step (t@s) (V thread cs)" by auto
        hence neq_th': "thread \<noteq> th'"
        proof(cases)
          assume "thread \<in> runing (t@s)"
          moreover have "th' \<notin> runing (t@s)" by fact
          ultimately show ?thesis by auto
        qed
        with not_holding have cnt_eq: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" 
          by (unfold V, simp add:cntP_def cntV_def count_def)
        moreover from in_thread
        have in_thread': "th' \<in> threads ((e # t) @ s)" by (unfold V, simp)
        ultimately show ?thesis by auto
      qed
    next
      case (P thread cs)
      from Cons and P have "step (t@s) (P thread cs)" by auto
      hence neq_th': "thread \<noteq> th'"
      proof(cases)
        assume "thread \<in> runing (t@s)"
        moreover note not_runing
        ultimately show ?thesis by auto
      qed
      with Cons and P have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
        by (auto simp:cntP_def cntV_def count_def)
      moreover from Cons and P have in_thread': "th' \<in> threads ((e # t) @ s)"
        by auto
      ultimately show ?thesis by auto
    next
      case (Create thread prio')
      from Cons and Create have "step (t@s) (Create thread prio')" by auto
      hence neq_th': "thread \<noteq> th'"
      proof(cases)
        assume "thread \<notin> threads (t @ s)"
        moreover have "th' \<in> threads (t@s)" by fact
        ultimately show ?thesis by auto
      qed
      with Cons and Create 
      have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
        by (auto simp:cntP_def cntV_def count_def)
      moreover from Cons and Create 
      have in_thread': "th' \<in> threads ((e # t) @ s)" by auto
      ultimately show ?thesis by auto
    next
      case (Exit thread)
      from Cons and Exit have "step (t@s) (Exit thread)" by auto
      hence neq_th': "thread \<noteq> th'"
      proof(cases)
        assume "thread \<in> runing (t @ s)"
        moreover note not_runing
        ultimately show ?thesis by auto
      qed
      with Cons and Exit 
      have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
        by (auto simp:cntP_def cntV_def count_def)
      moreover from Cons and Exit and neq_th' 
      have in_thread': "th' \<in> threads ((e # t) @ s)"
        by auto
      ultimately show ?thesis by auto
    next
      case (Set thread prio')
      with Cons
      show ?thesis 
        by (auto simp:cntP_def cntV_def count_def)
    qed
  next
    case Nil
    with assms
    show ?case by auto
  qed
qed

(*
lemma runing_precond:
  fixes th'
  assumes th'_in: "th' \<in> threads s"
  and eq_pv: "cntP s th' = cntV s th'"
  and neq_th': "th' \<noteq> th"
  shows "th' \<notin> runing (t@s)"
proof -
  from runing_precond_pre[OF th'_in eq_pv neq_th']
  have h1: "th' \<in> threads (t @ s)"  and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
  from pv_blocked[OF h1 neq_th' h2] 
  show ?thesis .
qed
*)

lemma runing_precond:
  fixes th'
  assumes th'_in: "th' \<in> threads s"
  and neq_th': "th' \<noteq> th"
  and is_runing: "th' \<in> runing (t@s)"
  shows "cntP s th' > cntV s th'"
proof -
  have "cntP s th' \<noteq> cntV s th'"
  proof
    assume eq_pv: "cntP s th' = cntV s th'"
    from runing_precond_pre[OF th'_in eq_pv neq_th']
    have h1: "th' \<in> threads (t @ s)"  
      and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
    from pv_blocked[OF h1 neq_th' h2] have " th' \<notin> runing (t @ s)" .
    with is_runing show "False" by simp
  qed
  moreover from cnp_cnv_cncs[OF vt_s, of th'] 
  have "cntV s th' \<le> cntP s th'" by auto
  ultimately show ?thesis by auto
qed

lemma moment_blocked_pre:
  assumes neq_th': "th' \<noteq> th"
  and th'_in: "th' \<in> threads ((moment i t)@s)"
  and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
  shows "cntP ((moment (i+j) t)@s) th' = cntV ((moment (i+j) t)@s) th' \<and>
         th' \<in> threads ((moment (i+j) t)@s)"
proof(induct j)
  case (Suc k)
  show ?case
  proof -
    { assume True: "Suc (i+k) \<le> length t"
      from moment_head [OF this] 
      obtain e where
        eq_me: "moment (Suc(i+k)) t = e#(moment (i+k) t)"
        by blast
      from red_moment[of "Suc(i+k)"]
      and eq_me have "extend_highest_gen s' th e' prio tm (e # moment (i + k) t)" by simp
      hence vt_e: "vt step (e#(moment (i + k) t)@s)"
        by (unfold extend_highest_gen_def extend_highest_gen_axioms_def 
                          highest_gen_def s_def, auto)
      have not_runing': "th' \<notin>  runing (moment (i + k) t @ s)"
      proof(unfold s_def)
        show "th' \<notin> runing (moment (i + k) t @ e' # s')"
        proof(rule extend_highest_gen.pv_blocked)
          from Suc show "th' \<in> threads (moment (i + k) t @ e' # s')"
            by (simp add:s_def)
        next
          from neq_th' show "th' \<noteq> th" .
        next
          from red_moment show "extend_highest_gen s' th e' prio tm (moment (i + k) t)" .
        next
          from Suc show "cntP (moment (i + k) t @ e' # s') th' = cntV (moment (i + k) t @ e' # s') th'"
            by (auto simp:s_def)
        qed
      qed
      from step_back_step[OF vt_e]
      have "step ((moment (i + k) t)@s) e" .
      hence "cntP (e#(moment (i + k) t)@s) th' = cntV (e#(moment (i + k) t)@s) th' \<and>
        th' \<in> threads (e#(moment (i + k) t)@s)
        "
      proof(cases)
        case (thread_create thread prio)
        with Suc show ?thesis by (auto simp:cntP_def cntV_def count_def)
      next
        case (thread_exit thread)
        moreover have "thread \<noteq> th'"
        proof -
          have "thread \<in> runing (moment (i + k) t @ s)" by fact
          moreover note not_runing'
          ultimately show ?thesis by auto
        qed
        moreover note Suc 
        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
      next
        case (thread_P thread cs)
        moreover have "thread \<noteq> th'"
        proof -
          have "thread \<in> runing (moment (i + k) t @ s)" by fact
          moreover note not_runing'
          ultimately show ?thesis by auto
        qed
        moreover note Suc 
        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
      next
        case (thread_V thread cs)
        moreover have "thread \<noteq> th'"
        proof -
          have "thread \<in> runing (moment (i + k) t @ s)" by fact
          moreover note not_runing'
          ultimately show ?thesis by auto
        qed
        moreover note Suc 
        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
      next
        case (thread_set thread prio')
        with Suc show ?thesis
          by (auto simp:cntP_def cntV_def count_def)
      qed
      with eq_me have ?thesis using eq_me by auto 
    } note h = this
    show ?thesis
    proof(cases "Suc (i+k) \<le> length t")
      case True
      from h [OF this] show ?thesis .
    next
      case False
      with moment_ge
      have eq_m: "moment (i + Suc k) t = moment (i+k) t" by auto
      with Suc show ?thesis by auto
    qed
  qed
next
  case 0
  from assms show ?case by auto
qed

lemma moment_blocked:
  assumes neq_th': "th' \<noteq> th"
  and th'_in: "th' \<in> threads ((moment i t)@s)"
  and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
  and le_ij: "i \<le> j"
  shows "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th' \<and>
         th' \<in> threads ((moment j t)@s) \<and>
         th' \<notin> runing ((moment j t)@s)"
proof -
  from moment_blocked_pre [OF neq_th' th'_in eq_pv, of "j-i"] and le_ij
  have h1: "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th'"
    and h2: "th' \<in> threads ((moment j t)@s)" by auto
  with extend_highest_gen.pv_blocked [OF  red_moment [of j], folded s_def, OF h2 neq_th' h1]
  show ?thesis by auto
qed

lemma runing_inversion_1:
  assumes neq_th': "th' \<noteq> th"
  and runing': "th' \<in> runing (t@s)"
  shows "th' \<in> threads s \<and> cntV s th' < cntP s th'"
proof(cases "th' \<in> threads s")
  case True
  with runing_precond [OF this neq_th' runing'] show ?thesis by simp
next
  case False
  let ?Q = "\<lambda> t. th' \<in> threads (t@s)"
  let ?q = "moment 0 t"
  from moment_eq and False have not_thread: "\<not> ?Q ?q" by simp
  from runing' have "th' \<in> threads (t@s)" by (simp add:runing_def readys_def)
  from p_split_gen [of ?Q, OF this not_thread]
  obtain i where lt_its: "i < length t"
    and le_i: "0 \<le> i"
    and pre: " th' \<notin> threads (moment i t @ s)" (is "th' \<notin> threads ?pre")
    and post: "(\<forall>i'>i. th' \<in> threads (moment i' t @ s))" by auto
  from lt_its have "Suc i \<le> length t" by auto
  from moment_head[OF this] obtain e where 
   eq_me: "moment (Suc i) t = e # moment i t" by blast
  from red_moment[of "Suc i"] and eq_me
  have "extend_highest_gen s' th e' prio tm (e # moment i t)" by simp
  hence vt_e: "vt step (e#(moment i t)@s)"
    by (unfold extend_highest_gen_def extend_highest_gen_axioms_def 
      highest_gen_def s_def, auto)
  from step_back_step[OF this] have stp_i: "step (moment i t @ s) e" .
  from post[rule_format, of "Suc i"] and eq_me 
  have not_in': "th' \<in> threads (e # moment i t@s)" by auto
  from create_pre[OF stp_i pre this] 
  obtain prio where eq_e: "e = Create th' prio" .
  have "cntP (moment i t@s) th' = cntV (moment i t@s) th'"
  proof(rule cnp_cnv_eq)
    from step_back_vt [OF vt_e] 
    show "vt step (moment i t @ s)" .
  next
    from eq_e and stp_i 
    have "step (moment i t @ s) (Create th' prio)" by simp
    thus "th' \<notin> threads (moment i t @ s)" by (cases, simp)
  qed
  with eq_e
  have "cntP ((e#moment i t)@s) th' = cntV ((e#moment i t)@s) th'" 
    by (simp add:cntP_def cntV_def count_def)
  with eq_me[symmetric]
  have h1: "cntP (moment (Suc i) t @ s) th' = cntV (moment (Suc i) t@ s) th'"
    by simp
  from eq_e have "th' \<in> threads ((e#moment i t)@s)" by simp
  with eq_me [symmetric]
  have h2: "th' \<in> threads (moment (Suc i) t @ s)" by simp
  from moment_blocked [OF neq_th' h2 h1, of "length t"] and lt_its
  and moment_ge
  have "th' \<notin> runing (t @ s)" by auto
  with runing'
  show ?thesis by auto
qed

lemma runing_inversion_2:
  assumes runing': "th' \<in> runing (t@s)"
  shows "th' = th \<or> (th' \<noteq> th \<and> th' \<in> threads s \<and> cntV s th' < cntP s th')"
proof -
  from runing_inversion_1[OF _ runing']
  show ?thesis by auto
qed

lemma live: "runing (t@s) \<noteq> {}"
proof(cases "th \<in> runing (t@s)")
  case True thus ?thesis by auto
next
  case False
  then have not_ready: "th \<notin> readys (t@s)"
    apply (unfold runing_def, 
            insert th_cp_max max_cp_readys_threads[OF vt_t, symmetric])
    by auto
  from th_kept have "th \<in> threads (t@s)" by auto
  from th_chain_to_ready[OF vt_t this] and not_ready
  obtain th' where th'_in: "th' \<in> readys (t@s)"
    and dp: "(Th th, Th th') \<in> (depend (t @ s))\<^sup>+" by auto
  have "th' \<in> runing (t@s)"
  proof -
    have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))"
    proof -
      have " Max ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')) = 
               preced th (t@s)"
      proof(rule Max_eqI)
        fix y
        assume "y \<in> (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')"
        then obtain th1 where
          h1: "th1 = th' \<or> th1 \<in>  dependents (wq (t @ s)) th'"
          and eq_y: "y = preced th1 (t@s)" by auto
        show "y \<le> preced th (t @ s)"
        proof -
          from max_preced
          have "preced th (t @ s) = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" .
          moreover have "y \<le> \<dots>"
          proof(rule Max_ge)
            from h1
            have "th1 \<in> threads (t@s)"
            proof
              assume "th1 = th'"
              with th'_in show ?thesis by (simp add:readys_def)
            next
              assume "th1 \<in> dependents (wq (t @ s)) th'"
              with dependents_threads [OF vt_t]
              show "th1 \<in> threads (t @ s)" by auto
            qed
            with eq_y show " y \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" by simp
          next
            from finite_threads[OF vt_t]
            show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" by simp
          qed
          ultimately show ?thesis by auto
        qed
      next
        from finite_threads[OF vt_t] dependents_threads [OF vt_t, of th']
        show "finite ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th'))"
          by (auto intro:finite_subset)
      next
        from dp
        have "th \<in> dependents (wq (t @ s)) th'" 
          by (unfold cs_dependents_def, auto simp:eq_depend)
        thus "preced th (t @ s) \<in> 
                (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')"
          by auto
      qed
      moreover have "\<dots> = Max (cp (t @ s) ` readys (t @ s))"
      proof -
        from max_preced and max_cp_eq[OF vt_t, symmetric]
        have "preced th (t @ s) = Max (cp (t @ s) ` threads (t @ s))" by simp
        with max_cp_readys_threads[OF vt_t] show ?thesis by simp
      qed
      ultimately show ?thesis by (unfold cp_eq_cpreced cpreced_def, simp)
    qed
    with th'_in show ?thesis by (auto simp:runing_def)
  qed
  thus ?thesis by auto
qed

end

end