| author | wenzelm | 
| Sun, 22 Jan 2006 18:46:00 +0100 | |
| changeset 18743 | 7ff2934480c9 | 
| parent 17288 | aa3833fb7bee | 
| permissions | -rw-r--r-- | 
(* Title: HOL/IOA/Asig.ML ID: $Id$ Author: Tobias Nipkow & Konrad Slind Copyright 1994 TU Muenchen *) bind_thms ("asig_projections", [asig_inputs_def, asig_outputs_def, asig_internals_def]); 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";