src/HOL/IOA/ABP/Spec.thy
author lcp
Tue, 25 Apr 1995 11:14:03 +0200
changeset 1072 0140ff702b23
parent 1050 0c36c6a52a1d
child 1139 993e475e70e2
permissions -rw-r--r--
updated version
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/IOA/ABP/Spec.thy
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     2
    ID:         $Id$
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow & Olaf Mueller
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     4
    Copyright   1995  TU Muenchen
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     5
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     6
The specification of reliable transmission
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     7
*)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     8
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     9
Spec = List + IOA + Action +
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    10
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    11
consts
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    12
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    13
spec_sig   :: "'m action signature"
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    14
spec_trans :: "('m action, 'm list)transition set"
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    15
spec_ioa   :: "('m action, 'm list)ioa"
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    16
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    17
defs
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    18
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    19
sig_def "spec_sig == (UN m.{S_msg(m)}, \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    20
\                     UN m.{R_msg(m)} Un {Next}, \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    21
\                     {})"
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    22
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    23
trans_def "spec_trans ==                           \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    24
\ {tr. let s = fst(tr);                            \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    25
\          t = snd(snd(tr))                        \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    26
\      in                                          \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    27
\      case fst(snd(tr))                           \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    28
\      of   \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    29
\      Next =>  t=s |\ (* Note that there is condition as in Sender *) 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    30
\      S_msg(m) => t = s@[m]  |                    \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    31
\      R_msg(m) => s = (m#t)  |                    \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    32
\      S_pkt(pkt) => False |                    \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    33
\      R_pkt(pkt) => False |                    \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    34
\      S_ack(b) => False |                      \
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    35
\      R_ack(b) => False}"
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    36
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    37
ioa_def "spec_ioa == (spec_sig, {[]}, spec_trans)"
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    38
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    39
end