src/HOLCF/IOA/ABP/Spec.thy
author nipkow
Thu, 30 Aug 2007 21:43:08 +0200
changeset 24490 a4c2a0ffa5be
parent 17244 0b2ff9541727
child 25135 4f8176c940cf
permissions -rw-r--r--
added lemma
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/Spec.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 specification of reliable transmission *}
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 Spec
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
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    11
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    12
consts
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    13
6468
a7b1669f5365 put types into "" because of signature clash;
mueller
parents: 3072
diff changeset
    14
spec_sig   :: "'m action signature"
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    15
spec_trans :: "('m action, 'm list)transition set"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    16
spec_ioa   :: "('m action, 'm list)ioa"
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    17
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    18
defs
a31419014be5 Old ABP 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
sig_def: "spec_sig == (UN m.{S_msg(m)}, 
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    21
                     UN m.{R_msg(m)} Un {Next}, 
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    22
                     {})"
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    23
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    24
trans_def: "spec_trans ==                           
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    25
 {tr. let s = fst(tr);                            
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    26
          t = snd(snd(tr))                        
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    27
      in                                          
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    28
      case fst(snd(tr))                           
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    29
      of   
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    30
      Next =>  t=s |            (* Note that there is condition as in Sender *)
3072
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    31
      S_msg(m) => t = s@[m]  |                    
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    32
      R_msg(m) => s = (m#t)  |                    
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    33
      S_pkt(pkt) => False |                    
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    34
      R_pkt(pkt) => False |                    
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    35
      S_ack(b) => False |                      
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    36
      R_ack(b) => False}"
a31419014be5 Old ABP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    37
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    38
ioa_def: "spec_ioa == (spec_sig, {[]}, spec_trans)"
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
end