src/HOLCF/IOA/meta_theory/TrivEx.thy
author oheimb
Tue, 21 Apr 1998 17:22:03 +0200
changeset 4816 64f075872f69
parent 4559 8e604d885b54
permissions -rw-r--r--
simplification of explicit theory usage and merges
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     1
(*  Title:      HOLCF/IOA/TrivEx.thy
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     2
    ID:         $Id$
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     3
    Author:     Olaf Mueller
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     4
    Copyright   1995  TU Muenchen
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     5
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     6
Trivial Abstraction Example
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     7
*)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     8
4816
64f075872f69 simplification of explicit theory usage and merges
oheimb
parents: 4559
diff changeset
     9
TrivEx = Abstraction + 
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    10
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    11
datatype action = INC
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    12
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    13
consts
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    14
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    15
C_asig   ::  action signature
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    16
C_trans  :: (action, nat)transition set
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    17
C_ioa    :: (action, nat)ioa
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    18
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    19
A_asig   :: action signature
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    20
A_trans  :: (action, bool)transition set
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    21
A_ioa    :: (action, bool)ioa
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    22
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    23
h_abs    :: "nat => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    24
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    25
defs
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    26
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    27
C_asig_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    28
  "C_asig == ({},{INC},{})"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    29
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    30
C_trans_def "C_trans ==                                           
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    31
 {tr. let s = fst(tr);                                               
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    32
          t = snd(snd(tr))                                           
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    33
      in case fst(snd(tr))                                           
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    34
      of                                                             
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    35
      INC       => t = Suc(s)}"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    36
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    37
C_ioa_def "C_ioa == 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    38
 (C_asig, {0}, C_trans,{},{})"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    39
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    40
A_asig_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    41
  "A_asig == ({},{INC},{})"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    42
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    43
A_trans_def "A_trans ==                                           
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    44
 {tr. let s = fst(tr);                                               
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    45
          t = snd(snd(tr))                                           
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    46
      in case fst(snd(tr))                                           
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    47
      of                                                             
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    48
      INC       => t = True}"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    49
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    50
A_ioa_def "A_ioa == 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    51
 (A_asig, {False}, A_trans,{},{})"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    52
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    53
h_abs_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    54
  "h_abs n == n~=0"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    55
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    56
rules
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    57
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    58
MC_result
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    59
  "validIOA A_ioa (<>[] <%(b,a,c). b>)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    60
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    61
end