| author | blanchet | 
| Wed, 18 Aug 2010 17:23:17 +0200 | |
| changeset 38591 | 7400530ab1d0 | 
| parent 35416 | d8d7d1b785af | 
| child 41842 | d8f76db6a207 | 
| permissions | -rw-r--r-- | 
| 13020 | 1  | 
header {* \chapter{Case Study: Single and Multi-Mutator Garbage Collection Algorithms}
 | 
2  | 
||
3  | 
\section {Formalization of the Memory} *}
 | 
|
4  | 
||
| 16417 | 5  | 
theory Graph imports Main begin  | 
| 13020 | 6  | 
|
7  | 
datatype node = Black | White  | 
|
8  | 
||
9  | 
types  | 
|
10  | 
nodes = "node list"  | 
|
11  | 
edge = "nat \<times> nat"  | 
|
12  | 
edges = "edge list"  | 
|
13  | 
||
14  | 
consts Roots :: "nat set"  | 
|
15  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32960 
diff
changeset
 | 
16  | 
definition Proper_Roots :: "nodes \<Rightarrow> bool" where  | 
| 13020 | 17  | 
  "Proper_Roots M \<equiv> Roots\<noteq>{} \<and> Roots \<subseteq> {i. i<length M}"
 | 
18  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32960 
diff
changeset
 | 
19  | 
definition Proper_Edges :: "(nodes \<times> edges) \<Rightarrow> bool" where  | 
| 13020 | 20  | 
"Proper_Edges \<equiv> (\<lambda>(M,E). \<forall>i<length E. fst(E!i)<length M \<and> snd(E!i)<length M)"  | 
21  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32960 
diff
changeset
 | 
22  | 
definition BtoW :: "(edge \<times> nodes) \<Rightarrow> bool" where  | 
| 13020 | 23  | 
"BtoW \<equiv> (\<lambda>(e,M). (M!fst e)=Black \<and> (M!snd e)\<noteq>Black)"  | 
24  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32960 
diff
changeset
 | 
25  | 
definition Blacks :: "nodes \<Rightarrow> nat set" where  | 
| 13020 | 26  | 
  "Blacks M \<equiv> {i. i<length M \<and> M!i=Black}"
 | 
27  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32960 
diff
changeset
 | 
28  | 
definition Reach :: "edges \<Rightarrow> nat set" where  | 
| 13020 | 29  | 
  "Reach E \<equiv> {x. (\<exists>path. 1<length path \<and> path!(length path - 1)\<in>Roots \<and> x=path!0
 | 
30  | 
\<and> (\<forall>i<length path - 1. (\<exists>j<length E. E!j=(path!(i+1), path!i))))  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32642 
diff
changeset
 | 
31  | 
\<or> x\<in>Roots}"  | 
| 13020 | 32  | 
|
33  | 
text{* Reach: the set of reachable nodes is the set of Roots together with the
 | 
|
34  | 
nodes reachable from some Root by a path represented by a list of  | 
|
35  | 
nodes (at least two since we traverse at least one edge), where two  | 
|
36  | 
consecutive nodes correspond to an edge in E. *}  | 
|
37  | 
||
38  | 
subsection {* Proofs about Graphs *}
 | 
|
39  | 
||
40  | 
lemmas Graph_defs= Blacks_def Proper_Roots_def Proper_Edges_def BtoW_def  | 
|
41  | 
declare Graph_defs [simp]  | 
|
42  | 
||
| 
13022
 
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
 
prensani 
parents: 
13020 
diff
changeset
 | 
43  | 
subsubsection{* Graph 1 *}
 | 
