src/HOLCF/IOA/ABP/Abschannel_finite.thy
author wenzelm
Thu, 22 Nov 2007 14:51:34 +0100
changeset 25456 6f79698f294d
parent 25135 4f8176c940cf
child 27361 24ec32bee347
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/Abschannels.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 -- finite version *}
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_finite
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     9
imports Abschannel IOA Action Lemmas
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
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    12
consts reverse :: "'a list => 'a list"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    13
primrec
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    14
  reverse_Nil:  "reverse([]) = []"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    15
  reverse_Cons: "reverse(x#xs) =  reverse(xs)@[x]"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    16
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    17
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    18
  ch_fin_asig :: "'a abs_action signature" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    19
  "ch_fin_asig = ch_asig"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    20
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    21
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    22
  ch_fin_trans :: "('a abs_action, 'a list)transition set" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    23
  "ch_fin_trans =
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    24
   {tr. let s = fst(tr);
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    25
            t = snd(snd(tr))
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    26
        in
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    27
        case fst(snd(tr))
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    28
          of S(b) => ((t = s) |
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    29
                     (if (b=hd(reverse(s)) & s~=[]) then  t=s else  t=s@[b])) |
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    30
             R(b) => s ~= [] &
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    31
                      b = hd(s) &
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    32
                      ((t = s) | (t = tl(s)))}"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    33
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    34
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    35
  ch_fin_ioa :: "('a abs_action, 'a list)ioa" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    36
  "ch_fin_ioa = (ch_fin_asig, {[]}, ch_fin_trans,{},{})"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    37
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    38
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    39
  srch_fin_ioa :: "('m action, 'm packet list)ioa" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    40
  "srch_fin_ioa = rename ch_fin_ioa  srch_actions"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    41
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    42
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    43
  rsch_fin_ioa :: "('m action, bool list)ioa" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    44
  "rsch_fin_ioa = rename ch_fin_ioa  rsch_actions"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    45
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    46
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    47
  srch_fin_asig :: "'m action signature" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    48
  "srch_fin_asig = asig_of(srch_fin_ioa)"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    49
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    50
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    51
  rsch_fin_asig :: "'m action signature" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    52
  "rsch_fin_asig = asig_of(rsch_fin_ioa)"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    53
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    54
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    55
  srch_fin_trans :: "('m action, 'm packet list)transition set" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    56
  "srch_fin_trans = trans_of(srch_fin_ioa)"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    57
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    58
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    59
  rsch_fin_trans :: "('m action, bool list)transition set" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19738
diff changeset
    60
  "rsch_fin_trans = trans_of(rsch_fin_ioa)"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    61
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    62
end