src/HOL/UNITY/Simple/Token.thy
author wenzelm
Tue, 09 Jan 2018 20:03:14 +0100
changeset 67394 b591933d39ec
parent 63648 f9f3006a5579
child 67443 3abf6a722518
permissions -rw-r--r--
uniform typesetting of \isamarkupcmt and \isasymcomment;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37936
1e4c5015a72e updated some headers;
wenzelm
parents: 35416
diff changeset
     1
(*  Title:      HOL/UNITY/Simple/Token.thy
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     3
    Copyright   1998  University of Cambridge
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     4
*)
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     5
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     6
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
     7
section \<open>The Token Ring\<close>
15618
paulson
parents: 13806
diff changeset
     8
paulson
parents: 13806
diff changeset
     9
theory Token
paulson
parents: 13806
diff changeset
    10
imports "../WFair"
paulson
parents: 13806
diff changeset
    11
paulson
parents: 13806
diff changeset
    12
begin
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    13
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    14
text\<open>From Misra, "A Logic for Concurrent Programming" (1994), sections 5.2 and 13.2.\<close>
15618
paulson
parents: 13806
diff changeset
    15
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    16
subsection\<open>Definitions\<close>
15618
paulson
parents: 13806
diff changeset
    17
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    18
datatype pstate = Hungry | Eating | Thinking
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    19
    \<comment>\<open>process states\<close>
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    20
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    21
record state =
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    22
  token :: "nat"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    23
  proc  :: "nat => pstate"
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    24
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    25
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 19769
diff changeset
    26
definition HasTok :: "nat => state set" where
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    27
    "HasTok i == {s. token s = i}"
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    28
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 19769
diff changeset
    29
definition H :: "nat => state set" where
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    30
    "H i == {s. proc s i = Hungry}"
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    31
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 19769
diff changeset
    32
definition E :: "nat => state set" where
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    33
    "E i == {s. proc s i = Eating}"
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    34
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 19769
diff changeset
    35
definition T :: "nat => state set" where
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    36
    "T i == {s. proc s i = Thinking}"
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    37
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    38
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    39
locale Token =
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    40
  fixes N and F and nodeOrder and "next"   
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    41
  defines nodeOrder_def:
15618
paulson
parents: 13806
diff changeset
    42
       "nodeOrder j == measure(%i. ((j+N)-i) mod N) \<inter> {..<N} \<times> {..<N}"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    43
      and next_def:
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    44
       "next i == (Suc i) mod N"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    45
  assumes N_positive [iff]: "0<N"
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    46
      and TR2:  "F \<in> (T i) co (T i \<union> H i)"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    47
      and TR3:  "F \<in> (H i) co (H i \<union> E i)"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    48
      and TR4:  "F \<in> (H i - HasTok i) co (H i)"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    49
      and TR5:  "F \<in> (HasTok i) co (HasTok i \<union> -(E i))"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    50
      and TR6:  "F \<in> (H i \<inter> HasTok i) leadsTo (E i)"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    51
      and TR7:  "F \<in> (HasTok i) leadsTo (HasTok (next i))"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    52
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    53
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    54
lemma HasToK_partition: "[| s \<in> HasTok i; s \<in> HasTok j |] ==> i=j"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    55
by (unfold HasTok_def, auto)
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    56
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    57
lemma not_E_eq: "(s \<notin> E i) = (s \<in> H i | s \<in> T i)"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    58
apply (simp add: H_def E_def T_def)
46911
6d2a2f0e904e tuned proofs;
wenzelm
parents: 37936
diff changeset
    59
apply (cases "proc s i", auto)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    60
done
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    61
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
    62
context Token
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
    63
begin
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
    64
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
    65
lemma token_stable: "F \<in> stable (-(E i) \<union> (HasTok i))"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    66
apply (unfold stable_def)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    67
apply (rule constrains_weaken)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    68
apply (rule constrains_Un [OF constrains_Un [OF TR2 TR4] TR5])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    69
apply (auto simp add: not_E_eq)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    70
apply (simp_all add: H_def E_def T_def)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    71
done
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    72
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    73
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    74
subsection\<open>Progress under Weak Fairness\<close>
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    75
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
    76
