| author | wenzelm | 
| Mon, 06 Feb 2023 12:58:45 +0100 | |
| changeset 77206 | 6784eaef7d0c | 
| parent 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 37936 | 1  | 
(* Title: HOL/UNITY/Comp/Priority.thy  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
2  | 
Author: Sidi O Ehmety, Cambridge University Computer Laboratory  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
3  | 
Copyright 2001 University of Cambridge  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
4  | 
*)  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
5  | 
|
| 63146 | 6  | 
section\<open>The priority system\<close>  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
7  | 
|
| 16417 | 8  | 
theory Priority imports PriorityAux begin  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
9  | 
|
| 63146 | 10  | 
text\<open>From Charpentier and Chandy,  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
11  | 
Examples of Program Composition Illustrating the Use of Universal Properties  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
12  | 
In J. Rolim (editor), Parallel and Distributed Processing,  | 
| 63146 | 13  | 
Spriner LNCS 1586 (1999), pages 1215-1227.\<close>  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
14  | 
|
| 42463 | 15  | 
type_synonym state = "(vertex*vertex)set"  | 
16  | 
type_synonym command = "vertex=>(state*state)set"  | 
|
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
18  | 
consts  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
19  | 
init :: "(vertex*vertex)set"  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
63146 
diff
changeset
 | 
20  | 
\<comment> \<open>the initial state\<close>  | 
| 15274 | 21  | 
|
| 63146 | 22  | 
text\<open>Following the definitions given in section 4.4\<close>  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
23  | 
|
| 36866 | 24  | 
definition highest :: "[vertex, (vertex*vertex)set]=>bool"  | 
| 61941 | 25  | 
  where "highest i r \<longleftrightarrow> A i r = {}"
 | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
63146 
diff
changeset
 | 
26  | 
\<comment> \<open>i has highest priority in r\<close>  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
27  | 
|
| 36866 | 28  | 
definition lowest :: "[vertex, (vertex*vertex)set]=>bool"  | 
| 61941 | 29  | 
  where "lowest i r \<longleftrightarrow> R i r = {}"
 | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
63146 
diff
changeset
 | 
30  | 
\<comment> \<open>i has lowest priority in r\<close>  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
31  | 
|
| 36866 | 32  | 
definition act :: command  | 
33  | 
  where "act i = {(s, s'). s'=reverse i s & highest i s}"
 | 
|
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
34  | 
|
| 36866 | 35  | 
definition Component :: "vertex=>state program"  | 
36  | 
  where "Component i = mk_total_program({init}, {act i}, UNIV)"
 | 
|
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
63146 
diff
changeset
 | 
37  | 
\<comment> \<open>All components start with the same initial state\<close>  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
38  | 
|
| 15274 | 39  | 
|
| 63146 | 40  | 
text\<open>Some Abbreviations\<close>  | 
| 36866 | 41  | 
definition Highest :: "vertex=>state set"  | 
42  | 
  where "Highest i = {s. highest i s}"
 | 
|
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
43  | 
|
| 36866 | 44  | 
definition Lowest :: "vertex=>state set"  | 
45  | 
  where "Lowest i = {s. lowest i s}"
 | 
|
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
46  | 
|
| 36866 | 47  | 
definition Acyclic :: "state set"  | 
48  | 
  where "Acyclic = {s. acyclic s}"
 | 
|
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
49  | 
|
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
50  | 
|
| 36866 | 51  | 
definition Maximal :: "state set"  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
63146 
diff
changeset
 | 
52  | 
\<comment> \<open>Every ``above'' set has a maximal vertex\<close>  | 
| 36866 | 53  | 
  where "Maximal = (\<Inter>i. {s. ~highest i s-->(\<exists>j \<in> above i  s. highest j s)})"
 | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
54  | 
|
| 36866 | 55  | 
definition Maximal' :: "state set"  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
63146 
diff
changeset
 | 
56  | 
\<comment> \<open>Maximal vertex: equivalent definition\<close>  | 
| 36866 | 57  | 
  where "Maximal' = (\<Inter>i. Highest i Un (\<Union>j. {s. j \<in> above i s} Int Highest j))"
 | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
59  | 
|
| 36866 | 60  | 
definition Safety :: "state set"  | 
61  | 
  where "Safety = (\<Inter>i. {s. highest i s --> (\<forall>j \<in> neighbors i s. ~highest j s)})"
 | 
