src/HOL/UNITY/Comp/PriorityAux.thy
author paulson
Mon, 08 Nov 2004 16:53:50 +0100
changeset 15274 c18f5b076e53
parent 14150 9a23e4eb5eb3
child 20625 1bb9a04f8c22
permissions -rw-r--r--
tidied comments
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/PriorityAux
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
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     5
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     6
Auxiliary definitions needed in Priority.thy
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     7
*)
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     8
15274
c18f5b076e53 tidied comments
paulson
parents: 14150
diff changeset
     9
theory PriorityAux 
c18f5b076e53 tidied comments
paulson
parents: 14150
diff changeset
    10
imports "../UNITY_Main"
c18f5b076e53 tidied comments
paulson
parents: 14150
diff changeset
    11
begin
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    12
14088
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    13
typedecl vertex
12338
de0f4a63baa5 renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents: 11194
diff changeset
    14
arities vertex :: type
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
constdefs
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    17
  symcl :: "(vertex*vertex)set=>(vertex*vertex)set"
14088
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    18
  "symcl r == r \<union> (r^-1)"
15274
c18f5b076e53 tidied comments
paulson
parents: 14150
diff changeset
    19
    --{* 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
    20
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    21
  neighbors :: "[vertex, (vertex*vertex)set]=>vertex set"
15274
c18f5b076e53 tidied comments
paulson
parents: 14150
diff changeset
    22
  "neighbors i r == ((r \<union> r^-1)``{i}) - {i}"
c18f5b076e53 tidied comments
paulson
parents: 14150
diff changeset
    23
    --{* Neighbors of a vertex i *}
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
  R :: "[vertex, (vertex*vertex)set]=>vertex set"
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    26
  "R i r == r``{i}"
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    27
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    28
  A :: "[vertex, (vertex*vertex)set]=>vertex set"
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    29
  "A i r == (r^-1)``{i}"
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    30
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    31
  reach :: "[vertex, (vertex*vertex)set]=> vertex set"
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    32
  "reach i r == (r^+)``{i}"
15274
c18f5b076e53 tidied comments
paulson
parents: 14150
diff changeset
    33
    --{* 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
    34
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    35
  above :: "[vertex, (vertex*vertex)set]=> vertex set"
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    36
  "above i r == ((r^-1)^+)``{i}"  
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    37
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    38
  reverse :: "[vertex, (vertex*vertex) set]=>(vertex*vertex)set"
14088
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    39
  "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
    40
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    41
  derive1 :: "[vertex, (vertex*vertex)set, (vertex*vertex)set]=>bool"
15274
c18f5b076e53 tidied comments
paulson
parents: 14150
diff changeset
    42
    --{* The original definition *}
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    43
  "derive1 i r q == symcl r = symcl q &
14088
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    44
                    (\<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
    45
                    A i r = {} & R i q = {}"
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
  derive :: "[vertex, (vertex*vertex)set, (vertex*vertex)set]=>bool"
15274
c18f5b076e53 tidied comments
paulson
parents: 14150
diff changeset
    48
    --{* Our alternative definition *}
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    49
  "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
    50
14088
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    51
axioms
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    52
  finite_vertex_univ:  "finite (UNIV :: vertex set)"
15274
c18f5b076e53 tidied comments
paulson
parents: 14150
diff changeset
    53
    --{* 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: 13796
diff changeset
    54
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    55
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: 13796
diff changeset
    56
        A_def [simp] R_def [simp] 
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    57
        above_def [simp] reach_def [simp] 
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    58
        reverse_def [simp] neighbors_def [simp]
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    59
14150
9a23e4eb5eb3 A document for UNITY
paulson
parents: 14088
diff changeset
    60
text{*All vertex sets are finite*}
14088
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    61
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: 13796
diff changeset
    62
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    63
text{* and relatons over vertex are finite too *}
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    64
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    65
lemmas finite_UNIV_Prod =
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    66
       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: 13796
diff changeset
    67
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    68
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: 13796
diff changeset
    69
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    70
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    71
(* The equalities (above i r = {}) = (A i r = {}) 
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    72
   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: 13796
diff changeset
    73
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    74
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: 13796
diff changeset
    75
apply auto
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    76
apply (erule trancl_induct, auto)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    77
done
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    78
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    79
(* Another form usefull in some situation *)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    80
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: 13796
diff changeset
    81
apply auto
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    82
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: 13796
diff changeset
    83
done
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    84
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    85
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    86
(* In finite universe acyclic coincides with wf *)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    87
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: 13796
diff changeset
    88
by (auto simp add: wf_iff_acyclic_if_finite)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    89
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    90
(* derive and derive1 are equivalent *)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    91
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: 13796
diff changeset
    92
by auto
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    93
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    94
(* Lemma 1 *)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    95
lemma lemma1_a: 
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    96
     "[| 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: 13796
diff changeset
    97
apply (unfold reach_def)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    98
apply (erule ImageE)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
    99
apply (erule trancl_induct) 
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   100
 apply (case_tac "i=k", simp_all) 
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   101
  apply (blast intro: r_into_trancl, blast, clarify) 
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   102
apply (drule_tac x = y in spec)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   103
apply (drule_tac x = z in spec)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   104
apply (blast dest: r_into_trancl intro: trancl_trans)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   105
done
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   106
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   107
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: 13796
diff changeset
   108
apply clarify 
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   109
apply (drule lemma1_a)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   110
apply (auto simp add: derive_derive1_eq 
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   111
            simp del: reach_def derive_def derive1_def)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   112
done
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   113
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   114
(* 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: 13796
diff changeset
   115
   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: 13796
diff changeset
   116
lemma reach_above_lemma:
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   117
      "(\<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: 13796
diff changeset
   118
       (\<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: 13796
diff changeset
   119
by (auto simp add: trancl_converse)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   120
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   121
(* Lemma 2 *)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   122
lemma maximal_converse_image0: 
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   123
     "(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: 13796
diff changeset
   124
apply auto
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   125
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: 13796
diff changeset
   126
done
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   127
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   128
lemma above_lemma_a: 
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   129
     "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: 13796
diff changeset
   130
apply (simp add: acyclic_eq_wf wf_eq_minimal) 
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   131
apply (drule_tac x = " ((r^-1) ^+) ``{i}" in spec)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   132
apply auto
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   133
apply (simp add: maximal_converse_image0 trancl_converse)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   134
done
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   135
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   136
lemma above_lemma_b: 
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   137
     "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: 13796
diff changeset
   138
apply (drule above_lemma_a)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   139
apply (auto simp add: image0_trancl_iff_image0_r)
61bd46feb919 converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
paulson
parents: 13796
diff changeset
   140
done
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   141
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   142
end