src/HOL/UNITY/Network.ML
author berghofe
Thu, 07 Oct 1999 15:40:32 +0200
changeset 7786 cf9d07ad62af
parent 7054 dfd96be49bd9
permissions -rw-r--r--
Replaced update_new by update.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Network
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     5
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     6
The Communication Network
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     7
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     8
From Misra, "A Logic for Concurrent Programming" (1994), section 5.7
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     9
*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    10
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    11
val [rsA, rsB, sent_nondec, rcvd_nondec, rcvd_idle, sent_idle] = 
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    12
Goalw [stable_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5232
diff changeset
    13
   "[| !! m. F : stable {s. s(Bproc,Rcvd) <= s(Aproc,Sent)};  \
fe887910e32e specifications as sets of programs
paulson
parents: 5232
diff changeset
    14
\      !! m. F : stable {s. s(Aproc,Rcvd) <= s(Bproc,Sent)};  \
fe887910e32e specifications as sets of programs
paulson
parents: 5232
diff changeset
    15
\      !! m proc. F : stable {s. m <= s(proc,Sent)};  \
fe887910e32e specifications as sets of programs
paulson
parents: 5232
diff changeset
    16
\      !! n proc. F : stable {s. n <= s(proc,Rcvd)};  \
6536
281d44905cab made many specification operators infix
paulson
parents: 6115
diff changeset
    17
\      !! m proc. F : {s. s(proc,Idle) = 1 & s(proc,Rcvd) = m} co \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    18
\                                 {s. s(proc,Rcvd) = m --> s(proc,Idle) = 1}; \
6536
281d44905cab made many specification operators infix
paulson
parents: 6115
diff changeset
    19
\      !! n proc. F : {s. s(proc,Idle) = 1 & s(proc,Sent) = n} co \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    20
\                                 {s. s(proc,Sent) = n} \
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5232
diff changeset
    21
\   |] ==> F : stable {s. s(Aproc,Idle) = 1 & s(Bproc,Idle) = 1 & \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    22
\                         s(Aproc,Sent) = s(Bproc,Rcvd) & \
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    23
\                         s(Bproc,Sent) = s(Aproc,Rcvd) & \
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    24
\                         s(Aproc,Rcvd) = m & s(Bproc,Rcvd) = n}";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    25
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    26
val sent_nondec_A = read_instantiate [("proc","Aproc")] sent_nondec;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    27
val sent_nondec_B = read_instantiate [("proc","Bproc")] sent_nondec;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    28
val rcvd_nondec_A = read_instantiate [("proc","Aproc")] rcvd_nondec;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    29
val rcvd_nondec_B = read_instantiate [("proc","Bproc")] rcvd_nondec;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    30
val rcvd_idle_A = read_instantiate [("proc","Aproc")] rcvd_idle;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    31
val rcvd_idle_B = read_instantiate [("proc","Bproc")] rcvd_idle;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    32
val sent_idle_A = read_instantiate [("proc","Aproc")] sent_idle;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    33
val sent_idle_B = read_instantiate [("proc","Bproc")] sent_idle;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    34
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    35
val rs_AB = [rsA, rsB] MRS constrains_Int;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    36
val sent_nondec_AB = [sent_nondec_A, sent_nondec_B] MRS constrains_Int;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    37
val rcvd_nondec_AB = [rcvd_nondec_A, rcvd_nondec_B] MRS constrains_Int;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    38
val rcvd_idle_AB = [rcvd_idle_A, rcvd_idle_B] MRS constrains_Int;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    39
val sent_idle_AB = [sent_idle_A, sent_idle_B] MRS constrains_Int;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    40
val nondec_AB = [sent_nondec_AB, rcvd_nondec_AB] MRS constrains_Int;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    41
val idle_AB = [rcvd_idle_AB, sent_idle_AB] MRS constrains_Int;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    42
val nondec_idle = [nondec_AB, idle_AB] MRS constrains_Int;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    43
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    44
by (rtac constrainsI 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    45
by (dtac ([rs_AB, nondec_idle] MRS constrains_Int RS constrainsD) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    46
by (assume_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    47
by (ALLGOALS Asm_full_simp_tac);
7054
dfd96be49bd9 a more robust proof
paulson
parents: 6676
diff changeset
    48
by (blast_tac (HOL_cs addIs [order_refl]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    49
by (Clarify_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    50
by (subgoals_tac ["s' (Aproc, Rcvd) = s (Aproc, Rcvd)",
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    51
		  "s' (Bproc, Rcvd) = s (Bproc, Rcvd)"] 1);
6676
62d1e642da30 preferring generic rules to specific ones...
paulson
parents: 6536
diff changeset
    52
by (REPEAT 
62d1e642da30 preferring generic rules to specific ones...
paulson
parents: 6536
diff changeset
    53
    (blast_tac (claset() addIs [order_antisym, le_trans, eq_imp_le]) 2));
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    54
by (Asm_simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    55
result();
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    56
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    57
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    58
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    59