src/HOL/HOLCF/IOA/ABP/Sender.thy
author wenzelm
Fri, 18 Aug 2017 20:47:47 +0200
changeset 66453 cc19f7ca2ed6
parent 62009 ecb5212d5885
permissions -rw-r--r--
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 41476
diff changeset
     1
(*  Title:      HOL/HOLCF/IOA/ABP/Sender.thy
40945
b8703f63bfb2 recoded latin1 as utf8;
wenzelm
parents: 40774
diff changeset
     2
    Author:     Olaf Müller
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     3
*)
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     4
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 61032
diff changeset
     5
section \<open>The implementation: sender\<close>
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     6
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     7
theory Sender
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 62009
diff changeset
     8
imports IOA.IOA Action Lemmas
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     9
begin
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    10
41476
0fa9629aa399 types -> type_synonym
huffman
parents: 40945
diff changeset
    11
type_synonym
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 61032
diff changeset
    12
  'm sender_state = "'m list  *  bool"  \<comment> \<open>messages, Alternating Bit\<close>
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    13
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    14
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    15
  sq :: "'m sender_state => 'm list" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    16
  "sq = fst"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    17
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    18
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    19
  sbit :: "'m sender_state => bool" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    20
  "sbit = snd"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    21
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    22
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    23
  sender_asig :: "'m action signature" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    24
  "sender_asig = ((UN m. {S_msg(m)}) Un (UN b. {R_ack(b)}),
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    25
                   UN pkt. {S_pkt(pkt)},
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    26
                   {})"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    27
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    28
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    29
  sender_trans :: "('m action, 'm sender_state)transition set" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    30
  "sender_trans =
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    31
   {tr. let s = fst(tr);
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    32
            t = snd(snd(tr))
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    33
        in case fst(snd(tr))
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    34
        of
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    35
        Next     => if sq(s)=[] then t=s else False |
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    36
        S_msg(m) => sq(t)=sq(s)@[m]   &
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    37
                    sbit(t)=sbit(s)  |
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    38
        R_msg(m) => False |
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    39
        S_pkt(pkt) => sq(s) ~= []  &
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    40
                       hdr(pkt) = sbit(s)      &
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    41
                      msg(pkt) = hd(sq(s))    &
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    42
                      sq(t) = sq(s)           &
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    43
                      sbit(t) = sbit(s) |
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    44
        R_pkt(pkt) => False |
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    45
        S_ack(b)   => False |
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    46
        R_ack(b)   => if b = sbit(s) then
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    47
                       sq(t)=tl(sq(s)) & sbit(t)=(~sbit(s)) else
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    48
                       sq(t)=sq(s) & sbit(t)=sbit(s)}"
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    49
  
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    50
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    51
  sender_ioa :: "('m action, 'm sender_state)ioa" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    52
  "sender_ioa =
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19738
diff changeset
    53
   (sender_asig, {([],True)}, sender_trans,{},{})"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    54
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    55
end