Implementation.thy
author zhangx
Fri, 12 Feb 2016 12:32:57 +0800
changeset 115 74fc1eae4605
parent 113 ce85c3c4e5bf
child 116 a7441db6f4e1
permissions -rw-r--r--
Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
107
30ed212f268a updated Correctness, Implementation and PIPBasics so that they work with Isabelle 2014 and 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
     1
theory Implementation
30ed212f268a updated Correctness, Implementation and PIPBasics so that they work with Isabelle 2014 and 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
     2
imports PIPBasics
30ed212f268a updated Correctness, Implementation and PIPBasics so that they work with Isabelle 2014 and 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
     3
begin
30ed212f268a updated Correctness, Implementation and PIPBasics so that they work with Isabelle 2014 and 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
     4
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 45
diff changeset
     5
section {*
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 45
diff changeset
     6
  This file contains lemmas used to guide the recalculation of current precedence 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 45
diff changeset
     7
  after every system call (or system operation)
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 45
diff changeset
     8
*}
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
     9
115
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    10
section {* Recursive calculation of @{term "cp"} *}
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    11
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    12
definition "cp_gen s x =
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    13
                  Max ((the_preced s \<circ> the_thread) ` subtree (tRAG s) x)"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    14
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    15
lemma cp_gen_alt_def:
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    16
  "cp_gen s = (Max \<circ> (\<lambda>x. (the_preced s \<circ> the_thread) ` subtree (tRAG s) x))"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    17
    by (auto simp:cp_gen_def)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    18
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    19
lemma cp_gen_def_cond: 
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    20
  assumes "x = Th th"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    21
  shows "cp s th = cp_gen s (Th th)"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    22
by (unfold cp_alt_def1 cp_gen_def, simp)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    23
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    24
lemma cp_gen_over_set:
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    25
  assumes "\<forall> x \<in> A. \<exists> th. x = Th th"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    26
  shows "cp_gen s ` A = (cp s \<circ> the_thread) ` A"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    27
