| author | hoelzl | 
| Fri, 22 Mar 2013 10:41:42 +0100 | |
| changeset 51471 | cad22a3cc09c | 
| parent 46911 | 6d2a2f0e904e | 
| child 57283 | 1f133cd8d3eb | 
| permissions | -rw-r--r-- | 
| 37936 | 1 | (* Title: HOL/UNITY/Comp/PriorityAux.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 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 4 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 5 | Auxiliary definitions needed in Priority.thy | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 6 | *) | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 7 | |
| 15274 | 8 | theory PriorityAux | 
| 9 | imports "../UNITY_Main" | |
| 10 | begin | |
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 11 | |
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 12 | typedecl vertex | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 13 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
20625diff
changeset | 14 | definition symcl :: "(vertex*vertex)set=>(vertex*vertex)set" where | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 15 | "symcl r == r \<union> (r^-1)" | 
| 15274 | 16 |     --{* symmetric closure: removes the orientation of a relation*}
 | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 17 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
20625diff
changeset | 18 | definition neighbors :: "[vertex, (vertex*vertex)set]=>vertex set" where | 
| 15274 | 19 |   "neighbors i r == ((r \<union> r^-1)``{i}) - {i}"
 | 
| 20 |     --{* Neighbors of a vertex i *}
 | |
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 21 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
20625diff
changeset | 22 | definition R :: "[vertex, (vertex*vertex)set]=>vertex set" where | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 23 |   "R i r == r``{i}"
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 24 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
20625diff
changeset | 25 | definition A :: "[vertex, (vertex*vertex)set]=>vertex set" where | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 26 |   "A i r == (r^-1)``{i}"
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 27 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
20625diff
changeset | 28 | definition reach :: "[vertex, (vertex*vertex)set]=> vertex set" where | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 29 |   "reach i r == (r^+)``{i}"
 | 
| 15274 | 30 |     --{* reachable and above vertices: the original notation was R* and A* *}
 | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 31 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
20625diff
changeset | 32 | definition above :: "[vertex, (vertex*vertex)set]=> vertex set" where | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 33 |   "above i r == ((r^-1)^+)``{i}"  
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 34 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
20625diff
changeset | 35 | definition reverse :: "[vertex, (vertex*vertex) set]=>(vertex*vertex)set" where | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 36 |   "reverse i r == (r - {(x,y). x=i | y=i} \<inter> r) \<union> ({(x,y). x=i|y=i} \<inter> r)^-1"
 | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 37 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
20625diff
changeset | 38 | definition derive1 :: "[vertex, (vertex*vertex)set, (vertex*vertex)set]=>bool" where | 
| 15274 | 39 |     --{* The original definition *}
 | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 40 | "derive1 i r q == symcl r = symcl q & | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 41 | (\<forall>k k'. k\<noteq>i & k'\<noteq>i -->((k,k'):r) = ((k,k'):q)) & | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 42 |                     A i r = {} & R i q = {}"
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 43 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
20625diff
changeset | 44 | definition derive :: "[vertex, (vertex*vertex)set, (vertex*vertex)set]=>bool" where | 
| 15274 | 45 |     --{* Our alternative definition *}
 | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 46 |   "derive i r q == A i r = {} & (q = reverse i r)"
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 47 | |
| 45827 | 48 | axiomatization where | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 49 | finite_vertex_univ: "finite (UNIV :: vertex set)" | 
| 15274 | 50 |     --{* we assume that the universe of vertices is finite  *}
 | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 51 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 52 | declare derive_def [simp] derive1_def [simp] symcl_def [simp] | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 53 | A_def [simp] R_def [simp] | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 54 | above_def [simp] reach_def [simp] | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 55 | reverse_def [simp] neighbors_def [simp] | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 56 | |
