author | wenzelm |
Thu, 23 Jul 2009 20:05:20 +0200 | |
changeset 32152 | 53716a67c3b1 |
parent 32151 | 2f65c45c2e7e |
child 32153 | a0e57fb1b930 |
child 32161 | abda97d2deea |
--- a/src/HOLCF/IOA/meta_theory/Automata.thy Thu Jul 23 18:44:10 2009 +0200 +++ b/src/HOLCF/IOA/meta_theory/Automata.thy Thu Jul 23 20:05:20 2009 +0200 @@ -399,7 +399,7 @@ ==> input_enabled (A||B)" apply (unfold input_enabled_def) apply (simp add: Let_def inputs_of_par trans_of_par) -apply (tactic "safe_tac (claset_of @{theory Fun})") +apply (tactic "safe_tac (global_claset_of @{theory Fun})") apply (simp add: inp_is_act) prefer 2 apply (simp add: inp_is_act)