Test.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Tue, 07 Jun 2016 13:51:39 +0100
changeset 126 a88af0e4731f
parent 122 420e03a2d9cc
child 127 38c6acf03f68
permissions -rw-r--r--
minor update
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
theory Test 
40
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
     2
imports Precedence_ord Graphs
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     3
begin
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
type_synonym thread = nat -- {* Type for thread identifiers. *}
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
type_synonym priority = nat  -- {* Type for priorities. *}
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
     7
type_synonym cs = nat -- {* Type for critical sections (or resources). *}
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     8
46
331137d43625 added one more reference to an incorrect specification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 45
diff changeset
     9
-- {* Schedulling Events *}
331137d43625 added one more reference to an incorrect specification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 45
diff changeset
    10
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
datatype event = 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    12
  Create thread priority 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    13
| Exit thread 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    14
| P thread cs 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    15
| V thread cs 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    16
| Set thread priority 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    17
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
type_synonym state = "event list"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
fun threads :: "state \<Rightarrow> thread set"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
  where 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
  "threads [] = {}" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
| "threads (Create th prio#s) = {th} \<union> threads s" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
| "threads (Exit th # s) = (threads s) - {th}" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
| "threads (_#s) = threads s" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
fun priority :: "thread \<Rightarrow> state \<Rightarrow> priority"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
  where
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    29
  "priority th [] = 0" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    30
| "priority th (Create th' prio#s) = (if th' = th then prio else priority th s)" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    31
| "priority th (Set th' prio#s) = (if th' = th then prio else priority th s)" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
| "priority th (_#s) = priority th s"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    34
fun last_set :: "thread \<Rightarrow> state \<Rightarrow> nat"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    35
  where
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    36
  "last_set th [] = 0" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    37
| "last_set th ((Create th' prio)#s) = (if (th = th') then length s else last_set th s)" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    38
| "last_set th ((Set th' prio)#s) = (if (th = th') then length s else last_set th s)" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    39
| "last_set th (_#s) = last_set th s"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    40
46
331137d43625 added one more reference to an incorrect specification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 45
diff changeset
    41
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    42
definition preced :: "thread \<Rightarrow> state \<Rightarrow> precedence"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    43
  where "preced th s \<equiv> Prc (priority th s) (last_set th s)"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    44
44
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
    45
abbreviation 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
    46
  "preceds s ths \<equiv> {preced th s | th. th \<in> ths}"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
    47
 
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    48
definition
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    49
  "holds wq th cs \<equiv> th \<in> set (wq cs) \<and> th = hd (wq cs)"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    50
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    51
definition
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    52
  "waits wq th cs \<equiv> th \<in> set (wq cs) \<and> th \<noteq> hd (wq cs)"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    53
46
331137d43625 added one more reference to an incorrect specification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 45
diff changeset
    54
--{* Nodes in Resource Graph *}
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    55
datatype node = 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    56
  Th "thread" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    57
| Cs "cs" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    58
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    59
definition
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    60
  "RAG wq \<equiv> {(Th th, Cs cs) | th cs. waits wq th cs} \<union> {(Cs cs, Th th) | cs th. holds wq th cs}"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    61
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    62
definition
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    63
  "dependants wq th \<equiv> {th' . (Th th', Th th) \<in> (RAG wq)^+}"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    64
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    65
record schedule_state = 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    66
  wq_fun :: "cs \<Rightarrow> thread list" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    67
  cprec_fun :: "thread \<Rightarrow> precedence" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    68
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    69
definition cpreced :: "(cs \<Rightarrow> thread list) \<Rightarrow> state \<Rightarrow> thread \<Rightarrow> precedence"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    70
  where 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    71
  "cpreced wq s th \<equiv> Max ({preced th s} \<union> {preced th' s | th'. th' \<in> dependants wq th})"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    72
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    73
abbreviation
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    74
  "all_unlocked \<equiv> \<lambda>_::cs. ([]::thread list)"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    75
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    76
abbreviation 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    77
  "initial_cprec \<equiv> \<lambda>_::thread. Prc 0 0"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    78
 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    79
abbreviation
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    80
  "release qs \<equiv> case qs of
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    81
             [] => [] 