| 14150 | 57 | text{*All vertex sets are finite*}
 | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 58 | declare finite_subset [OF subset_UNIV finite_vertex_univ, iff] | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 59 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 60 | text{* and relatons over vertex are finite too *}
 | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 61 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 62 | lemmas finite_UNIV_Prod = | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 63 | finite_Prod_UNIV [OF finite_vertex_univ finite_vertex_univ] | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 64 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 65 | declare finite_subset [OF subset_UNIV finite_UNIV_Prod, iff] | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 66 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 67 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 68 | (* The equalities (above i r = {}) = (A i r = {}) 
 | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 69 |    and (reach i r = {}) = (R i r) rely on the following theorem  *)
 | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 70 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 71 | lemma image0_trancl_iff_image0_r: "((r^+)``{i} = {}) = (r``{i} = {})"
 | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 72 | apply auto | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 73 | apply (erule trancl_induct, auto) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 74 | done | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 75 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 76 | (* Another form usefull in some situation *) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 77 | lemma image0_r_iff_image0_trancl: "(r``{i}={}) = (ALL x. ((i,x):r^+) = False)"
 | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 78 | apply auto | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 79 | apply (drule image0_trancl_iff_image0_r [THEN ssubst], auto) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 80 | done | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 81 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 82 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 83 | (* In finite universe acyclic coincides with wf *) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 84 | lemma acyclic_eq_wf: "!!r::(vertex*vertex)set. acyclic r = wf r" | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 85 | by (auto simp add: wf_iff_acyclic_if_finite) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 86 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 87 | (* derive and derive1 are equivalent *) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 88 | lemma derive_derive1_eq: "derive i r q = derive1 i r q" | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 89 | by auto | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 90 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 91 | (* Lemma 1 *) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 92 | lemma lemma1_a: | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 93 | "[| x \<in> reach i q; derive1 k r q |] ==> x\<noteq>k --> x \<in> reach i r" | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 94 | apply (unfold reach_def) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 95 | apply (erule ImageE) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 96 | apply (erule trancl_induct) | 
| 46911 | 97 | apply (cases "i=k", simp_all) | 
| 98 | apply (blast, blast, clarify) | |
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 99 | apply (drule_tac x = y in spec) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 100 | apply (drule_tac x = z in spec) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 101 | apply (blast dest: r_into_trancl intro: trancl_trans) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 102 | done | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 103 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 104 | lemma reach_lemma: "derive k r q ==> reach i q \<subseteq> (reach i r \<union> {k})"
 | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 105 | apply clarify | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 106 | apply (drule lemma1_a) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 107 | apply (auto simp add: derive_derive1_eq | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 108 | simp del: reach_def derive_def derive1_def) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 109 | done | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 110 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 111 | (* An other possible formulation of the above theorem based on | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 112 | the equivalence x \<in> reach y r = y \<in> above x r *) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 113 | lemma reach_above_lemma: | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 114 |       "(\<forall>i. reach i q \<subseteq> (reach i r \<union> {k})) = 
 | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 115 | (\<forall>x. x\<noteq>k --> (\<forall>i. i \<notin> above x r --> i \<notin> above x q))" | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 116 | by (auto simp add: trancl_converse) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 117 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 118 | (* Lemma 2 *) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 119 | lemma maximal_converse_image0: | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 120 |      "(z, i):r^+ ==> (\<forall>y. (y, z):r --> (y,i) \<notin> r^+) = ((r^-1)``{z}={})"
 | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 121 | apply auto | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 122 | apply (frule_tac r = r in trancl_into_trancl2, auto) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 123 | done | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 124 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 125 | lemma above_lemma_a: | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 126 |      "acyclic r ==> A i r\<noteq>{}-->(\<exists>j \<in> above i r. A j r = {})"
 | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 127 | apply (simp add: acyclic_eq_wf wf_eq_minimal) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 128 | apply (drule_tac x = " ((r^-1) ^+) ``{i}" in spec)
 | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 129 | apply auto | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 130 | apply (simp add: maximal_converse_image0 trancl_converse) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 131 | done | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 132 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 133 | lemma above_lemma_b: | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 134 |      "acyclic r ==> above i r\<noteq>{}-->(\<exists>j \<in> above i r. above j r = {})";
 | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 135 | apply (drule above_lemma_a) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 136 | apply (auto simp add: image0_trancl_iff_image0_r) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13796diff
changeset | 137 | done | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 138 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 139 | end |