| 13020 | 44  | 
|
45  | 
lemma Graph1_aux [rule_format]:  | 
|
46  | 
"\<lbrakk> Roots\<subseteq>Blacks M; \<forall>i<length E. \<not>BtoW(E!i,M)\<rbrakk>  | 
|
47  | 
\<Longrightarrow> 1< length path \<longrightarrow> (path!(length path - 1))\<in>Roots \<longrightarrow>  | 
|
48  | 
(\<forall>i<length path - 1. (\<exists>j. j < length E \<and> E!j=(path!(Suc i), path!i)))  | 
|
49  | 
\<longrightarrow> M!(path!0) = Black"  | 
|
50  | 
apply(induct_tac "path")  | 
|
51  | 
apply force  | 
|
52  | 
apply clarify  | 
|
53  | 
apply simp  | 
|
54  | 
apply(case_tac "list")  | 
|
55  | 
apply force  | 
|
56  | 
apply simp  | 
|
| 13601 | 57  | 
apply(rotate_tac -2)  | 
| 13020 | 58  | 
apply(erule_tac x = "0" in all_dupE)  | 
59  | 
apply simp  | 
|
60  | 
apply clarify  | 
|
61  | 
apply(erule allE , erule (1) notE impE)  | 
|
62  | 
apply simp  | 
|
63  | 
apply(erule mp)  | 
|
64  | 
apply(case_tac "lista")  | 
|
65  | 
apply force  | 
|
66  | 
apply simp  | 
|
67  | 
apply(erule mp)  | 
|
68  | 
apply clarify  | 
|
69  | 
apply(erule_tac x = "Suc i" in allE)  | 
|
70  | 
apply force  | 
|
71  | 
done  | 
|
72  | 
||
73  | 
lemma Graph1:  | 
|
74  | 
"\<lbrakk>Roots\<subseteq>Blacks M; Proper_Edges(M, E); \<forall>i<length E. \<not>BtoW(E!i,M) \<rbrakk>  | 
|
75  | 
\<Longrightarrow> Reach E\<subseteq>Blacks M"  | 
|
76  | 
apply (unfold Reach_def)  | 
|
77  | 
apply simp  | 
|
78  | 
apply clarify  | 
|
79  | 
apply(erule disjE)  | 
|
80  | 
apply clarify  | 
|
81  | 
apply(rule conjI)  | 
|
82  | 
apply(subgoal_tac "0< length path - Suc 0")  | 
|
83  | 
apply(erule allE , erule (1) notE impE)  | 
|
84  | 
apply force  | 
|
85  | 
apply simp  | 
|
86  | 
apply(rule Graph1_aux)  | 
|
87  | 
apply auto  | 
|
88  | 
done  | 
|
89  | 
||
| 
13022
 
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
 
prensani 
parents: 
13020 
diff
changeset
 | 
90  | 
subsubsection{* Graph 2 *}
 | 
