src/HOL/HOLCF/IOA/Asig.thy
author Manuel Eberl <eberlm@in.tum.de>
Tue, 12 Dec 2017 10:01:14 +0100
changeset 67167 88d1c9d86f48
parent 62116 bc178c0fe1a1
permissions -rw-r--r--
Moved analysis material from AFP
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62008
cbedaddc9351 clarified directory structure;
wenzelm
parents: 62002
diff changeset
     1
(*  Title:      HOL/HOLCF/IOA/Asig.thy
40945
b8703f63bfb2 recoded latin1 as utf8;
wenzelm
parents: 40774
diff changeset
     2
    Author:     Olaf Müller, Tobias Nipkow & Konrad Slind
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     3
*)
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     4
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
     5
section \<open>Action signatures\<close>
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     6
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     7
theory Asig
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     8
imports Main
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     9
begin
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    10
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    11
type_synonym 'a signature = "'a set \<times> 'a set \<times> 'a set"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    12
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    13
definition inputs :: "'action signature \<Rightarrow> 'action set"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    14
  where asig_inputs_def: "inputs = fst"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    15
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    16
definition outputs :: "'action signature \<Rightarrow> 'action set"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    17
  where asig_outputs_def: "outputs = fst \<circ> snd"
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    18
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    19
definition internals :: "'action signature \<Rightarrow> 'action set"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    20
  where asig_internals_def: "internals = snd \<circ> snd"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    21
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    22
definition actions :: "'action signature \<Rightarrow> 'action set"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    23
  where "actions asig = inputs asig \<union> outputs asig \<union> internals asig"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    24
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    25
definition externals :: "'action signature \<Rightarrow> 'action set"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    26
  where "externals asig = inputs asig \<union> outputs asig"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    27
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    28
definition locals :: "'action signature \<Rightarrow> 'action set"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    29
  where "locals asig = internals asig \<union> outputs asig"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    30
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    31
definition is_asig :: "'action signature \<Rightarrow> bool"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    32
  where "is_asig triple \<longleftrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    33
    inputs triple \<inter> outputs triple = {} \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    34
    outputs triple \<inter> internals triple = {} \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    35
    inputs triple \<inter> internals triple = {}"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    36
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    37
definition mk_ext_asig :: "'action signature \<Rightarrow> 'action signature"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    38
  where "mk_ext_asig triple = (inputs triple, outputs triple, {})"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    39
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    40
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    41
lemmas asig_projections = asig_inputs_def asig_outputs_def asig_internals_def
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    42
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 19741
diff changeset
    43
lemma asig_triple_proj:
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    44
  "outputs (a, b, c) = b \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    45
   inputs (a, b, c) = a \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    46
   internals (a, b, c) = c"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    47
  by (simp add: asig_projections)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    48
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    49
lemma int_and_ext_is_act: "a \<notin> internals S \<Longrightarrow> a \<notin> externals S \<Longrightarrow> a \<notin> actions S"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    50
  by (simp add: externals_def actions_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    51
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    52
lemma ext_is_act: "a \<in> externals S \<Longrightarrow> a \<in> actions S"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    53
  by (simp add: externals_def actions_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    54
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    55
lemma int_is_act: "a \<in> internals S \<Longrightarrow> a \<in> actions S"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    56
  by (simp add: asig_internals_def actions_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    57
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    58
lemma inp_is_act: "a \<in> inputs S \<Longrightarrow> a \<in> actions S"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    59
  by (simp add: asig_inputs_def actions_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    60
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    61
lemma out_is_act: "a \<in> outputs S \<Longrightarrow> a \<in> actions S"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    62
  by (simp add: asig_outputs_def actions_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    63
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    64
lemma ext_and_act: "x \<in> actions S \<and> x \<in> externals S \<longleftrightarrow> x \<in> externals S"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    65
  by (fast intro!: ext_is_act)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    66
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    67
lemma not_ext_is_int: "is_asig S \<Longrightarrow> x \<in> actions S \<Longrightarrow> x \<notin> externals S \<longleftrightarrow> x \<in> internals S"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    68
  by (auto simp add: actions_def is_asig_def externals_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    69
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    70
lemma not_ext_is_int_or_not_act: "is_asig S \<Longrightarrow> x \<notin> externals S \<longleftrightarrow> x \<in> internals S \<or> x \<notin> actions S"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    71
  by (auto simp add: actions_def is_asig_def externals_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    72
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    73
lemma int_is_not_ext:"is_asig S \<Longrightarrow> x \<in> internals S \<Longrightarrow> x \<notin> externals S"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    74
  by (auto simp add: externals_def actions_def is_asig_def)
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    75
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    76
end