src/HOLCF/IOA/meta_theory/Asig.thy
author wenzelm
Sat, 20 Oct 2007 18:54:33 +0200
changeset 25120 23fbc38f6432
parent 19741 f65265d71426
child 25135 4f8176c940cf
permissions -rw-r--r--
tuned abbrev interface;
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
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    50
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    51
lemmas asig_projections = asig_inputs_def asig_outputs_def asig_internals_def
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    52
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    53
lemma asig_triple_proj: 
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    54
 "(outputs    (a,b,c) = b)   &  
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    55
  (inputs     (a,b,c) = a) &  
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    56
  (internals  (a,b,c) = c)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    57
  apply (simp add: asig_projections)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    58
  done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    59
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    60
lemma int_and_ext_is_act: "[| a~:internals(S) ;a~:externals(S)|] ==> a~:actions(S)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    61
apply (simp add: externals_def actions_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    62
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    63
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    64
lemma ext_is_act: "[|a:externals(S)|] ==> a:actions(S)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    65
apply (simp add: externals_def actions_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    66
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    67
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    68
lemma int_is_act: "[|a:internals S|] ==> a:actions S"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    69
apply (simp add: asig_internals_def actions_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    70
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    71
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    72
lemma inp_is_act: "[|a:inputs S|] ==> a:actions S"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    73
apply (simp add: asig_inputs_def actions_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    74
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    75
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    76
lemma out_is_act: "[|a:outputs S|] ==> a:actions S"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    77
apply (simp add: asig_outputs_def actions_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    78
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    79
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    80
lemma ext_and_act: "(x: actions S & x : externals S) = (x: externals S)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    81
apply (fast intro!: ext_is_act)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    82
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    83
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    84
lemma not_ext_is_int: "[|is_asig S;x: actions S|] ==> (x~:externals S) = (x: internals S)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    85
apply (simp add: actions_def is_asig_def externals_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    86
apply blast
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    87
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    88
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    89
lemma not_ext_is_int_or_not_act: "is_asig S ==> (x~:externals S) = (x: internals S | x~:actions S)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    90
apply (simp add: actions_def is_asig_def externals_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    91
apply blast
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    92
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    93
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    94
lemma int_is_not_ext: 
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    95
 "[| is_asig (S); x:internals S |] ==> x~:externals S"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    96
apply (unfold externals_def actions_def is_asig_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    97
apply simp
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    98
apply blast
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    99
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   100
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   101
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   102
end