37
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
    82
          |  (_ # qs) => SOME q. distinct q \<and> set q = set qs"
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    83
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
    84
lemma [simp]: 
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
    85
  "(SOME q. distinct q \<and> q = []) = []"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
    86
by auto
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
    87
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
    88
lemma [simp]: 
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
    89
  "(x \<in> set (SOME q. distinct q \<and> set q = set p)) = (x \<in> set p)"
46
331137d43625 added one more reference to an incorrect specification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 45
diff changeset
    90
apply(rule iffI)
331137d43625 added one more reference to an incorrect specification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 45
diff changeset
    91
apply (metis (mono_tags, lifting) List.finite_set finite_distinct_list some_eq_ex)+
331137d43625 added one more reference to an incorrect specification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 45
diff changeset
    92
done
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
    93
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
    94
abbreviation
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
    95
  "next_to_run ths \<equiv> hd (SOME q::thread list. distinct q \<and> set q = set ths)"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
    96
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
    97
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    98
fun schs :: "state \<Rightarrow> schedule_state"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    99
  where 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   100
  "schs [] = (| wq_fun = \<lambda> cs. [],  cprec_fun = (\<lambda>_. Prc 0 0) |)" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   101
| "schs (Create th prio # s) = 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   102
       (let wq = wq_fun (schs s) in
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   103
          (|wq_fun = wq, cprec_fun = cpreced wq (Create th prio # s)|))"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   104
|  "schs (Exit th # s) = 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   105
       (let wq = wq_fun (schs s) in
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   106
          (|wq_fun = wq, cprec_fun = cpreced wq (Exit th # s)|))"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   107
|  "schs (Set th prio # s) = 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   108
       (let wq = wq_fun (schs s) in
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   109
          (|wq_fun = wq, cprec_fun = cpreced wq (Set th prio # s)|))"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   110
|  "schs (P th cs # s) = 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   111
       (let wq = wq_fun (schs s) in
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   112
        let new_wq = wq(cs := (wq cs @ [th])) in
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   113
          (|wq_fun = new_wq, cprec_fun = cpreced new_wq (P th cs # s)|))"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   114
|  "schs (V th cs # s) = 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   115
       (let wq = wq_fun (schs s) in
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   116
        let new_wq = wq(cs := release (wq cs)) in
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   117
          (|wq_fun = new_wq, cprec_fun = cpreced new_wq (V th cs # s)|))"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   118
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   119
definition wq :: "state \<Rightarrow> cs \<Rightarrow> thread list" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   120
  where "wq s = wq_fun (schs s)"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   121
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   122
definition cpreced2 :: "state \<Rightarrow> thread \<Rightarrow> precedence"
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   123
  where "cpreced2 s \<equiv> cprec_fun (schs s)"
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   124
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   125
abbreviation
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   126
  "cpreceds2 s ths \<equiv> {cpreced2 s th | th. th \<in> ths}"
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   127
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   128
definition
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   129
  "holds2 s \<equiv> holds (wq_fun (schs s))"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   130
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   131
definition
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   132
  "waits2 s \<equiv> waits (wq_fun (schs s))"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   133
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   134
definition
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   135
  "RAG2 s \<equiv> RAG (wq_fun (schs s))"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   136
  
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   137
definition
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   138
  "dependants2 s \<equiv> dependants (wq_fun (schs s))"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   139
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   140
(* ready -> is a thread that is not waiting for any resource *) 
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   141
definition readys :: "state \<Rightarrow> thread set"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   142
  where "readys s \<equiv> {th . th \<in> threads s \<and> (\<forall> cs. \<not> waits2 s th cs)}"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   143
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   144
definition runing :: "state \<Rightarrow> thread set"
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   145
  where "runing s \<equiv> {th . th \<in> readys s \<and> cpreced2 s th = Max (cpreceds2 s (readys s))}"
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   146
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   147
(* all resources a thread hols in a state *)
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
definition holding :: "state \<Rightarrow> thread \<Rightarrow> cs set"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   149
  where "holding s th \<equiv> {cs . holds2 s th cs}"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   150
37
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   151
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   152
lemma exists_distinct:
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   153
  obtains ys where "distinct ys" "set ys = set xs"
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   154
by (metis List.finite_set finite_distinct_list)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   155
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   156
lemma next_to_run_set [simp]:
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   157
  "wts \<noteq> [] \<Longrightarrow> next_to_run wts \<in> set wts"
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   158
apply(rule exists_distinct[of wts])
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   159
by (metis (mono_tags, lifting) hd_in_set set_empty some_eq_ex)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   160
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   161
lemma holding_RAG: 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   162
  "holding s th = {cs . (Cs cs, Th th) \<in> RAG2 s}"
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   163
unfolding holding_def RAG2_def RAG_def
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   164
unfolding holds2_def holds_def waits_def
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   165
by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   166
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   167
inductive step :: "state \<Rightarrow> event \<Rightarrow> bool"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   168
  where
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   169
  step_Create: "\<lbrakk>th \<notin> threads s\<rbrakk> \<Longrightarrow> step s (Create th prio)" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
| step_Exit: "\<lbrakk>th \<in> runing s; holding s th = {}\<rbrakk> \<Longrightarrow> step s (Exit th)" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   171
| step_P: "\<lbrakk>th \<in> runing s;  (Cs cs, Th th)  \<notin> (RAG2 s)^+\<rbrakk> \<Longrightarrow> step s (P th cs)" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   172
| step_V: "\<lbrakk>th \<in> runing s;  holds2 s th cs\<rbrakk> \<Longrightarrow> step s (V th cs)" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   173
| step_Set: "\<lbrakk>th \<in> runing s\<rbrakk> \<Longrightarrow> step s (Set th prio)"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   174
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   175
(* valid states *)
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   176
inductive vt :: "state \<Rightarrow> bool"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   177
  where
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   178
  vt_nil[intro]: "vt []" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   179
| vt_cons[intro]: "\<lbrakk>vt s; step s e\<rbrakk> \<Longrightarrow> vt (e#s)"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   180
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   181
lemma runing_ready: 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   182
  shows "runing s \<subseteq> readys s"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   183
  unfolding runing_def readys_def
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   184
  by auto 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   185
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   186
lemma readys_threads:
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   187
  shows "readys s \<subseteq> threads s"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   188
  unfolding readys_def
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   189
  by auto
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   190
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   191
lemma wq_threads: 
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   192
  assumes vt: "vt s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   193
  and h: "th \<in> set (wq s cs)"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   194
  shows "th \<in> threads s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   195
using assms
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   196
apply(induct)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   197
apply(simp add: wq_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   198
apply(erule step.cases)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   199
apply(auto simp add: wq_def Let_def holding_def holds2_def holds_def waits2_def runing_def readys_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   200
apply(simp add: waits_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   201
apply(auto simp add: waits_def split: if_splits)[1]
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   202
apply(auto split: if_splits)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   203
apply(simp only: waits_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   204
by (metis insert_iff set_simps(2))
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   205
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   206
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   207
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   208
lemma Domain_RAG_threads:
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   209
  assumes vt: "vt s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   210
  and in_dom: "(Th th) \<in> Domain (RAG2 s)"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   211
  shows "th \<in> threads s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   212
proof -
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   213
  from in_dom obtain n where "(Th th, n) \<in> RAG2 s" by auto
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   214
  then obtain cs where "n = Cs cs" "(Th th, Cs cs) \<in> RAG2 s"  
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   215
    unfolding RAG2_def RAG_def by auto
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   216
  then have "th \<in> set (wq s cs)"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   217
    unfolding wq_def RAG_def RAG2_def waits_def by auto
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   218
  with wq_threads [OF vt] show ?thesis .
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   219
qed
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   220
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   221
lemma dependants_threads:
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   222
  assumes vt: "vt s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   223
  shows "dependants2 s th \<subseteq> threads s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   224
proof
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   225
  fix th1 
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   226
  assume "th1 \<in> dependants2 s th"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   227
  then have h: "(Th th1, Th th) \<in> (RAG2 s)\<^sup>+"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   228
    unfolding dependants2_def dependants_def RAG2_def by simp
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   229
  then have "Th th1 \<in> Domain ((RAG2 s)\<^sup>+)" unfolding Domain_def by auto
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   230
  then have "Th th1 \<in> Domain (RAG2 s)" using trancl_domain by simp
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   231
  then show "th1 \<in> threads s" using vt by (rule_tac Domain_RAG_threads)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   232
qed
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   233
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   234
lemma finite_threads:
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   235
  assumes vt: "vt s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   236
  shows "finite (threads s)"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   237
using vt by (induct) (auto elim: step.cases)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   238
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   239
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   240
section {* Distinctness of @{const wq} *}
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   241
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   242
lemma wq_distinct_step: 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   243
  assumes "step s e" "distinct (wq s cs)" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   244
  shows "distinct (wq (e # s) cs)"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   245
using assms
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   246
unfolding wq_def
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   247
apply(erule_tac step.cases)
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   248
apply(auto simp add: RAG2_def RAG_def Let_def)[1]
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   249
apply(auto simp add: wq_def Let_def RAG2_def RAG_def holds_def runing_def waits2_def waits_def readys_def)
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   250
apply(auto split: list.split)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   251
apply(rule someI2)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   252
apply(auto)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   253
done
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   254
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   255
lemma wq_distinct: 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   256
  assumes "vt s" 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   257
  shows "distinct (wq s cs)"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   258
using assms
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   259
apply(induct)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   260
apply(simp add: wq_def)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   261
apply(simp add: wq_distinct_step)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   262
done
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   263
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   264
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   265
section {* Single_Valuedness of @{const waits2}, @{const holds2}, @{const RAG2} *}
40
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   266
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   267
lemma waits2_unique:
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   268
  assumes "vt s"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   269
  and "waits2 s th cs1"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   270
  and "waits2 s th cs2"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   271
  shows "cs1 = cs2"
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   272
using assms
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   273
apply(induct)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   274
apply(simp add: waits2_def waits_def)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   275
apply(erule step.cases)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   276
apply(auto simp add: Let_def waits2_def waits_def holds_def RAG2_def RAG_def 
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   277
 readys_def runing_def split: if_splits)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   278
apply (metis Nil_is_append_conv hd_append2 list.distinct(1) split_list)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   279
apply (metis Nil_is_append_conv hd_append2 list.distinct(1) split_list)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   280
apply (metis distinct.simps(2) distinct_length_2_or_more list.sel(1) wq_def wq_distinct)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   281
by (metis (full_types, hide_lams) distinct.simps(2) distinct_length_2_or_more list.sel(1) wq_def wq_distinct)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   282
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   283
lemma single_valued_waits2:
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   284
  assumes "vt s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   285
  shows "single_valuedP (waits2 s)"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   286
using assms
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   287
unfolding single_valued_def
122
420e03a2d9cc all updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 46
diff changeset
   288
by (simp add: Product_Type.Collect_case_prodD waits2_unique)
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   289
40
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   290
lemma single_valued_holds2:
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   291
  assumes "vt s"
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   292
  shows "single_valuedP (\<lambda>cs th. holds2 s th cs)"
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   293
unfolding single_valued_def holds2_def holds_def by simp
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   294
40
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   295
lemma single_valued_RAG2:
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   296
  assumes "vt s"
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   297
  shows "single_valued (RAG2 s)"
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   298
using single_valued_waits2[OF assms] single_valued_holds2[OF assms] 
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   299
unfolding RAG2_def RAG_def
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   300
apply(rule_tac single_valued_union)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   301
unfolding holds2_def[symmetric] waits2_def[symmetric]
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   302
apply(rule single_valued_Collect)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   303
apply(simp)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   304
apply(simp add: inj_on_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   305
apply(rule single_valued_Collect)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   306
apply(simp)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   307
apply(simp add: inj_on_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   308
apply(auto)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   309
done
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   310
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   311
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   312
section {* Properties of @{const RAG2} under events *}
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   313
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   314
lemma RAG_Set [simp]: 
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   315
  shows "RAG2 (Set th prio # s) = RAG2 s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   316
unfolding RAG2_def
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   317
by (simp add: Let_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   318
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   319
lemma RAG_Create [simp]: 
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   320
  "RAG2 (Create th prio # s) = RAG2 s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   321
unfolding RAG2_def
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   322
by (simp add: Let_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   323
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   324
lemma RAG_Exit [simp]: 
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   325
  shows "RAG2 (Exit th # s) = RAG2 s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   326
unfolding RAG2_def
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   327
by (simp add: Let_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   328
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   329
lemma RAG_P1:
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   330
  assumes "wq s cs = []"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   331
  shows "RAG2 (P th cs # s) \<subseteq> RAG2 s \<union> {(Cs cs, Th th)}"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   332
using assms
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   333
unfolding  RAG2_def RAG_def wq_def Let_def waits_def holds_def
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   334
by (auto simp add: Let_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   335
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   336
lemma RAG_P2:
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   337
  assumes "(Cs cs, Th th) \<notin> (RAG2 s)\<^sup>+" "wq s cs \<noteq> []"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   338
  shows "RAG2 (P th cs # s) \<subseteq> RAG2 s \<union> {(Th th, Cs cs)}"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   339
using assms
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   340
unfolding RAG2_def RAG_def wq_def Let_def waits_def holds_def
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   341
by (auto simp add: Let_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   342
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   343
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   344
lemma RAG_V1:
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   345
assumes vt: "wq s cs = [th]"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   346
shows "RAG2 (V th cs # s) \<subseteq> RAG2 s - {(Cs cs, Th th)}"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   347
using assms
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   348
unfolding RAG2_def RAG_def waits_def holds_def wq_def
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   349
by (auto simp add: Let_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   350
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   351
lemma RAG_V2:
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   352
assumes vt:"vt s" "wq s cs = th # wts \<and> wts \<noteq> []"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   353
shows "RAG2 (V th cs # s) \<subseteq>
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   354
  RAG2 s - {(Cs cs, Th th), (Th (next_to_run wts), Cs cs)} \<union> {(Cs cs, Th (next_to_run wts))}"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   355
unfolding RAG2_def RAG_def waits_def holds_def
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   356
using assms wq_distinct[OF vt(1), of"cs"]
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   357
by (auto simp add: Let_def wq_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   358
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   359
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   360
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   361
section {* Acyclicity of @{const RAG2} *}
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   362
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   363
lemma acyclic_RAG_step: 
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   364
  assumes vt: "vt s"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   365
  and     stp: "step s e"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   366
  and     ac: "acyclic (RAG2 s)"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   367
  shows "acyclic (RAG2 (e # s))"
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   368
using stp vt ac
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   369
proof (induct)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   370
  case (step_P th s cs)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   371
  have ac: "acyclic (RAG2 s)" by fact
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   372
  have ds: "(Cs cs, Th th) \<notin> (RAG2 s)\<^sup>+" by fact
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   373
  { assume wq_empty: "wq s cs = []" -- "case waiting queue is empty"
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   374
    then have "(Th th, Cs cs) \<notin> (RAG2 s)\<^sup>+"
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   375
    proof (rule_tac notI)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   376
      assume "(Th th, Cs cs) \<in> (RAG2 s)\<^sup>+"
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   377
      then obtain x where "(x, Cs cs) \<in> RAG2 s" using tranclD2 by metis 
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   378
      with wq_empty show False by (auto simp: RAG2_def RAG_def wq_def waits_def)
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   379
    qed
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   380
    with ac have "acyclic (RAG2 s \<union> {(Cs cs, Th th)})" by simp
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   381
    then have "acyclic (RAG2 (P th cs # s))" using RAG_P1[OF wq_empty] 
37
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   382
      by (rule acyclic_subset)
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   383
  }
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   384
  moreover
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   385
  { assume wq_not_empty: "wq s cs \<noteq> []" -- "case waiting queue is not empty"
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   386
    from ac ds
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   387
    have "acyclic (RAG2 s \<union> {(Th th, Cs cs)})" by simp
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   388
    then have "acyclic (RAG2 (P th cs # s))" using RAG_P2[OF ds wq_not_empty] 
37
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   389
      by (rule acyclic_subset)
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   390
  }
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   391
  ultimately show "acyclic (RAG2 (P th cs # s))" by metis
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   392
next    
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   393
  case (step_V th s cs) -- "case for release of a lock" 
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   394
  have vt: "vt s" by fact
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   395
  have ac: "acyclic (RAG2 s)" by fact
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   396
  have hd: "holds2 s th cs" by fact
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   397
  from vt have wq_distinct:"distinct (wq s cs)" by (rule wq_distinct)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   398
  from hd have "th \<in> set (wq s cs)" "th = hd (wq s cs)" unfolding holds2_def holds_def wq_def by auto
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   399
  then obtain wts where eq_wq: "wq s cs = th # wts"  by (cases "wq s cs") (auto)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   400
  -- "case no thread present in the waiting queue to take over"
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   401
  { assume "wts = []" 
37
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   402
    with eq_wq have "wq s cs = [th]" by simp
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   403
    then have "RAG2 (V th cs # s) \<subseteq> RAG2 s - {(Cs cs, Th th)}" by (rule RAG_V1)
37
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   404
    moreover have "acyclic (RAG2 s - {(Cs cs, Th th)})" using ac by (auto intro: acyclic_subset)
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   405
    ultimately 
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   406
    have "acyclic (RAG2 (V th cs # s))" by (auto intro: acyclic_subset)
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   407
  }
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   408
  moreover
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   409
  -- "at least one thread present to take over"
37
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   410
  { def nth \<equiv> "next_to_run wts"
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   411
    assume wq_not_empty: "wts \<noteq> []" 
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   412
    have "waits2 s nth cs" 
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   413
      using eq_wq wq_not_empty wq_distinct
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   414
      unfolding nth_def waits2_def waits_def wq_def[symmetric] by auto
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   415
    then have cs_in_RAG: "(Th nth, Cs cs) \<in> RAG2 s" 
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   416
      unfolding RAG2_def RAG_def waits2_def by auto
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   417
    have "RAG2 (V th cs # s) \<subseteq> RAG2 s - {(Cs cs, Th th), (Th nth, Cs cs)} \<union> {(Cs cs, Th nth)}" 
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   418
      unfolding nth_def using  vt wq_not_empty eq_wq by (rule_tac RAG_V2) (auto)
37
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   419
    moreover 
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   420
    have "acyclic (RAG2 s - {(Cs cs, Th th), (Th nth, Cs cs)} \<union> {(Cs cs, Th nth)})" 
37
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   421
    proof -
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   422
      have "acyclic (RAG2 s - {(Cs cs, Th th), (Th nth, Cs cs)})" using ac by (auto intro: acyclic_subset)
37
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   423
      moreover 
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   424
      have "(Th nth, Cs cs) \<notin> (RAG2 s - {(Cs cs, Th th), (Th nth, Cs cs)})\<^sup>+" 
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   425
      proof (rule notI)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   426
        assume "(Th nth, Cs cs) \<in> (RAG2 s - {(Cs cs, Th th), (Th nth, Cs cs)})\<^sup>+"
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   427
        then obtain z where a: "(Th nth, z) \<in> (RAG2 s - {(Cs cs, Th th), (Th nth, Cs cs)})"
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   428
          by (metis converse_tranclE)
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   429
        then have "(Th nth, z) \<in> RAG2 s" by simp
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   430
        then have "z = Cs cs" using cs_in_RAG single_valued_RAG2[OF vt]
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   431
          by (simp add: single_valued_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   432
        then show "False" using a by simp
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   433
      qed
37
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   434
      ultimately 
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   435
      show "acyclic (RAG2 s - {(Cs cs, Th th), (Th nth, Cs cs)} \<union> {(Cs cs, Th nth) })" by simp
37
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   436
    qed
38
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   437
    ultimately have "acyclic (RAG2 (V th cs # s))" 
c89013dca1aa finished proof of acyclity
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   438
      by (rule_tac acyclic_subset)
37
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   439
  }
c820ac0f3088 simplified RAG_acyclic proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   440
  ultimately show "acyclic (RAG2 (V th cs # s))" by metis
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   441
qed (simp_all)
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   442
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   443
39
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   444
lemma finite_RAG:
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   445
  assumes "vt s"
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   446
  shows "finite (RAG2 s)"
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   447
using assms
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   448
apply(induct)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   449
apply(simp add: RAG2_def RAG_def waits_def holds_def)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   450
apply(erule step.cases)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   451
apply(auto)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   452
apply(case_tac "wq sa cs = []")
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   453
apply(rule finite_subset)
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   454
apply(rule RAG_P1)
39
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   455
apply(simp)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   456
apply(simp)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   457
apply(rule finite_subset)
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   458
apply(rule RAG_P2)
39
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   459
apply(simp)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   460
apply(simp)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   461
apply(simp)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   462
apply(subgoal_tac "\<exists>wts. wq sa cs = th # wts")
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   463
apply(erule exE)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   464
apply(case_tac "wts = []")
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   465
apply(rule finite_subset)
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   466
apply(rule RAG_V1)
39
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   467
apply(simp)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   468
apply(simp)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   469
apply(rule finite_subset)
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   470
apply(rule RAG_V2)
39
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   471
apply(simp)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   472
apply(simp)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   473
apply(simp)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   474
apply(subgoal_tac "th \<in> set (wq sa cs) \<and> th = hd (wq sa cs)") 
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   475
apply(case_tac "wq sa cs") 
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   476
apply(auto)[2]
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   477
apply(auto simp add: holds2_def holds_def wq_def)
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   478
done
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   479
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   480
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   481
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   482
lemma dchain_unique:
39
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   483
  assumes vt: "vt s"
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   484
  and th1_d: "(n, Th th1) \<in> (RAG2 s)^+"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   485
  and th1_r: "th1 \<in> readys s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   486
  and th2_d: "(n, Th th2) \<in> (RAG2 s)^+"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   487
  and th2_r: "th2 \<in> readys s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   488
  shows "th1 = th2"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   489
proof(rule ccontr)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   490
  assume neq: "th1 \<noteq> th2"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   491
   with single_valued_confluent2 [OF single_valued_RAG2 [OF vt]] th1_d th2_d
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   492
  have "(Th th1, Th th2) \<in> (RAG2 s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (RAG2 s)\<^sup>+" by auto
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   493
  moreover
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   494
  { assume "(Th th1, Th th2) \<in> (RAG2 s)\<^sup>+"      
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   495
    then obtain n where dd: "(Th th1, n) \<in> RAG2 s" by (metis converse_tranclE)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   496
    then obtain cs where eq_n: "n = Cs cs"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   497
      unfolding RAG2_def RAG_def by (case_tac n) (auto)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   498
    from dd eq_n have "th1 \<notin> readys s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   499
      unfolding RAG2_def RAG_def waits2_def readys_def by (auto)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   500
    with th1_r have "False" by auto
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   501
  }
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   502
  moreover
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   503
  { assume "(Th th2, Th th1) \<in> (RAG2 s)\<^sup>+"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   504
    then obtain n where dd: "(Th th2, n) \<in> RAG2 s" by (metis converse_tranclE)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   505
    then obtain cs where eq_n: "n = Cs cs"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   506
      unfolding RAG2_def RAG_def by (case_tac n) (auto)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   507
    from dd eq_n have "th2 \<notin> readys s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   508
      unfolding RAG2_def RAG_def waits2_def readys_def by (auto)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   509
    with th2_r have "False" by auto
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   510
  }
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   511
  ultimately show "False" by metis
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   512
qed
39
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   513
44
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   514
lemma cpreced2_cpreced: "cpreced2 s th = cpreced (wq s) s th"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   515
unfolding cpreced2_def wq_def
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   516
apply(induct s rule: schs.induct)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   517
apply(simp add: Let_def cpreced_def dependants_def RAG_def waits_def holds_def preced_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   518
apply(simp add: Let_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   519
apply(simp add: Let_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   520
apply(simp add: Let_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   521
apply(subst (2) schs.simps)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   522
apply(simp add: Let_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   523
apply(subst (2) schs.simps)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   524
apply(simp add: Let_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   525
done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   526
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   527
lemma cpreced_Exit:
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   528
  shows "cpreced2 (Exit th # s) th' = cpreced2 s th'"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   529
by (simp add: cpreced2_cpreced cpreced_def preced_def wq_def Let_def)
36
af38526275f8 added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   530
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   531
lemma readys_Exit:
44
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   532
  shows "readys (Exit th # s) = readys s - {th}"
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   533
by (auto simp add: readys_def waits2_def Let_def)
40
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   534
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   535
lemma readys_Create:
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   536
  shows "readys (Create th prio # s) \<subseteq> {th} \<union> readys s"
44
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   537
apply (auto simp add: readys_def waits2_def Let_def waits_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   538
done
40
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   539
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   540
lemma readys_Set:
44
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   541
  shows "readys (Set th prio # s) = readys s"
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   542
by (auto simp add: readys_def waits2_def Let_def)
40
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   543
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   544
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   545
lemma readys_P:
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   546
  shows "readys (P th cs # s) \<subseteq> readys s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   547
apply(auto simp add: readys_def waits2_def Let_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   548
apply(simp add: waits_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   549
apply(case_tac "csa = cs")
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   550
apply(simp)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   551
apply(drule_tac x="cs" in spec)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   552
apply(simp)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   553
apply (metis hd_append2 in_set_insert insert_Nil list.sel(1))
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   554
apply(drule_tac x="csa" in spec)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   555
apply(simp)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   556
done
40
0781a2fc93f1 added a library about graphs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 39
diff changeset
   557
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   558
lemma readys_V:
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   559
  shows "readys (V th cs # s) \<subseteq> readys s \<union> set (wq s cs)"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   560
apply(auto simp add: readys_def waits2_def waits_def Let_def wq_def)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   561
done
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   562
44
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   563
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   564
fun the_th :: "node \<Rightarrow> thread"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   565
  where "the_th (Th th) = th"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   566
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   567
lemma image_Collect2:
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   568
  "f ` A = {f x | x. x \<in> A}"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   569
apply(auto)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   570
done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   571
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   572
lemma Collect_disj_eq2:
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   573
  "{f x | x. x = y \<or> x \<in> A} = {f y} \<union> {f x | x. x \<in> A}"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   574
by (auto)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   575
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   576
lemma last_set_lt: 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   577
  "th \<in> threads s \<Longrightarrow> last_set th s < length s"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   578
  apply(induct rule: threads.induct)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   579
  apply(auto)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   580
  done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   581
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   582
lemma last_set_eq_iff: 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   583
  assumes "th1 \<in> threads s" "th2 \<in> threads s"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   584
  shows "last_set th1 s = last_set th2 s \<longleftrightarrow> th1 = th2"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   585
  using assms
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   586
  apply(induct s rule: threads.induct) 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   587
  apply(auto split:if_splits dest:last_set_lt)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   588
  done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   589
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   590
lemma preced_eq_iff: 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   591
  assumes th_in1: "th1 \<in> threads s"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   592
  and th_in2: "th2 \<in> threads s"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   593
  shows "preced th1 s = preced th2 s \<longleftrightarrow> th1 = th2"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   594
using assms
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   595
by (auto simp add: preced_def last_set_eq_iff)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   596
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   597
lemma dm_RAG_threads:
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   598
  assumes vt: "vt s"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   599
  and in_dom: "(Th th) \<in> Domain (RAG2 s)"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   600
  shows "th \<in> threads s"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   601
proof -
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   602
  from in_dom obtain n where a: "(Th th, n) \<in> RAG2 s" by auto
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   603
  then obtain cs where "n = Cs cs" 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   604
    unfolding RAG2_def RAG_def
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   605
    by auto
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   606
  then have "(Th th, Cs cs) \<in> RAG2 s" using a by simp
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   607
  hence "th \<in> set (wq s cs)"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   608
    unfolding RAG2_def wq_def RAG_def waits_def
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   609
    by (auto)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   610
  then show ?thesis
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   611
    apply(rule_tac wq_threads)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   612
    apply(rule assms)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   613
    apply(simp)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   614
    done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   615
qed
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   616
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   617
lemma cpreced_eq_iff:
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   618
  assumes "th1 \<in> readys s" "th2 \<in> readys s" "vt s"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   619
  shows "cpreced2 s th1 = cpreced2 s th2 \<longleftrightarrow> th1 = th2" 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   620
proof 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   621
  def S1\<equiv>"({th1} \<union> dependants (wq s) th1)"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   622
  def S2\<equiv>"({th2} \<union> dependants (wq s) th2)"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   623
  have fin: "finite ((the_th o fst) ` ((RAG (wq s))\<^sup>+))" 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   624
      apply(rule)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   625
      apply(simp add: finite_trancl)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   626
      apply(simp add: wq_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   627
      apply(rule finite_RAG[simplified RAG2_def])
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   628
      apply(rule assms)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   629
      done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   630
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   631
  from fin have h: "finite (preceds s S1)" "finite (preceds s S2)"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   632
      apply(simp_all add: S2_def S1_def Collect_disj_eq2 image_Collect[symmetric])
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   633
      apply(rule)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   634
      apply(simp add: dependants_def)  
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   635
      apply(rule rev_finite_subset)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   636
      apply(assumption)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   637
      apply(auto simp add: image_def)[1]
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   638
      apply(metis fst_conv the_th.simps)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   639
      apply(rule)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   640
      apply(simp add: dependants_def)  
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   641
      apply(rule rev_finite_subset)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   642
      apply(assumption)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   643
      apply(auto simp add: image_def)[1]
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   644
      apply(metis fst_conv the_th.simps)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   645
      done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   646
    moreover have "S1 \<noteq> {}" "S2 \<noteq> {}" by (simp_all add: S1_def S2_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   647
    then have "(preceds s S1) \<noteq> {}" "(preceds s S2) \<noteq> {}" by simp_all
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   648
    ultimately have m: "Max (preceds s S1) \<in> (preceds s S1)" "Max (preceds s S2) \<in> (preceds s S2)"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   649
      apply(rule_tac [!] Max_in)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   650
      apply(simp_all)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   651
      done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   652
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   653
  assume q: "cpreced2 s th1 = cpreced2 s th2"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   654
  then have eq_max: "Max (preceds s S1) = Max (preceds s S2)"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   655
    unfolding cpreced2_cpreced cpreced_def
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   656
    apply(simp only: S1_def S2_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   657
    apply(simp add: Collect_disj_eq2)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   658
    done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   659
 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   660
  obtain th0 where th0_in: "th0 \<in> S1" "th0 \<in> S2" and 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   661
      eq_f_th1: "preced th0 s = Max (preceds s S1)"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   662
                "preced th0 s = Max (preceds s S2)"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   663
    using m 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   664
      apply(clarify)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   665
      apply(simp add: eq_max)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   666
      apply(subst (asm) (2)  preced_eq_iff)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   667
      apply(insert assms(2))[1]
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   668
      apply(simp add: S2_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   669
      apply(auto)[1]
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   670
      apply (metis contra_subsetD readys_threads)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   671
      apply(simp add: dependants_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   672
      apply(subgoal_tac "Th tha \<in> Domain ((RAG2 s)^+)")
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   673
      apply(simp add: trancl_domain)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   674
      apply (metis Domain_RAG_threads assms(3))
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   675
      apply(simp only: RAG2_def wq_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   676
      apply (metis Domain_iff)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   677
      apply(insert assms(1))[1]
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   678
      apply(simp add: S1_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   679
      apply(auto)[1]
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   680
      apply (metis contra_subsetD readys_threads)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   681
      apply(simp add: dependants_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   682
      apply(subgoal_tac "Th th \<in> Domain ((RAG2 s)^+)")
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   683
      apply(simp add: trancl_domain)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   684
      apply (metis Domain_RAG_threads assms(3))
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   685
      apply(simp only: RAG2_def wq_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   686
      apply (metis Domain_iff)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   687
      apply(simp)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   688
    done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   689
  then show "th1 = th2"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   690
    apply -
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   691
    apply(insert th0_in assms(1, 2))[1]
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   692
    apply(simp add: S1_def S2_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   693
    apply(auto)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   694
    --"first case"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   695
    prefer 2
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   696
    apply(subgoal_tac "Th th2 \<in> Domain (RAG2 s)")
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   697
    apply(subgoal_tac "\<exists>cs. (Th th2, Cs cs) \<in> RAG2 s")
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   698
    apply(erule exE)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   699
    apply(simp add: runing_def RAG2_def RAG_def readys_def waits2_def)[1]
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   700
    apply(auto simp add: RAG2_def RAG_def)[1]
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   701
    apply(subgoal_tac "Th th2 \<in> Domain ((RAG2 s)^+)")
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   702
    apply (metis trancl_domain)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   703
    apply(subgoal_tac "(Th th2, Th th1) \<in> (RAG2 s)^+")
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   704
    apply (metis Domain_iff)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   705
    apply(simp add: dependants_def RAG2_def wq_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   706
    --"second case"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   707
    apply(subgoal_tac "Th th1 \<in> Domain (RAG2 s)")
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   708
    apply(subgoal_tac "\<exists>cs. (Th th1, Cs cs) \<in> RAG2 s")
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   709
    apply(erule exE)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   710
    apply(insert assms(1))[1]
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   711
    apply(simp add: runing_def RAG2_def RAG_def readys_def waits2_def)[1]
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   712
    apply(auto simp add: RAG2_def RAG_def)[1]
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   713
    apply(subgoal_tac "Th th1 \<in> Domain ((RAG2 s)^+)")
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   714
    apply (metis trancl_domain)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   715
    apply(subgoal_tac "(Th th1, Th th2) \<in> (RAG2 s)^+")
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   716
    apply (metis Domain_iff)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   717
    apply(simp add: dependants_def RAG2_def wq_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   718
    --"third case"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   719
    apply(rule dchain_unique)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   720
    apply(rule assms(3))
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   721
    apply(simp add: dependants_def RAG2_def wq_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   722
    apply(simp)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   723
    apply(simp add: dependants_def RAG2_def wq_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   724
    apply(simp)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   725
    done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   726
next
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   727
  assume "th1 = th2"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   728
  then show "cpreced2 s th1 = cpreced2 s th2" by simp
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   729
qed
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   730
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   731
lemma at_most_one_running:
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   732
  assumes "vt s"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   733
  shows "card (runing s) \<le> 1"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   734
proof (rule ccontr)
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   735
  assume "\<not> card (runing s) \<le> 1"
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   736
  then have "2 \<le> card (runing s)" by auto
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   737
  moreover 
44
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   738
  have "finite (runing s)"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   739
    by (metis `\<not> card (runing s) \<le> 1` card_infinite le0)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   740
  ultimately obtain th1 th2 where a: 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   741
    "th1 \<noteq> th2" "th1 \<in> runing s" "th2 \<in> runing s" 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   742
    "cpreced2 s th1 = cpreced2 s th2" 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   743
    apply(auto simp add: numerals card_le_Suc_iff runing_def) 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   744
    apply(blast)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   745
    done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   746
  then have "th1 = th2"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   747
    apply(subst (asm) cpreced_eq_iff)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   748
    apply(auto intro: assms a)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   749
    apply (metis contra_subsetD runing_ready)+
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   750
    done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   751
  then show "False" using a(1) by auto
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   752
qed
41
66ed924aaa5c added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 40
diff changeset
   753
44
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   754
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   755
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   756
  (*
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   757
  obtain th0 where th0_in: "th0 \<in> S1 \<and> th0 \<in> S2" 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   758
    and eq_f_th0: "preced th0 s = Max ((\<lambda>th. preced th s) ` (S1 \<inter> S2))"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   759
  proof -
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   760
    from fin have h1: "finite ((\<lambda>th. preced th s) ` (S1 \<inter> S2))"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   761
      apply(simp only: S1_def S2_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   762
      apply(rule)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   763
      apply(rule)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   764
      apply(rule)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   765
      apply(simp add: dependants_def) 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   766
      apply(rule rev_finite_subset)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   767
      apply(assumption)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   768
      apply(auto simp add: image_def)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   769
      apply (metis fst_conv the_th.simps)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   770
      done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   771
    moreover 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   772
    have "S1 \<inter> S2 \<noteq> {}" apply (simp add: S1_def S2_def) 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   773
      apply(auto) 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   774
      
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   775
      done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   776
    then have h2: "((\<lambda>th. preced th s) ` (S1 \<union> S2)) \<noteq> {}" by simp
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   777
    ultimately have "Max ((\<lambda>th. preced th s) ` (S1 \<union> S2)) \<in> ((\<lambda>th. preced th s) ` (S1 \<union> S2))"
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   778
      apply(rule Max_in)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   779
      done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   780
    then show ?thesis using that[intro] apply(auto) 
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   781
      
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   782
      apply(erule_tac preced_unique)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   783
      done
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   784
  qed
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   785
  *)
f676a68935a0 updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 41
diff changeset
   786
45
fc83f79009bd updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 44
diff changeset
   787
thm waits_def waits2_def
fc83f79009bd updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 44
diff changeset
   788
39
7ea6b019ce24 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   789
end