src/HOLCF/IOA/ABP/Abschannel.thy
author wenzelm
Thu, 22 Nov 2007 14:51:34 +0100
changeset 25456 6f79698f294d
parent 25135 4f8176c940cf
child 35174 e15040ae75d7
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     1
(*  Title:      HOLCF/IOA/ABP/Abschannel.thy
a31419014be5 Old ABP 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
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     4
*)
a31419014be5 Old ABP 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 transmission channel *}
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     7
17244
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 Lemmas
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    10
begin
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    11
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    12
datatype 'a abs_action = S 'a | R 'a
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    13
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    14
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    15
(**********************************************************
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    16
       G e n e r i c   C h a n n e l
a31419014be5 Old ABP 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
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    19
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    20
  ch_asig :: "'a abs_action signature" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    21
  "ch_asig = (UN b. {S(b)}, UN b. {R(b)}, {})"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    22
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    23
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    24
  ch_trans :: "('a abs_action, 'a list)transition set" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    25
  "ch_trans =
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    26
   {tr. let s = fst(tr);
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    27
            t = snd(snd(tr))
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    28
        in
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    29
        case fst(snd(tr))
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    30
          of S(b) => ((t = s) | (t = s @ [b]))  |
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    31
             R(b) => s ~= [] &
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    32
                      b = hd(s) &
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    33
                      ((t = s) | (t = tl(s)))}"
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    34
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    35
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    36
  ch_ioa :: "('a abs_action, 'a list)ioa" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    37
  "ch_ioa = (ch_asig, {[]}, ch_trans,{},{})"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    38
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    39
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    40
(**********************************************************
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    41
  C o n c r e t e  C h a n n e l s  b y   R e n a m i n g
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    42
 *********************************************************)
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    43
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    44
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    45
  rsch_actions :: "'m action => bool abs_action option" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    46
  "rsch_actions (akt) =
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    47
          (case akt of
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    48
           Next    =>  None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    49
           S_msg(m) => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    50
            R_msg(m) => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    51
           S_pkt(packet) => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    52
            R_pkt(packet) => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    53
            S_ack(b) => Some(S(b)) |
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    54
            R_ack(b) => Some(R(b)))"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    55
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    56
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    57
  srch_actions :: "'m action =>(bool * 'm) abs_action option" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    58
  "srch_actions akt =
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    59
          (case akt of
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    60
            Next    =>  None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    61
           S_msg(m) => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    62
            R_msg(m) => None |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    63
           S_pkt(p) => Some(S(p)) |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    64
            R_pkt(p) => Some(R(p)) |
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    65
            S_ack(b) => None |
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    66
            R_ack(b) => None)"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    67
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    68
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    69
  srch_ioa :: "('m action, 'm packet list)ioa" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    70
  "srch_ioa = rename ch_ioa srch_actions"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    71
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    72
  rsch_ioa :: "('m action, bool list)ioa" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    73
  "rsch_ioa = rename ch_ioa rsch_actions"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    74
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    75
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    76
  srch_asig :: "'m action signature" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    77
  "srch_asig = asig_of(srch_ioa)"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    78
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    79
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    80
  rsch_asig :: "'m action signature" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    81
  "rsch_asig = asig_of(rsch_ioa)"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    82
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    83
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    84
  srch_trans :: "('m action, 'm packet list)transition set" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    85
  "srch_trans = trans_of(srch_ioa)"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    86
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    87
  rsch_trans :: "('m action, bool list)transition set" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    88
  "rsch_trans = trans_of(rsch_ioa)"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    89
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    90
end