| 13020 | 91  | 
|
92  | 
lemma Ex_first_occurrence [rule_format]:  | 
|
93  | 
"P (n::nat) \<longrightarrow> (\<exists>m. P m \<and> (\<forall>i. i<m \<longrightarrow> \<not> P i))";  | 
|
94  | 
apply(rule nat_less_induct)  | 
|
95  | 
apply clarify  | 
|
96  | 
apply(case_tac "\<forall>m. m<n \<longrightarrow> \<not> P m")  | 
|
97  | 
apply auto  | 
|
98  | 
done  | 
|
99  | 
||
100  | 
lemma Compl_lemma: "(n::nat)\<le>l \<Longrightarrow> (\<exists>m. m\<le>l \<and> n=l - m)"  | 
|
101  | 
apply(rule_tac x = "l - n" in exI)  | 
|
102  | 
apply arith  | 
|
103  | 
done  | 
|
104  | 
||
105  | 
lemma Ex_last_occurrence:  | 
|
106  | 
"\<lbrakk>P (n::nat); n\<le>l\<rbrakk> \<Longrightarrow> (\<exists>m. P (l - m) \<and> (\<forall>i. i<m \<longrightarrow> \<not>P (l - i)))"  | 
|
107  | 
apply(drule Compl_lemma)  | 
|
108  | 
apply clarify  | 
|
109  | 
apply(erule Ex_first_occurrence)  | 
|
110  | 
done  | 
|
111  | 
||
112  | 
lemma Graph2:  | 
|
113  | 
"\<lbrakk>T \<in> Reach E; R<length E\<rbrakk> \<Longrightarrow> T \<in> Reach (E[R:=(fst(E!R), T)])"  | 
|
114  | 
apply (unfold Reach_def)  | 
|
115  | 
apply clarify  | 
|
116  | 
apply simp  | 
|
117  | 
apply(case_tac "\<forall>z<length path. fst(E!R)\<noteq>path!z")  | 
|
118  | 
apply(rule_tac x = "path" in exI)  | 
|
119  | 
apply simp  | 
|
120  | 
apply clarify  | 
|
121  | 
apply(erule allE , erule (1) notE impE)  | 
|
122  | 
apply clarify  | 
|
123  | 
apply(rule_tac x = "j" in exI)  | 
|
124  | 
apply(case_tac "j=R")  | 
|
125  | 
apply(erule_tac x = "Suc i" in allE)  | 
|
126  | 
apply simp  | 
|
127  | 
apply (force simp add:nth_list_update)  | 
|
128  | 
apply simp  | 
|
129  | 
apply(erule exE)  | 
|
130  | 
apply(subgoal_tac "z \<le> length path - Suc 0")  | 
|
131  | 
prefer 2 apply arith  | 
|
132  | 
apply(drule_tac P = "\<lambda>m. m<length path \<and> fst(E!R)=path!m" in Ex_last_occurrence)  | 
|
133  | 
apply assumption  | 
|
134  | 
apply clarify  | 
|
135  | 
apply simp  | 
|
136  | 
apply(rule_tac x = "(path!0)#(drop (length path - Suc m) path)" in exI)  | 
|
137  | 
apply simp  | 
|
138  | 
apply(case_tac "length path - (length path - Suc m)")  | 
|
139  | 
apply arith  | 
|
140  | 
apply simp  | 
|
141  | 
apply(subgoal_tac "(length path - Suc m) + nat \<le> length path")  | 
|
142  | 
prefer 2 apply arith  | 
|
143  | 
apply(drule nth_drop)  | 
|
144  | 
apply simp  | 
|
145  | 
apply(subgoal_tac "length path - Suc m + nat = length path - Suc 0")  | 
|
146  | 
prefer 2 apply arith  | 
|
147  | 
apply simp  | 
|
148  | 
apply clarify  | 
|
149  | 
apply(case_tac "i")  | 
|
150  | 
apply(force simp add: nth_list_update)  | 
|
151  | 
apply simp  | 
|
152  | 
apply(subgoal_tac "(length path - Suc m) + nata \<le> length path")  | 
|
153  | 
prefer 2 apply arith  | 
|
154  | 
apply(subgoal_tac "(length path - Suc m) + (Suc nata) \<le> length path")  | 
|
155  | 
prefer 2 apply arith  | 
|
156  | 
apply simp  | 
|
157  | 
apply(erule_tac x = "length path - Suc m + nata" in allE)  | 
|
158  | 
apply simp  | 
|
159  | 
apply clarify  | 
|
160  | 
apply(rule_tac x = "j" in exI)  | 
|
161  | 
apply(case_tac "R=j")  | 
|
162  | 
prefer 2 apply force  | 
|
163  | 
apply simp  | 
|
164  | 
apply(drule_tac t = "path ! (length path - Suc m)" in sym)  | 
|
165  | 
apply simp  | 
|
166  | 
apply(case_tac " length path - Suc 0 < m")  | 
|
167  | 
apply(subgoal_tac "(length path - Suc m)=0")  | 
|
168  | 
prefer 2 apply arith  | 
|
169  | 
apply(simp del: diff_is_0_eq)  | 
|
170  | 
apply(subgoal_tac "Suc nata\<le>nat")  | 
|
171  | 
prefer 2 apply arith  | 
|
172  | 
apply(drule_tac n = "Suc nata" in Compl_lemma)  | 
|
173  | 
apply clarify  | 
|
| 31082 | 174  | 
using [[linarith_split_limit = 0]]  | 
| 13020 | 175  | 
apply force  | 
| 31082 | 176  | 
using [[linarith_split_limit = 9]]  | 
| 13020 | 177  | 
apply(drule leI)  | 
178  | 
apply(subgoal_tac "Suc (length path - Suc m + nata)=(length path - Suc 0) - (m - Suc nata)")  | 
|
179  | 
apply(erule_tac x = "m - (Suc nata)" in allE)  | 
|
180  | 
apply(case_tac "m")  | 
|
181  | 
apply simp  | 
|
182  | 
apply simp  | 
|
| 13601 | 183  | 
apply simp  | 
| 13020 | 184  | 
done  | 
185  | 
||
| 
20432
 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 
