src/HOLCF/IOA/meta_theory/Asig.ML
author paulson
Fri, 29 Jan 1999 16:23:56 +0100
changeset 6161 bc2a76ce1ea3
parent 5068 fb28eaa07e01
child 12218 6597093b77e7
permissions -rw-r--r--
tidied

(*  Title:      HOL/IOA/meta_theory/Asig.ML
    ID:         $Id$
    Author:     Olaf Mueller, Tobias Nipkow & Konrad Slind
    Copyright   1994,1996  TU Muenchen

Action signatures
*)

val asig_projections = [asig_inputs_def, asig_outputs_def, asig_internals_def];

Goal
 "(outputs    (a,b,c) = b)   & \
\ (inputs     (a,b,c) = a) & \
\ (internals  (a,b,c) = c)";
  by (simp_tac (simpset() addsimps asig_projections) 1);
qed "asig_triple_proj";

Goal "[| a~:internals(S) ;a~:externals(S)|] ==> a~:actions(S)";
by (asm_full_simp_tac (simpset() addsimps [externals_def,actions_def]) 1);
qed"int_and_ext_is_act";

Goal "[|a:externals(S)|] ==> a:actions(S)";
by (asm_full_simp_tac (simpset() addsimps [externals_def,actions_def]) 1);
qed"ext_is_act";

Goal "[|a:internals S|] ==> a:actions S";
by (asm_full_simp_tac (simpset() addsimps [asig_internals_def,actions_def]) 1);
qed"int_is_act";

Goal "[|a:inputs S|] ==> a:actions S";
by (asm_full_simp_tac (simpset() addsimps [asig_inputs_def,actions_def]) 1);
qed"inp_is_act";

Goal "[|a:outputs S|] ==> a:actions S";
by (asm_full_simp_tac (simpset() addsimps [asig_outputs_def,actions_def]) 1);
qed"out_is_act";

Goal "(x: actions S & x : externals S) = (x: externals S)";
by (fast_tac (claset() addSIs [ext_is_act]) 1 );
qed"ext_and_act";
 
Goal "[|is_asig S;x: actions S|] ==> (x~:externals S) = (x: internals S)";
by (asm_full_simp_tac (simpset() addsimps [actions_def,is_asig_def,externals_def]) 1);
by (Blast_tac 1);
qed"not_ext_is_int";

Goal "is_asig S ==> (x~:externals S) = (x: internals S | x~:actions S)";
by (asm_full_simp_tac (simpset() addsimps [actions_def,is_asig_def,externals_def]) 1);
by (Blast_tac 1);
qed"not_ext_is_int_or_not_act";

Goalw  [externals_def,actions_def,is_asig_def]
 "[| is_asig (S); x:internals S |] ==> x~:externals S";
by (Asm_full_simp_tac 1);
by (Blast_tac 1);
qed"int_is_not_ext";