src/HOLCF/IOA/NTP/Abschannel.thy
author aspinall
Fri, 30 Sep 2005 18:18:34 +0200
changeset 17740 fc385ce6187d
parent 17244 0b2ff9541727
child 19739 c58ef2aa5430
permissions -rw-r--r--
Add icon for interface.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     1
(*  Title:      HOL/IOA/NTP/Abschannel.thy
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     2
    ID:         $Id$
12218
wenzelm
parents: 6468
diff changeset
     3
    Author:     Olaf Müller
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     4
*)
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     5
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     6
header {* The (faulty) transmission channel (both directions) *}
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     7
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     8
theory Abschannel
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     9
imports IOA Action
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    10
begin
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    11
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    12
datatype 'a abs_action = S 'a | R 'a
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    13
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    14
consts
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    15
6468
a7b1669f5365 put types into "" because of signature clash;
mueller
parents: 3523
diff changeset
    16
ch_asig  :: "'a abs_action signature"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    17
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    18
ch_trans :: "('a abs_action, 'a multiset)transition set"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    19
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    20
ch_ioa   :: "('a abs_action, 'a multiset)ioa"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    21
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    22
rsch_actions  :: "'m action => bool abs_action option"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    23
srch_actions  :: "'m action =>(bool * 'm) abs_action option"
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    24
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    25
srch_asig  :: "'m action signature"
6468
a7b1669f5365 put types into "" because of signature clash;
mueller
parents: 3523
diff changeset
    26
rsch_asig  :: "'m action signature"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    27
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    28
srch_wfair :: "('m action)set set"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    29
srch_sfair :: "('m action)set set"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    30
rsch_sfair :: "('m action)set set"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    31
rsch_wfair :: "('m action)set set"
3523
23eae933c2d9 changes needed for introducing fairness
mueller
parents: 3073
diff changeset
    32
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    33
srch_trans :: "('m action, 'm packet multiset)transition set"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    34
rsch_trans :: "('m action, bool multiset)transition set"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    35
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    36
srch_ioa   :: "('m action, 'm packet multiset)ioa"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    37
rsch_ioa   :: "('m action, bool multiset)ioa"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    38
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    39
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    40
defs
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    41
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    42
srch_asig_def: "srch_asig == asig_of(srch_ioa)"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    43
rsch_asig_def: "rsch_asig == asig_of(rsch_ioa)"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    44
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    45
srch_trans_def: "srch_trans == trans_of(srch_ioa)"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    46
rsch_trans_def: "rsch_trans == trans_of(rsch_ioa)"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    47
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    48
srch_wfair_def: "srch_wfair == wfair_of(srch_ioa)"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    49
rsch_wfair_def: "rsch_wfair == wfair_of(rsch_ioa)"
3523
23eae933c2d9 changes needed for introducing fairness
mueller
parents: 3073
diff changeset
    50
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    51
srch_sfair_def: "srch_sfair == sfair_of(srch_ioa)"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    52
rsch_sfair_def: "rsch_sfair == sfair_of(rsch_ioa)"
3523
23eae933c2d9 changes needed for introducing fairness
mueller
parents: 3073
diff changeset
    53
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    54
srch_ioa_def: "srch_ioa == rename ch_ioa srch_actions"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    55
rsch_ioa_def: "rsch_ioa == rename ch_ioa rsch_actions"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    56
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    57
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    58
ch_asig_def: "ch_asig == (UN b. {S(b)}, UN b. {R(b)}, {})"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    59
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    60
ch_trans_def: "ch_trans ==
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    61
 {tr. let s = fst(tr);
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    62
          t = snd(snd(tr))
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    63
      in
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    64
      case fst(snd(tr))
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    65
        of S(b) => t = addm s b |
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    66
           R(b) => count s b ~= 0 & t = delm s b}"
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    67
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    68
ch_ioa_def: "ch_ioa == (ch_asig, {{|}}, ch_trans,{},{})"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    69
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    70
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    71
rsch_actions_def: "rsch_actions (akt) ==
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    72
            case akt of
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    73
           S_msg(m) => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    74
            R_msg(m) => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    75
           S_pkt(packet) => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    76
            R_pkt(packet) => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    77
            S_ack(b) => Some(S(b)) |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    78
            R_ack(b) => Some(R(b)) |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    79
           C_m_s =>  None  |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    80
           C_m_r =>  None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    81
           C_r_s =>  None  |
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    82
           C_r_r(m) => None"
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    83
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    84
srch_actions_def: "srch_actions (akt) ==
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    85
            case akt of
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    86
           S_msg(m) => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    87
            R_msg(m) => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    88
           S_pkt(p) => Some(S(p)) |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    89
            R_pkt(p) => Some(R(p)) |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    90
            S_ack(b) => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    91
            R_ack(b) => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    92
           C_m_s => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    93
           C_m_r => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    94
           C_r_s => None |
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    95
           C_r_r(m) => None"
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    96
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    97
ML {* use_legacy_bindings (the_context ()) *}
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    98
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    99
end