webertj 
parents: 
20279 
diff
changeset
 | 
186  | 
|
| 
13022
 
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
 
prensani 
parents: 
13020 
diff
changeset
 | 
187  | 
subsubsection{* Graph 3 *}
 | 
| 13020 | 188  | 
|
| 
32642
 
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
 
haftmann 
parents: 
32621 
diff
changeset
 | 
189  | 
declare min_max.inf_absorb1 [simp] min_max.inf_absorb2 [simp]  | 
| 
 
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
 
haftmann 
parents: 
32621 
diff
changeset
 | 
190  | 
|
| 13020 | 191  | 
lemma Graph3:  | 
192  | 
"\<lbrakk> T\<in>Reach E; R<length E \<rbrakk> \<Longrightarrow> Reach(E[R:=(fst(E!R),T)]) \<subseteq> Reach E"  | 
|
193  | 
apply (unfold Reach_def)  | 
|
194  | 
apply clarify  | 
|
195  | 
apply simp  | 
|
196  | 
apply(case_tac "\<exists>i<length path - 1. (fst(E!R),T)=(path!(Suc i),path!i)")  | 
|
197  | 
--{* the changed edge is part of the path *}
 | 
|
198  | 
apply(erule exE)  | 
|
199  | 
apply(drule_tac P = "\<lambda>i. i<length path - 1 \<and> (fst(E!R),T)=(path!Suc i,path!i)" in Ex_first_occurrence)  | 
|
200  | 
apply clarify  | 
|
201  | 
apply(erule disjE)  | 
|
202  | 
--{* T is NOT a root *}
 | 
|
203  | 
apply clarify  | 
|
204  | 
apply(rule_tac x = "(take m path)@patha" in exI)  | 
|
205  | 
apply(subgoal_tac "\<not>(length path\<le>m)")  | 
|
206  | 
prefer 2 apply arith  | 
|
| 32442 | 207  | 
apply(simp)  | 
| 13020 | 208  | 
apply(rule conjI)  | 
209  | 
apply(subgoal_tac "\<not>(m + length patha - 1 < m)")  | 
|
210  | 
prefer 2 apply arith  | 
|
| 32442 | 211  | 
apply(simp add: nth_append)  | 
| 13020 | 212  | 
apply(rule conjI)  | 
213  | 
apply(case_tac "m")  | 
|
214  | 
apply force  | 
|
215  | 
apply(case_tac "path")  | 
|
216  | 
apply force  | 
|
217  | 
apply force  | 
|
218  | 
apply clarify  | 
|
219  | 
apply(case_tac "Suc i\<le>m")  | 
|
220  | 
apply(erule_tac x = "i" in allE)  | 
|
221  | 
apply simp  | 
|
222  | 
apply clarify  | 
|
223  | 
apply(rule_tac x = "j" in exI)  | 
|
224  | 
apply(case_tac "Suc i<m")  | 
|
| 22230 | 225  | 
apply(simp add: nth_append)  | 
| 13020 | 226  | 
apply(case_tac "R=j")  | 
227  | 
apply(simp add: nth_list_update)  | 
|
228  | 
apply(case_tac "i=m")  | 
|
229  | 
apply force  | 
|
230  | 
apply(erule_tac x = "i" in allE)  | 
|
231  | 
apply force  | 
|
232  | 
apply(force simp add: nth_list_update)  | 
|
| 22230 | 233  | 
apply(simp add: nth_append)  | 
| 13020 | 234  | 
apply(subgoal_tac "i=m - 1")  | 
235  | 
prefer 2 apply arith  | 
|
236  | 
apply(case_tac "R=j")  | 
|
237  | 
apply(erule_tac x = "m - 1" in allE)  | 
|
238  | 
apply(simp add: nth_list_update)  | 
|
239  | 
apply(force simp add: nth_list_update)  | 
|
| 32442 | 240  | 
apply(simp add: nth_append)  | 
| 13020 | 241  | 
apply(rotate_tac -4)  | 
242  | 
apply(erule_tac x = "i - m" in allE)  | 
|
243  | 
apply(subgoal_tac "Suc (i - m)=(Suc i - m)" )  | 
|
244  | 
prefer 2 apply arith  | 
|
245  | 
apply simp  | 
|
246  | 
--{* T is a root *}
 | 
