src/HOLCF/IOA/meta_theory/LiveIOA.thy
author aspinall
Fri, 30 Sep 2005 18:18:34 +0200
changeset 17740 fc385ce6187d
parent 17233 41eee2e7b465
child 19741 f65265d71426
permissions -rw-r--r--
Add icon for interface.
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/meta_theory/LiveIOA.thy
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     2
    ID:         $Id$
12218
wenzelm
parents: 10835
diff changeset
     3
    Author:     Olaf Müller
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     4
*)
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     5
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     6
header {* Live I/O automata -- specified by temproal formulas *}
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     7
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     8
theory LiveIOA
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     9
imports TLS
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    10
begin
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    11
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    12
defaultsort type
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    13
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    14
types
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    15
  ('a, 's) live_ioa = "('a,'s)ioa * ('a,'s)ioa_temp"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    16
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    17
consts
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
validLIOA   :: "('a,'s)live_ioa => ('a,'s)ioa_temp  => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    20
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    21
WF         :: "('a,'s)ioa => 'a set => ('a,'s)ioa_temp"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    22
SF         :: "('a,'s)ioa => 'a set => ('a,'s)ioa_temp"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    23
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    24
liveexecutions    :: "('a,'s)live_ioa => ('a,'s)execution set"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    25
livetraces        :: "('a,'s)live_ioa => 'a trace set"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    26
live_implements   :: "('a,'s1)live_ioa => ('a,'s2)live_ioa => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    27
is_live_ref_map   :: "('s1 => 's2) => ('a,'s1)live_ioa => ('a,'s2)live_ioa => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    28
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    29
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    30
defs
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    31
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    32
validLIOA_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    33
  "validLIOA AL P == validIOA (fst AL) ((snd AL) .--> P)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    34
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    35
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    36
WF_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    37
  "WF A acts ==  <> [] <%(s,a,t) . Enabled A acts s> .--> [] <> <xt2 (plift (%a. a : acts))>"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    38
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    39
SF_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    40
  "SF A acts ==  [] <> <%(s,a,t) . Enabled A acts s> .--> [] <> <xt2 (plift (%a. a : acts))>"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    41
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    42
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    43
liveexecutions_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    44
   "liveexecutions AP == {exec. exec : executions (fst AP) & (exec |== (snd AP))}"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    45
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    46
livetraces_def:
10835
nipkow
parents: 4816
diff changeset
    47
  "livetraces AP == {mk_trace (fst AP)$(snd ex) | ex. ex:liveexecutions AP}"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    48
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    49
live_implements_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    50
  "live_implements CL AM == (inp (fst CL) = inp (fst AM)) &
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    51
                            (out (fst CL) = out (fst AM)) &
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    52
                            livetraces CL <= livetraces AM"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    53
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    54
is_live_ref_map_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    55
   "is_live_ref_map f CL AM ==
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    56
            is_ref_map f (fst CL ) (fst AM) &
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    57
            (! exec : executions (fst CL). (exec |== (snd CL)) -->
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    58
                                           ((corresp_ex (fst AM) f exec) |== (snd AM)))"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    59
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    60
ML {* use_legacy_bindings (the_context ()) *}
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    61
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    62
end