CpsG.thy~
author zhangx
Thu, 28 Jan 2016 21:14:17 +0800
changeset 90 ed938e2246b9
parent 80 17305a85493d
permissions -rw-r--r--
Retrofiting of: CpsG.thy (the parallel copy of PIPBasics.thy), ExtGG.thy (The paralell copy of Implemenation.thy), PrioG.thy (The paralell copy of Correctness.thy) has completed. The next step is to overwite original copies with the paralell ones.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
     1
theory CpsG
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
     2
imports PIPDefs
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
     3
begin
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
     4
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
     5
lemma f_image_eq:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
     6
  assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
     7
  shows "f ` A = g ` A"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
     8
proof
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
     9
  show "f ` A \<subseteq> g ` A"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    10
    by(rule image_subsetI, auto intro:h)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    11
next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    12
  show "g ` A \<subseteq> f ` A"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    13
   by (rule image_subsetI, auto intro:h[symmetric])
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    14
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    15
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    16
lemma Max_fg_mono:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    17
  assumes "finite A"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    18
  and "\<forall> a \<in> A. f a \<le> g a"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    19
  shows "Max (f ` A) \<le> Max (g ` A)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    20
proof(cases "A = {}")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    21
  case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    22
  thus ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    23
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    24
  case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    25
  show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    26
  proof(rule Max.boundedI)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    27
    from assms show "finite (f ` A)" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    28
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    29
    from False show "f ` A \<noteq> {}" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    30
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    31
    fix fa
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    32
    assume "fa \<in> f ` A"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    33
    then obtain a where h_fa: "a \<in> A" "fa = f a" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    34
    show "fa \<le> Max (g ` A)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    35
    proof(rule Max_ge_iff[THEN iffD2])
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    36
      from assms show "finite (g ` A)" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    37
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    38
      from False show "g ` A \<noteq> {}" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    39
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    40
      from h_fa have "g a \<in> g ` A" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    41
      moreover have "fa \<le> g a" using h_fa assms(2) by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    42
      ultimately show "\<exists>a\<in>g ` A. fa \<le> a" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    43
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    44
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    45
qed 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
    46
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
    47
lemma Max_f_mono:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
    48
  assumes seq: "A \<subseteq> B"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
    49
  and np: "A \<noteq> {}"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
    50
  and fnt: "finite B"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
    51
  shows "Max (f ` A) \<le> Max (f ` B)"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
    52
proof(rule Max_mono)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
    53
  from seq show "f ` A \<subseteq> f ` B" by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
    54
next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
    55
  from np show "f ` A \<noteq> {}" by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
    56
next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
    57
  from fnt and seq show "finite (f ` B)" by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
    58
qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
    59
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    60
lemma Max_UNION: 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    61
  assumes "finite A"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    62
  and "A \<noteq> {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    63
  and "\<forall> M \<in> f ` A. finite M"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    64
  and "\<forall> M \<in> f ` A. M \<noteq> {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    65
  shows "Max (\<Union>x\<in> A. f x) = Max (Max ` f ` A)" (is "?L = ?R")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    66
  using assms[simp]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    67
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    68
  have "?L = Max (\<Union>(f ` A))"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    69
    by (fold Union_image_eq, simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    70
  also have "... = ?R"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    71
    by (subst Max_Union, simp+)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    72
  finally show ?thesis .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    73
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    74
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    75
lemma max_Max_eq:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    76
  assumes "finite A"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    77
    and "A \<noteq> {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    78
    and "x = y"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    79
  shows "max x (Max A) = Max ({y} \<union> A)" (is "?L = ?R")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    80
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    81
  have "?R = Max (insert y A)" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    82
  also from assms have "... = ?L"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    83
      by (subst Max.insert, simp+)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    84
  finally show ?thesis by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    85
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    86
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    87
lemma birth_time_lt:  
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    88
  assumes "s \<noteq> []"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    89
  shows "last_set th s < length s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    90
  using assms
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    91
proof(induct s)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    92
  case (Cons a s)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    93
  show ?case
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    94
  proof(cases "s \<noteq> []")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    95
    case False
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    96
    thus ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    97
      by (cases a, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    98
  next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
    99
    case True
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
   100
    show ?thesis using Cons(1)[OF True]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
   101
      by (cases a, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
   102
  qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
   103
qed simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
   104
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
   105
lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
   106
  by (induct s, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
   107
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
   108
lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
   109
  by (drule_tac th_in_ne, unfold preced_def, auto intro: birth_time_lt)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
   110
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   111
lemma eq_RAG: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   112
  "RAG (wq s) = RAG s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   113
  by (unfold cs_RAG_def s_RAG_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   114
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   115
lemma waiting_holding:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   116
  assumes "waiting (s::state) th cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   117
  obtains th' where "holding s th' cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   118
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   119
  from assms[unfolded s_waiting_def, folded wq_def]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   120
  obtain th' where "th' \<in> set (wq s cs)" "th' = hd (wq s cs)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   121
    by (metis empty_iff hd_in_set list.set(1))
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   122
  hence "holding s th' cs" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   123
    by (unfold s_holding_def, fold wq_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   124
  from that[OF this] show ?thesis .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   125
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   126
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   127
lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   128
unfolding cp_def wq_def
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   129
apply(induct s rule: schs.induct)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   130
apply(simp add: Let_def cpreced_initial)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   131
apply(simp add: Let_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   132
apply(simp add: Let_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   133
apply(simp add: Let_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   134
apply(subst (2) schs.simps)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   135
apply(simp add: Let_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   136
apply(subst (2) schs.simps)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   137
apply(simp add: Let_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   138
done
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   139
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   140
lemma cp_alt_def:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   141
  "cp s th =  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   142
           Max ((the_preced s) ` {th'. Th th' \<in> (subtree (RAG s) (Th th))})"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   143
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   144
  have "Max (the_preced s ` ({th} \<union> dependants (wq s) th)) =
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   145
        Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   146
          (is "Max (_ ` ?L) = Max (_ ` ?R)")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   147
  proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   148
    have "?L = ?R" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   149
    by (auto dest:rtranclD simp:cs_dependants_def cs_RAG_def s_RAG_def subtree_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   150
    thus ?thesis by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   151
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   152
  thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   153
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   154
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   155
(* ccc *)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   156
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   157
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   158
locale valid_trace = 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   159
  fixes s
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   160
  assumes vt : "vt s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   161
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   162
locale valid_trace_e = valid_trace +
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   163
  fixes e
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   164
  assumes vt_e: "vt (e#s)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   165
begin
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   166
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   167
lemma pip_e: "PIP s e"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   168
  using vt_e by (cases, simp)  
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   169
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   170
end
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   171
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   172
locale valid_trace_create = valid_trace_e + 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   173
  fixes th prio
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   174
  assumes is_create: "e = Create th prio"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   175
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   176
locale valid_trace_exit = valid_trace_e + 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   177
  fixes th
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   178
  assumes is_exit: "e = Exit th"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   179
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   180
locale valid_trace_p = valid_trace_e + 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   181
  fixes th cs
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   182
  assumes is_p: "e = P th cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   183
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   184
locale valid_trace_v = valid_trace_e + 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   185
  fixes th cs
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   186
  assumes is_v: "e = V th cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   187
begin
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   188
  definition "rest = tl (wq s cs)"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   189
  definition "wq' = (SOME q. distinct q \<and> set q = set rest)"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   190
end
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   191
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   192
locale valid_trace_v_n = valid_trace_v +
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   193
  assumes rest_nnl: "rest \<noteq> []"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   194
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   195
locale valid_trace_v_e = valid_trace_v +
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   196
  assumes rest_nil: "rest = []"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   197
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   198
locale valid_trace_set= valid_trace_e + 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   199
  fixes th prio
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   200
  assumes is_set: "e = Set th prio"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   201
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   202
context valid_trace
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   203
begin
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   204
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   205
lemma ind [consumes 0, case_names Nil Cons, induct type]:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   206
  assumes "PP []"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   207
     and "(\<And>s e. valid_trace_e s e \<Longrightarrow>
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   208
                   PP s \<Longrightarrow> PIP s e \<Longrightarrow> PP (e # s))"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   209
     shows "PP s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   210
proof(induct rule:vt.induct[OF vt, case_names Init Step])
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   211
  case Init
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   212
  from assms(1) show ?case .
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   213
next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   214
  case (Step s e)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   215
  show ?case
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   216
  proof(rule assms(2))
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   217
    show "valid_trace_e s e" using Step by (unfold_locales, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   218
  next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   219
    show "PP s" using Step by simp
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   220
  next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   221
    show "PIP s e" using Step by simp
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   222
  qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   223
qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   224
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   225
lemma  vt_moment: "\<And> t. vt (moment t s)"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   226
proof(induct rule:ind)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   227
  case Nil
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   228
  thus ?case by (simp add:vt_nil)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   229
next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   230
  case (Cons s e t)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   231
  show ?case
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   232
  proof(cases "t \<ge> length (e#s)")
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   233
    case True
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   234
    from True have "moment t (e#s) = e#s" by simp
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   235
    thus ?thesis using Cons
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   236
      by (simp add:valid_trace_def valid_trace_e_def, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   237
  next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   238
    case False
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   239
    from Cons have "vt (moment t s)" by simp
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   240
    moreover have "moment t (e#s) = moment t s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   241
    proof -
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   242
      from False have "t \<le> length s" by simp
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   243
      from moment_app [OF this, of "[e]"] 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   244
      show ?thesis by simp
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   245
    qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   246
    ultimately show ?thesis by simp
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   247
  qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   248
qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   249
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   250
lemma finite_threads:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   251
  shows "finite (threads s)"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   252
using vt by (induct) (auto elim: step.cases)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   253
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   254
end
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   255
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   256
lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   257
  by (unfold s_RAG_def, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   258
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   259
locale valid_moment = valid_trace + 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   260
  fixes i :: nat
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   261
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   262
sublocale valid_moment < vat_moment: valid_trace "(moment i s)"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   263
  by (unfold_locales, insert vt_moment, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   264
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   265
lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   266
  by  (unfold s_waiting_def cs_waiting_def wq_def, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   267
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   268
lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   269
  by (unfold s_holding_def wq_def cs_holding_def, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   270
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   271
lemma runing_ready: 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   272
  shows "runing s \<subseteq> readys s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   273
  unfolding runing_def readys_def
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   274
  by auto 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   275
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   276
lemma readys_threads:
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   277
  shows "readys s \<subseteq> threads s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   278
  unfolding readys_def
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   279
  by auto
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   280
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   281
lemma wq_v_neq [simp]:
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   282
   "cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   283
  by (auto simp:wq_def Let_def cp_def split:list.splits)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   284
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   285
lemma runing_head:
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   286
  assumes "th \<in> runing s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   287
  and "th \<in> set (wq_fun (schs s) cs)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   288
  shows "th = hd (wq_fun (schs s) cs)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   289
  using assms
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   290
  by (simp add:runing_def readys_def s_waiting_def wq_def)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   291
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   292
context valid_trace
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   293
begin
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   294
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   295
lemma runing_wqE:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   296
  assumes "th \<in> runing s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   297
  and "th \<in> set (wq s cs)"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   298
  obtains rest where "wq s cs = th#rest"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   299
proof -
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   300
  from assms(2) obtain th' rest where eq_wq: "wq s cs = th'#rest"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   301
    by (meson list.set_cases)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   302
  have "th' = th"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   303
  proof(rule ccontr)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   304
    assume "th' \<noteq> th"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   305
    hence "th \<noteq> hd (wq s cs)" using eq_wq by auto 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   306
    with assms(2)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   307
    have "waiting s th cs" 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   308
      by (unfold s_waiting_def, fold wq_def, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   309
    with assms show False 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   310
      by (unfold runing_def readys_def, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   311
  qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   312
  with eq_wq that show ?thesis by metis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   313
qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   314
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   315
end
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   316
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   317
context valid_trace_create
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   318
begin
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   319
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   320
lemma wq_neq_simp [simp]:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   321
  shows "wq (e#s) cs' = wq s cs'"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   322
    using assms unfolding is_create wq_def
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   323
  by (auto simp:Let_def)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   324
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   325
lemma wq_distinct_kept:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   326
  assumes "distinct (wq s cs')"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   327
  shows "distinct (wq (e#s) cs')"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   328
  using assms by simp
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   329
end
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   330
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   331
context valid_trace_exit
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   332
begin
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   333
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   334
lemma wq_neq_simp [simp]:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   335
  shows "wq (e#s) cs' = wq s cs'"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   336
    using assms unfolding is_exit wq_def
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   337
  by (auto simp:Let_def)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   338
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   339
lemma wq_distinct_kept:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   340
  assumes "distinct (wq s cs')"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   341
  shows "distinct (wq (e#s) cs')"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   342
  using assms by simp
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   343
end
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   344
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   345
context valid_trace_p
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   346
begin
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   347
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   348
lemma wq_neq_simp [simp]:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   349
  assumes "cs' \<noteq> cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   350
  shows "wq (e#s) cs' = wq s cs'"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   351
    using assms unfolding is_p wq_def
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   352
  by (auto simp:Let_def)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   353
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   354
lemma runing_th_s:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   355
  shows "th \<in> runing s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   356
proof -
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   357
  from pip_e[unfolded is_p]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   358
  show ?thesis by (cases, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   359
qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   360
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   361
lemma ready_th_s: "th \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   362
  using runing_th_s
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   363
  by (unfold runing_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   364
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   365
lemma live_th_s: "th \<in> threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   366
  using readys_threads ready_th_s by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   367
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   368
lemma live_th_es: "th \<in> threads (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   369
  using live_th_s 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   370
  by (unfold is_p, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   371
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   372
lemma th_not_waiting: 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   373
  "\<not> waiting s th c"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   374
proof -
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   375
  have "th \<in> readys s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   376
    using runing_ready runing_th_s by blast 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   377
  thus ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   378
    by (unfold readys_def, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   379
qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   380
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   381
lemma waiting_neq_th: 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   382
  assumes "waiting s t c"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   383
  shows "t \<noteq> th"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   384
  using assms using th_not_waiting by blast 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   385
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   386
lemma th_not_in_wq: 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   387
  shows "th \<notin> set (wq s cs)"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   388
proof
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   389
  assume otherwise: "th \<in> set (wq s cs)"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   390
  from runing_wqE[OF runing_th_s this]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   391
  obtain rest where eq_wq: "wq s cs = th#rest" by blast
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   392
  with otherwise
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   393
  have "holding s th cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   394
    by (unfold s_holding_def, fold wq_def, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   395
  hence cs_th_RAG: "(Cs cs, Th th) \<in> RAG s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   396
    by (unfold s_RAG_def, fold holding_eq, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   397
  from pip_e[unfolded is_p]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   398
  show False
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   399
  proof(cases)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   400
    case (thread_P)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   401
    with cs_th_RAG show ?thesis by auto
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   402
  qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   403
qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   404
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   405
lemma wq_es_cs: 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   406
  "wq (e#s) cs =  wq s cs @ [th]"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   407
  by (unfold is_p wq_def, auto simp:Let_def)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   408
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   409
lemma wq_distinct_kept:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   410
  assumes "distinct (wq s cs')"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   411
  shows "distinct (wq (e#s) cs')"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   412
proof(cases "cs' = cs")
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   413
  case True
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   414
  show ?thesis using True assms th_not_in_wq
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   415
    by (unfold True wq_es_cs, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   416
qed (insert assms, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   417
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   418
end
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   419
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   420
context valid_trace_v
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   421
begin
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   422
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   423
lemma wq_neq_simp [simp]:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   424
  assumes "cs' \<noteq> cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   425
  shows "wq (e#s) cs' = wq s cs'"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   426
    using assms unfolding is_v wq_def
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   427
  by (auto simp:Let_def)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   428
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   429
lemma runing_th_s:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   430
  shows "th \<in> runing s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   431
proof -
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   432
  from pip_e[unfolded is_v]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   433
  show ?thesis by (cases, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   434
qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   435
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   436
lemma th_not_waiting: 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   437
  "\<not> waiting s th c"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   438
proof -
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   439
  have "th \<in> readys s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   440
    using runing_ready runing_th_s by blast 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   441
  thus ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   442
    by (unfold readys_def, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   443
qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   444
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   445
lemma waiting_neq_th: 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   446
  assumes "waiting s t c"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   447
  shows "t \<noteq> th"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   448
  using assms using th_not_waiting by blast 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   449
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   450
lemma wq_s_cs:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   451
  "wq s cs = th#rest"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   452
proof -
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   453
  from pip_e[unfolded is_v]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   454
  show ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   455
  proof(cases)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   456
    case (thread_V)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   457
    from this(2) show ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   458
      by (unfold rest_def s_holding_def, fold wq_def,
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   459
                 metis empty_iff list.collapse list.set(1))
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   460
  qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   461
qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   462
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   463
lemma wq_es_cs:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   464
  "wq (e#s) cs = wq'"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   465
 using wq_s_cs[unfolded wq_def]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   466
 by (auto simp:Let_def wq_def rest_def wq'_def is_v, simp) 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   467
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   468
lemma wq_distinct_kept:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   469
  assumes "distinct (wq s cs')"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   470
  shows "distinct (wq (e#s) cs')"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   471
proof(cases "cs' = cs")
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   472
  case True
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   473
  show ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   474
  proof(unfold True wq_es_cs wq'_def, rule someI2)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   475
    show "distinct rest \<and> set rest = set rest"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   476
        using assms[unfolded True wq_s_cs] by auto
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   477
  qed simp
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   478
qed (insert assms, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   479
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   480
end
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   481
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   482
context valid_trace_set
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   483
begin
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   484
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   485
lemma wq_neq_simp [simp]:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   486
  shows "wq (e#s) cs' = wq s cs'"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   487
    using assms unfolding is_set wq_def
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   488
  by (auto simp:Let_def)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   489
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   490
lemma wq_distinct_kept:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   491
  assumes "distinct (wq s cs')"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   492
  shows "distinct (wq (e#s) cs')"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   493
  using assms by simp
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   494
end
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   495
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   496
context valid_trace
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   497
begin
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   498
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   499
lemma actor_inv: 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   500
  assumes "PIP s e"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   501
  and "\<not> isCreate e"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   502
  shows "actor e \<in> runing s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   503
  using assms
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   504
  by (induct, auto)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   505
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   506
lemma isP_E:
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   507
  assumes "isP e"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   508
  obtains cs where "e = P (actor e) cs"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   509
  using assms by (cases e, auto)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   510
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   511
lemma isV_E:
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   512
  assumes "isV e"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   513
  obtains cs where "e = V (actor e) cs"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   514
  using assms by (cases e, auto) 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   515
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   516
lemma wq_distinct: "distinct (wq s cs)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   517
proof(induct rule:ind)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   518
  case (Cons s e)
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   519
  interpret vt_e: valid_trace_e s e using Cons by simp
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   520
  show ?case 
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   521
  proof(cases e)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   522
    case (Create th prio)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   523
    interpret vt_create: valid_trace_create s e th prio 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   524
      using Create by (unfold_locales, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   525
    show ?thesis using Cons by (simp add: vt_create.wq_distinct_kept) 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   526
  next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   527
    case (Exit th)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   528
    interpret vt_exit: valid_trace_exit s e th  
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   529
        using Exit by (unfold_locales, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   530
    show ?thesis using Cons by (simp add: vt_exit.wq_distinct_kept) 
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   531
  next
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   532
    case (P th cs)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   533
    interpret vt_p: valid_trace_p s e th cs using P by (unfold_locales, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   534
    show ?thesis using Cons by (simp add: vt_p.wq_distinct_kept) 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   535
  next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   536
    case (V th cs)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   537
    interpret vt_v: valid_trace_v s e th cs using V by (unfold_locales, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   538
    show ?thesis using Cons by (simp add: vt_v.wq_distinct_kept) 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   539
  next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   540
    case (Set th prio)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   541
    interpret vt_set: valid_trace_set s e th prio
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   542
        using Set by (unfold_locales, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   543
    show ?thesis using Cons by (simp add: vt_set.wq_distinct_kept) 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   544
  qed
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   545
qed (unfold wq_def Let_def, simp)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   546
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   547
end
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   548
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   549
context valid_trace_e
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   550
begin
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   551
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   552
text {*
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   553
  The following lemma shows that only the @{text "P"}
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   554
  operation can add new thread into waiting queues. 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   555
  Such kind of lemmas are very obvious, but need to be checked formally.
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   556
  This is a kind of confirmation that our modelling is correct.
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   557
*}
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   558
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   559
lemma wq_in_inv: 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   560
  assumes s_ni: "thread \<notin> set (wq s cs)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   561
  and s_i: "thread \<in> set (wq (e#s) cs)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   562
  shows "e = P thread cs"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   563
proof(cases e)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   564
  -- {* This is the only non-trivial case: *}
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   565
  case (V th cs1)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   566
  have False
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   567
  proof(cases "cs1 = cs")
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   568
    case True
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   569
    show ?thesis
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   570
    proof(cases "(wq s cs1)")
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   571
      case (Cons w_hd w_tl)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   572
      have "set (wq (e#s) cs) \<subseteq> set (wq s cs)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   573
      proof -
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   574
        have "(wq (e#s) cs) = (SOME q. distinct q \<and> set q = set w_tl)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   575
          using  Cons V by (auto simp:wq_def Let_def True split:if_splits)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   576
        moreover have "set ... \<subseteq> set (wq s cs)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   577
        proof(rule someI2)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   578
          show "distinct w_tl \<and> set w_tl = set w_tl"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   579
            by (metis distinct.simps(2) local.Cons wq_distinct)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   580
        qed (insert Cons True, auto)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   581
        ultimately show ?thesis by simp
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   582
      qed
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   583
      with assms show ?thesis by auto
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   584
    qed (insert assms V True, auto simp:wq_def Let_def split:if_splits)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   585
  qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   586
  thus ?thesis by auto
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   587
qed (insert assms, auto simp:wq_def Let_def split:if_splits)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   588
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   589
lemma wq_out_inv: 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   590
  assumes s_in: "thread \<in> set (wq s cs)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   591
  and s_hd: "thread = hd (wq s cs)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   592
  and s_i: "thread \<noteq> hd (wq (e#s) cs)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   593
  shows "e = V thread cs"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   594
proof(cases e)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   595
-- {* There are only two non-trivial cases: *}
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   596
  case (V th cs1)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   597
  show ?thesis
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   598
  proof(cases "cs1 = cs")
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   599
    case True
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   600
    have "PIP s (V th cs)" using pip_e[unfolded V[unfolded True]] .
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   601
    thus ?thesis
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   602
    proof(cases)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   603
      case (thread_V)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   604
      moreover have "th = thread" using thread_V(2) s_hd
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   605
          by (unfold s_holding_def wq_def, simp)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   606
      ultimately show ?thesis using V True by simp
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   607
    qed
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   608
  qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   609
next
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   610
  case (P th cs1)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   611
  show ?thesis
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   612
  proof(cases "cs1 = cs")
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   613
    case True
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   614
    with P have "wq (e#s) cs = wq_fun (schs s) cs @ [th]"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   615
      by (auto simp:wq_def Let_def split:if_splits)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   616
    with s_i s_hd s_in have False
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   617
      by (metis empty_iff hd_append2 list.set(1) wq_def) 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   618
    thus ?thesis by simp
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   619
  qed (insert assms P, auto simp:wq_def Let_def split:if_splits)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   620
qed (insert assms, auto simp:wq_def Let_def split:if_splits)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   621
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   622
end
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   623
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   624
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   625
context valid_trace
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   626
begin
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   627
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   628
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   629
text {* (* ddd *)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   630
  The nature of the work is like this: since it starts from a very simple and basic 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   631
  model, even intuitively very `basic` and `obvious` properties need to derived from scratch.
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   632
  For instance, the fact 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   633
  that one thread can not be blocked by two critical resources at the same time
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   634
  is obvious, because only running threads can make new requests, if one is waiting for 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   635
  a critical resource and get blocked, it can not make another resource request and get 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   636
  blocked the second time (because it is not running). 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   637
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   638
  To derive this fact, one needs to prove by contraction and 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   639
  reason about time (or @{text "moement"}). The reasoning is based on a generic theorem
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   640
  named @{text "p_split"}, which is about status changing along the time axis. It says if 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   641
  a condition @{text "Q"} is @{text "True"} at a state @{text "s"},
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   642
  but it was @{text "False"} at the very beginning, then there must exits a moment @{text "t"} 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   643
  in the history of @{text "s"} (notice that @{text "s"} itself is essentially the history 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   644
  of events leading to it), such that @{text "Q"} switched 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   645
  from being @{text "False"} to @{text "True"} and kept being @{text "True"}
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   646
  till the last moment of @{text "s"}.
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   647
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   648
  Suppose a thread @{text "th"} is blocked
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   649
  on @{text "cs1"} and @{text "cs2"} in some state @{text "s"}, 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   650
  since no thread is blocked at the very beginning, by applying 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   651
  @{text "p_split"} to these two blocking facts, there exist 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   652
  two moments @{text "t1"} and @{text "t2"}  in @{text "s"}, such that 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   653
  @{text "th"} got blocked on @{text "cs1"} and @{text "cs2"} 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   654
  and kept on blocked on them respectively ever since.
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   655
 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   656
  Without lost of generality, we assume @{text "t1"} is earlier than @{text "t2"}.
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   657
  However, since @{text "th"} was blocked ever since memonent @{text "t1"}, so it was still
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   658
  in blocked state at moment @{text "t2"} and could not
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   659
  make any request and get blocked the second time: Contradiction.
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   660
*}
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   661
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   662
lemma waiting_unique_pre: (* ddd *)
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   663
  assumes h11: "thread \<in> set (wq s cs1)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   664
  and h12: "thread \<noteq> hd (wq s cs1)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   665
  assumes h21: "thread \<in> set (wq s cs2)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   666
  and h22: "thread \<noteq> hd (wq s cs2)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   667
  and neq12: "cs1 \<noteq> cs2"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   668
  shows "False"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   669
proof -
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   670
  let "?Q" = "\<lambda> cs s. thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   671
  from h11 and h12 have q1: "?Q cs1 s" by simp
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   672
  from h21 and h22 have q2: "?Q cs2 s" by simp
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   673
  have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   674
  have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   675
  from p_split [of "?Q cs1", OF q1 nq1]
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   676
  obtain t1 where lt1: "t1 < length s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   677
    and np1: "\<not> ?Q cs1 (moment t1 s)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   678
    and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))" by auto
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   679
  from p_split [of "?Q cs2", OF q2 nq2]
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   680
  obtain t2 where lt2: "t2 < length s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   681
    and np2: "\<not> ?Q cs2 (moment t2 s)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   682
    and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))" by auto
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   683
  { fix s cs
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   684
    assume q: "?Q cs s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   685
    have "thread \<notin> runing s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   686
    proof
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   687
      assume "thread \<in> runing s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   688
      hence " \<forall>cs. \<not> (thread \<in> set (wq_fun (schs s) cs) \<and> 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   689
                 thread \<noteq> hd (wq_fun (schs s) cs))"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   690
        by (unfold runing_def s_waiting_def readys_def, auto)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   691
      from this[rule_format, of cs] q 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   692
      show False by (simp add: wq_def) 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   693
    qed
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   694
  } note q_not_runing = this
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   695
  { fix t1 t2 cs1 cs2
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   696
    assume  lt1: "t1 < length s"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   697
    and np1: "\<not> ?Q cs1 (moment t1 s)"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   698
    and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   699
    and lt2: "t2 < length s"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   700
    and np2: "\<not> ?Q cs2 (moment t2 s)"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   701
    and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   702
    and lt12: "t1 < t2"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   703
    let ?t3 = "Suc t2"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   704
    from lt2 have le_t3: "?t3 \<le> length s" by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   705
    from moment_plus [OF this] 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   706
    obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   707
    have "t2 < ?t3" by simp
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   708
    from nn2 [rule_format, OF this] and eq_m
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   709
    have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   710
         h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   711
    have "vt (e#moment t2 s)"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   712
    proof -
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   713
      from vt_moment 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   714
      have "vt (moment ?t3 s)" .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   715
      with eq_m show ?thesis by simp
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   716
    qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   717
    then interpret vt_e: valid_trace_e "moment t2 s" "e"
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   718
        by (unfold_locales, auto, cases, simp)
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   719
    have ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   720
    proof -
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   721
      have "thread \<in> runing (moment t2 s)"
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   722
      proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   723
        case True
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   724
        have "e = V thread cs2"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   725
        proof -
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   726
          have eq_th: "thread = hd (wq (moment t2 s) cs2)" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   727
              using True and np2  by auto 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   728
          from vt_e.wq_out_inv[OF True this h2]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   729
          show ?thesis .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   730
        qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   731
        thus ?thesis using vt_e.actor_inv[OF vt_e.pip_e] by auto
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   732
      next
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   733
        case False
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   734
        have "e = P thread cs2" using vt_e.wq_in_inv[OF False h1] .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   735
        with vt_e.actor_inv[OF vt_e.pip_e]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   736
        show ?thesis by auto
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   737
      qed
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   738
      moreover have "thread \<notin> runing (moment t2 s)"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   739
        by (rule q_not_runing[OF nn1[rule_format, OF lt12]])
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   740
      ultimately show ?thesis by simp
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   741
    qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   742
  } note lt_case = this
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   743
  show ?thesis
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   744
  proof -
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   745
    { assume "t1 < t2"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   746
      from lt_case[OF lt1 np1 nn1 lt2 np2 nn2 this]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   747
      have ?thesis .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   748
    } moreover {
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   749
      assume "t2 < t1"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   750
      from lt_case[OF lt2 np2 nn2 lt1 np1 nn1 this]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   751
      have ?thesis .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   752
    } moreover {
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   753
      assume eq_12: "t1 = t2"
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   754
      let ?t3 = "Suc t2"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   755
      from lt2 have le_t3: "?t3 \<le> length s" by auto
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   756
      from moment_plus [OF this] 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   757
      obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   758
      have lt_2: "t2 < ?t3" by simp
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   759
      from nn2 [rule_format, OF this] and eq_m
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   760
      have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   761
           h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   762
      from nn1[rule_format, OF lt_2[folded eq_12]] eq_m[folded eq_12]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   763
      have g1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   764
           g2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   765
      have "vt (e#moment t2 s)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   766
      proof -
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   767
        from vt_moment 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   768
        have "vt (moment ?t3 s)" .
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   769
        with eq_m show ?thesis by simp
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   770
      qed
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   771
      then interpret vt_e: valid_trace_e "moment t2 s" "e"
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   772
          by (unfold_locales, auto, cases, simp)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   773
      have "e = V thread cs2 \<or> e = P thread cs2"
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   774
      proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   775
        case True
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   776
        have "e = V thread cs2"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   777
        proof -
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   778
          have eq_th: "thread = hd (wq (moment t2 s) cs2)" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   779
              using True and np2  by auto 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   780
          from vt_e.wq_out_inv[OF True this h2]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   781
          show ?thesis .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   782
        qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   783
        thus ?thesis by auto
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   784
      next
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   785
        case False
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   786
        have "e = P thread cs2" using vt_e.wq_in_inv[OF False h1] .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   787
        thus ?thesis by auto
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   788
      qed
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   789
      moreover have "e = V thread cs1 \<or> e = P thread cs1"
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   790
      proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   791
        case True
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   792
        have eq_th: "thread = hd (wq (moment t1 s) cs1)" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   793
              using True and np1  by auto 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   794
        from vt_e.wq_out_inv[folded eq_12, OF True this g2]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   795
        have "e = V thread cs1" .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   796
        thus ?thesis by auto
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   797
      next
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   798
        case False
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   799
        have "e = P thread cs1" using vt_e.wq_in_inv[folded eq_12, OF False g1] .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   800
        thus ?thesis by auto
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   801
      qed
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   802
      ultimately have ?thesis using neq12 by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   803
    } ultimately show ?thesis using nat_neq_iff by blast 
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   804
  qed
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   805
qed
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   806
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   807
text {*
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   808
  This lemma is a simple corrolary of @{text "waiting_unique_pre"}.
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   809
*}
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   810
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   811
lemma waiting_unique:
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   812
  assumes "waiting s th cs1"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   813
  and "waiting s th cs2"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   814
  shows "cs1 = cs2"
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   815
  using waiting_unique_pre assms
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   816
  unfolding wq_def s_waiting_def
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   817
  by auto
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   818
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   819
end
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   820
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   821
(* not used *)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   822
text {*
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   823
  Every thread can only be blocked on one critical resource, 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   824
  symmetrically, every critical resource can only be held by one thread. 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   825
  This fact is much more easier according to our definition. 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   826
*}
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   827
lemma held_unique:
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   828
  assumes "holding (s::event list) th1 cs"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   829
  and "holding s th2 cs"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   830
  shows "th1 = th2"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   831
 by (insert assms, unfold s_holding_def, auto)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   832
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   833
lemma last_set_lt: "th \<in> threads s \<Longrightarrow> last_set th s < length s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   834
  apply (induct s, auto)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   835
  by (case_tac a, auto split:if_splits)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   836
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   837
lemma last_set_unique: 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   838
  "\<lbrakk>last_set th1 s = last_set th2 s; th1 \<in> threads s; th2 \<in> threads s\<rbrakk>
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   839
          \<Longrightarrow> th1 = th2"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   840
  apply (induct s, auto)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   841
  by (case_tac a, auto split:if_splits dest:last_set_lt)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   842
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   843
lemma preced_unique : 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   844
  assumes pcd_eq: "preced th1 s = preced th2 s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   845
  and th_in1: "th1 \<in> threads s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   846
  and th_in2: " th2 \<in> threads s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   847
  shows "th1 = th2"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   848
proof -
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   849
  from pcd_eq have "last_set th1 s = last_set th2 s" by (simp add:preced_def)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   850
  from last_set_unique [OF this th_in1 th_in2]
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   851
  show ?thesis .
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   852
qed
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   853
                      
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   854
lemma preced_linorder: 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   855
  assumes neq_12: "th1 \<noteq> th2"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   856
  and th_in1: "th1 \<in> threads s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   857
  and th_in2: " th2 \<in> threads s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   858
  shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   859
proof -
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   860
  from preced_unique [OF _ th_in1 th_in2] and neq_12 
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   861
  have "preced th1 s \<noteq> preced th2 s" by auto
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   862
  thus ?thesis by auto
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   863
qed
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   864
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   865
text {*
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   866
  The following three lemmas show that @{text "RAG"} does not change
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   867
  by the happening of @{text "Set"}, @{text "Create"} and @{text "Exit"}
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   868
  events, respectively.
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   869
*}
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   870
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   871
lemma RAG_set_unchanged: "(RAG (Set th prio # s)) = RAG s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   872
apply (unfold s_RAG_def s_waiting_def wq_def)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   873
by (simp add:Let_def)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   874
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   875
lemma (in valid_trace_set)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   876
   RAG_unchanged: "(RAG (e # s)) = RAG s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   877
   by (unfold is_set RAG_set_unchanged, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   878
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   879
lemma RAG_create_unchanged: "(RAG (Create th prio # s)) = RAG s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   880
apply (unfold s_RAG_def s_waiting_def wq_def)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   881
by (simp add:Let_def)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   882
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   883
lemma (in valid_trace_create)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   884
   RAG_unchanged: "(RAG (e # s)) = RAG s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   885
   by (unfold is_create RAG_create_unchanged, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   886
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   887
lemma RAG_exit_unchanged: "(RAG (Exit th # s)) = RAG s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   888
apply (unfold s_RAG_def s_waiting_def wq_def)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   889
by (simp add:Let_def)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   890
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   891
lemma (in valid_trace_exit)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   892
   RAG_unchanged: "(RAG (e # s)) = RAG s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   893
   by (unfold is_exit RAG_exit_unchanged, simp)
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   894
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   895
context valid_trace_v
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   896
begin
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   897
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   898
lemma distinct_rest: "distinct rest"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   899
  by (simp add: distinct_tl rest_def wq_distinct)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   900
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   901
lemma holding_cs_eq_th:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   902
  assumes "holding s t cs"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   903
  shows "t = th"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   904
proof -
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   905
  from pip_e[unfolded is_v]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   906
  show ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   907
  proof(cases)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   908
    case (thread_V)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   909
    from held_unique[OF this(2) assms]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   910
    show ?thesis by simp
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   911
  qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   912
qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   913
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   914
lemma distinct_wq': "distinct wq'"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   915
  by (metis (mono_tags, lifting) distinct_rest  some_eq_ex wq'_def)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   916
  
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   917
lemma set_wq': "set wq' = set rest"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   918
  by (metis (mono_tags, lifting) distinct_rest rest_def 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   919
      some_eq_ex wq'_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   920
    
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   921
lemma th'_in_inv:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   922
  assumes "th' \<in> set wq'"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   923
  shows "th' \<in> set rest"
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
   924
  using assms set_wq' by simp
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   925
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   926
lemma neq_t_th: 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   927
  assumes "waiting (e#s) t c"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   928
  shows "t \<noteq> th"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   929
proof
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   930
  assume otherwise: "t = th"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   931
  show False
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   932
  proof(cases "c = cs")
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   933
    case True
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   934
    have "t \<in> set wq'" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   935
     using assms[unfolded True s_waiting_def, folded wq_def, unfolded wq_es_cs]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   936
     by simp 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   937
    from th'_in_inv[OF this] have "t \<in> set rest" .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   938
    with wq_s_cs[folded otherwise] wq_distinct[of cs]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   939
    show ?thesis by simp
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   940
  next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   941
    case False
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   942
    have "wq (e#s) c = wq s c" using False
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   943
        by (unfold is_v, simp)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   944
    hence "waiting s t c" using assms 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   945
        by (simp add: cs_waiting_def waiting_eq)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   946
    hence "t \<notin> readys s" by (unfold readys_def, auto)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   947
    hence "t \<notin> runing s" using runing_ready by auto 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   948
    with runing_th_s[folded otherwise] show ?thesis by auto
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   949
  qed
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   950
qed
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
   951
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   952
lemma waiting_esI1:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   953
  assumes "waiting s t c"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   954
      and "c \<noteq> cs" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   955
  shows "waiting (e#s) t c" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   956
proof -
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   957
  have "wq (e#s) c = wq s c" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   958
    using assms(2) is_v by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   959
  with assms(1) show ?thesis 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   960
    using cs_waiting_def waiting_eq by auto 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   961
qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   962
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   963
lemma holding_esI2:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   964
  assumes "c \<noteq> cs" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   965
  and "holding s t c"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   966
  shows "holding (e#s) t c"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   967
proof -
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   968
  from assms(1) have "wq (e#s) c = wq s c" using is_v by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   969
  from assms(2)[unfolded s_holding_def, folded wq_def, 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   970
                folded this, unfolded wq_def, folded s_holding_def]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   971
  show ?thesis .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   972
qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   973
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   974
lemma holding_esI1:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   975
  assumes "holding s t c"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   976
  and "t \<noteq> th"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   977
  shows "holding (e#s) t c"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   978
proof -
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   979
  have "c \<noteq> cs" using assms using holding_cs_eq_th by blast 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   980
  from holding_esI2[OF this assms(1)]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   981
  show ?thesis .
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   982
qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   983
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   984
end
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   985
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
   986
context valid_trace_v_n
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   987
begin
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   988
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   989
lemma neq_wq': "wq' \<noteq> []" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   990
proof (unfold wq'_def, rule someI2)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   991
  show "distinct rest \<and> set rest = set rest"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   992
    by (simp add: distinct_rest) 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   993
next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   994
  fix x
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   995
  assume " distinct x \<and> set x = set rest" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   996
  thus "x \<noteq> []" using rest_nnl by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   997
qed 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   998
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
   999
definition "taker = hd wq'"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1000
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1001
definition "rest' = tl wq'"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1002
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1003
lemma eq_wq': "wq' = taker # rest'"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1004
  by (simp add: neq_wq' rest'_def taker_def)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1005
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1006
lemma next_th_taker: 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1007
  shows "next_th s th cs taker"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1008
  using rest_nnl taker_def wq'_def wq_s_cs 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1009
  by (auto simp:next_th_def)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1010
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1011
lemma taker_unique: 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1012
  assumes "next_th s th cs taker'"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1013
  shows "taker' = taker"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1014
proof -
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1015
  from assms
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1016
  obtain rest' where 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1017
    h: "wq s cs = th # rest'" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1018
       "taker' = hd (SOME q. distinct q \<and> set q = set rest')"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1019
          by (unfold next_th_def, auto)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1020
  with wq_s_cs have "rest' = rest" by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1021
  thus ?thesis using h(2) taker_def wq'_def by auto 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1022
qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1023
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1024
lemma waiting_set_eq:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1025
  "{(Th th', Cs cs) |th'. next_th s th cs th'} = {(Th taker, Cs cs)}"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1026
  by (smt all_not_in_conv bot.extremum insertI1 insert_subset 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1027
      mem_Collect_eq next_th_taker subsetI subset_antisym taker_def taker_unique)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1028
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1029
lemma holding_set_eq:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1030
  "{(Cs cs, Th th') |th'.  next_th s th cs th'} = {(Cs cs, Th taker)}"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1031
  using next_th_taker taker_def waiting_set_eq 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1032
  by fastforce
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1033
   
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1034
lemma holding_taker:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1035
  shows "holding (e#s) taker cs"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1036
    by (unfold s_holding_def, fold wq_def, unfold wq_es_cs, 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1037
        auto simp:neq_wq' taker_def)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1038
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1039
lemma waiting_esI2:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1040
  assumes "waiting s t cs"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1041
      and "t \<noteq> taker"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1042
  shows "waiting (e#s) t cs" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1043
proof -
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1044
  have "t \<in> set wq'" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1045
  proof(unfold wq'_def, rule someI2)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1046
    show "distinct rest \<and> set rest = set rest"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1047
          by (simp add: distinct_rest)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1048
  next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1049
    fix x
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1050
    assume "distinct x \<and> set x = set rest"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1051
    moreover have "t \<in> set rest"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1052
        using assms(1) cs_waiting_def waiting_eq wq_s_cs by auto 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1053
    ultimately show "t \<in> set x" by simp
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1054
  qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1055
  moreover have "t \<noteq> hd wq'"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1056
    using assms(2) taker_def by auto 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1057
  ultimately show ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1058
    by (unfold s_waiting_def, fold wq_def, unfold wq_es_cs, simp)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1059
qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1060
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1061
lemma waiting_esE:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1062
  assumes "waiting (e#s) t c" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1063
  obtains "c \<noteq> cs" "waiting s t c"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1064
     |    "c = cs" "t \<noteq> taker" "waiting s t cs" "t \<in> set rest'"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1065
proof(cases "c = cs")
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1066
  case False
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1067
  hence "wq (e#s) c = wq s c" using is_v by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1068
  with assms have "waiting s t c" using cs_waiting_def waiting_eq by auto 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1069
  from that(1)[OF False this] show ?thesis .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1070
next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1071
  case True
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1072
  from assms[unfolded s_waiting_def True, folded wq_def, unfolded wq_es_cs]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1073
  have "t \<noteq> hd wq'" "t \<in> set wq'" by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1074
  hence "t \<noteq> taker" by (simp add: taker_def) 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1075
  moreover hence "t \<noteq> th" using assms neq_t_th by blast 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1076
  moreover have "t \<in> set rest" by (simp add: `t \<in> set wq'` th'_in_inv) 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1077
  ultimately have "waiting s t cs"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1078
    by (metis cs_waiting_def list.distinct(2) list.sel(1) 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1079
                list.set_sel(2) rest_def waiting_eq wq_s_cs)  
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1080
  show ?thesis using that(2)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1081
  using True `t \<in> set wq'` `t \<noteq> taker` `waiting s t cs` eq_wq' by auto   
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1082
qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1083
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1084
lemma holding_esI1:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1085
  assumes "c = cs"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1086
  and "t = taker"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1087
  shows "holding (e#s) t c"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1088
  by (unfold assms, simp add: holding_taker)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1089
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1090
lemma holding_esE:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1091
  assumes "holding (e#s) t c" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1092
  obtains "c = cs" "t = taker"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1093
      | "c \<noteq> cs" "holding s t c"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1094
proof(cases "c = cs")
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1095
  case True
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1096
  from assms[unfolded True, unfolded s_holding_def, 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1097
             folded wq_def, unfolded wq_es_cs]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1098
  have "t = taker" by (simp add: taker_def) 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1099
  from that(1)[OF True this] show ?thesis .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1100
next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1101
  case False
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1102
  hence "wq (e#s) c = wq s c" using is_v by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1103
  from assms[unfolded s_holding_def, folded wq_def, 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1104
             unfolded this, unfolded wq_def, folded s_holding_def]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1105
  have "holding s t c"  .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1106
  from that(2)[OF False this] show ?thesis .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1107
qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1108
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1109
end 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1110
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1111
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1112
context valid_trace_v_e
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1113
begin
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1114
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1115
lemma nil_wq': "wq' = []" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1116
proof (unfold wq'_def, rule someI2)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1117
  show "distinct rest \<and> set rest = set rest"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1118
    by (simp add: distinct_rest) 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1119
next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1120
  fix x
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1121
  assume " distinct x \<and> set x = set rest" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1122
  thus "x = []" using rest_nil by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1123
qed 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1124
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1125
lemma no_taker: 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1126
  assumes "next_th s th cs taker"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1127
  shows "False"
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1128
proof -
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1129
  from assms[unfolded next_th_def]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1130
  obtain rest' where "wq s cs = th # rest'" "rest' \<noteq> []"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1131
    by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1132
  thus ?thesis using rest_def rest_nil by auto 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1133
qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1134
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1135
lemma waiting_set_eq:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1136
  "{(Th th', Cs cs) |th'. next_th s th cs th'} = {}"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1137
  using no_taker by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1138
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1139
lemma holding_set_eq:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1140
  "{(Cs cs, Th th') |th'.  next_th s th cs th'} = {}"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1141
  using no_taker by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1142
   
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1143
lemma no_holding:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1144
  assumes "holding (e#s) taker cs"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1145
  shows False
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1146
proof -
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1147
  from wq_es_cs[unfolded nil_wq']
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1148
  have " wq (e # s) cs = []" .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1149
  from assms[unfolded s_holding_def, folded wq_def, unfolded this]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1150
  show ?thesis by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1151
qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1152
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1153
lemma no_waiting:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1154
  assumes "waiting (e#s) t cs"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1155
  shows False
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1156
proof -
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1157
  from wq_es_cs[unfolded nil_wq']
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1158
  have " wq (e # s) cs = []" .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1159
  from assms[unfolded s_waiting_def, folded wq_def, unfolded this]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1160
  show ?thesis by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1161
qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1162
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1163
lemma waiting_esI2:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1164
  assumes "waiting s t c"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1165
  shows "waiting (e#s) t c"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1166
proof -
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1167
  have "c \<noteq> cs" using assms
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1168
    using cs_waiting_def rest_nil waiting_eq wq_s_cs by auto 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1169
  from waiting_esI1[OF assms this]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1170
  show ?thesis .
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1171
qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1172
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1173
lemma waiting_esE:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1174
  assumes "waiting (e#s) t c" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1175
  obtains "c \<noteq> cs" "waiting s t c"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1176
proof(cases "c = cs")
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1177
  case False
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1178
  hence "wq (e#s) c = wq s c" using is_v by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1179
  with assms have "waiting s t c" using cs_waiting_def waiting_eq by auto 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1180
  from that(1)[OF False this] show ?thesis .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1181
next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1182
  case True
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1183
  from no_waiting[OF assms[unfolded True]]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1184
  show ?thesis by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1185
qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1186
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1187
lemma holding_esE:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1188
  assumes "holding (e#s) t c" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1189
  obtains "c \<noteq> cs" "holding s t c"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1190
proof(cases "c = cs")
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1191
  case True
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1192
  from no_holding[OF assms[unfolded True]] 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1193
  show ?thesis by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1194
next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1195
  case False
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1196
  hence "wq (e#s) c = wq s c" using is_v by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1197
  from assms[unfolded s_holding_def, folded wq_def, 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1198
             unfolded this, unfolded wq_def, folded s_holding_def]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1199
  have "holding s t c"  .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1200
  from that[OF False this] show ?thesis .
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1201
qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1202
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1203
end 
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1204
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1205
lemma rel_eqI:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1206
  assumes "\<And> x y. (x,y) \<in> A \<Longrightarrow> (x,y) \<in> B"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1207
  and "\<And> x y. (x,y) \<in> B \<Longrightarrow> (x, y) \<in> A"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1208
  shows "A = B"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1209
  using assms by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1210
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1211
lemma in_RAG_E:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1212
  assumes "(n1, n2) \<in> RAG (s::state)"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1213
  obtains (waiting) th cs where "n1 = Th th" "n2 = Cs cs" "waiting s th cs"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1214
      | (holding) th cs where "n1 = Cs cs" "n2 = Th th" "holding s th cs"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1215
  using assms[unfolded s_RAG_def, folded waiting_eq holding_eq]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1216
  by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1217
  
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1218
context valid_trace_v
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1219
begin
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1220
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1221
lemma RAG_es:
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1222
  "RAG (e # s) =
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1223
   RAG s - {(Cs cs, Th th)} -
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1224
     {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1225
     {(Cs cs, Th th') |th'.  next_th s th cs th'}" (is "?L = ?R")
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1226
proof(rule rel_eqI)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1227
  fix n1 n2
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1228
  assume "(n1, n2) \<in> ?L"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1229
  thus "(n1, n2) \<in> ?R"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1230
  proof(cases rule:in_RAG_E)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1231
    case (waiting th' cs')
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1232
    show ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1233
    proof(cases "rest = []")
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1234
      case False
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1235
      interpret h_n: valid_trace_v_n s e th cs
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1236
        by (unfold_locales, insert False, simp)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1237
      from waiting(3)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1238
      show ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1239
      proof(cases rule:h_n.waiting_esE)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1240
        case 1
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1241
        with waiting(1,2)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1242
        show ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1243
        by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1244
             fold waiting_eq, auto)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1245
      next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1246
        case 2
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1247
        with waiting(1,2)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1248
        show ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1249
         by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1250
             fold waiting_eq, auto)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1251
      qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1252
    next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1253
      case True
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1254
      interpret h_e: valid_trace_v_e s e th cs
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1255
        by (unfold_locales, insert True, simp)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1256
      from waiting(3)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1257
      show ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1258
      proof(cases rule:h_e.waiting_esE)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1259
        case 1
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1260
        with waiting(1,2)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1261
        show ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1262
        by (unfold h_e.waiting_set_eq h_e.holding_set_eq s_RAG_def, 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1263
             fold waiting_eq, auto)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1264
      qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1265
    qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1266
  next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1267
    case (holding th' cs')
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1268
    show ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1269
    proof(cases "rest = []")
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1270
      case False
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1271
      interpret h_n: valid_trace_v_n s e th cs
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1272
        by (unfold_locales, insert False, simp)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1273
      from holding(3)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1274
      show ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1275
      proof(cases rule:h_n.holding_esE)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1276
        case 1
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1277
        with holding(1,2)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1278
        show ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1279
        by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1280
             fold waiting_eq, auto)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1281
      next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1282
        case 2
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1283
        with holding(1,2)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1284
        show ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1285
         by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1286
             fold holding_eq, auto)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1287
      qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1288
    next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1289
      case True
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1290
      interpret h_e: valid_trace_v_e s e th cs
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1291
        by (unfold_locales, insert True, simp)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1292
      from holding(3)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1293
      show ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1294
      proof(cases rule:h_e.holding_esE)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1295
        case 1
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1296
        with holding(1,2)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1297
        show ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1298
        by (unfold h_e.waiting_set_eq h_e.holding_set_eq s_RAG_def, 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1299
             fold holding_eq, auto)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1300
      qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1301
    qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1302
  qed
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1303
next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1304
  fix n1 n2
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1305
  assume h: "(n1, n2) \<in> ?R"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1306
  show "(n1, n2) \<in> ?L"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1307
  proof(cases "rest = []")
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1308
    case False
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1309
    interpret h_n: valid_trace_v_n s e th cs
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1310
        by (unfold_locales, insert False, simp)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1311
    from h[unfolded h_n.waiting_set_eq h_n.holding_set_eq]
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1312
    have "((n1, n2) \<in> RAG s \<and> (n1 \<noteq> Cs cs \<or> n2 \<noteq> Th th)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1313
                            \<and> (n1 \<noteq> Th h_n.taker \<or> n2 \<noteq> Cs cs)) \<or> 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1314
          (n2 = Th h_n.taker \<and> n1 = Cs cs)" 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1315
      by auto
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1316
   thus ?thesis
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1317
   proof
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1318
      assume "n2 = Th h_n.taker \<and> n1 = Cs cs"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1319
      with h_n.holding_taker
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1320
      show ?thesis 
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1321
        by (unfold s_RAG_def, fold holding_eq, auto)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1322
   next
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1323
    assume h: "(n1, n2) \<in> RAG s \<and>
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1324
        (n1 \<noteq> Cs cs \<or> n2 \<noteq> Th th) \<and> (n1 \<noteq> Th h_n.taker \<or> n2 \<noteq> Cs cs)"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1325
    hence "(n1, n2) \<in> RAG s" by simp
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1326
    thus ?thesis
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1327
    proof(cases rule:in_RAG_E)
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1328
      case (waiting th' cs')
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1329
      from h and this(1,2)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1330
      have "th' \<noteq> h_n.taker \<or> cs' \<noteq> cs" by auto
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1331
      hence "waiting (e#s) th' cs'" 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1332
      proof
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1333
        assume "cs' \<noteq> cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1334
        from waiting_esI1[OF waiting(3) this] 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1335
        show ?thesis .
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1336
      next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1337
        assume neq_th': "th' \<noteq> h_n.taker"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1338
        show ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1339
        proof(cases "cs' = cs")
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1340
          case False
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1341
          from waiting_esI1[OF waiting(3) this] 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1342
          show ?thesis .
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1343
        next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1344
          case True
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1345
          from h_n.waiting_esI2[OF waiting(3)[unfolded True] neq_th', folded True]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1346
          show ?thesis .
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1347
        qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1348
      qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1349
      thus ?thesis using waiting(1,2)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1350
        by (unfold s_RAG_def, fold waiting_eq, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1351
    next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1352
      case (holding th' cs')
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1353
      from h this(1,2)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1354
      have "cs' \<noteq> cs \<or> th' \<noteq> th" by auto
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1355
      hence "holding (e#s) th' cs'"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1356
      proof
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1357
        assume "cs' \<noteq> cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1358
        from holding_esI2[OF this holding(3)] 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1359
        show ?thesis .
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1360
      next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1361
        assume "th' \<noteq> th"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1362
        from holding_esI1[OF holding(3) this]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1363
        show ?thesis .
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1364
      qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1365
      thus ?thesis using holding(1,2)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1366
        by (unfold s_RAG_def, fold holding_eq, auto)
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1367
    qed
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1368
   qed
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1369
 next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1370
   case True
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1371
   interpret h_e: valid_trace_v_e s e th cs
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1372
        by (unfold_locales, insert True, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1373
   from h[unfolded h_e.waiting_set_eq h_e.holding_set_eq]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1374
   have h_s: "(n1, n2) \<in> RAG s" "(n1, n2) \<noteq> (Cs cs, Th th)" 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1375
      by auto
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1376
   from h_s(1)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1377
   show ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1378
   proof(cases rule:in_RAG_E)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1379
    case (waiting th' cs')
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1380
    from h_e.waiting_esI2[OF this(3)]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1381
    show ?thesis using waiting(1,2)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1382
      by (unfold s_RAG_def, fold waiting_eq, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1383
   next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1384
    case (holding th' cs')
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1385
    with h_s(2)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1386
    have "cs' \<noteq> cs \<or> th' \<noteq> th" by auto
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1387
    thus ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1388
    proof
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1389
      assume neq_cs: "cs' \<noteq> cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1390
      from holding_esI2[OF this holding(3)]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1391
      show ?thesis using holding(1,2)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1392
        by (unfold s_RAG_def, fold holding_eq, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1393
    next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1394
      assume "th' \<noteq> th"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1395
      from holding_esI1[OF holding(3) this]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1396
      show ?thesis using holding(1,2)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1397
        by (unfold s_RAG_def, fold holding_eq, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1398
    qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1399
   qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1400
 qed
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1401
qed
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1402
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1403
end
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1404
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1405
lemma step_RAG_v: 
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1406
assumes vt:
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1407
  "vt (V th cs#s)"
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1408
shows "
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1409
  RAG (V th cs # s) =
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1410
  RAG s - {(Cs cs, Th th)} -
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1411
  {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  1412
  {(Cs cs, Th th') |th'.  next_th s th cs th'}" (is "?L = ?R")
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1413
proof -
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1414
  interpret vt_v: valid_trace_v s "V th cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1415
    using assms step_back_vt by (unfold_locales, auto) 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1416
  show ?thesis using vt_v.RAG_es .
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1417
qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1418
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1419
lemma (in valid_trace_create)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1420
  th_not_in_threads: "th \<notin> threads s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1421
proof -
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1422
  from pip_e[unfolded is_create]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1423
  show ?thesis by (cases, simp)
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1424
qed
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1425
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1426
lemma (in valid_trace_create)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1427
  threads_es [simp]: "threads (e#s) = threads s \<union> {th}"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1428
  by (unfold is_create, simp)
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1429
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1430
lemma (in valid_trace_exit)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1431
  threads_es [simp]: "threads (e#s) = threads s - {th}"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1432
  by (unfold is_exit, simp)
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1433
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1434
lemma (in valid_trace_p)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1435
  threads_es [simp]: "threads (e#s) = threads s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1436
  by (unfold is_p, simp)
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1437
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1438
lemma (in valid_trace_v)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1439
  threads_es [simp]: "threads (e#s) = threads s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1440
  by (unfold is_v, simp)
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1441
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1442
lemma (in valid_trace_v)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1443
  th_not_in_rest[simp]: "th \<notin> set rest"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1444
proof
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1445
  assume otherwise: "th \<in> set rest"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1446
  have "distinct (wq s cs)" by (simp add: wq_distinct)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1447
  from this[unfolded wq_s_cs] and otherwise
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1448
  show False by auto
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1449
qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1450
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1451
lemma (in valid_trace_v)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1452
  set_wq_es_cs [simp]: "set (wq (e#s) cs) = set (wq s cs) - {th}"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1453
proof(unfold wq_es_cs wq'_def, rule someI2)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1454
  show "distinct rest \<and> set rest = set rest"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1455
    by (simp add: distinct_rest)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1456
next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1457
  fix x
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1458
  assume "distinct x \<and> set x = set rest"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1459
  thus "set x = set (wq s cs) - {th}" 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1460
      by (unfold wq_s_cs, simp)
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1461
qed
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1462
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1463
lemma (in valid_trace_exit)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1464
  th_not_in_wq: "th \<notin> set (wq s cs)"
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1465
proof -
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1466
  from pip_e[unfolded is_exit]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1467
  show ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1468
  by (cases, unfold holdents_def s_holding_def, fold wq_def, 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1469
             auto elim!:runing_wqE)
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1470
qed
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1471
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1472
lemma (in valid_trace) wq_threads: 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1473
  assumes "th \<in> set (wq s cs)"
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1474
  shows "th \<in> threads s"
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1475
  using assms
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1476
proof(induct rule:ind)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1477
  case (Nil)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1478
  thus ?case by (auto simp:wq_def)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1479
next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1480
  case (Cons s e)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1481
  interpret vt_e: valid_trace_e s e using Cons by simp
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1482
  show ?case
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1483
  proof(cases e)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1484
    case (Create th' prio')
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1485
    interpret vt: valid_trace_create s e th' prio'
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1486
      using Create by (unfold_locales, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1487
    show ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1488
      using Cons.hyps(2) Cons.prems by auto
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1489
  next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1490
    case (Exit th')
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1491
    interpret vt: valid_trace_exit s e th'
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1492
      using Exit by (unfold_locales, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1493
    show ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1494
      using Cons.hyps(2) Cons.prems vt.th_not_in_wq by auto 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1495
  next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1496
    case (P th' cs')
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1497
    interpret vt: valid_trace_p s e th' cs'
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1498
      using P by (unfold_locales, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1499
    show ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1500
      using Cons.hyps(2) Cons.prems readys_threads 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1501
        runing_ready vt.is_p vt.runing_th_s vt_e.wq_in_inv 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1502
        by fastforce 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1503
  next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1504
    case (V th' cs')
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1505
    interpret vt: valid_trace_v s e th' cs'
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1506
      using V by (unfold_locales, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1507
    show ?thesis using Cons
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1508
      using vt.is_v vt.threads_es vt_e.wq_in_inv by blast
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1509
  next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1510
    case (Set th' prio)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1511
    interpret vt: valid_trace_set s e th' prio
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1512
      using Set by (unfold_locales, simp)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1513
    show ?thesis using Cons.hyps(2) Cons.prems vt.is_set 
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1514
        by (auto simp:wq_def Let_def)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1515
  qed
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1516
qed 
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1517
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1518
context valid_trace
57
f1b39d77db00 Added generic theory "RTree.thy"
xingyuan zhang <xingyuanzhang@126.com>
parents:
diff changeset
  1519
begin
f1b39d77db00 Added generic theory "RTree.thy"
xingyuan zhang <xingyuanzhang@126.com>
parents:
diff changeset
  1520
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1521
lemma  dm_RAG_threads:
73
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1522
  assumes in_dom: "(Th th) \<in> Domain (RAG s)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1523
  shows "th \<in> threads s"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1524
proof -
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1525
  from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1526
  moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1527
  ultimately have "(Th th, Cs cs) \<in> RAG s" by simp
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1528
  hence "th \<in> set (wq s cs)"
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1529
    by (unfold s_RAG_def, auto simp:cs_waiting_def)
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1530
  from wq_threads [OF this] show ?thesis .
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1531
qed
b0054fb0d1ce Moment.thy further improved.
zhangx
parents: 65
diff changeset
  1532
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1533
lemma rg_RAG_threads: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1534
  assumes "(Th th) \<in> Range (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1535
  shows "th \<in> threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1536
  using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1537
  by (unfold s_RAG_def cs_waiting_def cs_holding_def, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1538
       auto intro:wq_threads)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1539
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  1540
lemma RAG_threads:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  1541
  assumes "(Th th) \<in> Field (RAG s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  1542
  shows "th \<in> threads s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  1543
  using assms
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  1544
  by (metis Field_def UnE dm_RAG_threads rg_RAG_threads)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  1545
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1546
end
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1547
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1548
lemma (in valid_trace_v)
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  1549
  preced_es [simp]: "preced th (e#s) = preced th s"
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1550
  by (unfold is_v preced_def, simp)
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 57
diff changeset
  1551
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1552
lemma the_preced_v[simp]: "the_preced (V th cs#s) = the_preced s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1553
proof
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1554
  fix th'
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1555
  show "the_preced (V th cs # s) th' = the_preced s th'"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1556
    by (unfold the_preced_def preced_def, simp)
61
f8194fd6214f CpsG.thy has been cleaned up.
zhangx
parents: 60
diff changeset
  1557
qed
f8194fd6214f CpsG.thy has been cleaned up.
zhangx
parents: 60
diff changeset
  1558
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1559
lemma (in valid_trace_v)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1560
  the_preced_es: "the_preced (e#s) = the_preced s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1561
  by (unfold is_v preced_def, simp)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1562
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1563
context valid_trace_p
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 57
diff changeset
  1564
begin
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 57
diff changeset
  1565
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1566
lemma not_holding_s_th_cs: "\<not> holding s th cs"
62
031d2ae9c9b8 In the middle of retrofiting ExtGG.thy.
zhangx
parents: 61
diff changeset
  1567
proof
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1568
  assume otherwise: "holding s th cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1569
  from pip_e[unfolded is_p]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1570
  show False
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1571
  proof(cases)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1572
    case (thread_P)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1573
    moreover have "(Cs cs, Th th) \<in> RAG s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1574
      using otherwise cs_holding_def 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1575
            holding_eq th_not_in_wq by auto
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1576
    ultimately show ?thesis by auto
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 57
diff changeset
  1577
  qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 57
diff changeset
  1578
qed
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 57
diff changeset
  1579
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1580
lemma waiting_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1581
  assumes "waiting s th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1582
  shows "waiting (e#s) th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1583
  using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1584
  by (metis cs_waiting_def hd_append2 list.sel(1) list.set_intros(2) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1585
      rotate1.simps(2) self_append_conv2 set_rotate1 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1586
        th_not_in_wq waiting_eq wq_es_cs wq_neq_simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1587
  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1588
lemma holding_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1589
  assumes "holding s th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1590
  shows "holding (e#s) th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1591
proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1592
  case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1593
  hence "wq (e#s) cs' = wq s cs'" by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1594
  with assms show ?thesis using cs_holding_def holding_eq by auto 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1595
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1596
  case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1597
  from assms[unfolded s_holding_def, folded wq_def]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1598
  obtain rest where eq_wq: "wq s cs' = th'#rest"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1599
    by (metis empty_iff list.collapse list.set(1)) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1600
  hence "wq (e#s) cs' = th'#(rest@[th])"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1601
    by (simp add: True wq_es_cs) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1602
  thus ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1603
    by (simp add: cs_holding_def holding_eq) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1604
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1605
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 57
diff changeset
  1606
end
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 57
diff changeset
  1607
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1608
locale valid_trace_p_h = valid_trace_p +
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1609
  assumes we: "wq s cs = []"
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 57
diff changeset
  1610
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1611
locale valid_trace_p_w = valid_trace_p +
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1612
  assumes wne: "wq s cs \<noteq> []"
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1613
begin
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1614
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1615
definition "holder = hd (wq s cs)"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1616
definition "waiters = tl (wq s cs)"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1617
definition "waiters' = waiters @ [th]"
57
f1b39d77db00 Added generic theory "RTree.thy"
xingyuan zhang <xingyuanzhang@126.com>
parents:
diff changeset
  1618
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1619
lemma wq_s_cs: "wq s cs = holder#waiters"
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1620
    by (simp add: holder_def waiters_def wne)
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1621
    
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1622
lemma wq_es_cs': "wq (e#s) cs = holder#waiters@[th]"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1623
  by (simp add: wq_es_cs wq_s_cs)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1624
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1625
lemma waiting_es_th_cs: "waiting (e#s) th cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1626
  using cs_waiting_def th_not_in_wq waiting_eq wq_es_cs' wq_s_cs by auto
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1627
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1628
lemma RAG_edge: "(Th th, Cs cs) \<in> RAG (e#s)"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1629
   by (unfold s_RAG_def, fold waiting_eq, insert waiting_es_th_cs, auto)
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1630
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1631
lemma holding_esE:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1632
  assumes "holding (e#s) th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1633
  obtains "holding s th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1634
  using assms 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1635
proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1636
  case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1637
  hence "wq (e#s) cs' = wq s cs'" by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1638
  with assms show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1639
    using cs_holding_def holding_eq that by auto 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1640
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1641
  case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1642
  with assms show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1643
  by (metis cs_holding_def holding_eq list.sel(1) list.set_intros(1) that 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1644
        wq_es_cs' wq_s_cs) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1645
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1646
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1647
lemma waiting_esE:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1648
  assumes "waiting (e#s) th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1649
  obtains "th' \<noteq> th" "waiting s th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1650
     |  "th' = th" "cs' = cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1651
proof(cases "waiting s th' cs'")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1652
  case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1653
  have "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1654
  proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1655
    assume otherwise: "th' = th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1656
    from True[unfolded this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1657
    show False by (simp add: th_not_waiting) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1658
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1659
  from that(1)[OF this True] show ?thesis .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1660
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1661
  case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1662
  hence "th' = th \<and> cs' = cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1663
      by (metis assms cs_waiting_def holder_def list.sel(1) rotate1.simps(2) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1664
        set_ConsD set_rotate1 waiting_eq wq_es_cs wq_es_cs' wq_neq_simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1665
  with that(2) show ?thesis by metis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1666
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1667
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1668
lemma RAG_es: "RAG (e # s) =  RAG s \<union> {(Th th, Cs cs)}" (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1669
proof(rule rel_eqI)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1670
  fix n1 n2
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1671
  assume "(n1, n2) \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1672
  thus "(n1, n2) \<in> ?R" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1673
  proof(cases rule:in_RAG_E)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1674
    case (waiting th' cs')
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1675
    from this(3)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1676
    show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1677
    proof(cases rule:waiting_esE)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1678
      case 1
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1679
      thus ?thesis using waiting(1,2)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1680
        by (unfold s_RAG_def, fold waiting_eq, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1681
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1682
      case 2
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1683
      thus ?thesis using waiting(1,2) by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1684
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1685
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1686
    case (holding th' cs')
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1687
    from this(3)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1688
    show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1689
    proof(cases rule:holding_esE)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1690
      case 1
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1691
      with holding(1,2)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1692
      show ?thesis by (unfold s_RAG_def, fold holding_eq, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1693
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1694
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1695
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1696
  fix n1 n2
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1697
  assume "(n1, n2) \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1698
  hence "(n1, n2) \<in> RAG s \<or> (n1 = Th th \<and> n2 = Cs cs)" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1699
  thus "(n1, n2) \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1700
  proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1701
    assume "(n1, n2) \<in> RAG s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1702
    thus ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1703
    proof(cases rule:in_RAG_E)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1704
      case (waiting th' cs')
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1705
      from waiting_kept[OF this(3)]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1706
      show ?thesis using waiting(1,2)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1707
         by (unfold s_RAG_def, fold waiting_eq, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1708
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1709
      case (holding th' cs')
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1710
      from holding_kept[OF this(3)]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1711
      show ?thesis using holding(1,2)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1712
         by (unfold s_RAG_def, fold holding_eq, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1713
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1714
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1715
    assume "n1 = Th th \<and> n2 = Cs cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1716
    thus ?thesis using RAG_edge by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1717
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1718
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1719
63
b620a2a0806a ExtGG.thy finished, but more comments are needed.
zhangx
parents: 62
diff changeset
  1720
end
58
ad57323fd4d6 Extended RTree.thy
zhangx
parents: 57
diff changeset
  1721
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1722
context valid_trace_p_h
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1723
begin
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1724
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1725
lemma wq_es_cs': "wq (e#s) cs = [th]"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1726
  using wq_es_cs[unfolded we] by simp
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1727
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1728
lemma holding_es_th_cs: 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1729
  shows "holding (e#s) th cs"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1730
proof -
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1731
  from wq_es_cs'
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1732
  have "th \<in> set (wq (e#s) cs)" "th = hd (wq (e#s) cs)" by auto
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1733
  thus ?thesis using cs_holding_def holding_eq by blast 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1734
qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1735
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1736
lemma RAG_edge: "(Cs cs, Th th) \<in> RAG (e#s)"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1737
  by (unfold s_RAG_def, fold holding_eq, insert holding_es_th_cs, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1738
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1739
lemma waiting_esE:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1740
  assumes "waiting (e#s) th' cs'"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1741
  obtains "waiting s th' cs'"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1742
  using assms
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1743
  by (metis cs_waiting_def event.distinct(15) is_p list.sel(1) 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1744
        set_ConsD waiting_eq we wq_es_cs' wq_neq_simp wq_out_inv)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1745
  
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1746
lemma holding_esE:
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1747
  assumes "holding (e#s) th' cs'"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1748
  obtains "cs' \<noteq> cs" "holding s th' cs'"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1749
    | "cs' = cs" "th' = th"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1750
proof(cases "cs' = cs")
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1751
  case True
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1752
  from held_unique[OF holding_es_th_cs assms[unfolded True]]
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1753
  have "th' = th" by simp
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1754
  from that(2)[OF True this] show ?thesis .
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1755
next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1756
  case False
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1757
  have "holding s th' cs'" using assms
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1758
    using False cs_holding_def holding_eq by auto
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1759
  from that(1)[OF False this] show ?thesis .
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1760
qed
57
f1b39d77db00 Added generic theory "RTree.thy"
xingyuan zhang <xingyuanzhang@126.com>
parents:
diff changeset
  1761
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1762
lemma RAG_es: "RAG (e # s) =  RAG s \<union> {(Cs cs, Th th)}" (is "?L = ?R")
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1763
proof(rule rel_eqI)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1764
  fix n1 n2
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1765
  assume "(n1, n2) \<in> ?L"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1766
  thus "(n1, n2) \<in> ?R" 
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1767
  proof(cases rule:in_RAG_E)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1768
    case (waiting th' cs')
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1769
    from this(3)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1770
    show ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1771
    proof(cases rule:waiting_esE)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1772
      case 1
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1773
      thus ?thesis using waiting(1,2)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1774
        by (unfold s_RAG_def, fold waiting_eq, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1775
    qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1776
  next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1777
    case (holding th' cs')
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1778
    from this(3)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1779
    show ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1780
    proof(cases rule:holding_esE)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1781
      case 1
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1782
      with holding(1,2)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1783
      show ?thesis by (unfold s_RAG_def, fold holding_eq, auto)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1784
    next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1785
      case 2
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1786
      with holding(1,2) show ?thesis by auto
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1787
    qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1788
  qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1789
next
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1790
  fix n1 n2
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1791
  assume "(n1, n2) \<in> ?R"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1792
  hence "(n1, n2) \<in> RAG s \<or> (n1 = Cs cs \<and> n2 = Th th)" by auto
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1793
  thus "(n1, n2) \<in> ?L"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1794
  proof
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1795
    assume "(n1, n2) \<in> RAG s"
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1796
    thus ?thesis
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1797
    proof(cases rule:in_RAG_E)
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  1798
      case (waiting th' cs')
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1799
      from waiting_kept[OF this(3)]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1800
      show ?thesis using waiting(1,2)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1801
         by (unfold s_RAG_def, fold waiting_eq, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1802
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1803
      case (holding th' cs')
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1804
      from holding_kept[OF this(3)]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1805
      show ?thesis using holding(1,2)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1806
         by (unfold s_RAG_def, fold holding_eq, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1807
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1808
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1809
    assume "n1 = Cs cs \<and> n2 = Th th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1810
    with holding_es_th_cs
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1811
    show ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1812
      by (unfold s_RAG_def, fold holding_eq, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1813
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1814
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1815
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1816
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1817
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1818
context valid_trace_p
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1819
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1820
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1821
lemma RAG_es': "RAG (e # s) =  (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)}
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1822
                                                  else RAG s \<union> {(Th th, Cs cs)})"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1823
proof(cases "wq s cs = []")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1824
  case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1825
  interpret vt_p: valid_trace_p_h using True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1826
    by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1827
  show ?thesis by (simp add: vt_p.RAG_es vt_p.we) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1828
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1829
  case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1830
  interpret vt_p: valid_trace_p_w using False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1831
    by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1832
  show ?thesis by (simp add: vt_p.RAG_es vt_p.wne) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1833
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1834
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1835
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1836
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1837
lemma (in valid_trace_v_n) finite_waiting_set:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1838
  "finite {(Th th', Cs cs) |th'. next_th s th cs th'}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1839
    by (simp add: waiting_set_eq)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1840
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1841
lemma (in valid_trace_v_n) finite_holding_set:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1842
  "finite {(Cs cs, Th th') |th'. next_th s th cs th'}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1843
    by (simp add: holding_set_eq)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1844
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1845
lemma (in valid_trace_v_e) finite_waiting_set:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1846
  "finite {(Th th', Cs cs) |th'. next_th s th cs th'}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1847
    by (simp add: waiting_set_eq)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1848
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1849
lemma (in valid_trace_v_e) finite_holding_set:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1850
  "finite {(Cs cs, Th th') |th'. next_th s th cs th'}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1851
    by (simp add: holding_set_eq)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1852
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1853
context valid_trace_v
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1854
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1855
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1856
lemma 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1857
  finite_RAG_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1858
  assumes "finite (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1859
  shows "finite (RAG (e#s))"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1860
proof(cases "rest = []")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1861
  case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1862
  interpret vt: valid_trace_v_e using True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1863
    by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1864
  show ?thesis using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1865
    by  (unfold RAG_es vt.waiting_set_eq vt.holding_set_eq, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1866
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1867
  case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1868
  interpret vt: valid_trace_v_n using False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1869
    by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1870
  show ?thesis using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1871
    by  (unfold RAG_es vt.waiting_set_eq vt.holding_set_eq, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1872
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1873
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1874
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1875
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1876
context valid_trace_v_e
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1877
begin 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1878
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1879
lemma 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1880
  acylic_RAG_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1881
  assumes "acyclic (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1882
  shows "acyclic (RAG (e#s))"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1883
proof(rule acyclic_subset[OF assms])
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1884
  show "RAG (e # s) \<subseteq> RAG s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1885
      by (unfold RAG_es waiting_set_eq holding_set_eq, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1886
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1887
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1888
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1889
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1890
context valid_trace_v_n
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1891
begin 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1892
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1893
lemma waiting_taker: "waiting s taker cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1894
  apply (unfold s_waiting_def, fold wq_def, unfold wq_s_cs taker_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1895
  using eq_wq' th'_in_inv wq'_def by fastforce
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1896
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1897
lemma 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1898
  acylic_RAG_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1899
  assumes "acyclic (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1900
  shows "acyclic (RAG (e#s))"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1901
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1902
  have "acyclic ((RAG s - {(Cs cs, Th th)} - {(Th taker, Cs cs)}) \<union> 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1903
                 {(Cs cs, Th taker)})" (is "acyclic (?A \<union> _)")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1904
  proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1905
    from assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1906
    have "acyclic ?A"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1907
       by (rule acyclic_subset, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1908
    moreover have "(Th taker, Cs cs) \<notin> ?A^*"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1909
    proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1910
      assume otherwise: "(Th taker, Cs cs) \<in> ?A^*"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1911
      hence "(Th taker, Cs cs) \<in> ?A^+"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1912
        by (unfold rtrancl_eq_or_trancl, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1913
      from tranclD[OF this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1914
      obtain cs' where h: "(Th taker, Cs cs') \<in> ?A" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1915
                          "(Th taker, Cs cs') \<in> RAG s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1916
        by (unfold s_RAG_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1917
      from this(2) have "waiting s taker cs'" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1918
        by (unfold s_RAG_def, fold waiting_eq, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1919
      from waiting_unique[OF this waiting_taker]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1920
      have "cs' = cs" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1921
      from h(1)[unfolded this] show False by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1922
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1923
    ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1924
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1925
  thus ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1926
    by (unfold RAG_es waiting_set_eq holding_set_eq, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1927
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1928
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1929
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1930
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1931
context valid_trace_p_h
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1932
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1933
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1934
lemma 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1935
  acylic_RAG_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1936
  assumes "acyclic (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1937
  shows "acyclic (RAG (e#s))"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1938
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1939
  have "acyclic (RAG s \<union> {(Cs cs, Th th)})" (is "acyclic (?A \<union> _)") 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1940
  proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1941
    from assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1942
    have "acyclic ?A"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1943
       by (rule acyclic_subset, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1944
    moreover have "(Th th, Cs cs) \<notin> ?A^*"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1945
    proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1946
      assume otherwise: "(Th th, Cs cs) \<in> ?A^*"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1947
      hence "(Th th, Cs cs) \<in> ?A^+"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1948
        by (unfold rtrancl_eq_or_trancl, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1949
      from tranclD[OF this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1950
      obtain cs' where h: "(Th th, Cs cs') \<in> RAG s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1951
        by (unfold s_RAG_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1952
      hence "waiting s th cs'" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1953
        by (unfold s_RAG_def, fold waiting_eq, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1954
      with th_not_waiting show False by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1955
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1956
    ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1957
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1958
  thus ?thesis by (unfold RAG_es, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1959
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1960
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1961
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1962
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1963
context valid_trace_p_w
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1964
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1965
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1966
lemma 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1967
  acylic_RAG_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1968
  assumes "acyclic (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1969
  shows "acyclic (RAG (e#s))"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1970
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1971
  have "acyclic (RAG s \<union> {(Th th, Cs cs)})" (is "acyclic (?A \<union> _)") 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1972
  proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1973
    from assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1974
    have "acyclic ?A"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1975
       by (rule acyclic_subset, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1976
    moreover have "(Cs cs, Th th) \<notin> ?A^*"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1977
    proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1978
      assume otherwise: "(Cs cs, Th th) \<in> ?A^*"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1979
      from pip_e[unfolded is_p]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1980
      show False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1981
      proof(cases)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1982
        case (thread_P)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1983
        moreover from otherwise have "(Cs cs, Th th) \<in> ?A^+"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1984
            by (unfold rtrancl_eq_or_trancl, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1985
        ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1986
      qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1987
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1988
    ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1989
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1990
  thus ?thesis by (unfold RAG_es, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1991
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1992
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1993
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1994
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1995
context valid_trace
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1996
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1997
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1998
lemma finite_RAG:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  1999
  shows "finite (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2000
proof(induct rule:ind)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2001
  case Nil
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2002
  show ?case 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2003
  by (auto simp: s_RAG_def cs_waiting_def 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2004
                   cs_holding_def wq_def acyclic_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2005
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2006
  case (Cons s e)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2007
  interpret vt_e: valid_trace_e s e using Cons by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2008
  show ?case
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2009
  proof(cases e)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2010
    case (Create th prio)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2011
    interpret vt: valid_trace_create s e th prio using Create
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2012
      by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2013
    show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2014
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2015
    case (Exit th)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2016
    interpret vt: valid_trace_exit s e th using Exit
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2017
      by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2018
    show ?thesis using Cons by (simp add: vt.RAG_unchanged)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2019
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2020
    case (P th cs)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2021
    interpret vt: valid_trace_p s e th cs using P
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2022
      by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2023
    show ?thesis using Cons using vt.RAG_es' by auto 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2024
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2025
    case (V th cs)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2026
    interpret vt: valid_trace_v s e th cs using V
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2027
      by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2028
    show ?thesis using Cons by (simp add: vt.finite_RAG_kept) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2029
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2030
    case (Set th prio)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2031
    interpret vt: valid_trace_set s e th prio using Set
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2032
      by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2033
    show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2034
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2035
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2036
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2037
lemma acyclic_RAG:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2038
  shows "acyclic (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2039
proof(induct rule:ind)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2040
  case Nil
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2041
  show ?case 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2042
  by (auto simp: s_RAG_def cs_waiting_def 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2043
                   cs_holding_def wq_def acyclic_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2044
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2045
  case (Cons s e)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2046
  interpret vt_e: valid_trace_e s e using Cons by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2047
  show ?case
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2048
  proof(cases e)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2049
    case (Create th prio)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2050
    interpret vt: valid_trace_create s e th prio using Create
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2051
      by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2052
    show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2053
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2054
    case (Exit th)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2055
    interpret vt: valid_trace_exit s e th using Exit
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2056
      by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2057
    show ?thesis using Cons by (simp add: vt.RAG_unchanged)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2058
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2059
    case (P th cs)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2060
    interpret vt: valid_trace_p s e th cs using P
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2061
      by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2062
    show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2063
    proof(cases "wq s cs = []")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2064
      case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2065
      then interpret vt_h: valid_trace_p_h s e th cs
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2066
        by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2067
      show ?thesis using Cons by (simp add: vt_h.acylic_RAG_kept) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2068
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2069
      case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2070
      then interpret vt_w: valid_trace_p_w s e th cs
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2071
        by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2072
      show ?thesis using Cons by (simp add: vt_w.acylic_RAG_kept) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2073
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2074
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2075
    case (V th cs)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2076
    interpret vt: valid_trace_v s e th cs using V
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2077
      by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2078
    show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2079
    proof(cases "vt.rest = []")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2080
      case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2081
      then interpret vt_e: valid_trace_v_e s e th cs
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2082
        by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2083
      show ?thesis by (simp add: Cons.hyps(2) vt_e.acylic_RAG_kept) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2084
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2085
      case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2086
      then interpret vt_n: valid_trace_v_n s e th cs
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2087
        by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2088
      show ?thesis by (simp add: Cons.hyps(2) vt_n.acylic_RAG_kept) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2089
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2090
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2091
    case (Set th prio)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2092
    interpret vt: valid_trace_set s e th prio using Set
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2093
      by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2094
    show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2095
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2096
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2097
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2098
lemma wf_RAG: "wf (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2099
proof(rule finite_acyclic_wf)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2100
  from finite_RAG show "finite (RAG s)" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2101
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2102
  from acyclic_RAG show "acyclic (RAG s)" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2103
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2104
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2105
lemma sgv_wRAG: "single_valued (wRAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2106
  using waiting_unique
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2107
  by (unfold single_valued_def wRAG_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2108
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2109
lemma sgv_hRAG: "single_valued (hRAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2110
  using held_unique 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2111
  by (unfold single_valued_def hRAG_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2112
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2113
lemma sgv_tRAG: "single_valued (tRAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2114
  by (unfold tRAG_def, rule single_valued_relcomp, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2115
              insert sgv_wRAG sgv_hRAG, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2116
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2117
lemma acyclic_tRAG: "acyclic (tRAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2118
proof(unfold tRAG_def, rule acyclic_compose)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2119
  show "acyclic (RAG s)" using acyclic_RAG .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2120
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2121
  show "wRAG s \<subseteq> RAG s" unfolding RAG_split by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2122
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2123
  show "hRAG s \<subseteq> RAG s" unfolding RAG_split by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2124
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2125
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2126
lemma unique_RAG: "\<lbrakk>(n, n1) \<in> RAG s; (n, n2) \<in> RAG s\<rbrakk> \<Longrightarrow> n1 = n2"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2127
  apply(unfold s_RAG_def, auto, fold waiting_eq holding_eq)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2128
  by(auto elim:waiting_unique held_unique)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2129
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2130
lemma sgv_RAG: "single_valued (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2131
  using unique_RAG by (auto simp:single_valued_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2132
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2133
lemma rtree_RAG: "rtree (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2134
  using sgv_RAG acyclic_RAG
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2135
  by (unfold rtree_def rtree_axioms_def sgv_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2136
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2137
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2138
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2139
sublocale valid_trace < rtree_RAG: rtree "RAG s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2140
proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2141
  show "single_valued (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2142
  apply (intro_locales)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2143
    by (unfold single_valued_def, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2144
        auto intro:unique_RAG)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2145
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2146
  show "acyclic (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2147
     by (rule acyclic_RAG)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2148
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2149
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2150
sublocale valid_trace < rtree_s: rtree "tRAG s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2151
proof(unfold_locales)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2152
  from sgv_tRAG show "single_valued (tRAG s)" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2153
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2154
  from acyclic_tRAG show "acyclic (tRAG s)" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2155
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2156
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2157
sublocale valid_trace < fsbtRAGs : fsubtree "RAG s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2158
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2159
  show "fsubtree (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2160
  proof(intro_locales)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2161
    show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG] .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2162
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2163
    show "fsubtree_axioms (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2164
    proof(unfold fsubtree_axioms_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2165
      from wf_RAG show "wf (RAG s)" .
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  2166
    qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  2167
  qed
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  2168
qed
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  2169
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2170
lemma tRAG_alt_def: 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2171
  "tRAG s = {(Th th1, Th th2) | th1 th2. 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2172
                  \<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2173
 by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2174
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2175
sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2176
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2177
  have "fsubtree (tRAG s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2178
  proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2179
    have "fbranch (tRAG s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2180
    proof(unfold tRAG_def, rule fbranch_compose)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2181
        show "fbranch (wRAG s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2182
        proof(rule finite_fbranchI)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2183
           from finite_RAG show "finite (wRAG s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2184
           by (unfold RAG_split, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2185
        qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2186
    next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2187
        show "fbranch (hRAG s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2188
        proof(rule finite_fbranchI)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2189
           from finite_RAG 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2190
           show "finite (hRAG s)" by (unfold RAG_split, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2191
        qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2192
    qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2193
    moreover have "wf (tRAG s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2194
    proof(rule wf_subset)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2195
      show "wf (RAG s O RAG s)" using wf_RAG
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2196
        by (fold wf_comp_self, simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2197
    next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2198
      show "tRAG s \<subseteq> (RAG s O RAG s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2199
        by (unfold tRAG_alt_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2200
    qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2201
    ultimately show ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2202
      by (unfold fsubtree_def fsubtree_axioms_def,auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2203
  qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2204
  from this[folded tRAG_def] show "fsubtree (tRAG s)" .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2205
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2206
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2207
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2208
context valid_trace
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2209
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2210
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2211
lemma finite_subtree_threads:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2212
    "finite {th'. Th th' \<in> subtree (RAG s) (Th th)}" (is "finite ?A")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2213
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2214
  have "?A = the_thread ` {Th th' | th' . Th th' \<in> subtree (RAG s) (Th th)}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2215
        by (auto, insert image_iff, fastforce)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2216
  moreover have "finite {Th th' | th' . Th th' \<in> subtree (RAG s) (Th th)}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2217
        (is "finite ?B")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2218
  proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2219
     have "?B = (subtree (RAG s) (Th th)) \<inter> {Th th' | th'. True}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2220
      by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2221
     moreover have "... \<subseteq> (subtree (RAG s) (Th th))" by auto
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2222
     moreover have "finite ..." by (simp add: fsbtRAGs.finite_subtree) 
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2223
     ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2224
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2225
  ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2226
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2227
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2228
lemma le_cp:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2229
  shows "preced th s \<le> cp s th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2230
  proof(unfold cp_alt_def, rule Max_ge)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2231
    show "finite (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2232
      by (simp add: finite_subtree_threads)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2233
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2234
    show "preced th s \<in> the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2235
      by (simp add: subtree_def the_preced_def)   
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2236
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2237
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2238
lemma cp_le:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2239
  assumes th_in: "th \<in> threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2240
  shows "cp s th \<le> Max (the_preced s ` threads s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2241
proof(unfold cp_alt_def, rule Max_f_mono)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2242
  show "finite (threads s)" by (simp add: finite_threads) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2243
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2244
  show " {th'. Th th' \<in> subtree (RAG s) (Th th)} \<noteq> {}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2245
    using subtree_def by fastforce
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2246
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2247
  show "{th'. Th th' \<in> subtree (RAG s) (Th th)} \<subseteq> threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2248
    using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2249
    by (smt Domain.DomainI dm_RAG_threads mem_Collect_eq 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2250
        node.inject(1) rtranclD subsetI subtree_def trancl_domain) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2251
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2252
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2253
lemma max_cp_eq: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2254
  shows "Max ((cp s) ` threads s) = Max (the_preced s ` threads s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2255
  (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2256
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2257
  have "?L \<le> ?R" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2258
  proof(cases "threads s = {}")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2259
    case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2260
    show ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2261
      by (rule Max.boundedI, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2262
          insert cp_le, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2263
          auto simp:finite_threads False)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2264
  qed auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2265
  moreover have "?R \<le> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2266
    by (rule Max_fg_mono, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2267
        simp add: finite_threads,
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2268
        simp add: le_cp the_preced_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2269
  ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2270
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2271
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2272
lemma wf_RAG_converse: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2273
  shows "wf ((RAG s)^-1)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2274
proof(rule finite_acyclic_wf_converse)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2275
  from finite_RAG 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2276
  show "finite (RAG s)" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2277
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2278
  from acyclic_RAG
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2279
  show "acyclic (RAG s)" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2280
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2281
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2282
lemma chain_building:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2283
  assumes "node \<in> Domain (RAG s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2284
  obtains th' where "th' \<in> readys s" "(node, Th th') \<in> (RAG s)^+"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2285
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2286
  from assms have "node \<in> Range ((RAG s)^-1)" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2287
  from wf_base[OF wf_RAG_converse this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2288
  obtain b where h_b: "(b, node) \<in> ((RAG s)\<inverse>)\<^sup>+" "\<forall>c. (c, b) \<notin> (RAG s)\<inverse>" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2289
  obtain th' where eq_b: "b = Th th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2290
  proof(cases b)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2291
    case (Cs cs)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2292
    from h_b(1)[unfolded trancl_converse] 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2293
    have "(node, b) \<in> ((RAG s)\<^sup>+)" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2294
    from tranclE[OF this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2295
    obtain n where "(n, b) \<in> RAG s" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2296
    from this[unfolded Cs]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2297
    obtain th1 where "waiting s th1 cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2298
      by (unfold s_RAG_def, fold waiting_eq, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2299
    from waiting_holding[OF this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2300
    obtain th2 where "holding s th2 cs" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2301
    hence "(Cs cs, Th th2) \<in> RAG s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2302
      by (unfold s_RAG_def, fold holding_eq, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2303
    with h_b(2)[unfolded Cs, rule_format]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2304
    have False by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2305
    thus ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2306
  qed auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2307
  have "th' \<in> readys s" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2308
  proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2309
    from h_b(2)[unfolded eq_b]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2310
    have "\<forall>cs. \<not> waiting s th' cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2311
      by (unfold s_RAG_def, fold waiting_eq, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2312
    moreover have "th' \<in> threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2313
    proof(rule rg_RAG_threads)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2314
      from tranclD[OF h_b(1), unfolded eq_b]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2315
      obtain z where "(z, Th th') \<in> (RAG s)" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2316
      thus "Th th' \<in> Range (RAG s)" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2317
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2318
    ultimately show ?thesis by (auto simp:readys_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2319
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2320
  moreover have "(node, Th th') \<in> (RAG s)^+" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2321
    using h_b(1)[unfolded trancl_converse] eq_b by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2322
  ultimately show ?thesis using that by metis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2323
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2324
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2325
text {* \noindent
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2326
  The following is just an instance of @{text "chain_building"}.
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2327
*}
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2328
lemma th_chain_to_ready:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2329
  assumes th_in: "th \<in> threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2330
  shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (RAG s)^+)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2331
proof(cases "th \<in> readys s")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2332
  case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2333
  thus ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2334
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2335
  case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2336
  from False and th_in have "Th th \<in> Domain (RAG s)" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2337
    by (auto simp:readys_def s_waiting_def s_RAG_def wq_def cs_waiting_def Domain_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2338
  from chain_building [rule_format, OF this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2339
  show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2340
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2341
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2342
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2343
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2344
lemma count_rec1 [simp]: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2345
  assumes "Q e"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2346
  shows "count Q (e#es) = Suc (count Q es)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2347
  using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2348
  by (unfold count_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2349
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2350
lemma count_rec2 [simp]: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2351
  assumes "\<not>Q e"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2352
  shows "count Q (e#es) = (count Q es)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2353
  using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2354
  by (unfold count_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2355
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2356
lemma count_rec3 [simp]: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2357
  shows "count Q [] =  0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2358
  by (unfold count_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2359
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2360
lemma cntP_simp1[simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2361
  "cntP (P th cs'#s) th = cntP s th + 1"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2362
  by (unfold cntP_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2363
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2364
lemma cntP_simp2[simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2365
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2366
  shows "cntP (P th cs'#s) th' = cntP s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2367
  using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2368
  by (unfold cntP_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2369
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2370
lemma cntP_simp3[simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2371
  assumes "\<not> isP e"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2372
  shows "cntP (e#s) th' = cntP s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2373
  using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2374
  by (unfold cntP_def, cases e, simp+)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2375
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2376
lemma cntV_simp1[simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2377
  "cntV (V th cs'#s) th = cntV s th + 1"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2378
  by (unfold cntV_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2379
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2380
lemma cntV_simp2[simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2381
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2382
  shows "cntV (V th cs'#s) th' = cntV s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2383
  using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2384
  by (unfold cntV_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2385
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2386
lemma cntV_simp3[simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2387
  assumes "\<not> isV e"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2388
  shows "cntV (e#s) th' = cntV s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2389
  using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2390
  by (unfold cntV_def, cases e, simp+)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2391
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2392
lemma cntP_diff_inv:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2393
  assumes "cntP (e#s) th \<noteq> cntP s th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2394
  shows "isP e \<and> actor e = th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2395
proof(cases e)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2396
  case (P th' pty)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2397
  show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2398
  by (cases "(\<lambda>e. \<exists>cs. e = P th cs) (P th' pty)", 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2399
        insert assms P, auto simp:cntP_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2400
qed (insert assms, auto simp:cntP_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2401
  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2402
lemma cntV_diff_inv:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2403
  assumes "cntV (e#s) th \<noteq> cntV s th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2404
  shows "isV e \<and> actor e = th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2405
proof(cases e)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2406
  case (V th' pty)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2407
  show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2408
  by (cases "(\<lambda>e. \<exists>cs. e = V th cs) (V th' pty)", 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2409
        insert assms V, auto simp:cntV_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2410
qed (insert assms, auto simp:cntV_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2411
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2412
lemma children_RAG_alt_def:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2413
  "children (RAG (s::state)) (Th th) = Cs ` {cs. holding s th cs}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2414
  by (unfold s_RAG_def, auto simp:children_def holding_eq)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2415
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2416
lemma holdents_alt_def:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2417
  "holdents s th = the_cs ` (children (RAG (s::state)) (Th th))"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2418
  by (unfold children_RAG_alt_def holdents_def, simp add: image_image)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2419
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2420
lemma cntCS_alt_def:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2421
  "cntCS s th = card (children (RAG s) (Th th))"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2422
  apply (unfold children_RAG_alt_def cntCS_def holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2423
  by (rule card_image[symmetric], auto simp:inj_on_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2424
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2425
context valid_trace
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2426
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2427
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2428
lemma finite_holdents: "finite (holdents s th)"
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2429
  by (unfold holdents_alt_def, insert fsbtRAGs.finite_children, auto)
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2430
  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2431
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2432
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2433
context valid_trace_p_w
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2434
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2435
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2436
lemma holding_s_holder: "holding s holder cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2437
  by (unfold s_holding_def, fold wq_def, unfold wq_s_cs, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2438
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2439
lemma holding_es_holder: "holding (e#s) holder cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2440
  by (unfold s_holding_def, fold wq_def, unfold wq_es_cs wq_s_cs, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2441
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2442
lemma holdents_es:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2443
  shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R") 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2444
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2445
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2446
    assume "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2447
    hence h: "holding (e#s) th' cs'" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2448
    have "holding s th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2449
    proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2450
      case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2451
      from held_unique[OF h[unfolded True] holding_es_holder]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2452
      have "th' = holder" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2453
      thus ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2454
        by (unfold True holdents_def, insert holding_s_holder, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2455
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2456
      case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2457
      hence "wq (e#s) cs' = wq s cs'" by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2458
      from h[unfolded s_holding_def, folded wq_def, unfolded this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2459
      show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2460
       by (unfold s_holding_def, fold wq_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2461
    qed 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2462
    hence "cs' \<in> ?R" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2463
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2464
    fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2465
    assume "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2466
    hence h: "holding s th' cs'" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2467
    have "holding (e#s) th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2468
    proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2469
      case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2470
      from held_unique[OF h[unfolded True] holding_s_holder]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2471
      have "th' = holder" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2472
      thus ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2473
        by (unfold True holdents_def, insert holding_es_holder, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2474
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2475
      case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2476
      hence "wq s cs' = wq (e#s) cs'" by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2477
      from h[unfolded s_holding_def, folded wq_def, unfolded this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2478
      show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2479
       by (unfold s_holding_def, fold wq_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2480
    qed 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2481
    hence "cs' \<in> ?L" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2482
  } ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2483
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2484
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2485
lemma cntCS_es_th[simp]: "cntCS (e#s) th' = cntCS s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2486
 by (unfold cntCS_def holdents_es, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2487
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2488
lemma th_not_ready_es: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2489
  shows "th \<notin> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2490
  using waiting_es_th_cs 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2491
  by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2492
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2493
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2494
  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2495
context valid_trace_p_h
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2496
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2497
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2498
lemma th_not_waiting':
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2499
  "\<not> waiting (e#s) th cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2500
proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2501
  case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2502
  show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2503
    by (unfold True s_waiting_def, fold wq_def, unfold wq_es_cs', auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2504
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2505
  case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2506
  from th_not_waiting[of cs', unfolded s_waiting_def, folded wq_def]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2507
  show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2508
    by (unfold s_waiting_def, fold wq_def, insert False, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2509
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2510
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2511
lemma ready_th_es: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2512
  shows "th \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2513
  using th_not_waiting'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2514
  by (unfold readys_def, insert live_th_es, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2515
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2516
lemma holdents_es_th:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2517
  "holdents (e#s) th = (holdents s th) \<union> {cs}" (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2518
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2519
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2520
    assume "cs' \<in> ?L" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2521
    hence "holding (e#s) th cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2522
      by (unfold holdents_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2523
    hence "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2524
     by (cases rule:holding_esE, auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2525
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2526
    fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2527
    assume "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2528
    hence "holding s th cs' \<or> cs' = cs" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2529
      by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2530
    hence "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2531
    proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2532
      assume "holding s th cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2533
      from holding_kept[OF this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2534
      show ?thesis by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2535
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2536
      assume "cs' = cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2537
      thus ?thesis using holding_es_th_cs
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2538
        by (unfold holdents_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2539
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2540
  } ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2541
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2542
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2543
lemma cntCS_es_th: "cntCS (e#s) th = cntCS s th + 1"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2544
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2545
  have "card (holdents s th \<union> {cs}) = card (holdents s th) + 1"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2546
  proof(subst card_Un_disjoint)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2547
    show "holdents s th \<inter> {cs} = {}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2548
      using not_holding_s_th_cs by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2549
  qed (auto simp:finite_holdents)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2550
  thus ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2551
   by (unfold cntCS_def holdents_es_th, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2552
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2553
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2554
lemma no_holder: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2555
  "\<not> holding s th' cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2556
proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2557
  assume otherwise: "holding s th' cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2558
  from this[unfolded s_holding_def, folded wq_def, unfolded we]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2559
  show False by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2560
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2561
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2562
lemma holdents_es_th':
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2563
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2564
  shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2565
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2566
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2567
    assume "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2568
    hence h_e: "holding (e#s) th' cs'" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2569
    have "cs' \<noteq> cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2570
    proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2571
      assume "cs' = cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2572
      from held_unique[OF h_e[unfolded this] holding_es_th_cs]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2573
      have "th' = th" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2574
      with assms show False by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2575
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2576
    from h_e[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp[OF this]]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2577
    have "th' \<in> set (wq s cs') \<and> th' = hd (wq s cs')" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2578
    hence "cs' \<in> ?R" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2579
      by (unfold holdents_def s_holding_def, fold wq_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2580
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2581
    fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2582
    assume "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2583
    hence "holding s th' cs'" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2584
    from holding_kept[OF this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2585
    have "holding (e # s) th' cs'" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2586
    hence "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2587
      by (unfold holdents_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2588
  } ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2589
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2590
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2591
lemma cntCS_es_th'[simp]: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2592
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2593
  shows "cntCS (e#s) th' = cntCS s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2594
  by (unfold cntCS_def holdents_es_th'[OF assms], simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2595
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2596
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2597
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2598
context valid_trace_p
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2599
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2600
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2601
lemma readys_kept1: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2602
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2603
  and "th' \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2604
  shows "th' \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2605
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2606
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2607
    assume wait: "waiting s th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2608
    have n_wait: "\<not> waiting (e#s) th' cs'" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2609
        using assms(2)[unfolded readys_def] by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2610
    have False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2611
    proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2612
      case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2613
      with n_wait wait
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2614
      show ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2615
        by (unfold s_waiting_def, fold wq_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2616
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2617
      case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2618
      show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2619
      proof(cases "wq s cs = []")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2620
        case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2621
        then interpret vt: valid_trace_p_h
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2622
          by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2623
        show ?thesis using n_wait wait waiting_kept by auto 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2624
      next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2625
        case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2626
        then interpret vt: valid_trace_p_w by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2627
        show ?thesis using n_wait wait waiting_kept by blast 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2628
      qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2629
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2630
  } with assms(2) show ?thesis  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2631
    by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2632
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2633
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2634
lemma readys_kept2: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2635
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2636
  and "th' \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2637
  shows "th' \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2638
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2639
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2640
    assume wait: "waiting (e#s) th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2641
    have n_wait: "\<not> waiting s th' cs'" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2642
        using assms(2)[unfolded readys_def] by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2643
    have False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2644
    proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2645
      case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2646
      with n_wait wait
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2647
      show ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2648
        by (unfold s_waiting_def, fold wq_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2649
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2650
      case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2651
      show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2652
      proof(cases "wq s cs = []")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2653
        case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2654
        then interpret vt: valid_trace_p_h
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2655
          by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2656
        show ?thesis using n_wait vt.waiting_esE wait by blast 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2657
      next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2658
        case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2659
        then interpret vt: valid_trace_p_w by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2660
        show ?thesis using assms(1) n_wait vt.waiting_esE wait by auto 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2661
      qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2662
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2663
  } with assms(2) show ?thesis  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2664
    by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2665
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2666
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2667
lemma readys_simp [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2668
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2669
  shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2670
  using readys_kept1[OF assms] readys_kept2[OF assms]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2671
  by metis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2672
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  2673
lemma cnp_cnv_cncs_kept: (* ddd *)
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2674
  assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2675
  shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2676
proof(cases "th' = th")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2677
  case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2678
  note eq_th' = this
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2679
  show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2680
  proof(cases "wq s cs = []")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2681
    case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2682
    then interpret vt: valid_trace_p_h by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2683
    show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2684
      using assms eq_th' is_p ready_th_s vt.cntCS_es_th vt.ready_th_es pvD_def by auto 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2685
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2686
    case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2687
    then interpret vt: valid_trace_p_w by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2688
    show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2689
      using add.commute add.left_commute assms eq_th' is_p live_th_s 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2690
            ready_th_s vt.th_not_ready_es pvD_def
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2691
      apply (auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2692
      by (fold is_p, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2693
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2694
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2695
  case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2696
  note h_False = False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2697
  thus ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2698
  proof(cases "wq s cs = []")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2699
    case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2700
    then interpret vt: valid_trace_p_h by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2701
    show ?thesis using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2702
      by (insert True h_False pvD_def, auto split:if_splits,unfold is_p, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2703
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2704
    case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2705
    then interpret vt: valid_trace_p_w by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2706
    show ?thesis using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2707
      by (insert False h_False pvD_def, auto split:if_splits,unfold is_p, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2708
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2709
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2710
57
f1b39d77db00 Added generic theory "RTree.thy"
xingyuan zhang <xingyuanzhang@126.com>
parents:
diff changeset
  2711
end
77
d37703e0c5c4 CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents: 73
diff changeset
  2712
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  2713
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2714
context valid_trace_v (* ccc *)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2715
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2716
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2717
lemma holding_th_cs_s: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2718
  "holding s th cs" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2719
 by  (unfold s_holding_def, fold wq_def, unfold wq_s_cs, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2720
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2721
lemma th_ready_s [simp]: "th \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2722
  using runing_th_s
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2723
  by (unfold runing_def readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2724
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2725
lemma th_live_s [simp]: "th \<in> threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2726
  using th_ready_s by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2727
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2728
lemma th_ready_es [simp]: "th \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2729
  using runing_th_s neq_t_th
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2730
  by (unfold is_v runing_def readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2731
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2732
lemma th_live_es [simp]: "th \<in> threads (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2733
  using th_ready_es by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2734
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2735
lemma pvD_th_s[simp]: "pvD s th = 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2736
  by (unfold pvD_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2737
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2738
lemma pvD_th_es[simp]: "pvD (e#s) th = 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2739
  by (unfold pvD_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2740
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2741
lemma cntCS_s_th [simp]: "cntCS s th > 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2742
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2743
  have "cs \<in> holdents s th" using holding_th_cs_s
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2744
    by (unfold holdents_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2745
  moreover have "finite (holdents s th)" using finite_holdents
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2746
    by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2747
  ultimately show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2748
    by (unfold cntCS_def, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2749
        auto intro!:card_gt_0_iff[symmetric, THEN iffD1])
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2750
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2751
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2752
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2753
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2754
context valid_trace_v_n
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2755
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2756
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2757
lemma not_ready_taker_s[simp]: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2758
  "taker \<notin> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2759
  using waiting_taker
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2760
  by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2761
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2762
lemma taker_live_s [simp]: "taker \<in> threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2763
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2764
  have "taker \<in> set wq'" by (simp add: eq_wq') 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2765
  from th'_in_inv[OF this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2766
  have "taker \<in> set rest" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2767
  hence "taker \<in> set (wq s cs)" by (simp add: wq_s_cs) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2768
  thus ?thesis using wq_threads by auto 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2769
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2770
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2771
lemma taker_live_es [simp]: "taker \<in> threads (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2772
  using taker_live_s threads_es by blast
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2773
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2774
lemma taker_ready_es [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2775
  shows "taker \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2776
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2777
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2778
    assume "waiting (e#s) taker cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2779
    hence False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2780
    proof(cases rule:waiting_esE)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2781
      case 1
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2782
      thus ?thesis using waiting_taker waiting_unique by auto 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2783
    qed simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2784
  } thus ?thesis by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2785
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2786
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2787
lemma neq_taker_th: "taker \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2788
  using th_not_waiting waiting_taker by blast
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2789
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2790
lemma not_holding_taker_s_cs:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2791
  shows "\<not> holding s taker cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2792
  using holding_cs_eq_th neq_taker_th by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2793
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2794
lemma holdents_es_taker:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2795
  "holdents (e#s) taker = holdents s taker \<union> {cs}" (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2796
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2797
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2798
    assume "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2799
    hence "holding (e#s) taker cs'" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2800
    hence "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2801
    proof(cases rule:holding_esE)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2802
      case 2
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2803
      thus ?thesis by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2804
    qed auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2805
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2806
    fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2807
    assume "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2808
    hence "holding s taker cs' \<or> cs' = cs" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2809
    hence "cs' \<in> ?L" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2810
    proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2811
      assume "holding s taker cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2812
      hence "holding (e#s) taker cs'" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2813
          using holding_esI2 holding_taker by fastforce 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2814
      thus ?thesis by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2815
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2816
      assume "cs' = cs"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2817
      with holding_taker
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2818
      show ?thesis by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2819
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2820
  } ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2821
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2822
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2823
lemma cntCS_es_taker [simp]: "cntCS (e#s) taker = cntCS s taker + 1"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2824
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2825
  have "card (holdents s taker \<union> {cs}) = card (holdents s taker) + 1"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2826
  proof(subst card_Un_disjoint)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2827
    show "holdents s taker \<inter> {cs} = {}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2828
      using not_holding_taker_s_cs by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2829
  qed (auto simp:finite_holdents)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2830
  thus ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2831
    by (unfold cntCS_def, insert holdents_es_taker, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2832
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2833
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2834
lemma pvD_taker_s[simp]: "pvD s taker = 1"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2835
  by (unfold pvD_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2836
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2837
lemma pvD_taker_es[simp]: "pvD (e#s) taker = 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2838
  by (unfold pvD_def, simp)  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2839
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2840
lemma pvD_th_s[simp]: "pvD s th = 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2841
  by (unfold pvD_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2842
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2843
lemma pvD_th_es[simp]: "pvD (e#s) th = 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2844
  by (unfold pvD_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2845
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2846
lemma holdents_es_th:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2847
  "holdents (e#s) th = holdents s th - {cs}" (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2848
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2849
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2850
    assume "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2851
    hence "holding (e#s) th cs'" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2852
    hence "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2853
    proof(cases rule:holding_esE)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2854
      case 2
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2855
      thus ?thesis by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2856
    qed (insert neq_taker_th, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2857
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2858
    fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2859
    assume "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2860
    hence "cs' \<noteq> cs" "holding s th cs'" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2861
    from holding_esI2[OF this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2862
    have "cs' \<in> ?L" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2863
  } ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2864
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2865
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2866
lemma cntCS_es_th [simp]: "cntCS (e#s) th = cntCS s th - 1"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2867
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2868
  have "card (holdents s th - {cs}) = card (holdents s th) - 1"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2869
  proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2870
    have "cs \<in> holdents s th" using holding_th_cs_s
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2871
      by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2872
    moreover have "finite (holdents s th)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2873
        by (simp add: finite_holdents) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2874
    ultimately show ?thesis by auto
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  2875
  qed
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2876
  thus ?thesis by (unfold cntCS_def holdents_es_th)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2877
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2878
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2879
lemma holdents_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2880
  assumes "th' \<noteq> taker"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2881
  and "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2882
  shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2883
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2884
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2885
    assume h: "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2886
    have "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2887
    proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2888
      case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2889
      hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2890
      from h have "holding (e#s) th' cs'" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2891
      from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2892
      show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2893
        by (unfold holdents_def s_holding_def, fold wq_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2894
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2895
      case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2896
      from h[unfolded this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2897
      have "holding (e#s) th' cs" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2898
      from held_unique[OF this holding_taker]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2899
      have "th' = taker" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2900
      with assms show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2901
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2902
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2903
    fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2904
    assume h: "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2905
    have "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2906
    proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2907
      case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2908
      hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2909
      from h have "holding s th' cs'" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2910
      from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2911
      show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2912
        by (unfold holdents_def s_holding_def, fold wq_def, insert eq_wq, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2913
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2914
      case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2915
      from h[unfolded this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2916
      have "holding s th' cs" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2917
      from held_unique[OF this holding_th_cs_s]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2918
      have "th' = th" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2919
      with assms show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2920
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2921
  } ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2922
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2923
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2924
lemma cntCS_kept [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2925
  assumes "th' \<noteq> taker"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2926
  and "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2927
  shows "cntCS (e#s) th' = cntCS s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2928
  by (unfold cntCS_def holdents_kept[OF assms], simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2929
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2930
lemma readys_kept1: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2931
  assumes "th' \<noteq> taker"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2932
  and "th' \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2933
  shows "th' \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2934
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2935
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2936
    assume wait: "waiting s th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2937
    have n_wait: "\<not> waiting (e#s) th' cs'" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2938
        using assms(2)[unfolded readys_def] by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2939
    have False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2940
    proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2941
      case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2942
      with n_wait wait
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2943
      show ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2944
        by (unfold s_waiting_def, fold wq_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2945
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2946
      case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2947
      have "th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest)" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2948
        using wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2949
      moreover have "\<not> (th' \<in> set rest \<and> th' \<noteq> hd (taker # rest'))" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2950
        using n_wait[unfolded True s_waiting_def, folded wq_def, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2951
                    unfolded wq_es_cs set_wq', unfolded eq_wq'] .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2952
      ultimately have "th' = taker" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2953
      with assms(1)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2954
      show ?thesis by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2955
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2956
  } with assms(2) show ?thesis  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2957
    by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2958
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2959
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2960
lemma readys_kept2: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2961
  assumes "th' \<noteq> taker"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2962
  and "th' \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2963
  shows "th' \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2964
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2965
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2966
    assume wait: "waiting (e#s) th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2967
    have n_wait: "\<not> waiting s th' cs'" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2968
        using assms(2)[unfolded readys_def] by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2969
    have False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2970
    proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2971
      case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2972
      with n_wait wait
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2973
      show ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2974
        by (unfold s_waiting_def, fold wq_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2975
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2976
      case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2977
      have "th' \<in> set rest \<and> th' \<noteq> hd (taker # rest')"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2978
          using  wait [unfolded True s_waiting_def, folded wq_def, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2979
                    unfolded wq_es_cs set_wq', unfolded eq_wq']  .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2980
      moreover have "\<not> (th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest))"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2981
          using n_wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2982
      ultimately have "th' = taker" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2983
      with assms(1)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2984
      show ?thesis by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2985
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2986
  } with assms(2) show ?thesis  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2987
    by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2988
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2989
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2990
lemma readys_simp [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2991
  assumes "th' \<noteq> taker"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2992
  shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2993
  using readys_kept1[OF assms] readys_kept2[OF assms]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2994
  by metis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2995
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2996
lemma cnp_cnv_cncs_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2997
  assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2998
  shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  2999
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3000
  { assume eq_th': "th' = taker"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3001
    have ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3002
      apply (unfold eq_th' pvD_taker_es cntCS_es_taker)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3003
      by (insert neq_taker_th assms[unfolded eq_th'], unfold is_v, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3004
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3005
    assume eq_th': "th' = th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3006
    have ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3007
      apply (unfold eq_th' pvD_th_es cntCS_es_th)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3008
      by (insert assms[unfolded eq_th'], unfold is_v, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3009
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3010
    assume h: "th' \<noteq> taker" "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3011
    have ?thesis using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3012
      apply (unfold cntCS_kept[OF h], insert h, unfold is_v, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3013
      by (fold is_v, unfold pvD_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3014
  } ultimately show ?thesis by metis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3015
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3016
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3017
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3018
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3019
context valid_trace_v_e
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3020
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3021
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3022
lemma holdents_es_th:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3023
  "holdents (e#s) th = holdents s th - {cs}" (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3024
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3025
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3026
    assume "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3027
    hence "holding (e#s) th cs'" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3028
    hence "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3029
    proof(cases rule:holding_esE)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3030
      case 1
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3031
      thus ?thesis by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3032
    qed 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3033
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3034
    fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3035
    assume "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3036
    hence "cs' \<noteq> cs" "holding s th cs'" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3037
    from holding_esI2[OF this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3038
    have "cs' \<in> ?L" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3039
  } ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3040
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3041
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3042
lemma cntCS_es_th [simp]: "cntCS (e#s) th = cntCS s th - 1"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3043
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3044
  have "card (holdents s th - {cs}) = card (holdents s th) - 1"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3045
  proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3046
    have "cs \<in> holdents s th" using holding_th_cs_s
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3047
      by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3048
    moreover have "finite (holdents s th)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3049
        by (simp add: finite_holdents) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3050
    ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3051
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3052
  thus ?thesis by (unfold cntCS_def holdents_es_th)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3053
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3054
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3055
lemma holdents_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3056
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3057
  shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3058
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3059
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3060
    assume h: "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3061
    have "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3062
    proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3063
      case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3064
      hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3065
      from h have "holding (e#s) th' cs'" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3066
      from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3067
      show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3068
        by (unfold holdents_def s_holding_def, fold wq_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3069
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3070
      case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3071
      from h[unfolded this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3072
      have "holding (e#s) th' cs" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3073
      from this[unfolded s_holding_def, folded wq_def, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3074
            unfolded wq_es_cs nil_wq']
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3075
      show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3076
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3077
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3078
    fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3079
    assume h: "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3080
    have "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3081
    proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3082
      case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3083
      hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3084
      from h have "holding s th' cs'" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3085
      from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3086
      show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3087
        by (unfold holdents_def s_holding_def, fold wq_def, insert eq_wq, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3088
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3089
      case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3090
      from h[unfolded this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3091
      have "holding s th' cs" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3092
      from held_unique[OF this holding_th_cs_s]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3093
      have "th' = th" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3094
      with assms show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3095
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3096
  } ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3097
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3098
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3099
lemma cntCS_kept [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3100
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3101
  shows "cntCS (e#s) th' = cntCS s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3102
  by (unfold cntCS_def holdents_kept[OF assms], simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3103
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3104
lemma readys_kept1: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3105
  assumes "th' \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3106
  shows "th' \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3107
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3108
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3109
    assume wait: "waiting s th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3110
    have n_wait: "\<not> waiting (e#s) th' cs'" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3111
        using assms(1)[unfolded readys_def] by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3112
    have False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3113
    proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3114
      case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3115
      with n_wait wait
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3116
      show ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3117
        by (unfold s_waiting_def, fold wq_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3118
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3119
      case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3120
      have "th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest)" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3121
        using wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] . 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3122
      hence "th' \<in> set rest" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3123
      with set_wq' have "th' \<in> set wq'" by metis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3124
      with nil_wq' show ?thesis by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3125
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3126
  } thus ?thesis using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3127
    by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3128
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3129
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3130
lemma readys_kept2: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3131
  assumes "th' \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3132
  shows "th' \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3133
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3134
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3135
    assume wait: "waiting (e#s) th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3136
    have n_wait: "\<not> waiting s th' cs'" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3137
        using assms[unfolded readys_def] by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3138
    have False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3139
    proof(cases "cs' = cs")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3140
      case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3141
      with n_wait wait
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3142
      show ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3143
        by (unfold s_waiting_def, fold wq_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3144
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3145
      case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3146
      have "th' \<in> set [] \<and> th' \<noteq> hd []"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3147
        using wait[unfolded True s_waiting_def, folded wq_def, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3148
              unfolded wq_es_cs nil_wq'] .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3149
      thus ?thesis by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3150
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3151
  } with assms show ?thesis  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3152
    by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3153
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3154
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3155
lemma readys_simp [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3156
  shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3157
  using readys_kept1[OF assms] readys_kept2[OF assms]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3158
  by metis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3159
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3160
lemma cnp_cnv_cncs_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3161
  assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3162
  shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3163
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3164
  {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3165
    assume eq_th': "th' = th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3166
    have ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3167
      apply (unfold eq_th' pvD_th_es cntCS_es_th)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3168
      by (insert assms[unfolded eq_th'], unfold is_v, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3169
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3170
    assume h: "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3171
    have ?thesis using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3172
      apply (unfold cntCS_kept[OF h], insert h, unfold is_v, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3173
      by (fold is_v, unfold pvD_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3174
  } ultimately show ?thesis by metis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3175
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3176
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3177
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3178
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3179
context valid_trace_v
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3180
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3181
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3182
lemma cnp_cnv_cncs_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3183
  assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3184
  shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3185
proof(cases "rest = []")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3186
  case True
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3187
  then interpret vt: valid_trace_v_e by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3188
  show ?thesis using assms using vt.cnp_cnv_cncs_kept by blast 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3189
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3190
  case False
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3191
  then interpret vt: valid_trace_v_n by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3192
  show ?thesis using assms using vt.cnp_cnv_cncs_kept by blast 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3193
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3194
79
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  3195
end
8067efcb43da Still improving CpsG.thy
zhangx
parents: 77
diff changeset
  3196
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3197
context valid_trace_create
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3198
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3199
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3200
lemma th_not_live_s [simp]: "th \<notin> threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3201
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3202
  from pip_e[unfolded is_create]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3203
  show ?thesis by (cases, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3204
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3205
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3206
lemma th_not_ready_s [simp]: "th \<notin> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3207
  using th_not_live_s by (unfold readys_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3208
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3209
lemma th_live_es [simp]: "th \<in> threads (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3210
  by (unfold is_create, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3211
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3212
lemma not_waiting_th_s [simp]: "\<not> waiting s th cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3213
proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3214
  assume "waiting s th cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3215
  from this[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3216
  have "th \<in> set (wq s cs')" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3217
  from wq_threads[OF this] have "th \<in> threads s" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3218
  with th_not_live_s show False by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3219
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3220
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3221
lemma not_holding_th_s [simp]: "\<not> holding s th cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3222
proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3223
  assume "holding s th cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3224
  from this[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3225
  have "th \<in> set (wq s cs')" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3226
  from wq_threads[OF this] have "th \<in> threads s" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3227
  with th_not_live_s show False by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3228
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3229
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3230
lemma not_waiting_th_es [simp]: "\<not> waiting (e#s) th cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3231
proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3232
  assume "waiting (e # s) th cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3233
  from this[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3234
  have "th \<in> set (wq s cs')" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3235
  from wq_threads[OF this] have "th \<in> threads s" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3236
  with th_not_live_s show False by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3237
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3238
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3239
lemma not_holding_th_es [simp]: "\<not> holding (e#s) th cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3240
proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3241
  assume "holding (e # s) th cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3242
  from this[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3243
  have "th \<in> set (wq s cs')" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3244
  from wq_threads[OF this] have "th \<in> threads s" .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3245
  with th_not_live_s show False by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3246
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3247
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3248
lemma ready_th_es [simp]: "th \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3249
  by (simp add:readys_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3250
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3251
lemma holdents_th_s: "holdents s th = {}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3252
  by (unfold holdents_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3253
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3254
lemma holdents_th_es: "holdents (e#s) th = {}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3255
  by (unfold holdents_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3256
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3257
lemma cntCS_th_s [simp]: "cntCS s th = 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3258
  by (unfold cntCS_def, simp add:holdents_th_s)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3259
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3260
lemma cntCS_th_es [simp]: "cntCS (e#s) th = 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3261
  by (unfold cntCS_def, simp add:holdents_th_es)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3262
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3263
lemma pvD_th_s [simp]: "pvD s th = 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3264
  by (unfold pvD_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3265
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3266
lemma pvD_th_es [simp]: "pvD (e#s) th = 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3267
  by (unfold pvD_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3268
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3269
lemma holdents_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3270
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3271
  shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3272
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3273
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3274
    assume h: "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3275
    hence "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3276
      by (unfold holdents_def s_holding_def, fold wq_def, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3277
             unfold wq_neq_simp, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3278
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3279
    fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3280
    assume h: "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3281
    hence "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3282
      by (unfold holdents_def s_holding_def, fold wq_def, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3283
             unfold wq_neq_simp, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3284
  } ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3285
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3286
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3287
lemma cntCS_kept [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3288
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3289
  shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3290
  using holdents_kept[OF assms]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3291
  by (unfold cntCS_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3292
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3293
lemma readys_kept1: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3294
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3295
  and "th' \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3296
  shows "th' \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3297
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3298
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3299
    assume wait: "waiting s th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3300
    have n_wait: "\<not> waiting (e#s) th' cs'" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3301
      using assms by (auto simp:readys_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3302
    from wait[unfolded s_waiting_def, folded wq_def]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3303
         n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3304
    have False by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3305
  } thus ?thesis using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3306
    by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3307
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3308
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3309
lemma readys_kept2: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3310
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3311
  and "th' \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3312
  shows "th' \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3313
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3314
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3315
    assume wait: "waiting (e#s) th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3316
    have n_wait: "\<not> waiting s th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3317
      using assms(2) by (auto simp:readys_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3318
    from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3319
         n_wait[unfolded s_waiting_def, folded wq_def]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3320
    have False by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3321
  } with assms show ?thesis  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3322
    by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3323
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3324
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3325
lemma readys_simp [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3326
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3327
  shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3328
  using readys_kept1[OF assms] readys_kept2[OF assms]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3329
  by metis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3330
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3331
lemma pvD_kept [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3332
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3333
  shows "pvD (e#s) th' = pvD s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3334
  using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3335
  by (unfold pvD_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3336
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3337
lemma cnp_cnv_cncs_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3338
  assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3339
  shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3340
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3341
  {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3342
    assume eq_th': "th' = th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3343
    have ?thesis using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3344
      by (unfold eq_th', simp, unfold is_create, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3345
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3346
    assume h: "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3347
    hence ?thesis using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3348
      by (simp, simp add:is_create)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3349
  } ultimately show ?thesis by metis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3350
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3351
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3352
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3353
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3354
context valid_trace_exit
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3355
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3356
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3357
lemma th_live_s [simp]: "th \<in> threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3358
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3359
  from pip_e[unfolded is_exit]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3360
  show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3361
  by (cases, unfold runing_def readys_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3362
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3363
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3364
lemma th_ready_s [simp]: "th \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3365
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3366
  from pip_e[unfolded is_exit]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3367
  show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3368
  by (cases, unfold runing_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3369
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3370
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3371
lemma th_not_live_es [simp]: "th \<notin> threads (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3372
  by (unfold is_exit, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3373
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3374
lemma not_holding_th_s [simp]: "\<not> holding s th cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3375
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3376
  from pip_e[unfolded is_exit]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3377
  show ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3378
   by (cases, unfold holdents_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3379
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3380
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3381
lemma cntCS_th_s [simp]: "cntCS s th = 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3382
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3383
  from pip_e[unfolded is_exit]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3384
  show ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3385
   by (cases, unfold cntCS_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3386
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3387
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3388
lemma not_holding_th_es [simp]: "\<not> holding (e#s) th cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3389
proof
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3390
  assume "holding (e # s) th cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3391
  from this[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3392
  have "holding s th cs'" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3393
    by (unfold s_holding_def, fold wq_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3394
  with not_holding_th_s 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3395
  show False by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3396
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3397
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3398
lemma ready_th_es [simp]: "th \<notin> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3399
  by (simp add:readys_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3400
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3401
lemma holdents_th_s: "holdents s th = {}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3402
  by (unfold holdents_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3403
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3404
lemma holdents_th_es: "holdents (e#s) th = {}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3405
  by (unfold holdents_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3406
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3407
lemma cntCS_th_es [simp]: "cntCS (e#s) th = 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3408
  by (unfold cntCS_def, simp add:holdents_th_es)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3409
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3410
lemma pvD_th_s [simp]: "pvD s th = 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3411
  by (unfold pvD_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3412
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3413
lemma pvD_th_es [simp]: "pvD (e#s) th = 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3414
  by (unfold pvD_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3415
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3416
lemma holdents_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3417
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3418
  shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3419
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3420
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3421
    assume h: "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3422
    hence "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3423
      by (unfold holdents_def s_holding_def, fold wq_def, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3424
             unfold wq_neq_simp, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3425
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3426
    fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3427
    assume h: "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3428
    hence "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3429
      by (unfold holdents_def s_holding_def, fold wq_def, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3430
             unfold wq_neq_simp, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3431
  } ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3432
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3433
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3434
lemma cntCS_kept [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3435
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3436
  shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3437
  using holdents_kept[OF assms]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3438
  by (unfold cntCS_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3439
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3440
lemma readys_kept1: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3441
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3442
  and "th' \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3443
  shows "th' \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3444
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3445
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3446
    assume wait: "waiting s th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3447
    have n_wait: "\<not> waiting (e#s) th' cs'" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3448
      using assms by (auto simp:readys_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3449
    from wait[unfolded s_waiting_def, folded wq_def]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3450
         n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3451
    have False by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3452
  } thus ?thesis using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3453
    by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3454
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3455
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3456
lemma readys_kept2: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3457
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3458
  and "th' \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3459
  shows "th' \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3460
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3461
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3462
    assume wait: "waiting (e#s) th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3463
    have n_wait: "\<not> waiting s th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3464
      using assms(2) by (auto simp:readys_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3465
    from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3466
         n_wait[unfolded s_waiting_def, folded wq_def]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3467
    have False by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3468
  } with assms show ?thesis  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3469
    by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3470
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3471
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3472
lemma readys_simp [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3473
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3474
  shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3475
  using readys_kept1[OF assms] readys_kept2[OF assms]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3476
  by metis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3477
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3478
lemma pvD_kept [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3479
  assumes "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3480
  shows "pvD (e#s) th' = pvD s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3481
  using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3482
  by (unfold pvD_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3483
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3484
lemma cnp_cnv_cncs_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3485
  assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3486
  shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3487
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3488
  {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3489
    assume eq_th': "th' = th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3490
    have ?thesis using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3491
      by (unfold eq_th', simp, unfold is_exit, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3492
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3493
    assume h: "th' \<noteq> th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3494
    hence ?thesis using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3495
      by (simp, simp add:is_exit)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3496
  } ultimately show ?thesis by metis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3497
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3498
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3499
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3500
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3501
context valid_trace_set
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3502
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3503
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3504
lemma th_live_s [simp]: "th \<in> threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3505
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3506
  from pip_e[unfolded is_set]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3507
  show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3508
  by (cases, unfold runing_def readys_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3509
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3510
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3511
lemma th_ready_s [simp]: "th \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3512
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3513
  from pip_e[unfolded is_set]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3514
  show ?thesis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3515
  by (cases, unfold runing_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3516
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3517
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3518
lemma th_not_live_es [simp]: "th \<in> threads (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3519
  by (unfold is_set, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3520
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3521
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3522
lemma holdents_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3523
  shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3524
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3525
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3526
    assume h: "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3527
    hence "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3528
      by (unfold holdents_def s_holding_def, fold wq_def, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3529
             unfold wq_neq_simp, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3530
  } moreover {
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3531
    fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3532
    assume h: "cs' \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3533
    hence "cs' \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3534
      by (unfold holdents_def s_holding_def, fold wq_def, 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3535
             unfold wq_neq_simp, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3536
  } ultimately show ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3537
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3538
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3539
lemma cntCS_kept [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3540
  shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3541
  using holdents_kept
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3542
  by (unfold cntCS_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3543
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3544
lemma threads_kept[simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3545
  "threads (e#s) = threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3546
  by (unfold is_set, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3547
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3548
lemma readys_kept1: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3549
  assumes "th' \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3550
  shows "th' \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3551
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3552
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3553
    assume wait: "waiting s th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3554
    have n_wait: "\<not> waiting (e#s) th' cs'" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3555
      using assms by (auto simp:readys_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3556
    from wait[unfolded s_waiting_def, folded wq_def]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3557
         n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3558
    have False by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3559
  } moreover have "th' \<in> threads s" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3560
    using assms[unfolded readys_def] by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3561
  ultimately show ?thesis 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3562
    by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3563
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3564
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3565
lemma readys_kept2: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3566
  assumes "th' \<in> readys s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3567
  shows "th' \<in> readys (e#s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3568
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3569
  { fix cs'
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3570
    assume wait: "waiting (e#s) th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3571
    have n_wait: "\<not> waiting s th' cs'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3572
      using assms by (auto simp:readys_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3573
    from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3574
         n_wait[unfolded s_waiting_def, folded wq_def]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3575
    have False by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3576
  } with assms show ?thesis  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3577
    by (unfold readys_def, auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3578
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3579
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3580
lemma readys_simp [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3581
  shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3582
  using readys_kept1 readys_kept2
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3583
  by metis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3584
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3585
lemma pvD_kept [simp]:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3586
  shows "pvD (e#s) th' = pvD s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3587
  by (unfold pvD_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3588
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3589
lemma cnp_cnv_cncs_kept:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3590
  assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3591
  shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3592
  using assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3593
  by (unfold is_set, simp, fold is_set, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3594
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3595
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3596
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3597
context valid_trace
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3598
begin
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3599
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3600
lemma cnp_cnv_cncs: "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3601
proof(induct rule:ind)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3602
  case Nil
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3603
  thus ?case 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3604
    by (unfold cntP_def cntV_def pvD_def cntCS_def holdents_def 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3605
              s_holding_def, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3606
next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3607
  case (Cons s e)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3608
  interpret vt_e: valid_trace_e s e using Cons by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3609
  show ?case
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3610
  proof(cases e)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3611
    case (Create th prio)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3612
    interpret vt_create: valid_trace_create s e th prio 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3613
      using Create by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3614
    show ?thesis using Cons by (simp add: vt_create.cnp_cnv_cncs_kept) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3615
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3616
    case (Exit th)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3617
    interpret vt_exit: valid_trace_exit s e th  
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3618
        using Exit by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3619
   show ?thesis using Cons by (simp add: vt_exit.cnp_cnv_cncs_kept) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3620
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3621
    case (P th cs)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3622
    interpret vt_p: valid_trace_p s e th cs using P by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3623
    show ?thesis using Cons by (simp add: vt_p.cnp_cnv_cncs_kept) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3624
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3625
    case (V th cs)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3626
    interpret vt_v: valid_trace_v s e th cs using V by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3627
    show ?thesis using Cons by (simp add: vt_v.cnp_cnv_cncs_kept) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3628
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3629
    case (Set th prio)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3630
    interpret vt_set: valid_trace_set s e th prio
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3631
        using Set by (unfold_locales, simp)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3632
    show ?thesis using Cons by (simp add: vt_set.cnp_cnv_cncs_kept) 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3633
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3634
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3635
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3636
lemma not_thread_holdents:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3637
  assumes not_in: "th \<notin> threads s" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3638
  shows "holdents s th = {}"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3639
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3640
  { fix cs
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3641
    assume "cs \<in> holdents s th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3642
    hence "holding s th cs" by (auto simp:holdents_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3643
    from this[unfolded s_holding_def, folded wq_def]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3644
    have "th \<in> set (wq s cs)" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3645
    with wq_threads have "th \<in> threads s" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3646
    with assms
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3647
    have False by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3648
  } thus ?thesis by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3649
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3650
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3651
lemma not_thread_cncs:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3652
  assumes not_in: "th \<notin> threads s" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3653
  shows "cntCS s th = 0"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3654
  using not_thread_holdents[OF assms]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3655
  by (simp add:cntCS_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3656
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3657
lemma cnp_cnv_eq:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3658
  assumes "th \<notin> threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3659
  shows "cntP s th = cntV s th"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3660
  using assms cnp_cnv_cncs not_thread_cncs pvD_def
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3661
  by (auto)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3662
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3663
lemma runing_unique:
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3664
  assumes runing_1: "th1 \<in> runing s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3665
  and runing_2: "th2 \<in> runing s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3666
  shows "th1 = th2"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3667
proof -
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3668
  from runing_1 and runing_2 have "cp s th1 = cp s th2"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3669
    unfolding runing_def by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3670
  from this[unfolded cp_alt_def]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3671
  have eq_max: 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3672
    "Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th1)}) =
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3673
     Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th2)})" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3674
        (is "Max ?L = Max ?R") .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3675
  have "Max ?L \<in> ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3676
  proof(rule Max_in)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3677
    show "finite ?L" by (simp add: finite_subtree_threads)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3678
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3679
    show "?L \<noteq> {}" using subtree_def by fastforce 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3680
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3681
  then obtain th1' where 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3682
    h_1: "Th th1' \<in> subtree (RAG s) (Th th1)" "the_preced s th1' = Max ?L"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3683
    by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3684
  have "Max ?R \<in> ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3685
  proof(rule Max_in)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3686
    show "finite ?R" by (simp add: finite_subtree_threads)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3687
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3688
    show "?R \<noteq> {}" using subtree_def by fastforce 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3689
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3690
  then obtain th2' where 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3691
    h_2: "Th th2' \<in> subtree (RAG s) (Th th2)" "the_preced s th2' = Max ?R"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3692
    by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3693
  have "th1' = th2'"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3694
  proof(rule preced_unique)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3695
    from h_1(1)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3696
    show "th1' \<in> threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3697
    proof(cases rule:subtreeE)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3698
      case 1
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3699
      hence "th1' = th1" by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3700
      with runing_1 show ?thesis by (auto simp:runing_def readys_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3701
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3702
      case 2
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3703
      from this(2)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3704
      have "(Th th1', Th th1) \<in> (RAG s)^+" by (auto simp:ancestors_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3705
      from tranclD[OF this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3706
      have "(Th th1') \<in> Domain (RAG s)" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3707
      from dm_RAG_threads[OF this] show ?thesis .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3708
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3709
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3710
    from h_2(1)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3711
    show "th2' \<in> threads s"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3712
    proof(cases rule:subtreeE)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3713
      case 1
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3714
      hence "th2' = th2" by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3715
      with runing_2 show ?thesis by (auto simp:runing_def readys_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3716
    next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3717
      case 2
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3718
      from this(2)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3719
      have "(Th th2', Th th2) \<in> (RAG s)^+" by (auto simp:ancestors_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3720
      from tranclD[OF this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3721
      have "(Th th2') \<in> Domain (RAG s)" by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3722
      from dm_RAG_threads[OF this] show ?thesis .
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3723
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3724
  next
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3725
    have "the_preced s th1' = the_preced s th2'" 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3726
     using eq_max h_1(2) h_2(2) by metis
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3727
    thus "preced th1' s = preced th2' s" by (simp add:the_preced_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3728
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3729
  from h_1(1)[unfolded this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3730
  have star1: "(Th th2', Th th1) \<in> (RAG s)^*" by (auto simp:subtree_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3731
  from h_2(1)[unfolded this]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3732
  have star2: "(Th th2', Th th2) \<in> (RAG s)^*" by (auto simp:subtree_def)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3733
  from star_rpath[OF star1] obtain xs1 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3734
    where rp1: "rpath (RAG s) (Th th2') xs1 (Th th1)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3735
    by auto
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3736
  from star_rpath[OF star2] obtain xs2 
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3737
    where rp2: "rpath (RAG s) (Th th2') xs2 (Th th2)"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3738
    by auto
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3739
  from rp1 rp2
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3740
  show ?thesis
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3741
  proof(cases)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3742
    case (less_1 xs')
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3743
    moreover have "xs' = []"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3744
    proof(rule ccontr)
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3745
      assume otherwise: "xs' \<noteq> []"
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3746
      from rpath_plus[OF less_1(3) this]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3747
      have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+" .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3748
      from tranclD[OF this]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3749
      obtain cs where "waiting s th1 cs"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3750
        by (unfold s_RAG_def, fold waiting_eq, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3751
      with runing_1 show False
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3752
        by (unfold runing_def readys_def, auto)
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3753
    qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3754
    ultimately have "xs2 = xs1" by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3755
    from rpath_dest_eq[OF rp1 rp2[unfolded this]]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3756
    show ?thesis by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3757
  next
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3758
    case (less_2 xs')
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3759
    moreover have "xs' = []"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3760
    proof(rule ccontr)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3761
      assume otherwise: "xs' \<noteq> []"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3762
      from rpath_plus[OF less_2(3) this]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3763
      have "(Th th2, Th th1) \<in> (RAG s)\<^sup>+" .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3764
      from tranclD[OF this]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3765
      obtain cs where "waiting s th2 cs"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3766
        by (unfold s_RAG_def, fold waiting_eq, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3767
      with runing_2 show False
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3768
        by (unfold runing_def readys_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3769
    qed
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3770
    ultimately have "xs2 = xs1" by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3771
    from rpath_dest_eq[OF rp1 rp2[unfolded this]]
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3772
    show ?thesis by simp
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3773
  qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3774
qed
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3775
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3776
lemma card_runing: "card (runing s) \<le> 1"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3777
proof(cases "runing s = {}")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3778
  case True
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3779
  thus ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3780
next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3781
  case False
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3782
  then obtain th where [simp]: "th \<in> runing s" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3783
  from runing_unique[OF this]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3784
  have "runing s = {th}" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3785
  thus ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3786
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3787
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3788
lemma create_pre:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3789
  assumes stp: "step s e"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3790
  and not_in: "th \<notin> threads s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3791
  and is_in: "th \<in> threads (e#s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3792
  obtains prio where "e = Create th prio"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3793
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3794
  from assms  
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3795
  show ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3796
  proof(cases)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3797
    case (thread_create thread prio)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3798
    with is_in not_in have "e = Create th prio" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3799
    from that[OF this] show ?thesis .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3800
  next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3801
    case (thread_exit thread)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3802
    with assms show ?thesis by (auto intro!:that)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3803
  next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3804
    case (thread_P thread)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3805
    with assms show ?thesis by (auto intro!:that)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3806
  next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3807
    case (thread_V thread)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3808
    with assms show ?thesis by (auto intro!:that)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3809
  next 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3810
    case (thread_set thread)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3811
    with assms show ?thesis by (auto intro!:that)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3812
  qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3813
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3814
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3815
lemma eq_pv_children:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3816
  assumes eq_pv: "cntP s th = cntV s th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3817
  shows "children (RAG s) (Th th) = {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3818
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3819
    from cnp_cnv_cncs and eq_pv
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3820
    have "cntCS s th = 0" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3821
      by (auto split:if_splits)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3822
    from this[unfolded cntCS_def holdents_alt_def]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3823
    have card_0: "card (the_cs ` children (RAG s) (Th th)) = 0" .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3824
    have "finite (the_cs ` children (RAG s) (Th th))"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3825
      by (simp add: fsbtRAGs.finite_children)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3826
    from card_0[unfolded card_0_eq[OF this]]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3827
    show ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3828
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3829
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3830
lemma eq_pv_holdents:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3831
  assumes eq_pv: "cntP s th = cntV s th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3832
  shows "holdents s th = {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3833
  by (unfold holdents_alt_def eq_pv_children[OF assms], simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3834
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3835
lemma eq_pv_subtree:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3836
  assumes eq_pv: "cntP s th = cntV s th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3837
  shows "subtree (RAG s) (Th th) = {Th th}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3838
  using eq_pv_children[OF assms]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3839
    by (unfold subtree_children, simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3840
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3841
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  3842
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3843
lemma cp_gen_alt_def:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3844
  "cp_gen s = (Max \<circ> (\<lambda>x. (the_preced s \<circ> the_thread) ` subtree (tRAG s) x))"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3845
    by (auto simp:cp_gen_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3846
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3847
lemma tRAG_nodeE:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3848
  assumes "(n1, n2) \<in> tRAG s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3849
  obtains th1 th2 where "n1 = Th th1" "n2 = Th th2"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3850
  using assms
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3851
  by (auto simp: tRAG_def wRAG_def hRAG_def tRAG_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3852
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3853
lemma subtree_nodeE:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3854
  assumes "n \<in> subtree (tRAG s) (Th th)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3855
  obtains th1 where "n = Th th1"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3856
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3857
  show ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3858
  proof(rule subtreeE[OF assms])
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3859
    assume "n = Th th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3860
    from that[OF this] show ?thesis .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3861
  next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3862
    assume "Th th \<in> ancestors (tRAG s) n"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3863
    hence "(n, Th th) \<in> (tRAG s)^+" by (auto simp:ancestors_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3864
    hence "\<exists> th1. n = Th th1"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3865
    proof(induct)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3866
      case (base y)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3867
      from tRAG_nodeE[OF this] show ?case by metis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3868
    next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3869
      case (step y z)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3870
      thus ?case by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3871
    qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3872
    with that show ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3873
  qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3874
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3875
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3876
lemma tRAG_star_RAG: "(tRAG s)^* \<subseteq> (RAG s)^*"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3877
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3878
  have "(wRAG s O hRAG s)^* \<subseteq> (RAG s O RAG s)^*" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3879
    by (rule rtrancl_mono, auto simp:RAG_split)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3880
  also have "... \<subseteq> ((RAG s)^*)^*"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3881
    by (rule rtrancl_mono, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3882
  also have "... = (RAG s)^*" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3883
  finally show ?thesis by (unfold tRAG_def, simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3884
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3885
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3886
lemma tRAG_subtree_RAG: "subtree (tRAG s) x \<subseteq> subtree (RAG s) x"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3887
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3888
  { fix a
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3889
    assume "a \<in> subtree (tRAG s) x"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3890
    hence "(a, x) \<in> (tRAG s)^*" by (auto simp:subtree_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3891
    with tRAG_star_RAG
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3892
    have "(a, x) \<in> (RAG s)^*" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3893
    hence "a \<in> subtree (RAG s) x" by (auto simp:subtree_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3894
  } thus ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3895
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3896
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3897
lemma tRAG_trancl_eq:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3898
   "{th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3899
    {th'. (Th th', Th th)  \<in> (RAG s)^+}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3900
   (is "?L = ?R")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3901
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3902
  { fix th'
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3903
    assume "th' \<in> ?L"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3904
    hence "(Th th', Th th) \<in> (tRAG s)^+" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3905
    from tranclD[OF this]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3906
    obtain z where h: "(Th th', z) \<in> tRAG s" "(z, Th th) \<in> (tRAG s)\<^sup>*" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3907
    from tRAG_subtree_RAG and this(2)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3908
    have "(z, Th th) \<in> (RAG s)^*" by (meson subsetCE tRAG_star_RAG) 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3909
    moreover from h(1) have "(Th th', z) \<in> (RAG s)^+" using tRAG_alt_def by auto 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3910
    ultimately have "th' \<in> ?R"  by auto 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3911
  } moreover 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3912
  { fix th'
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3913
    assume "th' \<in> ?R"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3914
    hence "(Th th', Th th) \<in> (RAG s)^+" by (auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3915
    from plus_rpath[OF this]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3916
    obtain xs where rp: "rpath (RAG s) (Th th') xs (Th th)" "xs \<noteq> []" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3917
    hence "(Th th', Th th) \<in> (tRAG s)^+"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3918
    proof(induct xs arbitrary:th' th rule:length_induct)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3919
      case (1 xs th' th)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3920
      then obtain x1 xs1 where Cons1: "xs = x1#xs1" by (cases xs, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3921
      show ?case
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3922
      proof(cases "xs1")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3923
        case Nil
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3924
        from 1(2)[unfolded Cons1 Nil]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3925
        have rp: "rpath (RAG s) (Th th') [x1] (Th th)" .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3926
        hence "(Th th', x1) \<in> (RAG s)" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3927
          by (cases, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3928
        then obtain cs where "x1 = Cs cs" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3929
              by (unfold s_RAG_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3930
        from rpath_nnl_lastE[OF rp[unfolded this]]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3931
        show ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3932
      next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3933
        case (Cons x2 xs2)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3934
        from 1(2)[unfolded Cons1[unfolded this]]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3935
        have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3936
        from rpath_edges_on[OF this]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3937
        have eds: "edges_on (Th th' # x1 # x2 # xs2) \<subseteq> RAG s" .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3938
        have "(Th th', x1) \<in> edges_on (Th th' # x1 # x2 # xs2)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3939
            by (simp add: edges_on_unfold)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3940
        with eds have rg1: "(Th th', x1) \<in> RAG s" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3941
        then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3942
        have "(x1, x2) \<in> edges_on (Th th' # x1 # x2 # xs2)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3943
            by (simp add: edges_on_unfold)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3944
        from this eds
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3945
        have rg2: "(x1, x2) \<in> RAG s" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3946
        from this[unfolded eq_x1] 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3947
        obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3948
        from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3949
        have rt1: "(Th th', Th th1) \<in> tRAG s" by (unfold tRAG_alt_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3950
        from rp have "rpath (RAG s) x2 xs2 (Th th)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3951
           by  (elim rpath_ConsE, simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3952
        from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3953
        show ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3954
        proof(cases "xs2 = []")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3955
          case True
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3956
          from rpath_nilE[OF rp'[unfolded this]]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3957
          have "th1 = th" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3958
          from rt1[unfolded this] show ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3959
        next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3960
          case False
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3961
          from 1(1)[rule_format, OF _ rp' this, unfolded Cons1 Cons]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3962
          have "(Th th1, Th th) \<in> (tRAG s)\<^sup>+" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3963
          with rt1 show ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3964
        qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3965
      qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3966
    qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3967
    hence "th' \<in> ?L" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3968
  } ultimately show ?thesis by blast
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3969
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3970
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3971
lemma tRAG_trancl_eq_Th:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3972
   "{Th th' | th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3973
    {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3974
    using tRAG_trancl_eq by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3975
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3976
lemma dependants_alt_def:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3977
  "dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3978
  by (metis eq_RAG s_dependants_def tRAG_trancl_eq)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3979
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3980
lemma dependants_alt_def1:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3981
  "dependants (s::state) th = {th'. (Th th', Th th) \<in> (RAG s)^+}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3982
  using dependants_alt_def tRAG_trancl_eq by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3983
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3984
context valid_trace
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3985
begin
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3986
lemma count_eq_RAG_plus:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3987
  assumes "cntP s th = cntV s th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3988
  shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3989
proof(rule ccontr)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3990
    assume otherwise: "{th'. (Th th', Th th) \<in> (RAG s)\<^sup>+} \<noteq> {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3991
    then obtain th' where "(Th th', Th th) \<in> (RAG s)^+" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3992
    from tranclD2[OF this]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3993
    obtain z where "z \<in> children (RAG s) (Th th)" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3994
      by (auto simp:children_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3995
    with eq_pv_children[OF assms]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3996
    show False by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3997
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3998
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  3999
lemma eq_pv_dependants:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4000
  assumes eq_pv: "cntP s th = cntV s th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4001
  shows "dependants s th = {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4002
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4003
  from count_eq_RAG_plus[OF assms, folded dependants_alt_def1]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4004
  show ?thesis .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4005
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4006
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4007
end
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4008
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4009
lemma eq_dependants: "dependants (wq s) = dependants s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4010
  by (simp add: s_dependants_abv wq_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4011
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4012
context valid_trace
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4013
begin
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4014
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4015
lemma count_eq_tRAG_plus:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4016
  assumes "cntP s th = cntV s th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4017
  shows "{th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4018
  using assms eq_pv_dependants dependants_alt_def eq_dependants by auto 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4019
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4020
lemma count_eq_RAG_plus_Th:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4021
  assumes "cntP s th = cntV s th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4022
  shows "{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4023
  using count_eq_RAG_plus[OF assms] by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4024
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4025
lemma count_eq_tRAG_plus_Th:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4026
  assumes "cntP s th = cntV s th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4027
  shows "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4028
   using count_eq_tRAG_plus[OF assms] by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4029
end
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4030
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4031
lemma inj_the_preced: 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4032
  "inj_on (the_preced s) (threads s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4033
  by (metis inj_onI preced_unique the_preced_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4034
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4035
lemma tRAG_Field:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4036
  "Field (tRAG s) \<subseteq> Field (RAG s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4037
  by (unfold tRAG_alt_def Field_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4038
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4039
lemma tRAG_ancestorsE:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4040
  assumes "x \<in> ancestors (tRAG s) u"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4041
  obtains th where "x = Th th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4042
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4043
  from assms have "(u, x) \<in> (tRAG s)^+" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4044
      by (unfold ancestors_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4045
  from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4046
  then obtain th where "x = Th th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4047
    by (unfold tRAG_alt_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4048
  from that[OF this] show ?thesis .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4049
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4050
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4051
lemma tRAG_mono:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4052
  assumes "RAG s' \<subseteq> RAG s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4053
  shows "tRAG s' \<subseteq> tRAG s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4054
  using assms 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4055
  by (unfold tRAG_alt_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4056
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4057
lemma holding_next_thI:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4058
  assumes "holding s th cs"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4059
  and "length (wq s cs) > 1"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4060
  obtains th' where "next_th s th cs th'"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4061
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4062
  from assms(1)[folded holding_eq, unfolded cs_holding_def]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4063
  have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4064
    by (unfold s_holding_def, fold wq_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4065
  then obtain rest where h1: "wq s cs = th#rest" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4066
    by (cases "wq s cs", auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4067
  with assms(2) have h2: "rest \<noteq> []" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4068
  let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4069
  have "next_th s th cs ?th'" using  h1(1) h2 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4070
    by (unfold next_th_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4071
  from that[OF this] show ?thesis .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4072
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4073
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4074
lemma RAG_tRAG_transfer:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4075
  assumes "vt s'"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4076
  assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4077
  and "(Cs cs, Th th'') \<in> RAG s'"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4078
  shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4079
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4080
  interpret vt_s': valid_trace "s'" using assms(1)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4081
    by (unfold_locales, simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4082
  { fix n1 n2
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4083
    assume "(n1, n2) \<in> ?L"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4084
    from this[unfolded tRAG_alt_def]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4085
    obtain th1 th2 cs' where 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4086
      h: "n1 = Th th1" "n2 = Th th2" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4087
         "(Th th1, Cs cs') \<in> RAG s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4088
         "(Cs cs', Th th2) \<in> RAG s" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4089
    from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \<in> RAG s'" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4090
    from h(3) and assms(2) 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4091
    have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4092
          (Th th1, Cs cs') \<in> RAG s'" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4093
    hence "(n1, n2) \<in> ?R"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4094
    proof
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4095
      assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4096
      hence eq_th1: "th1 = th" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4097
      moreover have "th2 = th''"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4098
      proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4099
        from h1 have "cs' = cs" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4100
        from assms(3) cs_in[unfolded this]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4101
        show ?thesis using vt_s'.unique_RAG by auto 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4102
      qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4103
      ultimately show ?thesis using h(1,2) by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4104
    next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4105
      assume "(Th th1, Cs cs') \<in> RAG s'"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4106
      with cs_in have "(Th th1, Th th2) \<in> tRAG s'"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4107
        by (unfold tRAG_alt_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4108
      from this[folded h(1, 2)] show ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4109
    qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4110
  } moreover {
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4111
    fix n1 n2
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4112
    assume "(n1, n2) \<in> ?R"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4113
    hence "(n1, n2) \<in>tRAG s' \<or> (n1, n2) = (Th th, Th th'')" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4114
    hence "(n1, n2) \<in> ?L" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4115
    proof
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4116
      assume "(n1, n2) \<in> tRAG s'"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4117
      moreover have "... \<subseteq> ?L"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4118
      proof(rule tRAG_mono)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4119
        show "RAG s' \<subseteq> RAG s" by (unfold assms(2), auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4120
      qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4121
      ultimately show ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4122
    next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4123
      assume eq_n: "(n1, n2) = (Th th, Th th'')"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4124
      from assms(2, 3) have "(Cs cs, Th th'') \<in> RAG s" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4125
      moreover have "(Th th, Cs cs) \<in> RAG s" using assms(2) by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4126
      ultimately show ?thesis 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4127
        by (unfold eq_n tRAG_alt_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4128
    qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4129
  } ultimately show ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4130
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4131
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4132
context valid_trace
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4133
begin
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4134
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4135
lemmas RAG_tRAG_transfer = RAG_tRAG_transfer[OF vt]
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  4136
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  4137
end
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 79
diff changeset
  4138
90
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4139
lemma tRAG_subtree_eq: 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4140
   "(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th'  \<in> (subtree (RAG s) (Th th))}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4141
   (is "?L = ?R")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4142
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4143
  { fix n 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4144
    assume h: "n \<in> ?L"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4145
    hence "n \<in> ?R"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4146
    by (smt mem_Collect_eq subsetCE subtree_def subtree_nodeE tRAG_subtree_RAG) 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4147
  } moreover {
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4148
    fix n
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4149
    assume "n \<in> ?R"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4150
    then obtain th' where h: "n = Th th'" "(Th th', Th th) \<in> (RAG s)^*"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4151
      by (auto simp:subtree_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4152
    from rtranclD[OF this(2)]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4153
    have "n \<in> ?L"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4154
    proof
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4155
      assume "Th th' \<noteq> Th th \<and> (Th th', Th th) \<in> (RAG s)\<^sup>+"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4156
      with h have "n \<in> {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4157
      thus ?thesis using subtree_def tRAG_trancl_eq by fastforce
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4158
    qed (insert h, auto simp:subtree_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4159
  } ultimately show ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4160
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4161
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4162
lemma threads_set_eq: 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4163
   "the_thread ` (subtree (tRAG s) (Th th)) = 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4164
                  {th'. Th th' \<in> (subtree (RAG s) (Th th))}" (is "?L = ?R")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4165
   by (auto intro:rev_image_eqI simp:tRAG_subtree_eq)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4166
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4167
lemma cp_alt_def1: 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4168
  "cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4169
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4170
  have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) =
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4171
       ((the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th))"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4172
       by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4173
  thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4174
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4175
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4176
lemma cp_gen_def_cond: 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4177
  assumes "x = Th th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4178
  shows "cp s th = cp_gen s (Th th)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4179
by (unfold cp_alt_def1 cp_gen_def, simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4180
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4181
lemma cp_gen_over_set:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4182
  assumes "\<forall> x \<in> A. \<exists> th. x = Th th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4183
  shows "cp_gen s ` A = (cp s \<circ> the_thread) ` A"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4184
proof(rule f_image_eq)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4185
  fix a
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4186
  assume "a \<in> A"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4187
  from assms[rule_format, OF this]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4188
  obtain th where eq_a: "a = Th th" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4189
  show "cp_gen s a = (cp s \<circ> the_thread) a"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4190
    by  (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4191
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4192
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4193
context valid_trace
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4194
begin
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4195
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4196
lemma subtree_tRAG_thread:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4197
  assumes "th \<in> threads s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4198
  shows "subtree (tRAG s) (Th th) \<subseteq> Th ` threads s" (is "?L \<subseteq> ?R")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4199
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4200
  have "?L = {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4201
    by (unfold tRAG_subtree_eq, simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4202
  also have "... \<subseteq> ?R"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4203
  proof
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4204
    fix x
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4205
    assume "x \<in> {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4206
    then obtain th' where h: "x = Th th'" "Th th' \<in> subtree (RAG s) (Th th)" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4207
    from this(2)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4208
    show "x \<in> ?R"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4209
    proof(cases rule:subtreeE)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4210
      case 1
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4211
      thus ?thesis by (simp add: assms h(1)) 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4212
    next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4213
      case 2
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4214
      thus ?thesis by (metis ancestors_Field dm_RAG_threads h(1) image_eqI) 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4215
    qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4216
  qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4217
  finally show ?thesis .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4218
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4219
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4220
lemma readys_root:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4221
  assumes "th \<in> readys s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4222
  shows "root (RAG s) (Th th)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4223
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4224
  { fix x
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4225
    assume "x \<in> ancestors (RAG s) (Th th)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4226
    hence h: "(Th th, x) \<in> (RAG s)^+" by (auto simp:ancestors_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4227
    from tranclD[OF this]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4228
    obtain z where "(Th th, z) \<in> RAG s" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4229
    with assms(1) have False
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4230
         apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4231
         by (fold wq_def, blast)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4232
  } thus ?thesis by (unfold root_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4233
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4234
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4235
lemma readys_in_no_subtree:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4236
  assumes "th \<in> readys s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4237
  and "th' \<noteq> th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4238
  shows "Th th \<notin> subtree (RAG s) (Th th')" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4239
proof
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4240
   assume "Th th \<in> subtree (RAG s) (Th th')"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4241
   thus False
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4242
   proof(cases rule:subtreeE)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4243
      case 1
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4244
      with assms show ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4245
   next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4246
      case 2
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4247
      with readys_root[OF assms(1)]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4248
      show ?thesis by (auto simp:root_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4249
   qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4250
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4251
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4252
lemma not_in_thread_isolated:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4253
  assumes "th \<notin> threads s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4254
  shows "(Th th) \<notin> Field (RAG s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4255
proof
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4256
  assume "(Th th) \<in> Field (RAG s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4257
  with dm_RAG_threads and rg_RAG_threads assms
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4258
  show False by (unfold Field_def, blast)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4259
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4260
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4261
end
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4262
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4263
definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4264
  where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4265
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4266
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4267
lemma detached_test:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4268
  shows "detached s th = (Th th \<notin> Field (RAG s))"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4269
apply(simp add: detached_def Field_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4270
apply(simp add: s_RAG_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4271
apply(simp add: s_holding_abv s_waiting_abv)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4272
apply(simp add: Domain_iff Range_iff)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4273
apply(simp add: wq_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4274
apply(auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4275
done
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4276
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4277
context valid_trace
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4278
begin
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4279
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4280
lemma detached_intro:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4281
  assumes eq_pv: "cntP s th = cntV s th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4282
  shows "detached s th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4283
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4284
  from eq_pv cnp_cnv_cncs
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4285
  have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:pvD_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4286
  thus ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4287
  proof
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4288
    assume "th \<notin> threads s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4289
    with rg_RAG_threads dm_RAG_threads
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4290
    show ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4291
      by (auto simp add: detached_def s_RAG_def s_waiting_abv 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4292
              s_holding_abv wq_def Domain_iff Range_iff)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4293
  next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4294
    assume "th \<in> readys s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4295
    moreover have "Th th \<notin> Range (RAG s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4296
    proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4297
      from eq_pv_children[OF assms]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4298
      have "children (RAG s) (Th th) = {}" .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4299
      thus ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4300
      by (unfold children_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4301
    qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4302
    ultimately show ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4303
      by (auto simp add: detached_def s_RAG_def s_waiting_abv 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4304
              s_holding_abv wq_def readys_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4305
  qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4306
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4307
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4308
lemma detached_elim:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4309
  assumes dtc: "detached s th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4310
  shows "cntP s th = cntV s th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4311
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4312
  have cncs_z: "cntCS s th = 0"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4313
  proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4314
    from dtc have "holdents s th = {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4315
      unfolding detached_def holdents_test s_RAG_def
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4316
      by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4317
    thus ?thesis by (auto simp:cntCS_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4318
  qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4319
  show ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4320
  proof(cases "th \<in> threads s")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4321
    case True
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4322
    with dtc 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4323
    have "th \<in> readys s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4324
      by (unfold readys_def detached_def Field_def Domain_def Range_def, 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4325
           auto simp:waiting_eq s_RAG_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4326
    with cncs_z  show ?thesis using cnp_cnv_cncs by (simp add:pvD_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4327
  next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4328
    case False
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4329
    with cncs_z and cnp_cnv_cncs show ?thesis by (simp add:pvD_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4330
  qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4331
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4332
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4333
lemma detached_eq:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4334
  shows "(detached s th) = (cntP s th = cntV s th)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4335
  by (insert vt, auto intro:detached_intro detached_elim)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4336
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4337
end
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4338
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4339
context valid_trace
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4340
begin
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4341
(* ddd *)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4342
lemma cp_gen_rec:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4343
  assumes "x = Th th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4344
  shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4345
proof(cases "children (tRAG s) x = {}")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4346
  case True
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4347
  show ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4348
    by (unfold True cp_gen_def subtree_children, simp add:assms)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4349
next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4350
  case False
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4351
  hence [simp]: "children (tRAG s) x \<noteq> {}" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4352
  note fsbttRAGs.finite_subtree[simp]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4353
  have [simp]: "finite (children (tRAG s) x)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4354
     by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4355
            rule children_subtree)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4356
  { fix r x
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4357
    have "subtree r x \<noteq> {}" by (auto simp:subtree_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4358
  } note this[simp]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4359
  have [simp]: "\<exists>x\<in>children (tRAG s) x. subtree (tRAG s) x \<noteq> {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4360
  proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4361
    from False obtain q where "q \<in> children (tRAG s) x" by blast
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4362
    moreover have "subtree (tRAG s) q \<noteq> {}" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4363
    ultimately show ?thesis by blast
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4364
  qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4365
  have h: "Max ((the_preced s \<circ> the_thread) `
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4366
                ({x} \<union> \<Union>(subtree (tRAG s) ` children (tRAG s) x))) =
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4367
        Max ({the_preced s th} \<union> cp_gen s ` children (tRAG s) x)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4368
                     (is "?L = ?R")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4369
  proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4370
    let "Max (?f ` (?A \<union> \<Union> (?g ` ?B)))" = ?L
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4371
    let "Max (_ \<union> (?h ` ?B))" = ?R
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4372
    let ?L1 = "?f ` \<Union>(?g ` ?B)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4373
    have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4374
    proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4375
      have "?L1 = ?f ` (\<Union> x \<in> ?B.(?g x))" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4376
      also have "... =  (\<Union> x \<in> ?B. ?f ` (?g x))" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4377
      finally have "Max ?L1 = Max ..." by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4378
      also have "... = Max (Max ` (\<lambda>x. ?f ` subtree (tRAG s) x) ` ?B)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4379
        by (subst Max_UNION, simp+)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4380
      also have "... = Max (cp_gen s ` children (tRAG s) x)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4381
          by (unfold image_comp cp_gen_alt_def, simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4382
      finally show ?thesis .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4383
    qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4384
    show ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4385
    proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4386
      have "?L = Max (?f ` ?A \<union> ?L1)" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4387
      also have "... = max (the_preced s (the_thread x)) (Max ?L1)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4388
            by (subst Max_Un, simp+)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4389
      also have "... = max (?f x) (Max (?h ` ?B))"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4390
        by (unfold eq_Max_L1, simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4391
      also have "... =?R"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4392
        by (rule max_Max_eq, (simp)+, unfold assms, simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4393
      finally show ?thesis .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4394
    qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4395
  qed  thus ?thesis 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4396
          by (fold h subtree_children, unfold cp_gen_def, simp) 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4397
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4398
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4399
lemma cp_rec:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4400
  "cp s th = Max ({the_preced s th} \<union> 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4401
                     (cp s o the_thread) ` children (tRAG s) (Th th))"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4402
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4403
  have "Th th = Th th" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4404
  note h =  cp_gen_def_cond[OF this] cp_gen_rec[OF this]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4405
  show ?thesis 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4406
  proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4407
    have "cp_gen s ` children (tRAG s) (Th th) = 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4408
                (cp s \<circ> the_thread) ` children (tRAG s) (Th th)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4409
    proof(rule cp_gen_over_set)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4410
      show " \<forall>x\<in>children (tRAG s) (Th th). \<exists>th. x = Th th"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4411
        by (unfold tRAG_alt_def, auto simp:children_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4412
    qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4413
    thus ?thesis by (subst (1) h(1), unfold h(2), simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4414
  qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4415
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4416
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4417
lemma next_th_holding:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4418
  assumes nxt: "next_th s th cs th'"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4419
  shows "holding (wq s) th cs"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4420
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4421
  from nxt[unfolded next_th_def]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4422
  obtain rest where h: "wq s cs = th # rest"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4423
                       "rest \<noteq> []" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4424
                       "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4425
  thus ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4426
    by (unfold cs_holding_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4427
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4428
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4429
lemma next_th_waiting:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4430
  assumes nxt: "next_th s th cs th'"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4431
  shows "waiting (wq s) th' cs"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4432
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4433
  from nxt[unfolded next_th_def]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4434
  obtain rest where h: "wq s cs = th # rest"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4435
                       "rest \<noteq> []" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4436
                       "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4437
  from wq_distinct[of cs, unfolded h]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4438
  have dst: "distinct (th # rest)" .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4439
  have in_rest: "th' \<in> set rest"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4440
  proof(unfold h, rule someI2)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4441
    show "distinct rest \<and> set rest = set rest" using dst by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4442
  next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4443
    fix x assume "distinct x \<and> set x = set rest"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4444
    with h(2)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4445
    show "hd x \<in> set (rest)" by (cases x, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4446
  qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4447
  hence "th' \<in> set (wq s cs)" by (unfold h(1), auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4448
  moreover have "th' \<noteq> hd (wq s cs)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4449
    by (unfold h(1), insert in_rest dst, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4450
  ultimately show ?thesis by (auto simp:cs_waiting_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4451
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4452
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4453
lemma next_th_RAG:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4454
  assumes nxt: "next_th (s::event list) th cs th'"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4455
  shows "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4456
  using vt assms next_th_holding next_th_waiting
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4457
  by (unfold s_RAG_def, simp)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4458
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4459
end
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4460
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4461
lemma next_th_unique: 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4462
  assumes nt1: "next_th s th cs th1"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4463
  and nt2: "next_th s th cs th2"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4464
  shows "th1 = th2"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4465
using assms by (unfold next_th_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4466
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4467
context valid_trace
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4468
begin
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4469
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4470
thm th_chain_to_ready
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4471
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4472
find_theorems subtree Th RAG
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4473
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4474
lemma "(threads s) = (\<Union> th \<in> readys s. {th'. Th th' \<in> subtree (RAG s) (Th th)})"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4475
    (is "?L = ?R")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4476
proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4477
  { fix th1
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4478
    assume "th1 \<in> ?L"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4479
    from th_chain_to_ready[OF this]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4480
    have "th1 \<in> readys s \<or> (\<exists>th'a. th'a \<in> readys s \<and> (Th th1, Th th'a) \<in> (RAG s)\<^sup>+)" .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4481
    hence "th1 \<in> ?R"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4482
    proof
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4483
      assume "th1 \<in> readys s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4484
      thus ?thesis by (auto simp:subtree_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4485
    next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4486
      assume "\<exists>th'a. th'a \<in> readys s \<and> (Th th1, Th th'a) \<in> (RAG s)\<^sup>+"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4487
      thus ?thesis 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4488
    qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4489
  } moreover 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4490
  { fix th'
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4491
    assume "th' \<in> ?R"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4492
    have "th' \<in> ?L" sorry
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4493
  } ultimately show ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4494
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4495
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4496
lemma max_cp_readys_threads_pre: (* ccc *)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4497
  assumes np: "threads s \<noteq> {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4498
  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4499
proof(unfold max_cp_eq)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4500
  show "Max (cp s ` readys s) = Max (the_preced s ` threads s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4501
  proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4502
    let ?p = "Max (the_preced s ` threads s)" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4503
    let ?f = "the_preced s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4504
    have "?p \<in> (?f ` threads s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4505
    proof(rule Max_in)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4506
      from finite_threads show "finite (?f ` threads s)" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4507
    next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4508
      from np show "?f ` threads s \<noteq> {}" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4509
    qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4510
    then obtain tm where tm_max: "?f tm = ?p" and tm_in: "tm \<in> threads s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4511
      by (auto simp:Image_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4512
    from th_chain_to_ready [OF tm_in]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4513
    have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+)" .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4514
    thus ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4515
    proof
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4516
      assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+ "
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4517
      then obtain th' where th'_in: "th' \<in> readys s" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4518
        and tm_chain:"(Th tm, Th th') \<in> (RAG s)\<^sup>+" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4519
      have "cp s th' = ?f tm"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4520
      proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4521
        from dependants_threads finite_threads
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4522
        show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th'))" 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4523
          by (auto intro:finite_subset)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4524
      next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4525
        fix p assume p_in: "p \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4526
        from tm_max have " preced tm s = Max ((\<lambda>th. preced th s) ` threads s)" .
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4527
        moreover have "p \<le> \<dots>"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4528
        proof(rule Max_ge)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4529
          from finite_threads
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4530
          show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4531
        next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4532
          from p_in and th'_in and dependants_threads[of th']
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4533
          show "p \<in> (\<lambda>th. preced th s) ` threads s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4534
            by (auto simp:readys_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4535
        qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4536
        ultimately show "p \<le> preced tm s" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4537
      next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4538
        show "preced tm s \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4539
        proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4540
          from tm_chain
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4541
          have "tm \<in> dependants (wq s) th'"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4542
            by (unfold cs_dependants_def s_RAG_def cs_RAG_def, auto)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4543
          thus ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4544
        qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4545
      qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4546
      with tm_max
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4547
      have h: "cp s th' = Max ((\<lambda>th. preced th s) ` threads s)" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4548
      show ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4549
      proof (fold h, rule Max_eqI)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4550
        fix q 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4551
        assume "q \<in> cp s ` readys s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4552
        then obtain th1 where th1_in: "th1 \<in> readys s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4553
          and eq_q: "q = cp s th1" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4554
        show "q \<le> cp s th'"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4555
          apply (unfold h eq_q)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4556
          apply (unfold cp_eq_cpreced cpreced_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4557
          apply (rule Max_mono)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4558
        proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4559
          from dependants_threads [of th1] th1_in
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4560
          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<subseteq> 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4561
                 (\<lambda>th. preced th s) ` threads s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4562
            by (auto simp:readys_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4563
        next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4564
          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4565
        next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4566
          from finite_threads 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4567
          show " finite ((\<lambda>th. preced th s) ` threads s)" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4568
        qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4569
      next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4570
        from finite_threads
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4571
        show "finite (cp s ` readys s)" by (auto simp:readys_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4572
      next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4573
        from th'_in
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4574
        show "cp s th' \<in> cp s ` readys s" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4575
      qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4576
    next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4577
      assume tm_ready: "tm \<in> readys s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4578
      show ?thesis
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4579
      proof(fold tm_max)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4580
        have cp_eq_p: "cp s tm = preced tm s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4581
        proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4582
          fix y 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4583
          assume hy: "y \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4584
          show "y \<le> preced tm s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4585
          proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4586
            { fix y'
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4587
              assume hy' : "y' \<in> ((\<lambda>th. preced th s) ` dependants (wq s) tm)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4588
              have "y' \<le> preced tm s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4589
              proof(unfold tm_max, rule Max_ge)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4590
                from hy' dependants_threads[of tm]
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4591
                show "y' \<in> (\<lambda>th. preced th s) ` threads s" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4592
              next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4593
                from finite_threads
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4594
                show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4595
              qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4596
            } with hy show ?thesis by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4597
          qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4598
        next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4599
          from dependants_threads[of tm] finite_threads
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4600
          show "finite ((\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm))"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4601
            by (auto intro:finite_subset)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4602
        next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4603
          show "preced tm s \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4604
            by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4605
        qed 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4606
        moreover have "Max (cp s ` readys s) = cp s tm"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4607
        proof(rule Max_eqI)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4608
          from tm_ready show "cp s tm \<in> cp s ` readys s" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4609
        next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4610
          from finite_threads
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4611
          show "finite (cp s ` readys s)" by (auto simp:readys_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4612
        next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4613
          fix y assume "y \<in> cp s ` readys s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4614
          then obtain th1 where th1_readys: "th1 \<in> readys s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4615
            and h: "y = cp s th1" by auto
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4616
          show "y \<le> cp s tm"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4617
            apply(unfold cp_eq_p h)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4618
            apply(unfold cp_eq_cpreced cpreced_def tm_max, rule Max_mono)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4619
          proof -
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4620
            from finite_threads
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4621
            show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4622
          next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4623
            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4624
              by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4625
          next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4626
            from dependants_threads[of th1] th1_readys
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4627
            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4628
                    \<subseteq> (\<lambda>th. preced th s) ` threads s"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4629
              by (auto simp:readys_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4630
          qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4631
        qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4632
        ultimately show " Max (cp s ` readys s) = preced tm s" by simp
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4633
      qed 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4634
    qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4635
  qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4636
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4637
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4638
text {* (* ccc *) \noindent
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4639
  Since the current precedence of the threads in ready queue will always be boosted,
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4640
  there must be one inside it has the maximum precedence of the whole system. 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4641
*}
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4642
lemma max_cp_readys_threads:
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4643
  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4644
proof(cases "threads s = {}")
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4645
  case True
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4646
  thus ?thesis 
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4647
    by (auto simp:readys_def)
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4648
next
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4649
  case False
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4650
  show ?thesis by (rule max_cp_readys_threads_pre[OF False])
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4651
qed
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4652
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4653
end
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4654
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4655
end
ed938e2246b9 Retrofiting of:
zhangx
parents: 80
diff changeset
  4656