|
247  | 
apply(case_tac "m=0")  | 
|
248  | 
apply force  | 
|
249  | 
apply(rule_tac x = "take (Suc m) path" in exI)  | 
|
250  | 
apply(subgoal_tac "\<not>(length path\<le>Suc m)" )  | 
|
251  | 
prefer 2 apply arith  | 
|
| 32442 | 252  | 
apply clarsimp  | 
| 13020 | 253  | 
apply(erule_tac x = "i" in allE)  | 
254  | 
apply simp  | 
|
255  | 
apply clarify  | 
|
256  | 
apply(case_tac "R=j")  | 
|
257  | 
apply(force simp add: nth_list_update)  | 
|
258  | 
apply(force simp add: nth_list_update)  | 
|
259  | 
--{* the changed edge is not part of the path *}
 | 
|
260  | 
apply(rule_tac x = "path" in exI)  | 
|
261  | 
apply simp  | 
|
262  | 
apply clarify  | 
|
263  | 
apply(erule_tac x = "i" in allE)  | 
|
264  | 
apply clarify  | 
|
265  | 
apply(case_tac "R=j")  | 
|
266  | 
apply(erule_tac x = "i" in allE)  | 
|
267  | 
apply simp  | 
|
268  | 
apply(force simp add: nth_list_update)  | 
|
269  | 
done  | 
|
270  | 
||
| 
13022
 
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
 
prensani 
parents: 
13020 
diff
changeset
 | 
271  | 
subsubsection{* Graph 4 *}
 | 
| 13020 | 272  | 
|
273  | 
lemma Graph4:  | 
|
274  | 
"\<lbrakk>T \<in> Reach E; Roots\<subseteq>Blacks M; I\<le>length E; T<length M; R<length E;  | 
|
275  | 
\<forall>i<I. \<not>BtoW(E!i,M); R<I; M!fst(E!R)=Black; M!T\<noteq>Black\<rbrakk> \<Longrightarrow>  | 
|
276  | 
(\<exists>r. I\<le>r \<and> r<length E \<and> BtoW(E[R:=(fst(E!R),T)]!r,M))"  | 
|
277  | 
apply (unfold Reach_def)  | 
|
278  | 
apply simp  | 
|
279  | 
apply(erule disjE)  | 
|
280  | 
prefer 2 apply force  | 
|
281  | 
apply clarify  | 
|
282  | 
--{* there exist a black node in the path to T *}
 | 
|
283  | 
apply(case_tac "\<exists>m<length path. M!(path!m)=Black")  | 
|
284  | 
apply(erule exE)  | 
|
285  | 
apply(drule_tac P = "\<lambda>m. m<length path \<and> M!(path!m)=Black" in Ex_first_occurrence)  | 
|
286  | 
apply clarify  | 
|
287  | 
apply(case_tac "ma")  | 
|
288  | 
apply force  | 
|
289  | 
apply simp  | 
|
290  | 
apply(case_tac "length path")  | 
|
291  | 
apply force  | 
|
292  | 
apply simp  | 
|
| 13601 | 293  | 
apply(erule_tac P = "\<lambda>i. i < nata \<longrightarrow> ?P i" and x = "nat" in allE)  | 
| 13020 | 294  | 
apply simp  | 
295  | 
apply clarify  | 
|
| 13601 | 296  | 
apply(erule_tac P = "\<lambda>i. i < Suc nat \<longrightarrow> ?P i" and x = "nat" in allE)  | 
| 13020 | 297  | 
apply simp  | 
298  | 
apply(case_tac "j<I")  | 
|
299  | 
apply(erule_tac x = "j" in allE)  | 
|
300  | 
apply force  | 
|
301  | 
apply(rule_tac x = "j" in exI)  | 
|
302  | 
apply(force simp add: nth_list_update)  | 
|
303  | 
apply simp  | 
|
304  | 
apply(rotate_tac -1)  | 
|
305  | 
apply(erule_tac x = "length path - 1" in allE)  | 
|
306  | 
apply(case_tac "length path")  | 
|
307  | 
apply force  | 
|
308  | 
apply force  | 
|
309  | 
done  | 
|
310  | 
||
| 
32642
 
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
 
