src/HOLCF/IOA/meta_theory/Asig.ML
changeset 3071 981258186b71
child 3275 3f53f2c876f4
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/IOA/meta_theory/Asig.ML	Wed Apr 30 11:20:15 1997 +0200
@@ -0,0 +1,38 @@
+(*  Title:      HOL/IOA/meta_theory/Asig.ML
+    ID:         $Id$
+    Author:     Olaf Mueller, Tobias Nipkow & Konrad Slind
+    Copyright   1994,1996  TU Muenchen
+
+Action signatures
+*)
+
+open Asig;
+
+val asig_projections = [asig_inputs_def, asig_outputs_def, asig_internals_def];
+
+goal thy "!!a.[| 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 thy "!!a.[|a:externals(S)|] ==> a:actions(S)";
+by (asm_full_simp_tac (!simpset addsimps [externals_def,actions_def]) 1);
+qed"ext_is_act";
+
+goal thy "!!a.[|a:internals S|] ==> a:actions S";
+by (asm_full_simp_tac (!simpset addsimps [asig_internals_def,actions_def]) 1);
+qed"int_is_act";
+
+goal thy "(x: actions S & x : externals S) = (x: externals S)";
+by (fast_tac (!claset addSIs [ext_is_act]) 1 );
+qed"ext_and_act";
+ 
+goal thy "!!S. [|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 (best_tac (set_cs addEs [equalityCE]) 1);
+qed"not_ext_is_int";
+
+goalw thy  [externals_def,actions_def,is_asig_def]
+ "!! x. [| is_asig (S); x:internals S |] ==> x~:externals S";
+by (Asm_full_simp_tac 1);
+by (best_tac (set_cs addEs [equalityCE]) 1);
+qed"int_is_not_ext";
\ No newline at end of file