|
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
62  | 
|
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
63  | 
|
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
64  | 
(* Composition of a finite set of component;  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
65  | 
the vertex 'UNIV' is finite by assumption *)  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
66  | 
|
| 36866 | 67  | 
definition system :: "state program"  | 
| 60773 | 68  | 
where "system = (\<Squnion>i. Component i)"  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
69  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
70  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
71  | 
declare highest_def [simp] lowest_def [simp]  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
72  | 
declare Highest_def [THEN def_set_simp, simp]  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
73  | 
and Lowest_def [THEN def_set_simp, simp]  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
74  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
75  | 
declare Component_def [THEN def_prg_Init, simp]  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
76  | 
declare act_def [THEN def_act_simp, simp]  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
77  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
78  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
79  | 
|
| 63146 | 80  | 
subsection\<open>Component correctness proofs\<close>  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
81  | 
|
| 63146 | 82  | 
text\<open>neighbors is stable\<close>  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
83  | 
lemma Component_neighbors_stable: "Component i \<in> stable {s. neighbors k s = n}"
 | 
| 
16184
 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 
paulson 
parents: 
15274 
diff
changeset
 | 
84  | 
by (simp add: Component_def, safety, auto)  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
85  | 
|
| 63146 | 86  | 
text\<open>property 4\<close>  | 
| 67613 | 87  | 
lemma Component_waits_priority: "Component i \<in> {s. ((i,j) \<in> s) = b} \<inter> (- Highest i) co {s. ((i,j) \<in> s)=b}"
 | 
| 
16184
 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 
paulson 
parents: 
15274 
diff
changeset
 | 
88  | 
by (simp add: Component_def, safety)  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
89  | 
|
| 63146 | 90  | 
text\<open>property 5: charpentier and Chandy mistakenly express it as  | 
91  | 
'transient Highest i'. Consider the case where i has neighbors\<close>  | 
|
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
92  | 
lemma Component_yields_priority:  | 
| 67613 | 93  | 
 "Component i \<in> {s. neighbors i s \<noteq> {}} Int Highest i  
 | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
94  | 
ensures - Highest i"  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
95  | 
apply (simp add: Component_def)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
96  | 
apply (ensures_tac "act i", blast+)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
97  | 
done  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
98  | 
|
| 63146 | 99  | 
text\<open>or better\<close>  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
100  | 
lemma Component_yields_priority': "Component i \<in> Highest i ensures Lowest i"  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
101  | 
apply (simp add: Component_def)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
102  | 
apply (ensures_tac "act i", blast+)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
103  | 
done  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
104  | 
|
| 63146 | 105  | 
text\<open>property 6: Component doesn't introduce cycle\<close>  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
106  | 
lemma Component_well_behaves: "Component i \<in> Highest i co Highest i Un Lowest i"  | 
| 
16184
 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 
paulson 
parents: 
15274 
diff
changeset
 | 
107  | 
by (simp add: Component_def, safety, fast)  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
108  | 
|
| 63146 | 109  | 
text\<open>property 7: local axiom\<close>  | 
| 67613 | 110  | 
lemma locality: "Component i \<in> stable {s. \<forall>j k. j\<noteq>i & k\<noteq>i--> ((j,k) \<in> s) = b j k}"
 | 
| 
16184
 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 
paulson 
parents: 
15274 
diff
changeset
 | 
111  | 
by (simp add: Component_def, safety)  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
112  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
113  | 
|
| 63146 | 114  | 
subsection\<open>System properties\<close>  | 
115  | 
text\<open>property 8: strictly universal\<close>  | 
|
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
116  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
117  | 
lemma Safety: "system \<in> stable Safety"  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
118  | 
apply (unfold Safety_def)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
119  | 
apply (rule stable_INT)  | 
| 
16184
 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 
paulson 
parents: 
15274 
diff
changeset
 | 
120  | 
apply (simp add: system_def, safety, fast)  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
121  | 
done  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
122  | 
|
| 63146 | 123  | 
text\<open>property 13: universal\<close>  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
124  | 
lemma p13: "system \<in> {s. s = q} co {s. s=q} Un {s. \<exists>i. derive i q s}"
 | 
| 
16184
 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 
paulson 
parents: 
15274 
diff
changeset
 | 
125  | 
by (simp add: system_def Component_def mk_total_program_def totalize_JN, safety, blast)  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
126  | 
|
| 63146 | 127  | 
text\<open>property 14: the 'above set' of a Component that hasn't got  | 
128  | 
priority doesn't increase\<close>  | 
|
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
129  | 
lemma above_not_increase:  | 
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
14087 
diff
changeset
 | 
130  | 
     "system \<in> -Highest i Int {s. j\<notin>above i s} co {s. j\<notin>above i s}"
 | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
