src/HOLCF/IOA/meta_theory/Asig.thy
author wenzelm
Sat, 13 May 2006 02:51:37 +0200
changeset 19628 de019ddcd89e
parent 17233 41eee2e7b465
child 19741 f65265d71426
permissions -rw-r--r--
actually reject malformed defs; added unchecked flag; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     1
(*  Title:      HOL/IOA/meta_theory/Asig.thy
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     2
    ID:         $Id$
12218
wenzelm
parents: 7661
diff changeset
     3
    Author:     Olaf Müller, Tobias Nipkow & Konrad Slind
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     4
*)
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     5
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     6
header {* Action signatures *}
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     7
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     8
theory Asig
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     9
imports Main
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    10
begin
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    11
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    12
types
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    13
  'a signature = "('a set * 'a set * 'a set)"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    14
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    15
consts
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    16
  actions       :: "'action signature => 'action set"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    17
  inputs        :: "'action signature => 'action set"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    18
  outputs       :: "'action signature => 'action set"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    19
  internals     :: "'action signature => 'action set"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    20
  externals     :: "'action signature => 'action set"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    21
  locals        :: "'action signature => 'action set"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    22
  is_asig       ::"'action signature => bool"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    23
  mk_ext_asig   ::"'action signature => 'action signature"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    24
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    25
defs
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    26
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    27
asig_inputs_def:    "inputs == fst"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    28
asig_outputs_def:   "outputs == (fst o snd)"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    29
asig_internals_def: "internals == (snd o snd)"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    30
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    31
actions_def:
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    32
   "actions(asig) == (inputs(asig) Un outputs(asig) Un internals(asig))"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    33
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    34
externals_def:
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    35
   "externals(asig) == (inputs(asig) Un outputs(asig))"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    36
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    37
locals_def:
3433
2de17c994071 added deadlock freedom, polished definitions and proofs
mueller
parents: 3275
diff changeset
    38
   "locals asig == ((internals asig) Un (outputs asig))"
2de17c994071 added deadlock freedom, polished definitions and proofs
mueller
parents: 3275
diff changeset
    39
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    40
is_asig_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    41
  "is_asig(triple) ==
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    42
     ((inputs(triple) Int outputs(triple) = {})    &
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    43
      (outputs(triple) Int internals(triple) = {}) &
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    44
      (inputs(triple) Int internals(triple) = {}))"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    45
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    46
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    47
mk_ext_asig_def:
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    48
  "mk_ext_asig(triple) == (inputs(triple), outputs(triple), {})"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    49
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    50
ML {* use_legacy_bindings (the_context ()) *}
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    51
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    52
end