haftmann 
parents: 
32621 
diff
changeset
 | 
311  | 
declare min_max.inf_absorb1 [simp del] min_max.inf_absorb2 [simp del]  | 
| 
 
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
 
haftmann 
parents: 
32621 
diff
changeset
 | 
312  | 
|
| 
13022
 
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
 
prensani 
parents: 
13020 
diff
changeset
 | 
313  | 
subsubsection {* Graph 5 *}
 | 
| 13020 | 314  | 
|
315  | 
lemma Graph5:  | 
|
316  | 
"\<lbrakk> T \<in> Reach E ; Roots \<subseteq> Blacks M; \<forall>i<R. \<not>BtoW(E!i,M); T<length M;  | 
|
317  | 
R<length E; M!fst(E!R)=Black; M!snd(E!R)=Black; M!T \<noteq> Black\<rbrakk>  | 
|
318  | 
\<Longrightarrow> (\<exists>r. R<r \<and> r<length E \<and> BtoW(E[R:=(fst(E!R),T)]!r,M))"  | 
|
319  | 
apply (unfold Reach_def)  | 
|
320  | 
apply simp  | 
|
321  | 
apply(erule disjE)  | 
|
322  | 
prefer 2 apply force  | 
|
323  | 
apply clarify  | 
|
324  | 
--{* there exist a black node in the path to T*}
 | 
|
325  | 
apply(case_tac "\<exists>m<length path. M!(path!m)=Black")  | 
|
326  | 
apply(erule exE)  | 
|
327  | 
apply(drule_tac P = "\<lambda>m. m<length path \<and> M!(path!m)=Black" in Ex_first_occurrence)  | 
|
328  | 
apply clarify  | 
|
329  | 
apply(case_tac "ma")  | 
|
330  | 
apply force  | 
|
331  | 
apply simp  | 
|
332  | 
apply(case_tac "length path")  | 
|
333  | 
apply force  | 
|
334  | 
apply simp  | 
|
| 13601 | 335  | 
apply(erule_tac P = "\<lambda>i. i < nata \<longrightarrow> ?P i" and x = "nat" in allE)  | 
| 13020 | 336  | 
apply simp  | 
337  | 
apply clarify  | 
|
| 13601 | 338  | 
apply(erule_tac P = "\<lambda>i. i < Suc nat \<longrightarrow> ?P i" and x = "nat" in allE)  | 
| 13020 | 339  | 
apply simp  | 
340  | 
apply(case_tac "j\<le>R")  | 
|
| 
26316
 
9e9e67e33557
removed redundant less_trans, less_linear, le_imp_less_or_eq, le_less_trans, less_le_trans (cf. Orderings.thy);
 
wenzelm 
parents: 
24742 
diff
changeset
 | 
341  | 
apply(drule le_imp_less_or_eq [of _ R])  | 
| 13020 | 342  | 
apply(erule disjE)  | 
343  | 
apply(erule allE , erule (1) notE impE)  | 
|
344  | 
apply force  | 
|
345  | 
apply force  | 
|
346  | 
apply(rule_tac x = "j" in exI)  | 
|
347  | 
apply(force simp add: nth_list_update)  | 
|
348  | 
apply simp  | 
|
349  | 
apply(rotate_tac -1)  | 
|
350  | 
apply(erule_tac x = "length path - 1" in allE)  | 
|
351  | 
apply(case_tac "length path")  | 
|
352  | 
apply force  | 
|
353  | 
apply force  | 
|
354  | 
done  | 
|
355  | 
||
| 
13022
 
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
 
