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