proof(rule f_image_eq)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    28
  fix a
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    29
  assume "a \<in> A"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    30
  from assms[rule_format, OF this]
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    31
  obtain th where eq_a: "a = Th th" by auto
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    32
  show "cp_gen s a = (cp s \<circ> the_thread) a"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    33
    by  (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    34
qed
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    35
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    36
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    37
context valid_trace
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    38
begin
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    39
(* ddd *)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    40
lemma cp_gen_rec:
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    41
  assumes "x = Th th"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    42
  shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    43
proof(cases "children (tRAG s) x = {}")
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    44
  case True
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    45
  show ?thesis
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    46
    by (unfold True cp_gen_def subtree_children, simp add:assms)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    47
next
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    48
  case False
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    49
  hence [simp]: "children (tRAG s) x \<noteq> {}" by auto
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    50
  note fsbttRAGs.finite_subtree[simp]
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    51
  have [simp]: "finite (children (tRAG s) x)"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    52
     by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], 
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    53
            rule children_subtree)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    54
  { fix r x
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    55
    have "subtree r x \<noteq> {}" by (auto simp:subtree_def)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    56
  } note this[simp]
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    57
  have [simp]: "\<exists>x\<in>children (tRAG s) x. subtree (tRAG s) x \<noteq> {}"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    58
  proof -
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    59
    from False obtain q where "q \<in> children (tRAG s) x" by blast
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    60
    moreover have "subtree (tRAG s) q \<noteq> {}" by simp
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    61
    ultimately show ?thesis by blast
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    62
  qed
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    63
  have h: "Max ((the_preced s \<circ> the_thread) `
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    64
                ({x} \<union> \<Union>(subtree (tRAG s) ` children (tRAG s) x))) =
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    65
        Max ({the_preced s th} \<union> cp_gen s ` children (tRAG s) x)"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    66
                     (is "?L = ?R")
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    67
  proof -
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    68
    let "Max (?f ` (?A \<union> \<Union> (?g ` ?B)))" = ?L
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    69
    let "Max (_ \<union> (?h ` ?B))" = ?R
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    70
    let ?L1 = "?f ` \<Union>(?g ` ?B)"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    71
    have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    72
    proof -
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    73
      have "?L1 = ?f ` (\<Union> x \<in> ?B.(?g x))" by simp
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    74
      also have "... =  (\<Union> x \<in> ?B. ?f ` (?g x))" by auto
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    75
      finally have "Max ?L1 = Max ..." by simp
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    76
      also have "... = Max (Max ` (\<lambda>x. ?f ` subtree (tRAG s) x) ` ?B)"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    77
        by (subst Max_UNION, simp+)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    78
      also have "... = Max (cp_gen s ` children (tRAG s) x)"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    79
          by (unfold image_comp cp_gen_alt_def, simp)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    80
      finally show ?thesis .
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    81
    qed
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    82
    show ?thesis
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    83
    proof -
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    84
      have "?L = Max (?f ` ?A \<union> ?L1)" by simp
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    85
      also have "... = max (the_preced s (the_thread x)) (Max ?L1)"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    86
            by (subst Max_Un, simp+)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    87
      also have "... = max (?f x) (Max (?h ` ?B))"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    88
        by (unfold eq_Max_L1, simp)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    89
      also have "... =?R"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    90
        by (rule max_Max_eq, (simp)+, unfold assms, simp)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    91
      finally show ?thesis .
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    92
    qed
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    93
  qed  thus ?thesis 
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    94
          by (fold h subtree_children, unfold cp_gen_def, simp) 
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    95
qed
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    96
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    97
lemma cp_rec:
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    98
  "cp s th = Max ({the_preced s th} \<union> 
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
    99
                     (cp s o the_thread) ` children (tRAG s) (Th th))"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   100
proof -
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   101
  have "Th th = Th th" by simp
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   102
  note h =  cp_gen_def_cond[OF this] cp_gen_rec[OF this]
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   103
  show ?thesis 
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   104
  proof -
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   105
    have "cp_gen s ` children (tRAG s) (Th th) = 
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   106
                (cp s \<circ> the_thread) ` children (tRAG s) (Th th)"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   107
    proof(rule cp_gen_over_set)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   108
      show " \<forall>x\<in>children (tRAG s) (Th th). \<exists>th. x = Th th"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   109
        by (unfold tRAG_alt_def, auto simp:children_def)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   110
    qed
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   111
    thus ?thesis by (subst (1) h(1), unfold h(2), simp)
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   112
  qed
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   113
qed
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   114
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   115
end
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 113
diff changeset
   116
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   117
text {* (* ddd *)
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   118
  One beauty of our modelling is that we follow the definitional extension tradition of HOL.
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   119
  The benefit of such a concise and miniature model is that  large number of intuitively 
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   120
  obvious facts are derived as lemmas, rather than asserted as axioms.
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   121
*}
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   122
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   123
text {*
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   124
  However, the lemmas in the forthcoming several locales are no longer 
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   125
  obvious. These lemmas show how the current precedences should be recalculated 
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   126
  after every execution step (in our model, every step is represented by an event, 
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   127
  which in turn, represents a system call, or operation). Each operation is 
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   128
  treated in a separate locale.
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   129
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   130
  The complication of current precedence recalculation comes 
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   131
  because the changing of RAG needs to be taken into account, 
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   132
  in addition to the changing of precedence. 
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   133
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   134
  The reason RAG changing affects current precedence is that,
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   135
  according to the definition, current precedence 
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   136
  of a thread is the maximum of the precedences of every threads in its subtree, 
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   137
  where the notion of sub-tree in RAG is defined in RTree.thy.
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   138
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   139
  Therefore, for each operation, lemmas about the change of precedences 
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   140
  and RAG are derived first, on which lemmas about current precedence 
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   141
  recalculation are based on.
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   142
*}
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   143
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   144
section {* The @{term Set} operation *}
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   145
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   146
context valid_trace_set
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   147
begin
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   149
text {* (* ddd *)
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   150
  The following two lemmas confirm that @{text "Set"}-operation
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   151
  only changes the precedence of the initiating thread (or actor)
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   152
  of the operation (or event).
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   153
*}
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   154
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   155
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   156
lemma eq_preced:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   157
  assumes "th' \<noteq> th"
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   158
  shows "preced th' (e#s) = preced th' s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   159
proof -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   160
  from assms show ?thesis 
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   161
    by (unfold is_set, auto simp:preced_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   162
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   163
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   164
lemma eq_the_preced: 
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   165
  assumes "th' \<noteq> th"
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   166
  shows "the_preced (e#s) th' = the_preced s th'"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   167
  using assms
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   168
  by (unfold the_preced_def, intro eq_preced, simp)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   169
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   171
text {* (* ddd *)
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   172
  Th following lemma @{text "eq_cp_pre"} says that the priority change of @{text "th"}
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   173
  only affects those threads, which as @{text "Th th"} in their sub-trees.
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   174
  
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   175
  The proof of this lemma is simplified by using the alternative definition 
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   176
  of @{text "cp"}. 
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   177
*}
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   178
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   179
lemma eq_cp_pre:
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   180
  assumes nd: "Th th \<notin> subtree (RAG s) (Th th')"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   181
  shows "cp (e#s) th' = cp s th'"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   182
proof -
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   183
  -- {* After unfolding using the alternative definition, elements 
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   184
        affecting the @{term "cp"}-value of threads become explicit. 
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   185
        We only need to prove the following: *}
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   186
  have "Max (the_preced (e#s) ` {th'a. Th th'a \<in> subtree (RAG (e#s)) (Th th')}) =
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   187
        Max (the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')})"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   188
        (is "Max (?f ` ?S1) = Max (?g ` ?S2)")
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   189
  proof -
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   190
    -- {* The base sets are equal. *}
113
ce85c3c4e5bf More improvements in PIPBasics.thy and Implemenation.thy.
zhangx
parents: 108
diff changeset
   191
    have "?S1 = ?S2" using RAG_es by simp
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   192
    -- {* The function values on the base set are equal as well. *}
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   193
    moreover have "\<forall> e \<in> ?S2. ?f e = ?g e"
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   194
    proof
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   195
      fix th1
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   196
      assume "th1 \<in> ?S2"
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   197
      with nd have "th1 \<noteq> th" by (auto)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   198
      from eq_the_preced[OF this]
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   199
      show "the_preced (e#s) th1 = the_preced s th1" .
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   200
    qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   201
    -- {* Therefore, the image of the functions are equal. *}
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   202
    ultimately have "(?f ` ?S1) = (?g ` ?S2)" by (auto intro!:f_image_eq)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   203
    thus ?thesis by simp
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   204
  qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   205
  thus ?thesis by (simp add:cp_alt_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   206
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   207
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   208
text {*
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   209
  The following lemma shows that @{term "th"} is not in the 
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   210
  sub-tree of any other thread. 
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   211
*}
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   212
lemma th_in_no_subtree:
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   213
  assumes "th' \<noteq> th"
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   214
  shows "Th th \<notin> subtree (RAG s) (Th th')"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   215
proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   216
  from readys_in_no_subtree[OF th_ready_s assms(1)]
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   217
  show ?thesis by blast
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   218
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   219
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   220
text {* 
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   221
  By combining @{thm "eq_cp_pre"} and @{thm "th_in_no_subtree"}, 
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   222
  it is obvious that the change of priority only affects the @{text "cp"}-value 
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   223
  of the initiating thread @{text "th"}.
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   224
*}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   225
lemma eq_cp:
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   226
  assumes "th' \<noteq> th"
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   227
  shows "cp (e#s) th' = cp s th'"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   228
  by (rule eq_cp_pre[OF th_in_no_subtree[OF assms]])
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   229
56
0fd478e14e87 Before switching to generic theory of relational trees.
xingyuan zhang <xingyuanzhang@126.com>
parents: 55
diff changeset
   230
end
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   231
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   232
section {* The @{term V} operation *}
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   233
56
0fd478e14e87 Before switching to generic theory of relational trees.
xingyuan zhang <xingyuanzhang@126.com>
parents: 55
diff changeset
   234
text {*
0fd478e14e87 Before switching to generic theory of relational trees.
xingyuan zhang <xingyuanzhang@126.com>
parents: 55
diff changeset
   235
  The following @{text "step_v_cps"} is the locale for @{text "V"}-operation.
55
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   236
*}
b85cfbd58f59 Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents: 53
diff changeset
   237
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   238
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   239
context valid_trace_v
61
f8194fd6214f CpsG.thy has been cleaned up.
zhangx
parents: 60
diff changeset
   240
begin
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   241
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   242
lemma ancestors_th: "ancestors (RAG s) (Th th) = {}"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   243
proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   244
  from readys_root[OF th_ready_s]
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   245
  show ?thesis
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   246
  by (unfold root_def, simp)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   247
qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   248
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   249
lemma edge_of_th:
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   250
    "(Cs cs, Th th) \<in> RAG s" 
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   251
proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   252
 from holding_th_cs_s
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   253
 show ?thesis 
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   254
    by (unfold s_RAG_def holding_eq, auto)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   255
qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   256
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   257
lemma ancestors_cs: 
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   258
  "ancestors (RAG s) (Cs cs) = {Th th}"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   259
proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   260
  have "ancestors (RAG s) (Cs cs) = ancestors (RAG s) (Th th)  \<union>  {Th th}"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   261
   by (rule rtree_RAG.ancestors_accum[OF edge_of_th])
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   262
  from this[unfolded ancestors_th] show ?thesis by simp
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   263
qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   264
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   265
end
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   266
56
0fd478e14e87 Before switching to generic theory of relational trees.
xingyuan zhang <xingyuanzhang@126.com>
parents: 55
diff changeset
   267
text {*
0fd478e14e87 Before switching to generic theory of relational trees.
xingyuan zhang <xingyuanzhang@126.com>
parents: 55
diff changeset
   268
  The following @{text "step_v_cps_nt"} is the sub-locale for @{text "V"}-operation, 
0fd478e14e87 Before switching to generic theory of relational trees.
xingyuan zhang <xingyuanzhang@126.com>
parents: 55
diff changeset
   269
  which represents the case when there is another thread @{text "th'"}
0fd478e14e87 Before switching to generic theory of relational trees.
xingyuan zhang <xingyuanzhang@126.com>
parents: 55
diff changeset
   270
  to take over the critical resource released by the initiating thread @{text "th"}.
0fd478e14e87 Before switching to generic theory of relational trees.
xingyuan zhang <xingyuanzhang@126.com>
parents: 55
diff changeset
   271
*}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   272
113
ce85c3c4e5bf More improvements in PIPBasics.thy and Implemenation.thy.
zhangx
parents: 108
diff changeset
   273
lemma (in valid_trace_v)
ce85c3c4e5bf More improvements in PIPBasics.thy and Implemenation.thy.
zhangx
parents: 108
diff changeset
   274
  the_preced_es: "the_preced (e#s) = the_preced s"
ce85c3c4e5bf More improvements in PIPBasics.thy and Implemenation.thy.
zhangx
parents: 108
diff changeset
   275
proof
ce85c3c4e5bf More improvements in PIPBasics.thy and Implemenation.thy.
zhangx
parents: 108
diff changeset
   276
  fix th'
ce85c3c4e5bf More improvements in PIPBasics.thy and Implemenation.thy.
zhangx
parents: 108
diff changeset
   277
  show "the_preced (e # s) th' = the_preced s th'"
ce85c3c4e5bf More improvements in PIPBasics.thy and Implemenation.thy.
zhangx
parents: 108
diff changeset
   278
    by (unfold the_preced_def preced_def is_v, simp)
ce85c3c4e5bf More improvements in PIPBasics.thy and Implemenation.thy.
zhangx
parents: 108
diff changeset
   279
qed
ce85c3c4e5bf More improvements in PIPBasics.thy and Implemenation.thy.
zhangx
parents: 108
diff changeset
   280
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   281
context valid_trace_v_n
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   282
begin
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   283
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   284
lemma sub_RAGs': 
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   285
  "{(Cs cs, Th th), (Th taker, Cs cs)} \<subseteq> RAG s"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   286
     using next_th_RAG[OF next_th_taker]  .
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   287
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   288
lemma ancestors_th': 
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   289
  "ancestors (RAG s) (Th taker) = {Th th, Cs cs}" 
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   290
proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   291
  have "ancestors (RAG s) (Th taker) = ancestors (RAG s) (Cs cs) \<union> {Cs cs}"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   292
  proof(rule  rtree_RAG.ancestors_accum)
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   293
    from sub_RAGs' show "(Th taker, Cs cs) \<in> RAG s" by auto
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   294
  qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   295
  thus ?thesis using ancestors_th ancestors_cs by auto
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   296
qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   297
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   298
lemma RAG_s:
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   299
  "RAG (e#s) = (RAG s - {(Cs cs, Th th), (Th taker, Cs cs)}) \<union>
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   300
                                         {(Cs cs, Th taker)}"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   301
 by (unfold RAG_es waiting_set_eq holding_set_eq, auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   302
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   303
lemma subtree_kept: (* ddd *)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   304
  assumes "th1 \<notin> {th, taker}"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   305
  shows "subtree (RAG (e#s)) (Th th1) = 
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   306
                     subtree (RAG s) (Th th1)" (is "_ = ?R")
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   307
proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   308
  let ?RAG' = "(RAG s - {(Cs cs, Th th), (Th taker, Cs cs)})"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   309
  let ?RAG'' = "?RAG' \<union> {(Cs cs, Th taker)}"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   310
  have "subtree ?RAG' (Th th1) = ?R" 
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   311
  proof(rule subset_del_subtree_outside)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   312
    show "Range {(Cs cs, Th th), (Th taker, Cs cs)} \<inter> subtree (RAG s) (Th th1) = {}"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   313
    proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   314
      have "(Th th) \<notin> subtree (RAG s) (Th th1)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   315
      proof(rule subtree_refute)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   316
        show "Th th1 \<notin> ancestors (RAG s) (Th th)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   317
          by (unfold ancestors_th, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   318
      next
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   319
        from assms show "Th th1 \<noteq> Th th" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   320
      qed
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   321
      moreover have "(Cs cs) \<notin>  subtree (RAG s) (Th th1)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   322
      proof(rule subtree_refute)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   323
        show "Th th1 \<notin> ancestors (RAG s) (Cs cs)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   324
          by (unfold ancestors_cs, insert assms, auto)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   325
      qed simp
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   326
      ultimately have "{Th th, Cs cs} \<inter> subtree (RAG s) (Th th1) = {}" by auto
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   327
      thus ?thesis by simp
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   328
     qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   329
  qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   330
  moreover have "subtree ?RAG'' (Th th1) =  subtree ?RAG' (Th th1)"
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   331
  proof(rule subtree_insert_next)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   332
    show "Th taker \<notin> subtree (RAG s - {(Cs cs, Th th), (Th taker, Cs cs)}) (Th th1)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   333
    proof(rule subtree_refute)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   334
      show "Th th1 \<notin> ancestors (RAG s - {(Cs cs, Th th), (Th taker, Cs cs)}) (Th taker)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   335
            (is "_ \<notin> ?R")
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   336
      proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   337
          have "?R \<subseteq> ancestors (RAG s) (Th taker)" by (rule ancestors_mono, auto)
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   338
          moreover have "Th th1 \<notin> ..." using ancestors_th' assms by simp
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   339
          ultimately show ?thesis by auto
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   340
      qed
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   341
    next
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   342
      from assms show "Th th1 \<noteq> Th taker" by simp
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   343
    qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   344
  qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   345
  ultimately show ?thesis by (unfold RAG_s, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   346
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   347
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   348
lemma cp_kept:
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   349
  assumes "th1 \<notin> {th, taker}"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   350
  shows "cp (e#s) th1 = cp s th1"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   351
    by (unfold cp_alt_def the_preced_es subtree_kept[OF assms], simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   352
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   353
end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   354
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   355
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   356
context valid_trace_v_e
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   357
begin
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   358
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   359
lemma RAG_s: "RAG (e#s) = RAG s - {(Cs cs, Th th)}"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   360
  by (unfold RAG_es waiting_set_eq holding_set_eq, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   361
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   362
lemma subtree_kept:
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   363
  assumes "th1 \<noteq> th"
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   364
  shows "subtree (RAG (e#s)) (Th th1) = subtree (RAG s) (Th th1)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   365
proof(unfold RAG_s, rule subset_del_subtree_outside)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   366
  show "Range {(Cs cs, Th th)} \<inter> subtree (RAG s) (Th th1) = {}"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   367
  proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   368
    have "(Th th) \<notin> subtree (RAG s) (Th th1)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   369
    proof(rule subtree_refute)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   370
      show "Th th1 \<notin> ancestors (RAG s) (Th th)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   371
          by (unfold ancestors_th, simp)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   372
    next
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   373
      from assms show "Th th1 \<noteq> Th th" by simp
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   374
    qed
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   375
    thus ?thesis by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   376
  qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   377
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   378
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   379
lemma cp_kept_1:
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   380
  assumes "th1 \<noteq> th"
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   381
  shows "cp (e#s) th1 = cp s th1"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   382
    by (unfold cp_alt_def the_preced_es subtree_kept[OF assms], simp)
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   383
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   384
lemma subtree_cs: "subtree (RAG s) (Cs cs) = {Cs cs}"
108
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   385
  (is "?L = ?R")
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   386
proof -
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   387
  { fix n
108
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   388
    assume "n \<in> ?L"
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   389
    hence "n \<in> ?R"
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   390
    proof(cases rule:subtreeE)
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   391
      case 2
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   392
      from this(2) have "(n, Cs cs) \<in> (RAG s)^+"
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   393
        by (auto simp:ancestors_def)
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   394
      from tranclD2[OF this]
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   395
      obtain th' where "waiting s th' cs"
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   396
        by (auto simp:s_RAG_def waiting_eq)
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   397
      with no_waiter_before 
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   398
      show ?thesis by simp
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   399
    qed simp
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   400
  } moreover {
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   401
    fix n
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   402
    assume "n \<in> ?R"
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   403
    hence "n \<in> ?L" by (auto simp:subtree_def)
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   404
  } ultimately show ?thesis by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   405
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   406
108
b769f43deb30 Several redundant lemmas removed.
zhangx
parents: 107
diff changeset
   407
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   408
lemma subtree_th: 
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   409
  "subtree (RAG (e#s)) (Th th) = subtree (RAG s) (Th th) - {Cs cs}"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   410
proof(unfold RAG_s, fold subtree_cs, rule rtree_RAG.subtree_del_inside)
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   411
  from edge_of_th
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   412
  show "(Cs cs, Th th) \<in> edges_in (RAG s) (Th th)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   413
    by (unfold edges_in_def, auto simp:subtree_def)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   414
qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   415
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   416
lemma cp_kept_2: 
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   417
  shows "cp (e#s) th = cp s th" 
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   418
 by (unfold cp_alt_def subtree_th the_preced_es, auto)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   419
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   420
lemma eq_cp:
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   421
  shows "cp (e#s) th' = cp s th'"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   422
  using cp_kept_1 cp_kept_2
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   423
  by (cases "th' = th", auto)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   424
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   425
end
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   426
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   427
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   428
section {* The @{term P} operation *}
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   429
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   430
context valid_trace_p
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   431
begin
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   432
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   433
lemma root_th: "root (RAG s) (Th th)"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   434
  by (simp add: ready_th_s readys_root)
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   435
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   436
lemma in_no_others_subtree:
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   437
  assumes "th' \<noteq> th"
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   438
  shows "Th th \<notin> subtree (RAG s) (Th th')"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   439
proof
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   440
  assume "Th th \<in> subtree (RAG s) (Th th')"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   441
  thus False
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   442
  proof(cases rule:subtreeE)
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   443
    case 1
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   444
    with assms show ?thesis by auto
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   445
  next
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   446
    case 2
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   447
    with root_th show ?thesis by (auto simp:root_def)
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   448
  qed
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   449
qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   450
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   451
lemma preced_kept: "the_preced (e#s) = the_preced s"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   452
proof
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   453
  fix th'
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   454
  show "the_preced (e # s) th' = the_preced s th'"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   455
    by (unfold the_preced_def is_p preced_def, simp)
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   456
qed
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   457
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   458
end
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   459
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   460
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   461
context valid_trace_p_h
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   462
begin
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   463
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   464
lemma subtree_kept:
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   465
  assumes "th' \<noteq> th"
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   466
  shows "subtree (RAG (e#s)) (Th th') = subtree (RAG s) (Th th')"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   467
proof(unfold RAG_es, rule subtree_insert_next)
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   468
  from in_no_others_subtree[OF assms] 
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   469
  show "Th th \<notin> subtree (RAG s) (Th th')" .
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   470
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   471
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   472
lemma cp_kept: 
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   473
  assumes "th' \<noteq> th"
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   474
  shows "cp (e#s) th' = cp s th'"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   475
proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   476
  have "(the_preced (e#s) ` {th'a. Th th'a \<in> subtree (RAG (e#s)) (Th th')}) =
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   477
        (the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')})"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   478
        by (unfold preced_kept subtree_kept[OF assms], simp)
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   479
  thus ?thesis by (unfold cp_alt_def, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   480
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   481
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   482
end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   483
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   484
context valid_trace_p_w
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   485
begin
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   486
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   487
lemma cs_held: "(Cs cs, Th holder) \<in> RAG s"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   488
  using holding_s_holder
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   489
  by (unfold s_RAG_def, fold holding_eq, auto)
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   490
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   491
lemma tRAG_s: 
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   492
  "tRAG (e#s) = tRAG s \<union> {(Th th, Th holder)}"
113
ce85c3c4e5bf More improvements in PIPBasics.thy and Implemenation.thy.
zhangx
parents: 108
diff changeset
   493
  using RAG_tRAG_transfer[OF RAG_es cs_held] .
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   494
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   495
lemma cp_kept:
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   496
  assumes "Th th'' \<notin> ancestors (tRAG (e#s)) (Th th)"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   497
  shows "cp (e#s) th'' = cp s th''"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   498
proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   499
  have h: "subtree (tRAG (e#s)) (Th th'') = subtree (tRAG s) (Th th'')"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   500
  proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   501
    have "Th holder \<notin> subtree (tRAG s) (Th th'')"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   502
    proof
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   503
      assume "Th holder \<in> subtree (tRAG s) (Th th'')"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   504
      thus False
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   505
      proof(rule subtreeE)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   506
         assume "Th holder = Th th''"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   507
         from assms[unfolded tRAG_s ancestors_def, folded this]
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   508
         show ?thesis by auto
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   509
      next
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   510
         assume "Th th'' \<in> ancestors (tRAG s) (Th holder)"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   511
         moreover have "... \<subseteq> ancestors (tRAG (e#s)) (Th holder)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   512
         proof(rule ancestors_mono)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   513
            show "tRAG s \<subseteq> tRAG (e#s)" by (unfold tRAG_s, auto)
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   514
         qed 
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   515
         ultimately have "Th th'' \<in> ancestors (tRAG (e#s)) (Th holder)" by auto
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   516
         moreover have "Th holder \<in> ancestors (tRAG (e#s)) (Th th)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   517
           by (unfold tRAG_s, auto simp:ancestors_def)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   518
         ultimately have "Th th'' \<in> ancestors (tRAG (e#s)) (Th th)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   519
                       by (auto simp:ancestors_def)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   520
         with assms show ?thesis by auto
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   521
      qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   522
    qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   523
    from subtree_insert_next[OF this]
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   524
    have "subtree (tRAG s \<union> {(Th th, Th holder)}) (Th th'') = subtree (tRAG s) (Th th'')" .
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   525
    from this[folded tRAG_s] show ?thesis .
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   526
  qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   527
  show ?thesis by (unfold cp_alt_def1 h preced_kept, simp)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   528
qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   529
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   530
lemma cp_gen_update_stop: (* ddd *)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   531
  assumes "u \<in> ancestors (tRAG (e#s)) (Th th)"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   532
  and "cp_gen (e#s) u = cp_gen s u"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   533
  and "y \<in> ancestors (tRAG (e#s)) u"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   534
  shows "cp_gen (e#s) y = cp_gen s y"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   535
  using assms(3)
102
3a801bbd2687 Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents: 93
diff changeset
   536
proof(induct rule:wf_induct[OF vat_es.fsbttRAGs.wf])
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   537
  case (1 x)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   538
  show ?case (is "?L = ?R")
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   539
  proof -
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   540
    from tRAG_ancestorsE[OF 1(2)]
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   541
    obtain th2 where eq_x: "x = Th th2" by blast
102
3a801bbd2687 Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents: 93
diff changeset
   542
    from vat_es.cp_gen_rec[OF this]
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   543
    have "?L = 
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   544
          Max ({the_preced (e#s) th2} \<union> cp_gen (e#s) ` RTree.children (tRAG (e#s)) x)" .
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   545
    also have "... = 
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   546
          Max ({the_preced s th2} \<union> cp_gen s ` RTree.children (tRAG s) x)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   547
    proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   548
      from preced_kept have "the_preced (e#s) th2 = the_preced s th2" by simp
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   549
      moreover have "cp_gen (e#s) ` RTree.children (tRAG (e#s)) x =
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   550
                     cp_gen s ` RTree.children (tRAG s) x"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   551
      proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   552
        have "RTree.children (tRAG (e#s)) x =  RTree.children (tRAG s) x"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   553
        proof(unfold tRAG_s, rule children_union_kept)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   554
          have start: "(Th th, Th holder) \<in> tRAG (e#s)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   555
            by (unfold tRAG_s, auto)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   556
          note x_u = 1(2)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   557
          show "x \<notin> Range {(Th th, Th holder)}"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   558
          proof
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   559
            assume "x \<in> Range {(Th th, Th holder)}"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   560
            hence eq_x: "x = Th holder" using RangeE by auto
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   561
            show False
102
3a801bbd2687 Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents: 93
diff changeset
   562
            proof(cases rule:vat_es.ancestors_headE[OF assms(1) start])
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   563
              case 1
102
3a801bbd2687 Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents: 93
diff changeset
   564
              from x_u[folded this, unfolded eq_x] vat_es.acyclic_tRAG
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   565
              show ?thesis by (auto simp:ancestors_def acyclic_def)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   566
            next
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   567
              case 2
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   568
              with x_u[unfolded eq_x]
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   569
              have "(Th holder, Th holder) \<in> (tRAG (e#s))^+" by (auto simp:ancestors_def)
102
3a801bbd2687 Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents: 93
diff changeset
   570
              with vat_es.acyclic_tRAG show ?thesis by (auto simp:acyclic_def)
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   571
            qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   572
          qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   573
        qed
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   574
        moreover have "cp_gen (e#s) ` RTree.children (tRAG (e#s)) x =
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   575
                       cp_gen s ` RTree.children (tRAG (e#s)) x" (is "?f ` ?A = ?g ` ?A")
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   576
        proof(rule f_image_eq)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   577
          fix a
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   578
          assume a_in: "a \<in> ?A"
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   579
          from 1(2)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   580
          show "?f a = ?g a"
102
3a801bbd2687 Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents: 93
diff changeset
   581
          proof(cases rule:vat_es.rtree_s.ancestors_childrenE[case_names in_ch out_ch])
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   582
             case in_ch
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   583
             show ?thesis
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   584
             proof(cases "a = u")
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   585
                case True
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   586
                from assms(2)[folded this] show ?thesis .
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   587
             next
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   588
                case False
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   589
                have a_not_in: "a \<notin> ancestors (tRAG (e#s)) (Th th)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   590
                proof
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   591
                  assume a_in': "a \<in> ancestors (tRAG (e#s)) (Th th)"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   592
                  have "a = u"
102
3a801bbd2687 Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents: 93
diff changeset
   593
                  proof(rule vat_es.rtree_s.ancestors_children_unique)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   594
                    from a_in' a_in show "a \<in> ancestors (tRAG (e#s)) (Th th) \<inter> 
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   595
                                          RTree.children (tRAG (e#s)) x" by auto
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   596
                  next 
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   597
                    from assms(1) in_ch show "u \<in> ancestors (tRAG (e#s)) (Th th) \<inter> 
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   598
                                      RTree.children (tRAG (e#s)) x" by auto
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   599
                  qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   600
                  with False show False by simp
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   601
                qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   602
                from a_in obtain th_a where eq_a: "a = Th th_a" 
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   603
                    by (unfold RTree.children_def tRAG_alt_def, auto)
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   604
                from cp_kept[OF a_not_in[unfolded eq_a]]
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   605
                have "cp (e#s) th_a = cp s th_a" .
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   606
                from this [unfolded cp_gen_def_cond[OF eq_a], folded eq_a]
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   607
                show ?thesis .
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   608
             qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   609
          next
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   610
            case (out_ch z)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   611
            hence h: "z \<in> ancestors (tRAG (e#s)) u" "z \<in> RTree.children (tRAG (e#s)) x" by auto
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   612
            show ?thesis
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   613
            proof(cases "a = z")
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   614
              case True
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   615
              from h(2) have zx_in: "(z, x) \<in> (tRAG (e#s))" by (auto simp:RTree.children_def)
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   616
              from 1(1)[rule_format, OF this h(1)]
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   617
              have eq_cp_gen: "cp_gen (e#s) z = cp_gen s z" .
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   618
              with True show ?thesis by metis
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   619
            next
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   620
              case False
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   621
              from a_in obtain th_a where eq_a: "a = Th th_a"
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   622
                by (auto simp:RTree.children_def tRAG_alt_def)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   623
              have "a \<notin> ancestors (tRAG (e#s)) (Th th)"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   624
              proof
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   625
                assume a_in': "a \<in> ancestors (tRAG (e#s)) (Th th)"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   626
                have "a = z"
102
3a801bbd2687 Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents: 93
diff changeset
   627
                proof(rule vat_es.rtree_s.ancestors_children_unique)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   628
                  from assms(1) h(1) have "z \<in> ancestors (tRAG (e#s)) (Th th)"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   629
                      by (auto simp:ancestors_def)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   630
                  with h(2) show " z \<in> ancestors (tRAG (e#s)) (Th th) \<inter> 
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   631
                                       RTree.children (tRAG (e#s)) x" by auto
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   632
                next
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   633
                  from a_in a_in'
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   634
                  show "a \<in> ancestors (tRAG (e#s)) (Th th) \<inter> RTree.children (tRAG (e#s)) x"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   635
                    by auto
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   636
                qed
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   637
                with False show False by auto
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   638
              qed
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   639
              from cp_kept[OF this[unfolded eq_a]]
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   640
              have "cp (e#s) th_a = cp s th_a" .
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   641
              from this[unfolded cp_gen_def_cond[OF eq_a], folded eq_a]
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   642
              show ?thesis .
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   643
            qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   644
          qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   645
        qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   646
        ultimately show ?thesis by metis
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   647
      qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   648
      ultimately show ?thesis by simp
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   649
    qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   650
    also have "... = ?R"
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   651
      by (fold cp_gen_rec[OF eq_x], simp)
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   652
    finally show ?thesis .
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   653
  qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   654
qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 56
diff changeset
   655
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   656
lemma cp_up:
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   657
  assumes "(Th th') \<in> ancestors (tRAG (e#s)) (Th th)"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   658
  and "cp (e#s) th' = cp s th'"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   659
  and "(Th th'') \<in> ancestors (tRAG (e#s)) (Th th')"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   660
  shows "cp (e#s) th'' = cp s th''"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   661
proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   662
  have "cp_gen (e#s) (Th th'') = cp_gen s (Th th'')"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   663
  proof(rule cp_gen_update_stop[OF assms(1) _ assms(3)])
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   664
    from assms(2) cp_gen_def_cond[OF refl[of "Th th'"]]
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   665
    show "cp_gen (e#s) (Th th') = cp_gen s (Th th')" by metis
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   666
  qed
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   667
  with cp_gen_def_cond[OF refl[of "Th th''"]]
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   668
  show ?thesis by metis
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   669
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   670
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   671
end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   672
68
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   673
section {* The @{term Create} operation *}
db196b066b97 Before retrofiting PIPBasics.thy
zhangx
parents: 65
diff changeset
   674
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   675
context valid_trace_create
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   676
begin 
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   677
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   678
lemma tRAG_kept: "tRAG (e#s) = tRAG s"
113
ce85c3c4e5bf More improvements in PIPBasics.thy and Implemenation.thy.
zhangx
parents: 108
diff changeset
   679
  by (unfold tRAG_alt_def RAG_es, auto)
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   680
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   681
lemma preced_kept:
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   682
  assumes "th' \<noteq> th"
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   683
  shows "the_preced (e#s) th' = the_preced s th'"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   684
  by (unfold the_preced_def preced_def is_create, insert assms, auto)
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   685
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   686
lemma th_not_in: "Th th \<notin> Field (tRAG s)"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   687
  by (meson not_in_thread_isolated subsetCE tRAG_Field th_not_live_s)
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   688
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   689
lemma eq_cp:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   690
  assumes neq_th: "th' \<noteq> th"
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   691
  shows "cp (e#s) th' = cp s th'"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   692
proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   693
  have "(the_preced (e#s) \<circ> the_thread) ` subtree (tRAG (e#s)) (Th th') =
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   694
        (the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th')"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   695
  proof(unfold tRAG_kept, rule f_image_eq)
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   696
    fix a
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   697
    assume a_in: "a \<in> subtree (tRAG s) (Th th')"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   698
    then obtain th_a where eq_a: "a = Th th_a" 
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   699
    proof(cases rule:subtreeE)
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   700
      case 2
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   701
      from ancestors_Field[OF 2(2)]
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   702
      and that show ?thesis by (unfold tRAG_alt_def, auto)
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   703
    qed auto
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   704
    have neq_th_a: "th_a \<noteq> th"
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   705
    proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   706
      have "(Th th) \<notin> subtree (tRAG s) (Th th')"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   707
      proof
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   708
        assume "Th th \<in> subtree (tRAG s) (Th th')"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   709
        thus False
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   710
        proof(cases rule:subtreeE)
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   711
          case 2
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   712
          from ancestors_Field[OF this(2)]
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   713
          and th_not_in[unfolded Field_def]
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   714
          show ?thesis by auto
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   715
        qed (insert assms, auto)
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   716
      qed
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   717
      with a_in[unfolded eq_a] show ?thesis by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   718
    qed
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   719
    from preced_kept[OF this]
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   720
    show "(the_preced (e#s) \<circ> the_thread) a = (the_preced s \<circ> the_thread) a"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   721
      by (unfold eq_a, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   722
  qed
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   723
  thus ?thesis by (unfold cp_alt_def1, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   724
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   725
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   726
lemma children_of_th: "RTree.children (tRAG (e#s)) (Th th) = {}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   727
proof -
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   728
  { fix a
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   729
    assume "a \<in> RTree.children (tRAG (e#s)) (Th th)"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   730
    hence "(a, Th th) \<in> tRAG (e#s)" by (auto simp:RTree.children_def)
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   731
    with th_not_in have False 
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   732
     by (unfold Field_def tRAG_kept, auto)
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   733
  } thus ?thesis by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   734
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   735
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   736
lemma eq_cp_th: "cp (e#s) th = preced th (e#s)"
102
3a801bbd2687 Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents: 93
diff changeset
   737
 by (unfold vat_es.cp_rec children_of_th, simp add:the_preced_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   738
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   739
end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   740
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   741
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   742
context valid_trace_exit
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   743
begin
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   744
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   745
lemma preced_kept:
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   746
  assumes "th' \<noteq> th"
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   747
  shows "the_preced (e#s) th' = the_preced s th'"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   748
  using assms
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   749
  by (unfold the_preced_def is_exit preced_def, simp)
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   750
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   751
lemma tRAG_kept: "tRAG (e#s) = tRAG s"
113
ce85c3c4e5bf More improvements in PIPBasics.thy and Implemenation.thy.
zhangx
parents: 108
diff changeset
   752
  by (unfold tRAG_alt_def RAG_es, auto)
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   753
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   754
lemma th_RAG: "Th th \<notin> Field (RAG s)"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   755
proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   756
  have "Th th \<notin> Range (RAG s)"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   757
  proof
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   758
    assume "Th th \<in> Range (RAG s)"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   759
    then obtain cs where "holding (wq s) th cs"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   760
      by (unfold Range_iff s_RAG_def, auto)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   761
    with holdents_th_s[unfolded holdents_def]
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   762
    show False by (unfold holding_eq, auto)
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   763
  qed
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   764
  moreover have "Th th \<notin> Domain (RAG s)"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   765
  proof
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   766
    assume "Th th \<in> Domain (RAG s)"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   767
    then obtain cs where "waiting (wq s) th cs"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   768
      by (unfold Domain_iff s_RAG_def, auto)
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   769
    with th_ready_s show False by (unfold readys_def waiting_eq, auto)
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   770
  qed
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   771
  ultimately show ?thesis by (auto simp:Field_def)
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   772
qed
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   773
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   774
lemma th_tRAG: "(Th th) \<notin> Field (tRAG s)"
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   775
  using th_RAG tRAG_Field by auto
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   776
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   777
lemma eq_cp:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   778
  assumes neq_th: "th' \<noteq> th"
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   779
  shows "cp (e#s) th' = cp s th'"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   780
proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   781
  have "(the_preced (e#s) \<circ> the_thread) ` subtree (tRAG (e#s)) (Th th') =
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   782
        (the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th')"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   783
  proof(unfold tRAG_kept, rule f_image_eq)
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   784
    fix a
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   785
    assume a_in: "a \<in> subtree (tRAG s) (Th th')"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   786
    then obtain th_a where eq_a: "a = Th th_a" 
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   787
    proof(cases rule:subtreeE)
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   788
      case 2
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   789
      from ancestors_Field[OF 2(2)]
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   790
      and that show ?thesis by (unfold tRAG_alt_def, auto)
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   791
    qed auto
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   792
    have neq_th_a: "th_a \<noteq> th"
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   793
    proof -
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   794
      from readys_in_no_subtree[OF th_ready_s assms]
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   795
      have "(Th th) \<notin> subtree (RAG s) (Th th')" .
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   796
      with tRAG_subtree_RAG[of s "Th th'"]
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   797
      have "(Th th) \<notin> subtree (tRAG s) (Th th')" by auto
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   798
      with a_in[unfolded eq_a] show ?thesis by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   799
    qed
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   800
    from preced_kept[OF this]
92
4763aa246dbd Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents: 68
diff changeset
   801
    show "(the_preced (e#s) \<circ> the_thread) a = (the_preced s \<circ> the_thread) a"
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   802
      by (unfold eq_a, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   803
  qed
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   804
  thus ?thesis by (unfold cp_alt_def1, simp)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   805
qed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   806
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   807
end
60
f98a95f3deae Main proofs in CpsG.thy completed.
zhangx
parents: 59
diff changeset
   808
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   809
end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   810