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