minimize imports
authorhuffman
Mon, 20 Aug 2007 22:48:24 +0200
changeset 24365 038b164edfef
parent 24364 31e359126ab6
child 24366 08b116049547
minimize imports
src/HOL/Word/BinGeneral.thy
src/HOL/Word/WordDefinition.thy
--- a/src/HOL/Word/BinGeneral.thy	Mon Aug 20 21:31:10 2007 +0200
+++ b/src/HOL/Word/BinGeneral.thy	Mon Aug 20 22:48:24 2007 +0200
@@ -9,7 +9,7 @@
 
 header {* Basic Definitions for Binary Integers *}
 
-theory BinGeneral imports TdThs Num_Lemmas
+theory BinGeneral imports Num_Lemmas
 
 begin
 
--- a/src/HOL/Word/WordDefinition.thy	Mon Aug 20 21:31:10 2007 +0200
+++ b/src/HOL/Word/WordDefinition.thy	Mon Aug 20 22:48:24 2007 +0200
@@ -8,7 +8,7 @@
 
 header {* Definition of Word Type *}
 
-theory WordDefinition imports Size BitSyntax BinBoolList begin 
+theory WordDefinition imports Size BinBoolList TdThs begin
 
 typedef (open word) 'a word
   = "{(0::int) ..< 2^len_of TYPE('a::len0)}" by auto