prensani 
parents: 
13020 
diff
changeset
 | 
356  | 
subsubsection {* Other lemmas about graphs *}
 | 
| 13020 | 357  | 
|
358  | 
lemma Graph6:  | 
|
359  | 
"\<lbrakk>Proper_Edges(M,E); R<length E ; T<length M\<rbrakk> \<Longrightarrow> Proper_Edges(M,E[R:=(fst(E!R),T)])"  | 
|
360  | 
apply (unfold Proper_Edges_def)  | 
|
361  | 
apply(force simp add: nth_list_update)  | 
|
362  | 
done  | 
|
363  | 
||
364  | 
lemma Graph7:  | 
|
365  | 
"\<lbrakk>Proper_Edges(M,E)\<rbrakk> \<Longrightarrow> Proper_Edges(M[T:=a],E)"  | 
|
366  | 
apply (unfold Proper_Edges_def)  | 
|
367  | 
apply force  | 
|
368  | 
done  | 
|
369  | 
||
370  | 
lemma Graph8:  | 
|
371  | 
"\<lbrakk>Proper_Roots(M)\<rbrakk> \<Longrightarrow> Proper_Roots(M[T:=a])"  | 
|
372  | 
apply (unfold Proper_Roots_def)  | 
|
373  | 
apply force  | 
|
374  | 
done  | 
|
375  | 
||
376  | 
text{* Some specific lemmata for the verification of garbage collection algorithms. *}
 | 
|
377  | 
||
378  | 
lemma Graph9: "j<length M \<Longrightarrow> Blacks M\<subseteq>Blacks (M[j := Black])"  | 
|
379  | 
apply (unfold Blacks_def)  | 
|
380  | 
apply(force simp add: nth_list_update)  | 
|
381  | 
done  | 
|
382  | 
||
383  | 
lemma Graph10 [rule_format (no_asm)]: "\<forall>i. M!i=a \<longrightarrow>M[i:=a]=M"  | 
|
384  | 
apply(induct_tac "M")  | 
|
385  | 
apply auto  | 
|
386  | 
apply(case_tac "i")  | 
|
387  | 
apply auto  | 
|
388  | 
done  | 
|
389  | 
||
390  | 
lemma Graph11 [rule_format (no_asm)]:  | 
|
391  | 
"\<lbrakk> M!j\<noteq>Black;j<length M\<rbrakk> \<Longrightarrow> Blacks M \<subset> Blacks (M[j := Black])"  | 
|
392  | 
apply (unfold Blacks_def)  | 
|
393  | 
apply(rule psubsetI)  | 
|
394  | 
apply(force simp add: nth_list_update)  | 
|
395  | 
apply safe  | 
|
396  | 
apply(erule_tac c = "j" in equalityCE)  | 
|
397  | 
apply auto  | 
|
398  | 
done  | 
|
399  | 
||
400  | 
lemma Graph12: "\<lbrakk>a\<subseteq>Blacks M;j<length M\<rbrakk> \<Longrightarrow> a\<subseteq>Blacks (M[j := Black])"  | 
|
401  | 
apply (unfold Blacks_def)  | 
|
402  | 
apply(force simp add: nth_list_update)  | 
|
403  | 
done  | 
|
404  | 
||
405  | 
lemma Graph13: "\<lbrakk>a\<subset> Blacks M;j<length M\<rbrakk> \<Longrightarrow> a \<subset> Blacks (M[j := Black])"  | 
|
406  | 
apply (unfold Blacks_def)  | 
|
407  | 
apply(erule psubset_subset_trans)  | 
|
408  | 
apply(force simp add: nth_list_update)  | 
|
409  | 
done  | 
|
410  | 
||
411  | 
declare Graph_defs [simp del]  | 
|
412  | 
||
413  | 
end  |