14087 
diff
changeset
 | 
131  | 
apply (insert reach_lemma [of concl: j])  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
132  | 
apply (simp add: system_def Component_def mk_total_program_def totalize_JN,  | 
| 
16184
 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 
paulson 
parents: 
15274 
diff
changeset
 | 
133  | 
safety)  | 
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
14087 
diff
changeset
 | 
134  | 
apply (simp add: trancl_converse, blast)  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
135  | 
done  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
136  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
137  | 
lemma above_not_increase':  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
138  | 
     "system \<in> -Highest i Int {s. above i s = x} co {s. above i s <= x}"
 | 
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
14087 
diff
changeset
 | 
139  | 
apply (insert above_not_increase [of i])  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
140  | 
apply (simp add: trancl_converse constrains_def, blast)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
141  | 
done  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
142  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
143  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
144  | 
|
| 63146 | 145  | 
text\<open>p15: universal property: all Components well behave\<close>  | 
| 46911 | 146  | 
lemma system_well_behaves: "system \<in> Highest i co Highest i Un Lowest i"  | 
147  | 
by (simp add: system_def Component_def mk_total_program_def totalize_JN, safety, auto)  | 
|
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
148  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
149  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
150  | 
lemma Acyclic_eq: "Acyclic = (\<Inter>i. {s. i\<notin>above i s})"
 | 
| 46911 | 151  | 
by (auto simp add: Acyclic_def acyclic_def trancl_converse)  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
152  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
153  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
154  | 
lemmas system_co =  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
155  | 
constrains_Un [OF above_not_increase [rule_format] system_well_behaves]  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
156  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
157  | 
lemma Acyclic_stable: "system \<in> stable Acyclic"  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
158  | 
apply (simp add: stable_def Acyclic_eq)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
159  | 
apply (auto intro!: constrains_INT system_co [THEN constrains_weaken]  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
160  | 
simp add: image0_r_iff_image0_trancl trancl_converse)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
161  | 
done  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
162  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
163  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
164  | 
lemma Acyclic_subset_Maximal: "Acyclic <= Maximal"  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
165  | 
apply (unfold Acyclic_def Maximal_def, clarify)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
166  | 
apply (drule above_lemma_b, auto)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
167  | 
done  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
168  | 
|
| 63146 | 169  | 
text\<open>property 17: original one is an invariant\<close>  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
170  | 
lemma Acyclic_Maximal_stable: "system \<in> stable (Acyclic Int Maximal)"  | 
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
14087 
diff
changeset
 | 
171  | 
by (simp add: Acyclic_subset_Maximal [THEN Int_absorb2] Acyclic_stable)  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
172  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
173  | 
|
| 63146 | 174  | 
text\<open>property 5: existential property\<close>  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
175  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
176  | 
lemma Highest_leadsTo_Lowest: "system \<in> Highest i leadsTo Lowest i"  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
177  | 
apply (simp add: system_def Component_def mk_total_program_def totalize_JN)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
178  | 
apply (ensures_tac "act i", auto)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
179  | 
done  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
180  | 
|
| 63146 | 181  | 
text\<open>a lowest i can never be in any abover set\<close>  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
182  | 
lemma Lowest_above_subset: "Lowest i <= (\<Inter>k. {s. i\<notin>above k s})"
 | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
183  | 
by (auto simp add: image0_r_iff_image0_trancl trancl_converse)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
184  | 
|
| 63146 | 185  | 
text\<open>property 18: a simpler proof than the original, one which uses psp\<close>  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
186  | 
lemma Highest_escapes_above: "system \<in> Highest i leadsTo (\<Inter>k. {s. i\<notin>above k s})"
 | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
187  | 
apply (rule leadsTo_weaken_R)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
188  | 
apply (rule_tac [2] Lowest_above_subset)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
189  | 
apply (rule Highest_leadsTo_Lowest)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
190  | 
done  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
191  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
192  | 
lemma Highest_escapes_above':  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
193  | 
     "system \<in> Highest j Int {s. j \<in> above i s} leadsTo {s. j\<notin>above i s}"
 | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
194  | 
by (blast intro: leadsTo_weaken [OF Highest_escapes_above Int_lower1 INT_lower])  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
195  | 
|
| 63146 | 196  | 
subsection\<open>The main result: above set decreases\<close>  | 
| 15274 | 197  | 
|
| 63146 | 198  | 
text\<open>The original proof of the following formula was wrong\<close>  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
199  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
200  | 
lemma Highest_iff_above0: "Highest i = {s. above i s ={}}"
 | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
