src/Sequents/LK.thy
changeset 48891 c0eafbd55de3
parent 45602 2a858377c3d2
child 51717 9e7d1c139569
--- a/src/Sequents/LK.thy	Wed Aug 22 22:47:16 2012 +0200
+++ b/src/Sequents/LK.thy	Wed Aug 22 22:55:41 2012 +0200
@@ -14,7 +14,6 @@
 
 theory LK
 imports LK0
-uses ("simpdata.ML")
 begin
 
 axiomatization where
@@ -215,7 +214,7 @@
   apply (tactic {* fast_tac (LK_pack add_safes @{thms subst}) 1 *})
   done
 
-use "simpdata.ML"
+ML_file "simpdata.ML"
 setup {* Simplifier.map_simpset_global (K LK_ss) *}