PrioG.thy
changeset 35 92f61f6a0fe7
parent 32 e861aff29655
child 36 af38526275f8
equal deleted inserted replaced
34:313acffe63b6 35:92f61f6a0fe7
    22   assume h1: "step s e"
    22   assume h1: "step s e"
    23   and h2: "distinct (wq s cs)"
    23   and h2: "distinct (wq s cs)"
    24   thus "distinct (wq (e # s) cs)"
    24   thus "distinct (wq (e # s) cs)"
    25   proof(induct rule:step.induct, auto simp: wq_def Let_def split:list.splits)
    25   proof(induct rule:step.induct, auto simp: wq_def Let_def split:list.splits)
    26     fix thread s
    26     fix thread s
    27     assume h1: "(Cs cs, Th thread) \<notin> (depend s)\<^sup>+"
    27     assume h1: "(Cs cs, Th thread) \<notin> (RAG s)\<^sup>+"
    28       and h2: "thread \<in> set (wq_fun (schs s) cs)"
    28       and h2: "thread \<in> set (wq_fun (schs s) cs)"
    29       and h3: "thread \<in> runing s"
    29       and h3: "thread \<in> runing s"
    30     show "False" 
    30     show "False" 
    31     proof -
    31     proof -
    32       from h3 have "\<And> cs. thread \<in>  set (wq_fun (schs s) cs) \<Longrightarrow> 
    32       from h3 have "\<And> cs. thread \<in>  set (wq_fun (schs s) cs) \<Longrightarrow> 
    33                              thread = hd ((wq_fun (schs s) cs))" 
    33                              thread = hd ((wq_fun (schs s) cs))" 
    34         by (simp add:runing_def readys_def s_waiting_def wq_def)
    34         by (simp add:runing_def readys_def s_waiting_def wq_def)
    35       from this [OF h2] have "thread = hd (wq_fun (schs s) cs)" .
    35       from this [OF h2] have "thread = hd (wq_fun (schs s) cs)" .
    36       with h2
    36       with h2
    37       have "(Cs cs, Th thread) \<in> (depend s)"
    37       have "(Cs cs, Th thread) \<in> (RAG s)"
    38         by (simp add:s_depend_def s_holding_def wq_def cs_holding_def)
    38         by (simp add:s_RAG_def s_holding_def wq_def cs_holding_def)
    39       with h1 show False by auto
    39       with h1 show False by auto
    40     qed
    40     qed
    41   next
    41   next
    42     fix thread s a list
    42     fix thread s a list
    43     assume dst: "distinct list"
    43     assume dst: "distinct list"
   110       qed
   110       qed
   111   qed
   111   qed
   112 qed
   112 qed
   113 
   113 
   114 lemma p_pre: "\<lbrakk>vt ((P thread cs)#s)\<rbrakk> \<Longrightarrow> 
   114 lemma p_pre: "\<lbrakk>vt ((P thread cs)#s)\<rbrakk> \<Longrightarrow> 
   115   thread \<in> runing s \<and> (Cs cs, Th thread)  \<notin> (depend s)^+"
   115   thread \<in> runing s \<and> (Cs cs, Th thread)  \<notin> (RAG s)^+"
   116 apply (ind_cases "vt ((P thread cs)#s)")
   116 apply (ind_cases "vt ((P thread cs)#s)")
   117 apply (ind_cases "step s (P thread cs)")
   117 apply (ind_cases "step s (P thread cs)")
   118 by auto
   118 by auto
   119 
   119 
   120 lemma abs1:
   120 lemma abs1:
   499       qed
   499       qed
   500     qed
   500     qed
   501   qed
   501   qed
   502 qed
   502 qed
   503 
   503 
   504 lemma depend_set_unchanged: "(depend (Set th prio # s)) = depend s"
   504 lemma RAG_set_unchanged: "(RAG (Set th prio # s)) = RAG s"
   505 apply (unfold s_depend_def s_waiting_def wq_def)
   505 apply (unfold s_RAG_def s_waiting_def wq_def)
   506 by (simp add:Let_def)
   506 by (simp add:Let_def)
   507 
   507 
   508 lemma depend_create_unchanged: "(depend (Create th prio # s)) = depend s"
   508 lemma RAG_create_unchanged: "(RAG (Create th prio # s)) = RAG s"
   509 apply (unfold s_depend_def s_waiting_def wq_def)
   509 apply (unfold s_RAG_def s_waiting_def wq_def)
   510 by (simp add:Let_def)
   510 by (simp add:Let_def)
   511 
   511 
   512 lemma depend_exit_unchanged: "(depend (Exit th # s)) = depend s"
   512 lemma RAG_exit_unchanged: "(RAG (Exit th # s)) = RAG s"
   513 apply (unfold s_depend_def s_waiting_def wq_def)
   513 apply (unfold s_RAG_def s_waiting_def wq_def)
   514 by (simp add:Let_def)
   514 by (simp add:Let_def)
   515 
   515 
   516 
   516 
   517 
   517 
   518 lemma step_v_hold_inv[elim_format]:
   518 lemma step_v_hold_inv[elim_format]:
   771       show False by auto
   771       show False by auto
   772     qed
   772     qed
   773   qed
   773   qed
   774 qed
   774 qed
   775 
   775 
   776 lemma step_depend_v:
   776 lemma step_RAG_v:
   777 fixes th::thread
   777 fixes th::thread
   778 assumes vt:
   778 assumes vt:
   779   "vt (V th cs#s)"
   779   "vt (V th cs#s)"
   780 shows "
   780 shows "
   781   depend (V th cs # s) =
   781   RAG (V th cs # s) =
   782   depend s - {(Cs cs, Th th)} -
   782   RAG s - {(Cs cs, Th th)} -
   783   {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
   783   {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
   784   {(Cs cs, Th th') |th'.  next_th s th cs th'}"
   784   {(Cs cs, Th th') |th'.  next_th s th cs th'}"
   785   apply (insert vt, unfold s_depend_def) 
   785   apply (insert vt, unfold s_RAG_def) 
   786   apply (auto split:if_splits list.splits simp:Let_def)
   786   apply (auto split:if_splits list.splits simp:Let_def)
   787   apply (auto elim: step_v_waiting_mono step_v_hold_inv 
   787   apply (auto elim: step_v_waiting_mono step_v_hold_inv 
   788               step_v_release step_v_wait_inv
   788               step_v_release step_v_wait_inv
   789               step_v_get_hold step_v_release_inv)
   789               step_v_get_hold step_v_release_inv)
   790   apply (erule_tac step_v_not_wait, auto)
   790   apply (erule_tac step_v_not_wait, auto)
   791   done
   791   done
   792 
   792 
   793 lemma step_depend_p:
   793 lemma step_RAG_p:
   794   "vt (P th cs#s) \<Longrightarrow>
   794   "vt (P th cs#s) \<Longrightarrow>
   795   depend (P th cs # s) =  (if (wq s cs = []) then depend s \<union> {(Cs cs, Th th)}
   795   RAG (P th cs # s) =  (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)}
   796                                              else depend s \<union> {(Th th, Cs cs)})"
   796                                              else RAG s \<union> {(Th th, Cs cs)})"
   797   apply(simp only: s_depend_def wq_def)
   797   apply(simp only: s_RAG_def wq_def)
   798   apply (auto split:list.splits prod.splits simp:Let_def wq_def cs_waiting_def cs_holding_def)
   798   apply (auto split:list.splits prod.splits simp:Let_def wq_def cs_waiting_def cs_holding_def)
   799   apply(case_tac "csa = cs", auto)
   799   apply(case_tac "csa = cs", auto)
   800   apply(fold wq_def)
   800   apply(fold wq_def)
   801   apply(drule_tac step_back_step)
   801   apply(drule_tac step_back_step)
   802   apply(ind_cases " step s (P (hd (wq s cs)) cs)")
   802   apply(ind_cases " step s (P (hd (wq s cs)) cs)")
   803   apply(auto simp:s_depend_def wq_def cs_holding_def)
   803   apply(auto simp:s_RAG_def wq_def cs_holding_def)
   804   done
   804   done
   805 
   805 
   806 lemma simple_A:
   806 lemma simple_A:
   807   fixes A
   807   fixes A
   808   assumes h: "\<And> x y. \<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> x = y"
   808   assumes h: "\<And> x y. \<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> x = y"
   813   case False then obtain a where "a \<in> A" by auto
   813   case False then obtain a where "a \<in> A" by auto
   814   with h have "A = {a}" by auto
   814   with h have "A = {a}" by auto
   815   thus ?thesis by simp
   815   thus ?thesis by simp
   816 qed
   816 qed
   817 
   817 
   818 lemma depend_target_th: "(Th th, x) \<in> depend (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
   818 lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
   819   by (unfold s_depend_def, auto)
   819   by (unfold s_RAG_def, auto)
   820 
   820 
   821 lemma acyclic_depend: 
   821 lemma acyclic_RAG: 
   822   fixes s
   822   fixes s
   823   assumes vt: "vt s"
   823   assumes vt: "vt s"
   824   shows "acyclic (depend s)"
   824   shows "acyclic (RAG s)"
   825 proof -
   825 proof -
   826   from vt show ?thesis
   826   from vt show ?thesis
   827   proof(induct)
   827   proof(induct)
   828     case (vt_cons s e)
   828     case (vt_cons s e)
   829     assume ih: "acyclic (depend s)"
   829     assume ih: "acyclic (RAG s)"
   830       and stp: "step s e"
   830       and stp: "step s e"
   831       and vt: "vt s"
   831       and vt: "vt s"
   832     show ?case
   832     show ?case
   833     proof(cases e)
   833     proof(cases e)
   834       case (Create th prio)
   834       case (Create th prio)
   835       with ih
   835       with ih
   836       show ?thesis by (simp add:depend_create_unchanged)
   836       show ?thesis by (simp add:RAG_create_unchanged)
   837     next
   837     next
   838       case (Exit th)
   838       case (Exit th)
   839       with ih show ?thesis by (simp add:depend_exit_unchanged)
   839       with ih show ?thesis by (simp add:RAG_exit_unchanged)
   840     next
   840     next
   841       case (V th cs)
   841       case (V th cs)
   842       from V vt stp have vtt: "vt (V th cs#s)" by auto
   842       from V vt stp have vtt: "vt (V th cs#s)" by auto
   843       from step_depend_v [OF this]
   843       from step_RAG_v [OF this]
   844       have eq_de: 
   844       have eq_de: 
   845         "depend (e # s) = 
   845         "RAG (e # s) = 
   846             depend s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
   846             RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
   847             {(Cs cs, Th th') |th'. next_th s th cs th'}"
   847             {(Cs cs, Th th') |th'. next_th s th cs th'}"
   848         (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
   848         (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
   849       from ih have ac: "acyclic (?A - ?B - ?C)" by (auto elim:acyclic_subset)
   849       from ih have ac: "acyclic (?A - ?B - ?C)" by (auto elim:acyclic_subset)
   850       from step_back_step [OF vtt]
   850       from step_back_step [OF vtt]
   851       have "step s (V th cs)" .
   851       have "step s (V th cs)" .
   869             assume "(Th ?th', Cs cs) \<in> ?E\<^sup>*"
   869             assume "(Th ?th', Cs cs) \<in> ?E\<^sup>*"
   870             hence " (Th ?th', Cs cs) \<in> ?E\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
   870             hence " (Th ?th', Cs cs) \<in> ?E\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
   871             from tranclD [OF this]
   871             from tranclD [OF this]
   872             obtain x where th'_e: "(Th ?th', x) \<in> ?E" by blast
   872             obtain x where th'_e: "(Th ?th', x) \<in> ?E" by blast
   873             hence th_d: "(Th ?th', x) \<in> ?A" by simp
   873             hence th_d: "(Th ?th', x) \<in> ?A" by simp
   874             from depend_target_th [OF this]
   874             from RAG_target_th [OF this]
   875             obtain cs' where eq_x: "x = Cs cs'" by auto
   875             obtain cs' where eq_x: "x = Cs cs'" by auto
   876             with th_d have "(Th ?th', Cs cs') \<in> ?A" by simp
   876             with th_d have "(Th ?th', Cs cs') \<in> ?A" by simp
   877             hence wt_th': "waiting s ?th' cs'"
   877             hence wt_th': "waiting s ?th' cs'"
   878               unfolding s_depend_def s_waiting_def cs_waiting_def wq_def by simp
   878               unfolding s_RAG_def s_waiting_def cs_waiting_def wq_def by simp
   879             hence "cs' = cs"
   879             hence "cs' = cs"
   880             proof(rule waiting_unique [OF vt])
   880             proof(rule waiting_unique [OF vt])
   881               from eq_wq wq_distinct[OF vt, of cs]
   881               from eq_wq wq_distinct[OF vt, of cs]
   882               show "waiting s ?th' cs" 
   882               show "waiting s ?th' cs" 
   883                 apply (unfold s_waiting_def wq_def, auto)
   883                 apply (unfold s_waiting_def wq_def, auto)
   943         qed 
   943         qed 
   944       qed
   944       qed
   945   next
   945   next
   946     case (P th cs)
   946     case (P th cs)
   947     from P vt stp have vtt: "vt (P th cs#s)" by auto
   947     from P vt stp have vtt: "vt (P th cs#s)" by auto
   948     from step_depend_p [OF this] P
   948     from step_RAG_p [OF this] P
   949     have "depend (e # s) = 
   949     have "RAG (e # s) = 
   950       (if wq s cs = [] then depend s \<union> {(Cs cs, Th th)} else 
   950       (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
   951       depend s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
   951       RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
   952       by simp
   952       by simp
   953     moreover have "acyclic ?R"
   953     moreover have "acyclic ?R"
   954     proof(cases "wq s cs = []")
   954     proof(cases "wq s cs = []")
   955       case True
   955       case True
   956       hence eq_r: "?R =  depend s \<union> {(Cs cs, Th th)}" by simp
   956       hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
   957       have "(Th th, Cs cs) \<notin> (depend s)\<^sup>*"
   957       have "(Th th, Cs cs) \<notin> (RAG s)\<^sup>*"
   958       proof
   958       proof
   959         assume "(Th th, Cs cs) \<in> (depend s)\<^sup>*"
   959         assume "(Th th, Cs cs) \<in> (RAG s)\<^sup>*"
   960         hence "(Th th, Cs cs) \<in> (depend s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
   960         hence "(Th th, Cs cs) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
   961         from tranclD2 [OF this]
   961         from tranclD2 [OF this]
   962         obtain x where "(x, Cs cs) \<in> depend s" by auto
   962         obtain x where "(x, Cs cs) \<in> RAG s" by auto
   963         with True show False by (auto simp:s_depend_def cs_waiting_def)
   963         with True show False by (auto simp:s_RAG_def cs_waiting_def)
   964       qed
   964       qed
   965       with acyclic_insert ih eq_r show ?thesis by auto
   965       with acyclic_insert ih eq_r show ?thesis by auto
   966     next
   966     next
   967       case False
   967       case False
   968       hence eq_r: "?R =  depend s \<union> {(Th th, Cs cs)}" by simp
   968       hence eq_r: "?R =  RAG s \<union> {(Th th, Cs cs)}" by simp
   969       have "(Cs cs, Th th) \<notin> (depend s)\<^sup>*"
   969       have "(Cs cs, Th th) \<notin> (RAG s)\<^sup>*"
   970       proof
   970       proof
   971         assume "(Cs cs, Th th) \<in> (depend s)\<^sup>*"
   971         assume "(Cs cs, Th th) \<in> (RAG s)\<^sup>*"
   972         hence "(Cs cs, Th th) \<in> (depend s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
   972         hence "(Cs cs, Th th) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
   973           moreover from step_back_step [OF vtt] have "step s (P th cs)" .
   973           moreover from step_back_step [OF vtt] have "step s (P th cs)" .
   974           ultimately show False
   974           ultimately show False
   975           proof -
   975           proof -
   976             show " \<lbrakk>(Cs cs, Th th) \<in> (depend s)\<^sup>+; step s (P th cs)\<rbrakk> \<Longrightarrow> False"
   976             show " \<lbrakk>(Cs cs, Th th) \<in> (RAG s)\<^sup>+; step s (P th cs)\<rbrakk> \<Longrightarrow> False"
   977               by (ind_cases "step s (P th cs)", simp)
   977               by (ind_cases "step s (P th cs)", simp)
   978           qed
   978           qed
   979         qed
   979         qed
   980         with acyclic_insert ih eq_r show ?thesis by auto
   980         with acyclic_insert ih eq_r show ?thesis by auto
   981       qed
   981       qed
   982       ultimately show ?thesis by simp
   982       ultimately show ?thesis by simp
   983     next
   983     next
   984       case (Set thread prio)
   984       case (Set thread prio)
   985       with ih
   985       with ih
   986       thm depend_set_unchanged
   986       thm RAG_set_unchanged
   987       show ?thesis by (simp add:depend_set_unchanged)
   987       show ?thesis by (simp add:RAG_set_unchanged)
   988     qed
   988     qed
   989   next
   989   next
   990     case vt_nil
   990     case vt_nil
   991     show "acyclic (depend ([]::state))"
   991     show "acyclic (RAG ([]::state))"
   992       by (auto simp: s_depend_def cs_waiting_def 
   992       by (auto simp: s_RAG_def cs_waiting_def 
   993                       cs_holding_def wq_def acyclic_def)
   993                       cs_holding_def wq_def acyclic_def)
   994   qed
   994   qed
   995 qed
   995 qed
   996 
   996 
   997 lemma finite_depend: 
   997 lemma finite_RAG: 
   998   fixes s
   998   fixes s
   999   assumes vt: "vt s"
   999   assumes vt: "vt s"
  1000   shows "finite (depend s)"
  1000   shows "finite (RAG s)"
  1001 proof -
  1001 proof -
  1002   from vt show ?thesis
  1002   from vt show ?thesis
  1003   proof(induct)
  1003   proof(induct)
  1004     case (vt_cons s e)
  1004     case (vt_cons s e)
  1005     assume ih: "finite (depend s)"
  1005     assume ih: "finite (RAG s)"
  1006       and stp: "step s e"
  1006       and stp: "step s e"
  1007       and vt: "vt s"
  1007       and vt: "vt s"
  1008     show ?case
  1008     show ?case
  1009     proof(cases e)
  1009     proof(cases e)
  1010       case (Create th prio)
  1010       case (Create th prio)
  1011       with ih
  1011       with ih
  1012       show ?thesis by (simp add:depend_create_unchanged)
  1012       show ?thesis by (simp add:RAG_create_unchanged)
  1013     next
  1013     next
  1014       case (Exit th)
  1014       case (Exit th)
  1015       with ih show ?thesis by (simp add:depend_exit_unchanged)
  1015       with ih show ?thesis by (simp add:RAG_exit_unchanged)
  1016     next
  1016     next
  1017       case (V th cs)
  1017       case (V th cs)
  1018       from V vt stp have vtt: "vt (V th cs#s)" by auto
  1018       from V vt stp have vtt: "vt (V th cs#s)" by auto
  1019       from step_depend_v [OF this]
  1019       from step_RAG_v [OF this]
  1020       have eq_de: "depend (e # s) = 
  1020       have eq_de: "RAG (e # s) = 
  1021                    depend s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
  1021                    RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
  1022                       {(Cs cs, Th th') |th'. next_th s th cs th'}
  1022                       {(Cs cs, Th th') |th'. next_th s th cs th'}
  1023 "
  1023 "
  1024         (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
  1024         (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
  1025       moreover from ih have ac: "finite (?A - ?B - ?C)" by simp
  1025       moreover from ih have ac: "finite (?A - ?B - ?C)" by simp
  1026       moreover have "finite ?D"
  1026       moreover have "finite ?D"
  1039       qed
  1039       qed
  1040       ultimately show ?thesis by simp
  1040       ultimately show ?thesis by simp
  1041     next
  1041     next
  1042       case (P th cs)
  1042       case (P th cs)
  1043       from P vt stp have vtt: "vt (P th cs#s)" by auto
  1043       from P vt stp have vtt: "vt (P th cs#s)" by auto
  1044       from step_depend_p [OF this] P
  1044       from step_RAG_p [OF this] P
  1045       have "depend (e # s) = 
  1045       have "RAG (e # s) = 
  1046               (if wq s cs = [] then depend s \<union> {(Cs cs, Th th)} else 
  1046               (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
  1047                                     depend s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
  1047                                     RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
  1048         by simp
  1048         by simp
  1049       moreover have "finite ?R"
  1049       moreover have "finite ?R"
  1050       proof(cases "wq s cs = []")
  1050       proof(cases "wq s cs = []")
  1051         case True
  1051         case True
  1052         hence eq_r: "?R =  depend s \<union> {(Cs cs, Th th)}" by simp
  1052         hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
  1053         with True and ih show ?thesis by auto
  1053         with True and ih show ?thesis by auto
  1054       next
  1054       next
  1055         case False
  1055         case False
  1056         hence "?R = depend s \<union> {(Th th, Cs cs)}" by simp
  1056         hence "?R = RAG s \<union> {(Th th, Cs cs)}" by simp
  1057         with False and ih show ?thesis by auto
  1057         with False and ih show ?thesis by auto
  1058       qed
  1058       qed
  1059       ultimately show ?thesis by auto
  1059       ultimately show ?thesis by auto
  1060     next
  1060     next
  1061       case (Set thread prio)
  1061       case (Set thread prio)
  1062       with ih
  1062       with ih
  1063       show ?thesis by (simp add:depend_set_unchanged)
  1063       show ?thesis by (simp add:RAG_set_unchanged)
  1064     qed
  1064     qed
  1065   next
  1065   next
  1066     case vt_nil
  1066     case vt_nil
  1067     show "finite (depend ([]::state))"
  1067     show "finite (RAG ([]::state))"
  1068       by (auto simp: s_depend_def cs_waiting_def 
  1068       by (auto simp: s_RAG_def cs_waiting_def 
  1069                    cs_holding_def wq_def acyclic_def)
  1069                    cs_holding_def wq_def acyclic_def)
  1070   qed
  1070   qed
  1071 qed
  1071 qed
  1072 
  1072 
  1073 text {* Several useful lemmas *}
  1073 text {* Several useful lemmas *}
  1074 
  1074 
  1075 lemma wf_dep_converse: 
  1075 lemma wf_dep_converse: 
  1076   fixes s
  1076   fixes s
  1077   assumes vt: "vt s"
  1077   assumes vt: "vt s"
  1078   shows "wf ((depend s)^-1)"
  1078   shows "wf ((RAG s)^-1)"
  1079 proof(rule finite_acyclic_wf_converse)
  1079 proof(rule finite_acyclic_wf_converse)
  1080   from finite_depend [OF vt]
  1080   from finite_RAG [OF vt]
  1081   show "finite (depend s)" .
  1081   show "finite (RAG s)" .
  1082 next
  1082 next
  1083   from acyclic_depend[OF vt]
  1083   from acyclic_RAG[OF vt]
  1084   show "acyclic (depend s)" .
  1084   show "acyclic (RAG s)" .
  1085 qed
  1085 qed
  1086 
  1086 
  1087 lemma hd_np_in: "x \<in> set l \<Longrightarrow> hd l \<in> set l"
  1087 lemma hd_np_in: "x \<in> set l \<Longrightarrow> hd l \<in> set l"
  1088 by (induct l, auto)
  1088 by (induct l, auto)
  1089 
  1089 
  1090 lemma th_chasing: "(Th th, Cs cs) \<in> depend (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> depend s"
  1090 lemma th_chasing: "(Th th, Cs cs) \<in> RAG (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> RAG s"
  1091   by (auto simp:s_depend_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
  1091   by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
  1092 
  1092 
  1093 lemma wq_threads: 
  1093 lemma wq_threads: 
  1094   fixes s cs
  1094   fixes s cs
  1095   assumes vt: "vt s"
  1095   assumes vt: "vt s"
  1096   and h: "th \<in> set (wq s cs)"
  1096   and h: "th \<in> set (wq s cs)"
  1112       case (Exit th')
  1112       case (Exit th')
  1113       with stp ih h show ?thesis
  1113       with stp ih h show ?thesis
  1114         apply (auto simp:wq_def Let_def)
  1114         apply (auto simp:wq_def Let_def)
  1115         apply (ind_cases "step s (Exit th')")
  1115         apply (ind_cases "step s (Exit th')")
  1116         apply (auto simp:runing_def readys_def s_holding_def s_waiting_def holdents_def
  1116         apply (auto simp:runing_def readys_def s_holding_def s_waiting_def holdents_def
  1117                s_depend_def s_holding_def cs_holding_def)
  1117                s_RAG_def s_holding_def cs_holding_def)
  1118         done
  1118         done
  1119     next
  1119     next
  1120       case (V th' cs')
  1120       case (V th' cs')
  1121       show ?thesis
  1121       show ?thesis
  1122       proof(cases "cs' = cs")
  1122       proof(cases "cs' = cs")
  1190     case vt_nil
  1190     case vt_nil
  1191     thus ?case by (auto simp:wq_def)
  1191     thus ?case by (auto simp:wq_def)
  1192   qed
  1192   qed
  1193 qed
  1193 qed
  1194 
  1194 
  1195 lemma range_in: "\<lbrakk>vt s; (Th th) \<in> Range (depend (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s"
  1195 lemma range_in: "\<lbrakk>vt s; (Th th) \<in> Range (RAG (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s"
  1196   apply(unfold s_depend_def cs_waiting_def cs_holding_def)
  1196   apply(unfold s_RAG_def cs_waiting_def cs_holding_def)
  1197   by (auto intro:wq_threads)
  1197   by (auto intro:wq_threads)
  1198 
  1198 
  1199 lemma readys_v_eq:
  1199 lemma readys_v_eq:
  1200   fixes th thread cs rest
  1200   fixes th thread cs rest
  1201   assumes vt: "vt s"
  1201   assumes vt: "vt s"
  1229     qed
  1229     qed
  1230 qed
  1230 qed
  1231 
  1231 
  1232 lemma chain_building:
  1232 lemma chain_building:
  1233   assumes vt: "vt s"
  1233   assumes vt: "vt s"
  1234   shows "node \<in> Domain (depend s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (depend s)^+)"
  1234   shows "node \<in> Domain (RAG s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (RAG s)^+)"
  1235 proof -
  1235 proof -
  1236   from wf_dep_converse [OF vt]
  1236   from wf_dep_converse [OF vt]
  1237   have h: "wf ((depend s)\<inverse>)" .
  1237   have h: "wf ((RAG s)\<inverse>)" .
  1238   show ?thesis
  1238   show ?thesis
  1239   proof(induct rule:wf_induct [OF h])
  1239   proof(induct rule:wf_induct [OF h])
  1240     fix x
  1240     fix x
  1241     assume ih [rule_format]: 
  1241     assume ih [rule_format]: 
  1242       "\<forall>y. (y, x) \<in> (depend s)\<inverse> \<longrightarrow> 
  1242       "\<forall>y. (y, x) \<in> (RAG s)\<inverse> \<longrightarrow> 
  1243            y \<in> Domain (depend s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (depend s)\<^sup>+)"
  1243            y \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (RAG s)\<^sup>+)"
  1244     show "x \<in> Domain (depend s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (depend s)\<^sup>+)"
  1244     show "x \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+)"
  1245     proof
  1245     proof
  1246       assume x_d: "x \<in> Domain (depend s)"
  1246       assume x_d: "x \<in> Domain (RAG s)"
  1247       show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (depend s)\<^sup>+"
  1247       show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+"
  1248       proof(cases x)
  1248       proof(cases x)
  1249         case (Th th)
  1249         case (Th th)
  1250         from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> depend s" by (auto simp:s_depend_def)
  1250         from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> RAG s" by (auto simp:s_RAG_def)
  1251         with Th have x_in_r: "(Cs cs, x) \<in> (depend s)^-1" by simp
  1251         with Th have x_in_r: "(Cs cs, x) \<in> (RAG s)^-1" by simp
  1252         from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> depend s" by blast
  1252         from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> RAG s" by blast
  1253         hence "Cs cs \<in> Domain (depend s)" by auto
  1253         hence "Cs cs \<in> Domain (RAG s)" by auto
  1254         from ih [OF x_in_r this] obtain th'
  1254         from ih [OF x_in_r this] obtain th'
  1255           where th'_ready: " th' \<in> readys s" and cs_in: "(Cs cs, Th th') \<in> (depend s)\<^sup>+" by auto
  1255           where th'_ready: " th' \<in> readys s" and cs_in: "(Cs cs, Th th') \<in> (RAG s)\<^sup>+" by auto
  1256         have "(x, Th th') \<in> (depend s)\<^sup>+" using Th x_in cs_in by auto
  1256         have "(x, Th th') \<in> (RAG s)\<^sup>+" using Th x_in cs_in by auto
  1257         with th'_ready show ?thesis by auto
  1257         with th'_ready show ?thesis by auto
  1258       next
  1258       next
  1259         case (Cs cs)
  1259         case (Cs cs)
  1260         from x_d Cs obtain th' where th'_d: "(Th th', x) \<in> (depend s)^-1" by (auto simp:s_depend_def)
  1260         from x_d Cs obtain th' where th'_d: "(Th th', x) \<in> (RAG s)^-1" by (auto simp:s_RAG_def)
  1261         show ?thesis
  1261         show ?thesis
  1262         proof(cases "th' \<in> readys s")
  1262         proof(cases "th' \<in> readys s")
  1263           case True
  1263           case True
  1264           from True and th'_d show ?thesis by auto
  1264           from True and th'_d show ?thesis by auto
  1265         next
  1265         next
  1266           case False
  1266           case False
  1267           from th'_d and range_in [OF vt] have "th' \<in> threads s" by auto
  1267           from th'_d and range_in [OF vt] have "th' \<in> threads s" by auto
  1268           with False have "Th th' \<in> Domain (depend s)" 
  1268           with False have "Th th' \<in> Domain (RAG s)" 
  1269             by (auto simp:readys_def wq_def s_waiting_def s_depend_def cs_waiting_def Domain_def)
  1269             by (auto simp:readys_def wq_def s_waiting_def s_RAG_def cs_waiting_def Domain_def)
  1270           from ih [OF th'_d this]
  1270           from ih [OF th'_d this]
  1271           obtain th'' where 
  1271           obtain th'' where 
  1272             th''_r: "th'' \<in> readys s" and 
  1272             th''_r: "th'' \<in> readys s" and 
  1273             th''_in: "(Th th', Th th'') \<in> (depend s)\<^sup>+" by auto
  1273             th''_in: "(Th th', Th th'') \<in> (RAG s)\<^sup>+" by auto
  1274           from th'_d and th''_in 
  1274           from th'_d and th''_in 
  1275           have "(x, Th th'') \<in> (depend s)\<^sup>+" by auto
  1275           have "(x, Th th'') \<in> (RAG s)\<^sup>+" by auto
  1276           with th''_r show ?thesis by auto
  1276           with th''_r show ?thesis by auto
  1277         qed
  1277         qed
  1278       qed
  1278       qed
  1279     qed
  1279     qed
  1280   qed
  1280   qed
  1282 
  1282 
  1283 lemma th_chain_to_ready:
  1283 lemma th_chain_to_ready:
  1284   fixes s th
  1284   fixes s th
  1285   assumes vt: "vt s"
  1285   assumes vt: "vt s"
  1286   and th_in: "th \<in> threads s"
  1286   and th_in: "th \<in> threads s"
  1287   shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (depend s)^+)"
  1287   shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (RAG s)^+)"
  1288 proof(cases "th \<in> readys s")
  1288 proof(cases "th \<in> readys s")
  1289   case True
  1289   case True
  1290   thus ?thesis by auto
  1290   thus ?thesis by auto
  1291 next
  1291 next
  1292   case False
  1292   case False
  1293   from False and th_in have "Th th \<in> Domain (depend s)" 
  1293   from False and th_in have "Th th \<in> Domain (RAG s)" 
  1294     by (auto simp:readys_def s_waiting_def s_depend_def wq_def cs_waiting_def Domain_def)
  1294     by (auto simp:readys_def s_waiting_def s_RAG_def wq_def cs_waiting_def Domain_def)
  1295   from chain_building [rule_format, OF vt this]
  1295   from chain_building [rule_format, OF vt this]
  1296   show ?thesis by auto
  1296   show ?thesis by auto
  1297 qed
  1297 qed
  1298 
  1298 
  1299 lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs"
  1299 lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs"
  1303   by (unfold s_holding_def wq_def cs_holding_def, simp)
  1303   by (unfold s_holding_def wq_def cs_holding_def, simp)
  1304 
  1304 
  1305 lemma holding_unique: "\<lbrakk>holding (s::state) th1 cs; holding s th2 cs\<rbrakk> \<Longrightarrow> th1 = th2"
  1305 lemma holding_unique: "\<lbrakk>holding (s::state) th1 cs; holding s th2 cs\<rbrakk> \<Longrightarrow> th1 = th2"
  1306   by (unfold s_holding_def cs_holding_def, auto)
  1306   by (unfold s_holding_def cs_holding_def, auto)
  1307 
  1307 
  1308 lemma unique_depend: "\<lbrakk>vt s; (n, n1) \<in> depend s; (n, n2) \<in> depend s\<rbrakk> \<Longrightarrow> n1 = n2"
  1308 lemma unique_RAG: "\<lbrakk>vt s; (n, n1) \<in> RAG s; (n, n2) \<in> RAG s\<rbrakk> \<Longrightarrow> n1 = n2"
  1309   apply(unfold s_depend_def, auto, fold waiting_eq holding_eq)
  1309   apply(unfold s_RAG_def, auto, fold waiting_eq holding_eq)
  1310   by(auto elim:waiting_unique holding_unique)
  1310   by(auto elim:waiting_unique holding_unique)
  1311 
  1311 
  1312 lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r"
  1312 lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r"
  1313 by (induct rule:trancl_induct, auto)
  1313 by (induct rule:trancl_induct, auto)
  1314 
  1314 
  1315 lemma dchain_unique:
  1315 lemma dchain_unique:
  1316   assumes vt: "vt s"
  1316   assumes vt: "vt s"
  1317   and th1_d: "(n, Th th1) \<in> (depend s)^+"
  1317   and th1_d: "(n, Th th1) \<in> (RAG s)^+"
  1318   and th1_r: "th1 \<in> readys s"
  1318   and th1_r: "th1 \<in> readys s"
  1319   and th2_d: "(n, Th th2) \<in> (depend s)^+"
  1319   and th2_d: "(n, Th th2) \<in> (RAG s)^+"
  1320   and th2_r: "th2 \<in> readys s"
  1320   and th2_r: "th2 \<in> readys s"
  1321   shows "th1 = th2"
  1321   shows "th1 = th2"
  1322 proof -
  1322 proof -
  1323   { assume neq: "th1 \<noteq> th2"
  1323   { assume neq: "th1 \<noteq> th2"
  1324     hence "Th th1 \<noteq> Th th2" by simp
  1324     hence "Th th1 \<noteq> Th th2" by simp
  1325     from unique_chain [OF _ th1_d th2_d this] and unique_depend [OF vt]
  1325     from unique_chain [OF _ th1_d th2_d this] and unique_RAG [OF vt]
  1326     have "(Th th1, Th th2) \<in> (depend s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (depend s)\<^sup>+" by auto
  1326     have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (RAG s)\<^sup>+" by auto
  1327     hence "False"
  1327     hence "False"
  1328     proof
  1328     proof
  1329       assume "(Th th1, Th th2) \<in> (depend s)\<^sup>+"
  1329       assume "(Th th1, Th th2) \<in> (RAG s)\<^sup>+"
  1330       from trancl_split [OF this]
  1330       from trancl_split [OF this]
  1331       obtain n where dd: "(Th th1, n) \<in> depend s" by auto
  1331       obtain n where dd: "(Th th1, n) \<in> RAG s" by auto
  1332       then obtain cs where eq_n: "n = Cs cs"
  1332       then obtain cs where eq_n: "n = Cs cs"
  1333         by (auto simp:s_depend_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
  1333         by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
  1334       from dd eq_n have "th1 \<notin> readys s"
  1334       from dd eq_n have "th1 \<notin> readys s"
  1335         by (auto simp:readys_def s_depend_def wq_def s_waiting_def cs_waiting_def)
  1335         by (auto simp:readys_def s_RAG_def wq_def s_waiting_def cs_waiting_def)
  1336       with th1_r show ?thesis by auto
  1336       with th1_r show ?thesis by auto
  1337     next
  1337     next
  1338       assume "(Th th2, Th th1) \<in> (depend s)\<^sup>+"
  1338       assume "(Th th2, Th th1) \<in> (RAG s)\<^sup>+"
  1339       from trancl_split [OF this]
  1339       from trancl_split [OF this]
  1340       obtain n where dd: "(Th th2, n) \<in> depend s" by auto
  1340       obtain n where dd: "(Th th2, n) \<in> RAG s" by auto
  1341       then obtain cs where eq_n: "n = Cs cs"
  1341       then obtain cs where eq_n: "n = Cs cs"
  1342         by (auto simp:s_depend_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
  1342         by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
  1343       from dd eq_n have "th2 \<notin> readys s"
  1343       from dd eq_n have "th2 \<notin> readys s"
  1344         by (auto simp:readys_def wq_def s_depend_def s_waiting_def cs_waiting_def)
  1344         by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def)
  1345       with th2_r show ?thesis by auto
  1345       with th2_r show ?thesis by auto
  1346     qed
  1346     qed
  1347   } thus ?thesis by auto
  1347   } thus ?thesis by auto
  1348 qed
  1348 qed
  1349              
  1349              
  1353   assumes vt: "vt (P th cs#s)"
  1353   assumes vt: "vt (P th cs#s)"
  1354   and "wq s cs = []"
  1354   and "wq s cs = []"
  1355   shows "holdents (P th cs#s) th = holdents s th \<union> {cs}"
  1355   shows "holdents (P th cs#s) th = holdents s th \<union> {cs}"
  1356 proof -
  1356 proof -
  1357   from assms show ?thesis
  1357   from assms show ?thesis
  1358   unfolding  holdents_test step_depend_p[OF vt] by (auto)
  1358   unfolding  holdents_test step_RAG_p[OF vt] by (auto)
  1359 qed
  1359 qed
  1360 
  1360 
  1361 lemma step_holdents_p_eq:
  1361 lemma step_holdents_p_eq:
  1362   fixes th cs s
  1362   fixes th cs s
  1363   assumes vt: "vt (P th cs#s)"
  1363   assumes vt: "vt (P th cs#s)"
  1364   and "wq s cs \<noteq> []"
  1364   and "wq s cs \<noteq> []"
  1365   shows "holdents (P th cs#s) th = holdents s th"
  1365   shows "holdents (P th cs#s) th = holdents s th"
  1366 proof -
  1366 proof -
  1367   from assms show ?thesis
  1367   from assms show ?thesis
  1368   unfolding  holdents_test step_depend_p[OF vt] by auto
  1368   unfolding  holdents_test step_RAG_p[OF vt] by auto
  1369 qed
  1369 qed
  1370 
  1370 
  1371 
  1371 
  1372 lemma finite_holding:
  1372 lemma finite_holding:
  1373   fixes s th cs
  1373   fixes s th cs
  1374   assumes vt: "vt s"
  1374   assumes vt: "vt s"
  1375   shows "finite (holdents s th)"
  1375   shows "finite (holdents s th)"
  1376 proof -
  1376 proof -
  1377   let ?F = "\<lambda> (x, y). the_cs x"
  1377   let ?F = "\<lambda> (x, y). the_cs x"
  1378   from finite_depend [OF vt]
  1378   from finite_RAG [OF vt]
  1379   have "finite (depend s)" .
  1379   have "finite (RAG s)" .
  1380   hence "finite (?F `(depend s))" by simp
  1380   hence "finite (?F `(RAG s))" by simp
  1381   moreover have "{cs . (Cs cs, Th th) \<in> depend s} \<subseteq> \<dots>" 
  1381   moreover have "{cs . (Cs cs, Th th) \<in> RAG s} \<subseteq> \<dots>" 
  1382   proof -
  1382   proof -
  1383     { have h: "\<And> a A f. a \<in> A \<Longrightarrow> f a \<in> f ` A" by auto
  1383     { have h: "\<And> a A f. a \<in> A \<Longrightarrow> f a \<in> f ` A" by auto
  1384       fix x assume "(Cs x, Th th) \<in> depend s"
  1384       fix x assume "(Cs x, Th th) \<in> RAG s"
  1385       hence "?F (Cs x, Th th) \<in> ?F `(depend s)" by (rule h)
  1385       hence "?F (Cs x, Th th) \<in> ?F `(RAG s)" by (rule h)
  1386       moreover have "?F (Cs x, Th th) = x" by simp
  1386       moreover have "?F (Cs x, Th th) = x" by simp
  1387       ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` depend s" by simp 
  1387       ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` RAG s" by simp 
  1388     } thus ?thesis by auto
  1388     } thus ?thesis by auto
  1389   qed
  1389   qed
  1390   ultimately show ?thesis by (unfold holdents_test, auto intro:finite_subset)
  1390   ultimately show ?thesis by (unfold holdents_test, auto intro:finite_subset)
  1391 qed
  1391 qed
  1392 
  1392 
  1395   assumes vtv: "vt (V thread cs#s)"
  1395   assumes vtv: "vt (V thread cs#s)"
  1396   shows "(cntCS (V thread cs#s) thread + 1) = cntCS s thread"
  1396   shows "(cntCS (V thread cs#s) thread + 1) = cntCS s thread"
  1397 proof -
  1397 proof -
  1398   from step_back_step[OF vtv]
  1398   from step_back_step[OF vtv]
  1399   have cs_in: "cs \<in> holdents s thread" 
  1399   have cs_in: "cs \<in> holdents s thread" 
  1400     apply (cases, unfold holdents_test s_depend_def, simp)
  1400     apply (cases, unfold holdents_test s_RAG_def, simp)
  1401     by (unfold cs_holding_def s_holding_def wq_def, auto)
  1401     by (unfold cs_holding_def s_holding_def wq_def, auto)
  1402   moreover have cs_not_in: 
  1402   moreover have cs_not_in: 
  1403     "(holdents (V thread cs#s) thread) = holdents s thread - {cs}"
  1403     "(holdents (V thread cs#s) thread) = holdents s thread - {cs}"
  1404     apply (insert wq_distinct[OF step_back_vt[OF vtv], of cs])
  1404     apply (insert wq_distinct[OF step_back_vt[OF vtv], of cs])
  1405     apply (unfold holdents_test, unfold step_depend_v[OF vtv],
  1405     apply (unfold holdents_test, unfold step_RAG_v[OF vtv],
  1406             auto simp:next_th_def)
  1406             auto simp:next_th_def)
  1407   proof -
  1407   proof -
  1408     fix rest
  1408     fix rest
  1409     assume dst: "distinct (rest::thread list)"
  1409     assume dst: "distinct (rest::thread list)"
  1410       and ne: "rest \<noteq> []"
  1410       and ne: "rest \<noteq> []"
  1423     next
  1423     next
  1424       fix x assume " distinct x \<and> set x = set rest" with ne
  1424       fix x assume " distinct x \<and> set x = set rest" with ne
  1425       show "x \<noteq> []" by auto
  1425       show "x \<noteq> []" by auto
  1426     qed
  1426     qed
  1427     ultimately 
  1427     ultimately 
  1428     show "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> depend s"
  1428     show "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
  1429       by auto
  1429       by auto
  1430   next
  1430   next
  1431     fix rest
  1431     fix rest
  1432     assume dst: "distinct (rest::thread list)"
  1432     assume dst: "distinct (rest::thread list)"
  1433       and ne: "rest \<noteq> []"
  1433       and ne: "rest \<noteq> []"
  1492             wq_def cs_waiting_def Let_def)
  1492             wq_def cs_waiting_def Let_def)
  1493         from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
  1493         from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
  1494         from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
  1494         from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
  1495         have eq_cncs: "cntCS (e#s) th = cntCS s th"
  1495         have eq_cncs: "cntCS (e#s) th = cntCS s th"
  1496           unfolding cntCS_def holdents_test
  1496           unfolding cntCS_def holdents_test
  1497           by (simp add:depend_create_unchanged eq_e)
  1497           by (simp add:RAG_create_unchanged eq_e)
  1498         { assume "th \<noteq> thread"
  1498         { assume "th \<noteq> thread"
  1499           with eq_readys eq_e
  1499           with eq_readys eq_e
  1500           have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
  1500           have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
  1501                       (th \<in> readys (s) \<or> th \<notin> threads (s))" 
  1501                       (th \<in> readys (s) \<or> th \<notin> threads (s))" 
  1502             by (simp add:threads.simps)
  1502             by (simp add:threads.simps)
  1517       and no_hold: "holdents s thread = {}"
  1517       and no_hold: "holdents s thread = {}"
  1518       from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
  1518       from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
  1519       from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
  1519       from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
  1520       have eq_cncs: "cntCS (e#s) th = cntCS s th"
  1520       have eq_cncs: "cntCS (e#s) th = cntCS s th"
  1521         unfolding cntCS_def holdents_test
  1521         unfolding cntCS_def holdents_test
  1522         by (simp add:depend_exit_unchanged eq_e)
  1522         by (simp add:RAG_exit_unchanged eq_e)
  1523       { assume "th \<noteq> thread"
  1523       { assume "th \<noteq> thread"
  1524         with eq_e
  1524         with eq_e
  1525         have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
  1525         have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
  1526           (th \<in> readys (s) \<or> th \<notin> threads (s))" 
  1526           (th \<in> readys (s) \<or> th \<notin> threads (s))" 
  1527           apply (simp add:threads.simps readys_def)
  1527           apply (simp add:threads.simps readys_def)
  1542       } ultimately show ?thesis by blast
  1542       } ultimately show ?thesis by blast
  1543     next
  1543     next
  1544       case (thread_P thread cs)
  1544       case (thread_P thread cs)
  1545       assume eq_e: "e = P thread cs"
  1545       assume eq_e: "e = P thread cs"
  1546         and is_runing: "thread \<in> runing s"
  1546         and is_runing: "thread \<in> runing s"
  1547         and no_dep: "(Cs cs, Th thread) \<notin> (depend s)\<^sup>+"
  1547         and no_dep: "(Cs cs, Th thread) \<notin> (RAG s)\<^sup>+"
  1548       from thread_P vt stp ih  have vtp: "vt (P thread cs#s)" by auto
  1548       from thread_P vt stp ih  have vtp: "vt (P thread cs#s)" by auto
  1549       show ?thesis 
  1549       show ?thesis 
  1550       proof -
  1550       proof -
  1551         { have hh: "\<And> A B C. (B = C) \<Longrightarrow> (A \<and> B) = (A \<and> C)" by blast
  1551         { have hh: "\<And> A B C. (B = C) \<Longrightarrow> (A \<and> B) = (A \<and> C)" by blast
  1552           assume neq_th: "th \<noteq> thread"
  1552           assume neq_th: "th \<noteq> thread"
  1559             apply (subgoal_tac "wq_fun (schs s) cs \<noteq> []", auto)
  1559             apply (subgoal_tac "wq_fun (schs s) cs \<noteq> []", auto)
  1560             apply (erule_tac x = cs in allE, auto)
  1560             apply (erule_tac x = cs in allE, auto)
  1561             by (case_tac "(wq_fun (schs s) cs)", auto)
  1561             by (case_tac "(wq_fun (schs s) cs)", auto)
  1562           moreover from neq_th eq_e have "cntCS (e # s) th = cntCS s th"
  1562           moreover from neq_th eq_e have "cntCS (e # s) th = cntCS s th"
  1563             apply (simp add:cntCS_def holdents_test)
  1563             apply (simp add:cntCS_def holdents_test)
  1564             by (unfold  step_depend_p [OF vtp], auto)
  1564             by (unfold  step_RAG_p [OF vtp], auto)
  1565           moreover from eq_e neq_th have "cntP (e # s) th = cntP s th"
  1565           moreover from eq_e neq_th have "cntP (e # s) th = cntP s th"
  1566             by (simp add:cntP_def count_def)
  1566             by (simp add:cntP_def count_def)
  1567           moreover from eq_e neq_th have "cntV (e#s) th = cntV s th"
  1567           moreover from eq_e neq_th have "cntV (e#s) th = cntV s th"
  1568             by (simp add:cntV_def count_def)
  1568             by (simp add:cntV_def count_def)
  1569           moreover from eq_e neq_th have "threads (e#s) = threads s" by simp
  1569           moreover from eq_e neq_th have "threads (e#s) = threads s" by simp
  1580             show ?thesis
  1580             show ?thesis
  1581             proof (cases "wq s cs = []")
  1581             proof (cases "wq s cs = []")
  1582               case True
  1582               case True
  1583               with is_runing
  1583               with is_runing
  1584               have "th \<in> readys (e#s)"
  1584               have "th \<in> readys (e#s)"
  1585                 apply (unfold eq_e wq_def, unfold readys_def s_depend_def)
  1585                 apply (unfold eq_e wq_def, unfold readys_def s_RAG_def)
  1586                 apply (simp add: wq_def[symmetric] runing_def eq_th s_waiting_def)
  1586                 apply (simp add: wq_def[symmetric] runing_def eq_th s_waiting_def)
  1587                 by (auto simp:readys_def wq_def Let_def s_waiting_def wq_def)
  1587                 by (auto simp:readys_def wq_def Let_def s_waiting_def wq_def)
  1588               moreover have "cntCS (e # s) th = 1 + cntCS s th"
  1588               moreover have "cntCS (e # s) th = 1 + cntCS s th"
  1589               proof -
  1589               proof -
  1590                 have "card {csa. csa = cs \<or> (Cs csa, Th thread) \<in> depend s} =
  1590                 have "card {csa. csa = cs \<or> (Cs csa, Th thread) \<in> RAG s} =
  1591                   Suc (card {cs. (Cs cs, Th thread) \<in> depend s})" (is "card ?L = Suc (card ?R)")
  1591                   Suc (card {cs. (Cs cs, Th thread) \<in> RAG s})" (is "card ?L = Suc (card ?R)")
  1592                 proof -
  1592                 proof -
  1593                   have "?L = insert cs ?R" by auto
  1593                   have "?L = insert cs ?R" by auto
  1594                   moreover have "card \<dots> = Suc (card (?R - {cs}))" 
  1594                   moreover have "card \<dots> = Suc (card (?R - {cs}))" 
  1595                   proof(rule card_insert)
  1595                   proof(rule card_insert)
  1596                     from finite_holding [OF vt, of thread]
  1596                     from finite_holding [OF vt, of thread]
  1597                     show " finite {cs. (Cs cs, Th thread) \<in> depend s}"
  1597                     show " finite {cs. (Cs cs, Th thread) \<in> RAG s}"
  1598                       by (unfold holdents_test, simp)
  1598                       by (unfold holdents_test, simp)
  1599                   qed
  1599                   qed
  1600                   moreover have "?R - {cs} = ?R"
  1600                   moreover have "?R - {cs} = ?R"
  1601                   proof -
  1601                   proof -
  1602                     have "cs \<notin> ?R"
  1602                     have "cs \<notin> ?R"
  1603                     proof
  1603                     proof
  1604                       assume "cs \<in> {cs. (Cs cs, Th thread) \<in> depend s}"
  1604                       assume "cs \<in> {cs. (Cs cs, Th thread) \<in> RAG s}"
  1605                       with no_dep show False by auto
  1605                       with no_dep show False by auto
  1606                     qed
  1606                     qed
  1607                     thus ?thesis by auto
  1607                     thus ?thesis by auto
  1608                   qed
  1608                   qed
  1609                   ultimately show ?thesis by auto
  1609                   ultimately show ?thesis by auto
  1610                 qed
  1610                 qed
  1611                 thus ?thesis
  1611                 thus ?thesis
  1612                   apply (unfold eq_e eq_th cntCS_def)
  1612                   apply (unfold eq_e eq_th cntCS_def)
  1613                   apply (simp add: holdents_test)
  1613                   apply (simp add: holdents_test)
  1614                   by (unfold step_depend_p [OF vtp], auto simp:True)
  1614                   by (unfold step_RAG_p [OF vtp], auto simp:True)
  1615               qed
  1615               qed
  1616               moreover from is_runing have "th \<in> readys s"
  1616               moreover from is_runing have "th \<in> readys s"
  1617                 by (simp add:runing_def eq_th)
  1617                 by (simp add:runing_def eq_th)
  1618               moreover note eq_cnp eq_cnv ih [of th]
  1618               moreover note eq_cnp eq_cnv ih [of th]
  1619               ultimately show ?thesis by auto
  1619               ultimately show ?thesis by auto
  1636                 show False by (fold eq_e, auto)
  1636                 show False by (fold eq_e, auto)
  1637               qed
  1637               qed
  1638               moreover from is_runing have "th \<in> threads (e#s)" 
  1638               moreover from is_runing have "th \<in> threads (e#s)" 
  1639                 by (unfold eq_e, auto simp:runing_def readys_def eq_th)
  1639                 by (unfold eq_e, auto simp:runing_def readys_def eq_th)
  1640               moreover have "cntCS (e # s) th = cntCS s th"
  1640               moreover have "cntCS (e # s) th = cntCS s th"
  1641                 apply (unfold cntCS_def holdents_test eq_e step_depend_p[OF vtp])
  1641                 apply (unfold cntCS_def holdents_test eq_e step_RAG_p[OF vtp])
  1642                 by (auto simp:False)
  1642                 by (auto simp:False)
  1643               moreover note eq_cnp eq_cnv ih[of th]
  1643               moreover note eq_cnp eq_cnv ih[of th]
  1644               moreover from is_runing have "th \<in> readys s"
  1644               moreover from is_runing have "th \<in> readys s"
  1645                 by (simp add:runing_def eq_th)
  1645                 by (simp add:runing_def eq_th)
  1646               ultimately show ?thesis by auto
  1646               ultimately show ?thesis by auto
  1733             case False
  1733             case False
  1734             have "(th \<in> readys (e # s)) = (th \<in> readys s)"
  1734             have "(th \<in> readys (e # s)) = (th \<in> readys s)"
  1735               apply (insert step_back_vt[OF vtv])
  1735               apply (insert step_back_vt[OF vtv])
  1736               by (unfold eq_e, rule readys_v_eq [OF _ neq_th eq_wq False], auto)
  1736               by (unfold eq_e, rule readys_v_eq [OF _ neq_th eq_wq False], auto)
  1737             moreover have "cntCS (e#s) th = cntCS s th"
  1737             moreover have "cntCS (e#s) th = cntCS s th"
  1738               apply (insert neq_th, unfold eq_e cntCS_def holdents_test step_depend_v[OF vtv], auto)
  1738               apply (insert neq_th, unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
  1739               proof -
  1739               proof -
  1740                 have "{csa. (Cs csa, Th th) \<in> depend s \<or> csa = cs \<and> next_th s thread cs th} =
  1740                 have "{csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
  1741                       {cs. (Cs cs, Th th) \<in> depend s}"
  1741                       {cs. (Cs cs, Th th) \<in> RAG s}"
  1742                 proof -
  1742                 proof -
  1743                   from False eq_wq
  1743                   from False eq_wq
  1744                   have " next_th s thread cs th \<Longrightarrow> (Cs cs, Th th) \<in> depend s"
  1744                   have " next_th s thread cs th \<Longrightarrow> (Cs cs, Th th) \<in> RAG s"
  1745                     apply (unfold next_th_def, auto)
  1745                     apply (unfold next_th_def, auto)
  1746                   proof -
  1746                   proof -
  1747                     assume ne: "rest \<noteq> []"
  1747                     assume ne: "rest \<noteq> []"
  1748                       and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
  1748                       and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
  1749                       and eq_wq: "wq s cs = thread # rest"
  1749                       and eq_wq: "wq s cs = thread # rest"
  1757                     next
  1757                     next
  1758                       fix x assume "distinct x \<and> set x = set rest"
  1758                       fix x assume "distinct x \<and> set x = set rest"
  1759                       with ne show "x \<noteq> []" by auto
  1759                       with ne show "x \<noteq> []" by auto
  1760                     qed
  1760                     qed
  1761                     ultimately show 
  1761                     ultimately show 
  1762                       "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> depend s"
  1762                       "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
  1763                       by auto
  1763                       by auto
  1764                   qed    
  1764                   qed    
  1765                   thus ?thesis by auto
  1765                   thus ?thesis by auto
  1766                 qed
  1766                 qed
  1767                 thus "card {csa. (Cs csa, Th th) \<in> depend s \<or> csa = cs \<and> next_th s thread cs th} =
  1767                 thus "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
  1768                              card {cs. (Cs cs, Th th) \<in> depend s}" by simp 
  1768                              card {cs. (Cs cs, Th th) \<in> RAG s}" by simp 
  1769               qed
  1769               qed
  1770             moreover note ih eq_cnp eq_cnv eq_threads
  1770             moreover note ih eq_cnp eq_cnv eq_threads
  1771             ultimately show ?thesis by auto
  1771             ultimately show ?thesis by auto
  1772           next
  1772           next
  1773             case True
  1773             case True
  1794               qed
  1794               qed
  1795               moreover have "cntCS (e#s) th = cntCS s th" 
  1795               moreover have "cntCS (e#s) th = cntCS s th" 
  1796               proof -
  1796               proof -
  1797                 from eq_wq and  th_in and neq_hd
  1797                 from eq_wq and  th_in and neq_hd
  1798                 have "(holdents (e # s) th) = (holdents s th)"
  1798                 have "(holdents (e # s) th) = (holdents s th)"
  1799                   apply (unfold eq_e step_depend_v[OF vtv], 
  1799                   apply (unfold eq_e step_RAG_v[OF vtv], 
  1800                          auto simp:next_th_def eq_set s_depend_def holdents_test wq_def
  1800                          auto simp:next_th_def eq_set s_RAG_def holdents_test wq_def
  1801                                    Let_def cs_holding_def)
  1801                                    Let_def cs_holding_def)
  1802                   by (insert wq_distinct[OF step_back_vt[OF vtv], of cs], auto simp:wq_def)
  1802                   by (insert wq_distinct[OF step_back_vt[OF vtv], of cs], auto simp:wq_def)
  1803                 thus ?thesis by (simp add:cntCS_def)
  1803                 thus ?thesis by (simp add:cntCS_def)
  1804               qed
  1804               qed
  1805               moreover note ih eq_cnp eq_cnv eq_threads
  1805               moreover note ih eq_cnp eq_cnv eq_threads
  1860                   show ?thesis .
  1860                   show ?thesis .
  1861                 qed
  1861                 qed
  1862                 ultimately show ?thesis using ih by auto
  1862                 ultimately show ?thesis using ih by auto
  1863               qed
  1863               qed
  1864               moreover from True neq_th have "cntCS (e # s) th = 1 + cntCS s th"
  1864               moreover from True neq_th have "cntCS (e # s) th = 1 + cntCS s th"
  1865                 apply (unfold cntCS_def holdents_test eq_e step_depend_v[OF vtv], auto)
  1865                 apply (unfold cntCS_def holdents_test eq_e step_RAG_v[OF vtv], auto)
  1866               proof -
  1866               proof -
  1867                 show "card {csa. (Cs csa, Th th) \<in> depend s \<or> csa = cs} =
  1867                 show "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs} =
  1868                                Suc (card {cs. (Cs cs, Th th) \<in> depend s})"
  1868                                Suc (card {cs. (Cs cs, Th th) \<in> RAG s})"
  1869                   (is "card ?A = Suc (card ?B)")
  1869                   (is "card ?A = Suc (card ?B)")
  1870                 proof -
  1870                 proof -
  1871                   have "?A = insert cs ?B" by auto
  1871                   have "?A = insert cs ?B" by auto
  1872                   hence "card ?A = card (insert cs ?B)" by simp
  1872                   hence "card ?A = card (insert cs ?B)" by simp
  1873                   also have "\<dots> = Suc (card ?B)"
  1873                   also have "\<dots> = Suc (card ?B)"
  1874                   proof(rule card_insert_disjoint)
  1874                   proof(rule card_insert_disjoint)
  1875                     have "?B \<subseteq> ((\<lambda> (x, y). the_cs x) ` depend s)" 
  1875                     have "?B \<subseteq> ((\<lambda> (x, y). the_cs x) ` RAG s)" 
  1876                       apply (auto simp:image_def)
  1876                       apply (auto simp:image_def)
  1877                       by (rule_tac x = "(Cs x, Th th)" in bexI, auto)
  1877                       by (rule_tac x = "(Cs x, Th th)" in bexI, auto)
  1878                     with finite_depend[OF step_back_vt[OF vtv]]
  1878                     with finite_RAG[OF step_back_vt[OF vtv]]
  1879                     show "finite {cs. (Cs cs, Th th) \<in> depend s}" by (auto intro:finite_subset)
  1879                     show "finite {cs. (Cs cs, Th th) \<in> RAG s}" by (auto intro:finite_subset)
  1880                   next
  1880                   next
  1881                     show "cs \<notin> {cs. (Cs cs, Th th) \<in> depend s}"
  1881                     show "cs \<notin> {cs. (Cs cs, Th th) \<in> RAG s}"
  1882                     proof
  1882                     proof
  1883                       assume "cs \<in> {cs. (Cs cs, Th th) \<in> depend s}"
  1883                       assume "cs \<in> {cs. (Cs cs, Th th) \<in> RAG s}"
  1884                       hence "(Cs cs, Th th) \<in> depend s" by simp
  1884                       hence "(Cs cs, Th th) \<in> RAG s" by simp
  1885                       with True neq_th eq_wq show False
  1885                       with True neq_th eq_wq show False
  1886                         by (auto simp:next_th_def s_depend_def cs_holding_def)
  1886                         by (auto simp:next_th_def s_RAG_def cs_holding_def)
  1887                     qed
  1887                     qed
  1888                   qed
  1888                   qed
  1889                   finally show ?thesis .
  1889                   finally show ?thesis .
  1890                 qed
  1890                 qed
  1891               qed
  1891               qed
  1903       proof -
  1903       proof -
  1904         from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
  1904         from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
  1905         from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
  1905         from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
  1906         have eq_cncs: "cntCS (e#s) th = cntCS s th"
  1906         have eq_cncs: "cntCS (e#s) th = cntCS s th"
  1907           unfolding cntCS_def holdents_test
  1907           unfolding cntCS_def holdents_test
  1908           by (simp add:depend_set_unchanged eq_e)
  1908           by (simp add:RAG_set_unchanged eq_e)
  1909         from eq_e have eq_readys: "readys (e#s) = readys s" 
  1909         from eq_e have eq_readys: "readys (e#s) = readys s" 
  1910           by (simp add:readys_def cs_waiting_def s_waiting_def wq_def,
  1910           by (simp add:readys_def cs_waiting_def s_waiting_def wq_def,
  1911                   auto simp:Let_def)
  1911                   auto simp:Let_def)
  1912         { assume "th \<noteq> thread"
  1912         { assume "th \<noteq> thread"
  1913           with eq_readys eq_e
  1913           with eq_readys eq_e
  1929     qed
  1929     qed
  1930   next
  1930   next
  1931     case vt_nil
  1931     case vt_nil
  1932     show ?case 
  1932     show ?case 
  1933       by (unfold cntP_def cntV_def cntCS_def, 
  1933       by (unfold cntP_def cntV_def cntCS_def, 
  1934         auto simp:count_def holdents_test s_depend_def wq_def cs_holding_def)
  1934         auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
  1935   qed
  1935   qed
  1936 qed
  1936 qed
  1937 
  1937 
  1938 lemma not_thread_cncs:
  1938 lemma not_thread_cncs:
  1939   fixes th s
  1939   fixes th s
  1953       case (thread_create thread prio)
  1953       case (thread_create thread prio)
  1954       assume eq_e: "e = Create thread prio"
  1954       assume eq_e: "e = Create thread prio"
  1955         and not_in': "thread \<notin> threads s"
  1955         and not_in': "thread \<notin> threads s"
  1956       have "cntCS (e # s) th = cntCS s th"
  1956       have "cntCS (e # s) th = cntCS s th"
  1957         apply (unfold eq_e cntCS_def holdents_test)
  1957         apply (unfold eq_e cntCS_def holdents_test)
  1958         by (simp add:depend_create_unchanged)
  1958         by (simp add:RAG_create_unchanged)
  1959       moreover have "th \<notin> threads s" 
  1959       moreover have "th \<notin> threads s" 
  1960       proof -
  1960       proof -
  1961         from not_in eq_e show ?thesis by simp
  1961         from not_in eq_e show ?thesis by simp
  1962       qed
  1962       qed
  1963       moreover note ih ultimately show ?thesis by auto
  1963       moreover note ih ultimately show ?thesis by auto
  1965       case (thread_exit thread)
  1965       case (thread_exit thread)
  1966       assume eq_e: "e = Exit thread"
  1966       assume eq_e: "e = Exit thread"
  1967       and nh: "holdents s thread = {}"
  1967       and nh: "holdents s thread = {}"
  1968       have eq_cns: "cntCS (e # s) th = cntCS s th"
  1968       have eq_cns: "cntCS (e # s) th = cntCS s th"
  1969         apply (unfold eq_e cntCS_def holdents_test)
  1969         apply (unfold eq_e cntCS_def holdents_test)
  1970         by (simp add:depend_exit_unchanged)
  1970         by (simp add:RAG_exit_unchanged)
  1971       show ?thesis
  1971       show ?thesis
  1972       proof(cases "th = thread")
  1972       proof(cases "th = thread")
  1973         case True
  1973         case True
  1974         have "cntCS s th = 0" by (unfold cntCS_def, auto simp:nh True)
  1974         have "cntCS s th = 0" by (unfold cntCS_def, auto simp:nh True)
  1975         with eq_cns show ?thesis by simp
  1975         with eq_cns show ?thesis by simp
  1991           by (simp add:runing_def readys_def)
  1991           by (simp add:runing_def readys_def)
  1992         ultimately show ?thesis by auto
  1992         ultimately show ?thesis by auto
  1993       qed
  1993       qed
  1994       hence "cntCS (e # s) th  = cntCS s th "
  1994       hence "cntCS (e # s) th  = cntCS s th "
  1995         apply (unfold cntCS_def holdents_test eq_e)
  1995         apply (unfold cntCS_def holdents_test eq_e)
  1996         by (unfold step_depend_p[OF vtp], auto)
  1996         by (unfold step_RAG_p[OF vtp], auto)
  1997       moreover have "cntCS s th = 0"
  1997       moreover have "cntCS s th = 0"
  1998       proof(rule ih)
  1998       proof(rule ih)
  1999         from not_in eq_e show "th \<notin> threads s" by simp
  1999         from not_in eq_e show "th \<notin> threads s" by simp
  2000       qed
  2000       qed
  2001       ultimately show ?thesis by simp
  2001       ultimately show ?thesis by simp
  2033         from wq_threads[OF step_back_vt[OF vtv], OF this] and ni
  2033         from wq_threads[OF step_back_vt[OF vtv], OF this] and ni
  2034         show False by auto
  2034         show False by auto
  2035       qed
  2035       qed
  2036       moreover note neq_th eq_wq
  2036       moreover note neq_th eq_wq
  2037       ultimately have "cntCS (e # s) th  = cntCS s th"
  2037       ultimately have "cntCS (e # s) th  = cntCS s th"
  2038         by (unfold eq_e cntCS_def holdents_test step_depend_v[OF vtv], auto)
  2038         by (unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
  2039       moreover have "cntCS s th = 0"
  2039       moreover have "cntCS s th = 0"
  2040       proof(rule ih)
  2040       proof(rule ih)
  2041         from not_in eq_e show "th \<notin> threads s" by simp
  2041         from not_in eq_e show "th \<notin> threads s" by simp
  2042       qed
  2042       qed
  2043       ultimately show ?thesis by simp
  2043       ultimately show ?thesis by simp
  2048         and is_runing: "thread \<in> runing s"
  2048         and is_runing: "thread \<in> runing s"
  2049       from not_in and eq_e have "th \<notin> threads s" by auto
  2049       from not_in and eq_e have "th \<notin> threads s" by auto
  2050       from ih [OF this] and eq_e
  2050       from ih [OF this] and eq_e
  2051       show ?thesis 
  2051       show ?thesis 
  2052         apply (unfold eq_e cntCS_def holdents_test)
  2052         apply (unfold eq_e cntCS_def holdents_test)
  2053         by (simp add:depend_set_unchanged)
  2053         by (simp add:RAG_set_unchanged)
  2054     qed
  2054     qed
  2055     next
  2055     next
  2056       case vt_nil
  2056       case vt_nil
  2057       show ?case
  2057       show ?case
  2058       by (unfold cntCS_def, 
  2058       by (unfold cntCS_def, 
  2059         auto simp:count_def holdents_test s_depend_def wq_def cs_holding_def)
  2059         auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
  2060   qed
  2060   qed
  2061 qed
  2061 qed
  2062 
  2062 
  2063 lemma eq_waiting: "waiting (wq (s::state)) th cs = waiting s th cs"
  2063 lemma eq_waiting: "waiting (wq (s::state)) th cs = waiting s th cs"
  2064   by (auto simp:s_waiting_def cs_waiting_def wq_def)
  2064   by (auto simp:s_waiting_def cs_waiting_def wq_def)
  2065 
  2065 
  2066 lemma dm_depend_threads:
  2066 lemma dm_RAG_threads:
  2067   fixes th s
  2067   fixes th s
  2068   assumes vt: "vt s"
  2068   assumes vt: "vt s"
  2069   and in_dom: "(Th th) \<in> Domain (depend s)"
  2069   and in_dom: "(Th th) \<in> Domain (RAG s)"
  2070   shows "th \<in> threads s"
  2070   shows "th \<in> threads s"
  2071 proof -
  2071 proof -
  2072   from in_dom obtain n where "(Th th, n) \<in> depend s" by auto
  2072   from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto
  2073   moreover from depend_target_th[OF this] obtain cs where "n = Cs cs" by auto
  2073   moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto
  2074   ultimately have "(Th th, Cs cs) \<in> depend s" by simp
  2074   ultimately have "(Th th, Cs cs) \<in> RAG s" by simp
  2075   hence "th \<in> set (wq s cs)"
  2075   hence "th \<in> set (wq s cs)"
  2076     by (unfold s_depend_def, auto simp:cs_waiting_def)
  2076     by (unfold s_RAG_def, auto simp:cs_waiting_def)
  2077   from wq_threads [OF vt this] show ?thesis .
  2077   from wq_threads [OF vt this] show ?thesis .
  2078 qed
  2078 qed
  2079 
  2079 
  2080 lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th"
  2080 lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th"
  2081 unfolding cp_def wq_def
  2081 unfolding cp_def wq_def
  2110     proof -
  2110     proof -
  2111       have "finite ?A" 
  2111       have "finite ?A" 
  2112       proof -
  2112       proof -
  2113         have "finite (dependants (wq s) th1)"
  2113         have "finite (dependants (wq s) th1)"
  2114         proof-
  2114         proof-
  2115           have "finite {th'. (Th th', Th th1) \<in> (depend (wq s))\<^sup>+}"
  2115           have "finite {th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+}"
  2116           proof -
  2116           proof -
  2117             let ?F = "\<lambda> (x, y). the_th x"
  2117             let ?F = "\<lambda> (x, y). the_th x"
  2118             have "{th'. (Th th', Th th1) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (wq s))\<^sup>+)"
  2118             have "{th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
  2119               apply (auto simp:image_def)
  2119               apply (auto simp:image_def)
  2120               by (rule_tac x = "(Th x, Th th1)" in bexI, auto)
  2120               by (rule_tac x = "(Th x, Th th1)" in bexI, auto)
  2121             moreover have "finite \<dots>"
  2121             moreover have "finite \<dots>"
  2122             proof -
  2122             proof -
  2123               from finite_depend[OF vt] have "finite (depend s)" .
  2123               from finite_RAG[OF vt] have "finite (RAG s)" .
  2124               hence "finite ((depend (wq s))\<^sup>+)"
  2124               hence "finite ((RAG (wq s))\<^sup>+)"
  2125                 apply (unfold finite_trancl)
  2125                 apply (unfold finite_trancl)
  2126                 by (auto simp: s_depend_def cs_depend_def wq_def)
  2126                 by (auto simp: s_RAG_def cs_RAG_def wq_def)
  2127               thus ?thesis by auto
  2127               thus ?thesis by auto
  2128             qed
  2128             qed
  2129             ultimately show ?thesis by (auto intro:finite_subset)
  2129             ultimately show ?thesis by (auto intro:finite_subset)
  2130           qed
  2130           qed
  2131           thus ?thesis by (simp add:cs_dependants_def)
  2131           thus ?thesis by (simp add:cs_dependants_def)
  2149     proof -
  2149     proof -
  2150       have "finite ?B" 
  2150       have "finite ?B" 
  2151       proof -
  2151       proof -
  2152         have "finite (dependants (wq s) th2)"
  2152         have "finite (dependants (wq s) th2)"
  2153         proof-
  2153         proof-
  2154           have "finite {th'. (Th th', Th th2) \<in> (depend (wq s))\<^sup>+}"
  2154           have "finite {th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+}"
  2155           proof -
  2155           proof -
  2156             let ?F = "\<lambda> (x, y). the_th x"
  2156             let ?F = "\<lambda> (x, y). the_th x"
  2157             have "{th'. (Th th', Th th2) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (wq s))\<^sup>+)"
  2157             have "{th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
  2158               apply (auto simp:image_def)
  2158               apply (auto simp:image_def)
  2159               by (rule_tac x = "(Th x, Th th2)" in bexI, auto)
  2159               by (rule_tac x = "(Th x, Th th2)" in bexI, auto)
  2160             moreover have "finite \<dots>"
  2160             moreover have "finite \<dots>"
  2161             proof -
  2161             proof -
  2162               from finite_depend[OF vt] have "finite (depend s)" .
  2162               from finite_RAG[OF vt] have "finite (RAG s)" .
  2163               hence "finite ((depend (wq s))\<^sup>+)"
  2163               hence "finite ((RAG (wq s))\<^sup>+)"
  2164                 apply (unfold finite_trancl)
  2164                 apply (unfold finite_trancl)
  2165                 by (auto simp: s_depend_def cs_depend_def wq_def)
  2165                 by (auto simp: s_RAG_def cs_RAG_def wq_def)
  2166               thus ?thesis by auto
  2166               thus ?thesis by auto
  2167             qed
  2167             qed
  2168             ultimately show ?thesis by (auto intro:finite_subset)
  2168             ultimately show ?thesis by (auto intro:finite_subset)
  2169           qed
  2169           qed
  2170           thus ?thesis by (simp add:cs_dependants_def)
  2170           thus ?thesis by (simp add:cs_dependants_def)
  2188   proof (rule preced_unique)
  2188   proof (rule preced_unique)
  2189     from th1_in have "th1' = th1 \<or> (th1' \<in> dependants (wq s) th1)" by simp
  2189     from th1_in have "th1' = th1 \<or> (th1' \<in> dependants (wq s) th1)" by simp
  2190     thus "th1' \<in> threads s"
  2190     thus "th1' \<in> threads s"
  2191     proof
  2191     proof
  2192       assume "th1' \<in> dependants (wq s) th1"
  2192       assume "th1' \<in> dependants (wq s) th1"
  2193       hence "(Th th1') \<in> Domain ((depend s)^+)"
  2193       hence "(Th th1') \<in> Domain ((RAG s)^+)"
  2194         apply (unfold cs_dependants_def cs_depend_def s_depend_def)
  2194         apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
  2195         by (auto simp:Domain_def)
  2195         by (auto simp:Domain_def)
  2196       hence "(Th th1') \<in> Domain (depend s)" by (simp add:trancl_domain)
  2196       hence "(Th th1') \<in> Domain (RAG s)" by (simp add:trancl_domain)
  2197       from dm_depend_threads[OF vt this] show ?thesis .
  2197       from dm_RAG_threads[OF vt this] show ?thesis .
  2198     next
  2198     next
  2199       assume "th1' = th1"
  2199       assume "th1' = th1"
  2200       with runing_1 show ?thesis
  2200       with runing_1 show ?thesis
  2201         by (unfold runing_def readys_def, auto)
  2201         by (unfold runing_def readys_def, auto)
  2202     qed
  2202     qed
  2203   next
  2203   next
  2204     from th2_in have "th2' = th2 \<or> (th2' \<in> dependants (wq s) th2)" by simp
  2204     from th2_in have "th2' = th2 \<or> (th2' \<in> dependants (wq s) th2)" by simp
  2205     thus "th2' \<in> threads s"
  2205     thus "th2' \<in> threads s"
  2206     proof
  2206     proof
  2207       assume "th2' \<in> dependants (wq s) th2"
  2207       assume "th2' \<in> dependants (wq s) th2"
  2208       hence "(Th th2') \<in> Domain ((depend s)^+)"
  2208       hence "(Th th2') \<in> Domain ((RAG s)^+)"
  2209         apply (unfold cs_dependants_def cs_depend_def s_depend_def)
  2209         apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
  2210         by (auto simp:Domain_def)
  2210         by (auto simp:Domain_def)
  2211       hence "(Th th2') \<in> Domain (depend s)" by (simp add:trancl_domain)
  2211       hence "(Th th2') \<in> Domain (RAG s)" by (simp add:trancl_domain)
  2212       from dm_depend_threads[OF vt this] show ?thesis .
  2212       from dm_RAG_threads[OF vt this] show ?thesis .
  2213     next
  2213     next
  2214       assume "th2' = th2"
  2214       assume "th2' = th2"
  2215       with runing_2 show ?thesis
  2215       with runing_2 show ?thesis
  2216         by (unfold runing_def readys_def, auto)
  2216         by (unfold runing_def readys_def, auto)
  2217     qed
  2217     qed
  2225     proof
  2225     proof
  2226       assume "th2' = th2" thus ?thesis using eq_th' eq_th12 by simp
  2226       assume "th2' = th2" thus ?thesis using eq_th' eq_th12 by simp
  2227     next
  2227     next
  2228       assume "th2' \<in> dependants (wq s) th2"
  2228       assume "th2' \<in> dependants (wq s) th2"
  2229       with eq_th12 eq_th' have "th1 \<in> dependants (wq s) th2" by simp
  2229       with eq_th12 eq_th' have "th1 \<in> dependants (wq s) th2" by simp
  2230       hence "(Th th1, Th th2) \<in> (depend s)^+"
  2230       hence "(Th th1, Th th2) \<in> (RAG s)^+"
  2231         by (unfold cs_dependants_def s_depend_def cs_depend_def, simp)
  2231         by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
  2232       hence "Th th1 \<in> Domain ((depend s)^+)" 
  2232       hence "Th th1 \<in> Domain ((RAG s)^+)" 
  2233         apply (unfold cs_dependants_def cs_depend_def s_depend_def)
  2233         apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
  2234         by (auto simp:Domain_def)
  2234         by (auto simp:Domain_def)
  2235       hence "Th th1 \<in> Domain (depend s)" by (simp add:trancl_domain)
  2235       hence "Th th1 \<in> Domain (RAG s)" by (simp add:trancl_domain)
  2236       then obtain n where d: "(Th th1, n) \<in> depend s" by (auto simp:Domain_def)
  2236       then obtain n where d: "(Th th1, n) \<in> RAG s" by (auto simp:Domain_def)
  2237       from depend_target_th [OF this]
  2237       from RAG_target_th [OF this]
  2238       obtain cs' where "n = Cs cs'" by auto
  2238       obtain cs' where "n = Cs cs'" by auto
  2239       with d have "(Th th1, Cs cs') \<in> depend s" by simp
  2239       with d have "(Th th1, Cs cs') \<in> RAG s" by simp
  2240       with runing_1 have "False"
  2240       with runing_1 have "False"
  2241         apply (unfold runing_def readys_def s_depend_def)
  2241         apply (unfold runing_def readys_def s_RAG_def)
  2242         by (auto simp:eq_waiting)
  2242         by (auto simp:eq_waiting)
  2243       thus ?thesis by simp
  2243       thus ?thesis by simp
  2244     qed
  2244     qed
  2245   next
  2245   next
  2246     assume th1'_in: "th1' \<in> dependants (wq s) th1"
  2246     assume th1'_in: "th1' \<in> dependants (wq s) th1"
  2247     from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp
  2247     from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp
  2248     thus ?thesis 
  2248     thus ?thesis 
  2249     proof
  2249     proof
  2250       assume "th2' = th2"
  2250       assume "th2' = th2"
  2251       with th1'_in eq_th12 have "th2 \<in> dependants (wq s) th1" by simp
  2251       with th1'_in eq_th12 have "th2 \<in> dependants (wq s) th1" by simp
  2252       hence "(Th th2, Th th1) \<in> (depend s)^+"
  2252       hence "(Th th2, Th th1) \<in> (RAG s)^+"
  2253         by (unfold cs_dependants_def s_depend_def cs_depend_def, simp)
  2253         by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
  2254       hence "Th th2 \<in> Domain ((depend s)^+)" 
  2254       hence "Th th2 \<in> Domain ((RAG s)^+)" 
  2255         apply (unfold cs_dependants_def cs_depend_def s_depend_def)
  2255         apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
  2256         by (auto simp:Domain_def)
  2256         by (auto simp:Domain_def)
  2257       hence "Th th2 \<in> Domain (depend s)" by (simp add:trancl_domain)
  2257       hence "Th th2 \<in> Domain (RAG s)" by (simp add:trancl_domain)
  2258       then obtain n where d: "(Th th2, n) \<in> depend s" by (auto simp:Domain_def)
  2258       then obtain n where d: "(Th th2, n) \<in> RAG s" by (auto simp:Domain_def)
  2259       from depend_target_th [OF this]
  2259       from RAG_target_th [OF this]
  2260       obtain cs' where "n = Cs cs'" by auto
  2260       obtain cs' where "n = Cs cs'" by auto
  2261       with d have "(Th th2, Cs cs') \<in> depend s" by simp
  2261       with d have "(Th th2, Cs cs') \<in> RAG s" by simp
  2262       with runing_2 have "False"
  2262       with runing_2 have "False"
  2263         apply (unfold runing_def readys_def s_depend_def)
  2263         apply (unfold runing_def readys_def s_RAG_def)
  2264         by (auto simp:eq_waiting)
  2264         by (auto simp:eq_waiting)
  2265       thus ?thesis by simp
  2265       thus ?thesis by simp
  2266     next
  2266     next
  2267       assume "th2' \<in> dependants (wq s) th2"
  2267       assume "th2' \<in> dependants (wq s) th2"
  2268       with eq_th12 have "th1' \<in> dependants (wq s) th2" by simp
  2268       with eq_th12 have "th1' \<in> dependants (wq s) th2" by simp
  2269       hence h1: "(Th th1', Th th2) \<in> (depend s)^+"
  2269       hence h1: "(Th th1', Th th2) \<in> (RAG s)^+"
  2270         by (unfold cs_dependants_def s_depend_def cs_depend_def, simp)
  2270         by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
  2271       from th1'_in have h2: "(Th th1', Th th1) \<in> (depend s)^+"
  2271       from th1'_in have h2: "(Th th1', Th th1) \<in> (RAG s)^+"
  2272         by (unfold cs_dependants_def s_depend_def cs_depend_def, simp)
  2272         by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
  2273       show ?thesis
  2273       show ?thesis
  2274       proof(rule dchain_unique[OF vt h1 _ h2, symmetric])
  2274       proof(rule dchain_unique[OF vt h1 _ h2, symmetric])
  2275         from runing_1 show "th1 \<in> readys s" by (simp add:runing_def)
  2275         from runing_1 show "th1 \<in> readys s" by (simp add:runing_def)
  2276         from runing_2 show "th2 \<in> readys s" by (simp add:runing_def) 
  2276         from runing_2 show "th2 \<in> readys s" by (simp add:runing_def) 
  2277       qed
  2277       qed
  2409     case vt_nil
  2409     case vt_nil
  2410     show ?case by (auto simp:cntP_def cntV_def count_def)
  2410     show ?case by (auto simp:cntP_def cntV_def count_def)
  2411   qed
  2411   qed
  2412 qed
  2412 qed
  2413 
  2413 
  2414 lemma eq_depend: 
  2414 lemma eq_RAG: 
  2415   "depend (wq s) = depend s"
  2415   "RAG (wq s) = RAG s"
  2416 by (unfold cs_depend_def s_depend_def, auto)
  2416 by (unfold cs_RAG_def s_RAG_def, auto)
  2417 
  2417 
  2418 lemma count_eq_dependants:
  2418 lemma count_eq_dependants:
  2419   assumes vt: "vt s"
  2419   assumes vt: "vt s"
  2420   and eq_pv: "cntP s th = cntV s th"
  2420   and eq_pv: "cntP s th = cntV s th"
  2421   shows "dependants (wq s) th = {}"
  2421   shows "dependants (wq s) th = {}"
  2422 proof -
  2422 proof -
  2423   from cnp_cnv_cncs[OF vt] and eq_pv
  2423   from cnp_cnv_cncs[OF vt] and eq_pv
  2424   have "cntCS s th = 0" 
  2424   have "cntCS s th = 0" 
  2425     by (auto split:if_splits)
  2425     by (auto split:if_splits)
  2426   moreover have "finite {cs. (Cs cs, Th th) \<in> depend s}"
  2426   moreover have "finite {cs. (Cs cs, Th th) \<in> RAG s}"
  2427   proof -
  2427   proof -
  2428     from finite_holding[OF vt, of th] show ?thesis
  2428     from finite_holding[OF vt, of th] show ?thesis
  2429       by (simp add:holdents_test)
  2429       by (simp add:holdents_test)
  2430   qed
  2430   qed
  2431   ultimately have h: "{cs. (Cs cs, Th th) \<in> depend s} = {}"
  2431   ultimately have h: "{cs. (Cs cs, Th th) \<in> RAG s} = {}"
  2432     by (unfold cntCS_def holdents_test cs_dependants_def, auto)
  2432     by (unfold cntCS_def holdents_test cs_dependants_def, auto)
  2433   show ?thesis
  2433   show ?thesis
  2434   proof(unfold cs_dependants_def)
  2434   proof(unfold cs_dependants_def)
  2435     { assume "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<noteq> {}"
  2435     { assume "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}"
  2436       then obtain th' where "(Th th', Th th) \<in> (depend (wq s))\<^sup>+" by auto
  2436       then obtain th' where "(Th th', Th th) \<in> (RAG (wq s))\<^sup>+" by auto
  2437       hence "False"
  2437       hence "False"
  2438       proof(cases)
  2438       proof(cases)
  2439         assume "(Th th', Th th) \<in> depend (wq s)"
  2439         assume "(Th th', Th th) \<in> RAG (wq s)"
  2440         thus "False" by (auto simp:cs_depend_def)
  2440         thus "False" by (auto simp:cs_RAG_def)
  2441       next
  2441       next
  2442         fix c
  2442         fix c
  2443         assume "(c, Th th) \<in> depend (wq s)"
  2443         assume "(c, Th th) \<in> RAG (wq s)"
  2444         with h and eq_depend show "False"
  2444         with h and eq_RAG show "False"
  2445           by (cases c, auto simp:cs_depend_def)
  2445           by (cases c, auto simp:cs_RAG_def)
  2446       qed
  2446       qed
  2447     } thus "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} = {}" by auto
  2447     } thus "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} = {}" by auto
  2448   qed
  2448   qed
  2449 qed
  2449 qed
  2450 
  2450 
  2451 lemma dependants_threads:
  2451 lemma dependants_threads:
  2452   fixes s th
  2452   fixes s th
  2453   assumes vt: "vt s"
  2453   assumes vt: "vt s"
  2454   shows "dependants (wq s) th \<subseteq> threads s"
  2454   shows "dependants (wq s) th \<subseteq> threads s"
  2455 proof
  2455 proof
  2456   { fix th th'
  2456   { fix th th'
  2457     assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (depend (wq s))\<^sup>+}"
  2457     assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (RAG (wq s))\<^sup>+}"
  2458     have "Th th \<in> Domain (depend s)"
  2458     have "Th th \<in> Domain (RAG s)"
  2459     proof -
  2459     proof -
  2460       from h obtain th' where "(Th th, Th th') \<in> (depend (wq s))\<^sup>+" by auto
  2460       from h obtain th' where "(Th th, Th th') \<in> (RAG (wq s))\<^sup>+" by auto
  2461       hence "(Th th) \<in> Domain ( (depend (wq s))\<^sup>+)" by (auto simp:Domain_def)
  2461       hence "(Th th) \<in> Domain ( (RAG (wq s))\<^sup>+)" by (auto simp:Domain_def)
  2462       with trancl_domain have "(Th th) \<in> Domain (depend (wq s))" by simp
  2462       with trancl_domain have "(Th th) \<in> Domain (RAG (wq s))" by simp
  2463       thus ?thesis using eq_depend by simp
  2463       thus ?thesis using eq_RAG by simp
  2464     qed
  2464     qed
  2465     from dm_depend_threads[OF vt this]
  2465     from dm_RAG_threads[OF vt this]
  2466     have "th \<in> threads s" .
  2466     have "th \<in> threads s" .
  2467   } note hh = this
  2467   } note hh = this
  2468   fix th1 
  2468   fix th1 
  2469   assume "th1 \<in> dependants (wq s) th"
  2469   assume "th1 \<in> dependants (wq s) th"
  2470   hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (depend (wq s))\<^sup>+}"
  2470   hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (RAG (wq s))\<^sup>+}"
  2471     by (unfold cs_dependants_def, simp)
  2471     by (unfold cs_dependants_def, simp)
  2472   from hh [OF this] show "th1 \<in> threads s" .
  2472   from hh [OF this] show "th1 \<in> threads s" .
  2473 qed
  2473 qed
  2474 
  2474 
  2475 lemma finite_threads:
  2475 lemma finite_threads:
  2494 lemma cp_le:
  2494 lemma cp_le:
  2495   assumes vt: "vt s"
  2495   assumes vt: "vt s"
  2496   and th_in: "th \<in> threads s"
  2496   and th_in: "th \<in> threads s"
  2497   shows "cp s th \<le> Max ((\<lambda> th. (preced th s)) ` threads s)"
  2497   shows "cp s th \<le> Max ((\<lambda> th. (preced th s)) ` threads s)"
  2498 proof(unfold cp_eq_cpreced cpreced_def cs_dependants_def)
  2498 proof(unfold cp_eq_cpreced cpreced_def cs_dependants_def)
  2499   show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+}))
  2499   show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}))
  2500          \<le> Max ((\<lambda>th. preced th s) ` threads s)"
  2500          \<le> Max ((\<lambda>th. preced th s) ` threads s)"
  2501     (is "Max (?f ` ?A) \<le> Max (?f ` ?B)")
  2501     (is "Max (?f ` ?A) \<le> Max (?f ` ?B)")
  2502   proof(rule Max_f_mono)
  2502   proof(rule Max_f_mono)
  2503     show "{th} \<union> {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<noteq> {}" by simp
  2503     show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}" by simp
  2504   next
  2504   next
  2505     from finite_threads [OF vt]
  2505     from finite_threads [OF vt]
  2506     show "finite (threads s)" .
  2506     show "finite (threads s)" .
  2507   next
  2507   next
  2508     from th_in
  2508     from th_in
  2509     show "{th} \<union> {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<subseteq> threads s"
  2509     show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> threads s"
  2510       apply (auto simp:Domain_def)
  2510       apply (auto simp:Domain_def)
  2511       apply (rule_tac dm_depend_threads[OF vt])
  2511       apply (rule_tac dm_RAG_threads[OF vt])
  2512       apply (unfold trancl_domain [of "depend s", symmetric])
  2512       apply (unfold trancl_domain [of "RAG s", symmetric])
  2513       by (unfold cs_depend_def s_depend_def, auto simp:Domain_def)
  2513       by (unfold cs_RAG_def s_RAG_def, auto simp:Domain_def)
  2514   qed
  2514   qed
  2515 qed
  2515 qed
  2516 
  2516 
  2517 lemma le_cp:
  2517 lemma le_cp:
  2518   assumes vt: "vt s"
  2518   assumes vt: "vt s"
  2526     case False
  2526     case False
  2527     have "finite ?A" (is "finite (?f ` ?B)")
  2527     have "finite ?A" (is "finite (?f ` ?B)")
  2528     proof -
  2528     proof -
  2529       have "finite ?B" 
  2529       have "finite ?B" 
  2530       proof-
  2530       proof-
  2531         have "finite {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+}"
  2531         have "finite {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}"
  2532         proof -
  2532         proof -
  2533           let ?F = "\<lambda> (x, y). the_th x"
  2533           let ?F = "\<lambda> (x, y). the_th x"
  2534           have "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (wq s))\<^sup>+)"
  2534           have "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
  2535             apply (auto simp:image_def)
  2535             apply (auto simp:image_def)
  2536             by (rule_tac x = "(Th x, Th th)" in bexI, auto)
  2536             by (rule_tac x = "(Th x, Th th)" in bexI, auto)
  2537           moreover have "finite \<dots>"
  2537           moreover have "finite \<dots>"
  2538           proof -
  2538           proof -
  2539             from finite_depend[OF vt] have "finite (depend s)" .
  2539             from finite_RAG[OF vt] have "finite (RAG s)" .
  2540             hence "finite ((depend (wq s))\<^sup>+)"
  2540             hence "finite ((RAG (wq s))\<^sup>+)"
  2541               apply (unfold finite_trancl)
  2541               apply (unfold finite_trancl)
  2542               by (auto simp: s_depend_def cs_depend_def wq_def)
  2542               by (auto simp: s_RAG_def cs_RAG_def wq_def)
  2543             thus ?thesis by auto
  2543             thus ?thesis by auto
  2544           qed
  2544           qed
  2545           ultimately show ?thesis by (auto intro:finite_subset)
  2545           ultimately show ?thesis by (auto intro:finite_subset)
  2546         qed
  2546         qed
  2547         thus ?thesis by (simp add:cs_dependants_def)
  2547         thus ?thesis by (simp add:cs_dependants_def)
  2619       from np show "?f ` threads s \<noteq> {}" by simp
  2619       from np show "?f ` threads s \<noteq> {}" by simp
  2620     qed
  2620     qed
  2621     then obtain tm where tm_max: "?f tm = ?p" and tm_in: "tm \<in> threads s"
  2621     then obtain tm where tm_max: "?f tm = ?p" and tm_in: "tm \<in> threads s"
  2622       by (auto simp:Image_def)
  2622       by (auto simp:Image_def)
  2623     from th_chain_to_ready [OF vt tm_in]
  2623     from th_chain_to_ready [OF vt tm_in]
  2624     have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (depend s)\<^sup>+)" .
  2624     have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+)" .
  2625     thus ?thesis
  2625     thus ?thesis
  2626     proof
  2626     proof
  2627       assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (depend s)\<^sup>+ "
  2627       assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+ "
  2628       then obtain th' where th'_in: "th' \<in> readys s" 
  2628       then obtain th' where th'_in: "th' \<in> readys s" 
  2629         and tm_chain:"(Th tm, Th th') \<in> (depend s)\<^sup>+" by auto
  2629         and tm_chain:"(Th tm, Th th') \<in> (RAG s)\<^sup>+" by auto
  2630       have "cp s th' = ?f tm"
  2630       have "cp s th' = ?f tm"
  2631       proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
  2631       proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
  2632         from dependants_threads[OF vt] finite_threads[OF vt]
  2632         from dependants_threads[OF vt] finite_threads[OF vt]
  2633         show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th'))" 
  2633         show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th'))" 
  2634           by (auto intro:finite_subset)
  2634           by (auto intro:finite_subset)
  2648       next
  2648       next
  2649         show "preced tm s \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
  2649         show "preced tm s \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
  2650         proof -
  2650         proof -
  2651           from tm_chain
  2651           from tm_chain
  2652           have "tm \<in> dependants (wq s) th'"
  2652           have "tm \<in> dependants (wq s) th'"
  2653             by (unfold cs_dependants_def s_depend_def cs_depend_def, auto)
  2653             by (unfold cs_dependants_def s_RAG_def cs_RAG_def, auto)
  2654           thus ?thesis by auto
  2654           thus ?thesis by auto
  2655         qed
  2655         qed
  2656       qed
  2656       qed
  2657       with tm_max
  2657       with tm_max
  2658       have h: "cp s th' = Max ((\<lambda>th. preced th s) ` threads s)" by simp
  2658       have h: "cp s th' = Max ((\<lambda>th. preced th s) ` threads s)" by simp
  2778 definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool"
  2778 definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool"
  2779   where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))"
  2779   where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))"
  2780 
  2780 
  2781 
  2781 
  2782 lemma detached_test:
  2782 lemma detached_test:
  2783   shows "detached s th = (Th th \<notin> Field (depend s))"
  2783   shows "detached s th = (Th th \<notin> Field (RAG s))"
  2784 apply(simp add: detached_def Field_def)
  2784 apply(simp add: detached_def Field_def)
  2785 apply(simp add: s_depend_def)
  2785 apply(simp add: s_RAG_def)
  2786 apply(simp add: s_holding_abv s_waiting_abv)
  2786 apply(simp add: s_holding_abv s_waiting_abv)
  2787 apply(simp add: Domain_iff Range_iff)
  2787 apply(simp add: Domain_iff Range_iff)
  2788 apply(simp add: wq_def)
  2788 apply(simp add: wq_def)
  2789 apply(auto)
  2789 apply(auto)
  2790 done
  2790 done
  2803   with eq_cnt
  2803   with eq_cnt
  2804   have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:eq_pv)
  2804   have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:eq_pv)
  2805   thus ?thesis
  2805   thus ?thesis
  2806   proof
  2806   proof
  2807     assume "th \<notin> threads s"
  2807     assume "th \<notin> threads s"
  2808     with range_in[OF vt] dm_depend_threads[OF vt]
  2808     with range_in[OF vt] dm_RAG_threads[OF vt]
  2809     show ?thesis
  2809     show ?thesis
  2810       by (auto simp add: detached_def s_depend_def s_waiting_abv s_holding_abv wq_def Domain_iff Range_iff)
  2810       by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def Domain_iff Range_iff)
  2811   next
  2811   next
  2812     assume "th \<in> readys s"
  2812     assume "th \<in> readys s"
  2813     moreover have "Th th \<notin> Range (depend s)"
  2813     moreover have "Th th \<notin> Range (RAG s)"
  2814     proof -
  2814     proof -
  2815       from card_0_eq [OF finite_holding [OF vt]] and cncs_zero
  2815       from card_0_eq [OF finite_holding [OF vt]] and cncs_zero
  2816       have "holdents s th = {}"
  2816       have "holdents s th = {}"
  2817         by (simp add:cntCS_def)
  2817         by (simp add:cntCS_def)
  2818       thus ?thesis
  2818       thus ?thesis
  2819         apply(auto simp:holdents_test)
  2819         apply(auto simp:holdents_test)
  2820         apply(case_tac a)
  2820         apply(case_tac a)
  2821         apply(auto simp:holdents_test s_depend_def)
  2821         apply(auto simp:holdents_test s_RAG_def)
  2822         done
  2822         done
  2823     qed
  2823     qed
  2824     ultimately show ?thesis
  2824     ultimately show ?thesis
  2825       by (auto simp add: detached_def s_depend_def s_waiting_abv s_holding_abv wq_def readys_def)
  2825       by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def readys_def)
  2826   qed
  2826   qed
  2827 qed
  2827 qed
  2828 
  2828 
  2829 lemma detached_elim:
  2829 lemma detached_elim:
  2830   fixes s th
  2830   fixes s th
  2836   have eq_pv: " cntP s th =
  2836   have eq_pv: " cntP s th =
  2837     cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" .
  2837     cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" .
  2838   have cncs_z: "cntCS s th = 0"
  2838   have cncs_z: "cntCS s th = 0"
  2839   proof -
  2839   proof -
  2840     from dtc have "holdents s th = {}"
  2840     from dtc have "holdents s th = {}"
  2841       unfolding detached_def holdents_test s_depend_def
  2841       unfolding detached_def holdents_test s_RAG_def
  2842       by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff)
  2842       by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff)
  2843     thus ?thesis by (auto simp:cntCS_def)
  2843     thus ?thesis by (auto simp:cntCS_def)
  2844   qed
  2844   qed
  2845   show ?thesis
  2845   show ?thesis
  2846   proof(cases "th \<in> threads s")
  2846   proof(cases "th \<in> threads s")
  2847     case True
  2847     case True
  2848     with dtc 
  2848     with dtc 
  2849     have "th \<in> readys s"
  2849     have "th \<in> readys s"
  2850       by (unfold readys_def detached_def Field_def Domain_def Range_def, 
  2850       by (unfold readys_def detached_def Field_def Domain_def Range_def, 
  2851            auto simp:eq_waiting s_depend_def)
  2851            auto simp:eq_waiting s_RAG_def)
  2852     with cncs_z and eq_pv show ?thesis by simp
  2852     with cncs_z and eq_pv show ?thesis by simp
  2853   next
  2853   next
  2854     case False
  2854     case False
  2855     with cncs_z and eq_pv show ?thesis by simp
  2855     with cncs_z and eq_pv show ?thesis by simp
  2856   qed
  2856   qed