src/HOL/UNITY/Simple/Network.thy
author wenzelm
Sat, 23 Apr 2011 13:00:19 +0200
changeset 42463 f270e3e18be5
parent 37936 1e4c5015a72e
child 44871 fbfdc5ac86be
permissions -rw-r--r--
modernized specifications;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37936
1e4c5015a72e updated some headers;
wenzelm
parents: 32960
diff changeset
     1
(*  Title:      HOL/UNITY/Simple/Network.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
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 16417
diff changeset
     5
The Communication Network.
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     6
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 16417
diff changeset
     7
From Misra, "A Logic for Concurrent Programming" (1994), section 5.7.
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     8
*)
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     9
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13806
diff changeset
    10
theory Network imports UNITY begin
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    11
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    12
(*The state assigns a number to each process variable*)
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    13
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    14
datatype pvar = Sent | Rcvd | Idle
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    15
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    16
datatype pname = Aproc | Bproc
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    17
42463
f270e3e18be5 modernized specifications;
wenzelm
parents: 37936
diff changeset
    18
type_synonym state = "pname * pvar => nat"
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    19
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    20
locale F_props =
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    21
  fixes F 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    22
  assumes rsA: "F \<in> stable {s. s(Bproc,Rcvd) \<le> s(Aproc,Sent)}"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    23
      and rsB: "F \<in> stable {s. s(Aproc,Rcvd) \<le> s(Bproc,Sent)}"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    24
    and sent_nondec: "F \<in> stable {s. m \<le> s(proc,Sent)}"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    25
    and rcvd_nondec: "F \<in> stable {s. n \<le> s(proc,Rcvd)}"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    26
    and rcvd_idle: "F \<in> {s. s(proc,Idle) = Suc 0 & s(proc,Rcvd) = m}
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    27
                        co {s. s(proc,Rcvd) = m --> s(proc,Idle) = Suc 0}"
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    28
    and sent_idle: "F \<in> {s. s(proc,Idle) = Suc 0 & s(proc,Sent) = n}
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    29
                        co {s. s(proc,Sent) = n}"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    30
  
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    31
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    32
lemmas (in F_props) 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    33
        sent_nondec_A = sent_nondec [of _ Aproc]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    34
    and sent_nondec_B = sent_nondec [of _ Bproc]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    35
    and rcvd_nondec_A = rcvd_nondec [of _ Aproc]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    36
    and rcvd_nondec_B = rcvd_nondec [of _ Bproc]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    37
    and rcvd_idle_A = rcvd_idle [of Aproc]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    38
    and rcvd_idle_B = rcvd_idle [of Bproc]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    39
    and sent_idle_A = sent_idle [of Aproc]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    40
    and sent_idle_B = sent_idle [of Bproc]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    41
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    42
    and rs_AB = stable_Int [OF rsA rsB]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    43
    and sent_nondec_AB = stable_Int [OF sent_nondec_A sent_nondec_B]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    44
    and rcvd_nondec_AB = stable_Int [OF rcvd_nondec_A rcvd_nondec_B]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    45
    and rcvd_idle_AB = constrains_Int [OF rcvd_idle_A rcvd_idle_B]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    46
    and sent_idle_AB = constrains_Int [OF sent_idle_A sent_idle_B]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    47
    and nondec_AB = stable_Int [OF sent_nondec_AB rcvd_nondec_AB]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    48
    and idle_AB = constrains_Int [OF rcvd_idle_AB sent_idle_AB]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    49
    and nondec_idle = constrains_Int [OF nondec_AB [unfolded stable_def] 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    50
                                         idle_AB]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    51
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    52
lemma (in F_props)
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    53
  shows "F \<in> stable {s. s(Aproc,Idle) = Suc 0 & s(Bproc,Idle) = Suc 0 &  
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 16417
diff changeset
    54
                        s(Aproc,Sent) = s(Bproc,Rcvd) &  
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 16417
diff changeset
    55
                        s(Bproc,Sent) = s(Aproc,Rcvd) &  
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 16417
diff changeset
    56
                        s(Aproc,Rcvd) = m & s(Bproc,Rcvd) = n}"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    57
apply (unfold stable_def) 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    58
apply (rule constrainsI)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    59
apply (drule constrains_Int [OF rs_AB [unfolded stable_def] nondec_idle, 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    60
                             THEN constrainsD], assumption)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    61
apply simp_all
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    62
apply (blast del: le0, clarify) 
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    63
apply (subgoal_tac "s' (Aproc, Rcvd) = s (Aproc, Rcvd)")
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    64
apply (subgoal_tac "s' (Bproc, Rcvd) = s (Bproc, Rcvd)") 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    65
apply simp 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    66
apply (blast intro: order_antisym le_trans eq_imp_le)+
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    67
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    68
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    69
end