lemma wf_nodeOrder: "wf(nodeOrder j)"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    77
apply (unfold nodeOrder_def)
15618
paulson
parents: 13806
diff changeset
    78
apply (rule wf_measure [THEN wf_subset], blast)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    79
done
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    80
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
    81
lemma nodeOrder_eq: 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    82
     "[| i<N; j<N |] ==> ((next i, i) \<in> nodeOrder j) = (i \<noteq> j)"
19769
c40ce2de2020 Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
krauss
parents: 15618
diff changeset
    83
apply (unfold nodeOrder_def next_def)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    84
apply (auto simp add: mod_Suc mod_geq)
63648
f9f3006a5579 "split add" -> "split"
nipkow
parents: 63146
diff changeset
    85
apply (auto split: nat_diff_split simp add: linorder_neq_iff mod_geq)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    86
done
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    87
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    88
text\<open>From "A Logic for Concurrent Programming", but not used in Chapter 4.
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    89
  Note the use of \<open>cases\<close>.  Reasoning about leadsTo takes practice!\<close>
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
    90
lemma TR7_nodeOrder:
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    91
     "[| i<N; j<N |] ==>    
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    92
      F \<in> (HasTok i) leadsTo ({s. (token s, i) \<in> nodeOrder j} \<union> HasTok j)"
46911
6d2a2f0e904e tuned proofs;
wenzelm
parents: 37936
diff changeset
    93
apply (cases "i=j")
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    94
apply (blast intro: subset_imp_leadsTo)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    95
apply (rule TR7 [THEN leadsTo_weaken_R])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    96
apply (auto simp add: HasTok_def nodeOrder_eq)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    97
done
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    98
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    99
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   100
text\<open>Chapter 4 variant, the one actually used below.\<close>
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   101
lemma TR7_aux: "[| i<N; j<N; i\<noteq>j |]     
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   102
      ==> F \<in> (HasTok i) leadsTo {s. (token s, i) \<in> nodeOrder j}"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   103
apply (rule TR7 [THEN leadsTo_weaken_R])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   104
apply (auto simp add: HasTok_def nodeOrder_eq)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   105
done
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   106
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   107
lemma token_lemma:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   108
     "({s. token s < N} \<inter> token -` {m}) = (if m<N then token -` {m} else {})"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   109
by auto
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   110
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   111
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   112
text\<open>Misra's TR9: the token reaches an arbitrary node\<close>
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   113
lemma leadsTo_j: "j<N ==> F \<in> {s. token s < N} leadsTo (HasTok j)"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   114
apply (rule leadsTo_weaken_R)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   115
apply (rule_tac I = "-{j}" and f = token and B = "{}" 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   116
       in wf_nodeOrder [THEN bounded_induct])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   117
apply (simp_all (no_asm_simp) add: token_lemma vimage_Diff HasTok_def)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   118
 prefer 2 apply blast
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   119
apply clarify
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   120
apply (rule TR7_aux [THEN leadsTo_weaken])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   121
apply (auto simp add: HasTok_def nodeOrder_def)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   122
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   123
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   124
text\<open>Misra's TR8: a hungry process eventually eats\<close>
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   125
lemma token_progress:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   126
     "j<N ==> F \<in> ({s. token s < N} \<inter> H j) leadsTo (E j)"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   127
apply (rule leadsTo_cancel1 [THEN leadsTo_Un_duplicate])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   128
apply (rule_tac [2] TR6)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   129
apply (rule psp [OF leadsTo_j TR3, THEN leadsTo_weaken], blast+)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   130
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   131
46912
e0cd5c4df8e6 tuned context specifications and proofs;
wenzelm
parents: 46911
diff changeset
   132
end
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   133
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   134
end