src/HOLCF/IOA/NTP/Spec.thy
author wenzelm
Fri, 15 Aug 2008 15:50:52 +0200
changeset 27885 76b51cd0a37c
parent 27361 24ec32bee347
child 35174 e15040ae75d7
permissions -rw-r--r--
renamed T.source_of' to T.source_position_of;
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/Spec.thy
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     2
    ID:         $Id$
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
     3
    Author:     Tobias Nipkow & Konrad Slind
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 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
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    11
27361
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    12
definition
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    13
  spec_sig :: "'m action signature" where
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    14
  sig_def: "spec_sig = (UN m.{S_msg(m)}, 
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    15
                        UN m.{R_msg(m)}, 
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    16
                        {})"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    17
27361
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    18
definition
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    19
  spec_trans :: "('m action, 'm list)transition set" where
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    20
  trans_def: "spec_trans =
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    21
   {tr. let s = fst(tr);                            
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    22
            t = snd(snd(tr))                        
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    23
        in                                          
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    24
        case fst(snd(tr))                           
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    25
        of                                          
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    26
        S_msg(m) => t = s@[m]  |                    
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    27
        R_msg(m) => s = (m#t)  |                    
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    28
        S_pkt(pkt) => False |                    
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    29
        R_pkt(pkt) => False |                    
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    30
        S_ack(b) => False |                      
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    31
        R_ack(b) => False |                      
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    32
        C_m_s => False |                            
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    33
        C_m_r => False |                            
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    34
        C_r_s => False |                            
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    35
        C_r_r(m) => False}"
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    36
27361
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    37
definition
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    38
  spec_ioa :: "('m action, 'm list)ioa" where
24ec32bee347 modernized specifications;
wenzelm
parents: 19739
diff changeset
    39
  ioa_def: "spec_ioa = (spec_sig, {[]}, spec_trans,{},{})"
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    40
3073
88366253a09a Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff changeset
    41
end