201  | 
by (auto simp add: image0_trancl_iff_image0_r)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
202  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
203  | 
lemmas above_decreases_lemma =  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
204  | 
psp [THEN leadsTo_weaken, OF Highest_escapes_above' above_not_increase']  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
205  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
206  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
207  | 
lemma above_decreases:  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
208  | 
     "system \<in> (\<Union>j. {s. above i s = x} Int {s. j \<in> above i s} Int Highest j)  
 | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
209  | 
               leadsTo {s. above i s < x}"
 | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
210  | 
apply (rule leadsTo_UN)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
211  | 
apply (rule single_leadsTo_I, clarify)  | 
| 
57492
 
74bf65a1910a
Hypsubst preserves equality hypotheses
 
Thomas Sewell <thomas.sewell@nicta.com.au> 
parents: 
46911 
diff
changeset
 | 
212  | 
apply (rule_tac x = "above i xa" in above_decreases_lemma)  | 
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
14087 
diff
changeset
 | 
213  | 
apply (simp_all (no_asm_use) add: Highest_iff_above0)  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
214  | 
apply blast+  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
215  | 
done  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
216  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
217  | 
(** Just a massage of conditions to have the desired form ***)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
218  | 
lemma Maximal_eq_Maximal': "Maximal = Maximal'"  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
219  | 
by (unfold Maximal_def Maximal'_def Highest_def, blast)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
220  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
221  | 
lemma Acyclic_subset:  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
222  | 
   "x\<noteq>{} ==>  
 | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
223  | 
    Acyclic Int {s. above i s = x} <=  
 | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
224  | 
    (\<Union>j. {s. above i s = x} Int {s. j \<in> above i s} Int Highest j)"
 | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
225  | 
apply (rule_tac B = "Maximal' Int {s. above i s = x}" in subset_trans)
 | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
226  | 
apply (simp (no_asm) add: Maximal_eq_Maximal' [symmetric])  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
227  | 
apply (blast intro: Acyclic_subset_Maximal [THEN subsetD])  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
228  | 
apply (simp (no_asm) del: above_def add: Maximal'_def Highest_iff_above0)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
229  | 
apply blast  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
230  | 
done  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
231  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
232  | 
lemmas above_decreases' = leadsTo_weaken_L [OF above_decreases Acyclic_subset]  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
233  | 
lemmas above_decreases_psp = psp_stable [OF above_decreases' Acyclic_stable]  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
234  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
235  | 
lemma above_decreases_psp':  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
236  | 
"x\<noteq>{}==> system \<in> Acyclic Int {s. above i s = x} leadsTo 
 | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
237  | 
                   Acyclic Int {s. above i s < x}"
 | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
238  | 
by (erule above_decreases_psp [THEN leadsTo_weaken], blast, auto)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
239  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
240  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
241  | 
lemmas finite_psubset_induct = wf_finite_psubset [THEN leadsTo_wf_induct]  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
242  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
243  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
244  | 
lemma Progress: "system \<in> Acyclic leadsTo Highest i"  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
245  | 
apply (rule_tac f = "%s. above i s" in finite_psubset_induct)  | 
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
14087 
diff
changeset
 | 
246  | 
apply (simp del: above_def  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
247  | 
add: Highest_iff_above0 vimage_def finite_psubset_def, clarify)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
248  | 
apply (case_tac "m={}")
 | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
249  | 
apply (rule Int_lower2 [THEN [2] leadsTo_weaken_L])  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
250  | 
apply (force simp add: leadsTo_refl)  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
251  | 
apply (rule_tac A' = "Acyclic Int {x. above i x < m}" in leadsTo_weaken_R)
 | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
252  | 
apply (blast intro: above_decreases_psp')+  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
253  | 
done  | 
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
254  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
255  | 
|
| 63146 | 256  | 
text\<open>We have proved all (relevant) theorems given in the paper. We didn't  | 
| 69597 | 257  | 
assume any thing about the relation \<^term>\<open>r\<close>. It is not necessary that  | 
258  | 
\<^term>\<open>r\<close> be a priority relation as assumed in the original proof. It  | 
|
| 63146 | 259  | 
suffices that we start from a state which is finite and acyclic.\<close>  | 
| 
14087
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
260  | 
|
| 
 
cb07c3948668
Conversion of UNITY/Comp/Priority.thy to a linear Isar script
 
paulson 
parents: 
13812 
diff
changeset
 | 
261  | 
|
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
14087 
diff
changeset
 | 
262  | 
end  |