src/HOL/IOA/ABP/Correctness.thy
author paulson
Tue, 16 Jul 1996 15:47:07 +0200
changeset 1866 a1a41b4b02e7
parent 1476 608483c2122a
permissions -rw-r--r--
Acknowledged Stefan Berghofer for finding errors
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1138
mueller
parents: 1050
diff changeset
     1
(*  Title:      HOL/IOA/example/Correctness.thy
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1138
mueller
parents: 1050
diff changeset
     3
    Author:     Tobias Nipkow & Konrad Slind
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     4
    Copyright   1994  TU Muenchen
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     5
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     6
The main correctness proof: System_fin implements System
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     7
*)
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     8
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
     9
Correctness = Solve + Env + Impl + Impl_finite + 
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    10
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    11
consts
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    12
1376
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    13
reduce           :: 'a list => 'a list
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    14
1376
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    15
abs              :: 'c
1138
mueller
parents: 1050
diff changeset
    16
system_ioa       :: "('m action, bool * 'm impl_state)ioa"
mueller
parents: 1050
diff changeset
    17
system_fin_ioa   :: "('m action, bool * 'm impl_state)ioa"
mueller
parents: 1050
diff changeset
    18
  
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    19
primrec
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    20
  reduce List.list  
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    21
  reduce_Nil  "reduce [] = []"
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 1138
diff changeset
    22
  reduce_Cons "reduce(x#xs) =   
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1376
diff changeset
    23
                 (case xs of   
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1376
diff changeset
    24
                     [] => [x]   
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1376
diff changeset
    25
               |   y#ys => (if (x=y)   
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1376
diff changeset
    26
                              then reduce xs   
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1376
diff changeset
    27
                              else (x#(reduce xs))))"
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    28
1138
mueller
parents: 1050
diff changeset
    29
  
mueller
parents: 1050
diff changeset
    30
defs
mueller
parents: 1050
diff changeset
    31
  
mueller
parents: 1050
diff changeset
    32
system_def
mueller
parents: 1050
diff changeset
    33
  "system_ioa == (env_ioa || impl_ioa)"
mueller
parents: 1050
diff changeset
    34
mueller
parents: 1050
diff changeset
    35
system_fin_def
mueller
parents: 1050
diff changeset
    36
  "system_fin_ioa == (env_ioa || impl_fin_ioa)"
mueller
parents: 1050
diff changeset
    37
  
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 1138
diff changeset
    38
abs_def "abs  ==   
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1376
diff changeset
    39
        (%p.(fst(p),(fst(snd(p)),(fst(snd(snd(p))),   
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1376
diff changeset
    40
         (reduce(fst(snd(snd(snd(p))))),reduce(snd(snd(snd(snd(p))))))))))"
1138
mueller
parents: 1050
diff changeset
    41
mueller
parents: 1050
diff changeset
    42
rules
mueller
parents: 1050
diff changeset
    43
mueller
parents: 1050
diff changeset
    44
  sys_IOA     "IOA system_ioa"
mueller
parents: 1050
diff changeset
    45
  sys_fin_IOA "IOA system_fin_ioa"
mueller
parents: 1050
diff changeset
    46
  
1050
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    47
end
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    48
0c36c6a52a1d ABP: Alternating bit protocol example
nipkow
parents:
diff changeset
    49