src/HOL/IOA/Asig.ML
author paulson
Tue, 29 Mar 2005 12:30:48 +0200
changeset 15635 8408a06590a6
parent 5143 b94cd208f073
child 17288 aa3833fb7bee
permissions -rw-r--r--
converted HOL-Subst to tactic scripts
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4530
ac1821645636 corrected Title
oheimb
parents: 4089
diff changeset
     1
(*  Title:      HOL/IOA/Asig.ML
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
     2
    ID:         $Id$
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
     3
    Author:     Tobias Nipkow & Konrad Slind
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
     4
    Copyright   1994  TU Muenchen
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
     5
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
     6
Action signatures
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
     7
*)
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
     8
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
     9
open Asig;
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    10
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    11
val asig_projections = [asig_inputs_def, asig_outputs_def, asig_internals_def];
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    12
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
    13
Goal "[| a~:internals(S) ;a~:externals(S)|] ==> a~:actions(S)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3078
diff changeset
    14
by (asm_full_simp_tac (simpset() addsimps [externals_def,actions_def]) 1);
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    15
qed"int_and_ext_is_act";
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    16
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
    17
Goal "[|a:externals(S)|] ==> a:actions(S)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3078
diff changeset
    18
by (asm_full_simp_tac (simpset() addsimps [externals_def,actions_def]) 1);
3078
984866a8f905 old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff changeset
    19
qed"ext_is_act";