tuned signature;
authorwenzelm
Thu, 27 Oct 2011 19:41:08 +0200
changeset 45278 7c6c8e950636
parent 45277 85b0ca9dd82f
child 45279 89a17197cb98
tuned signature;
src/HOL/Tools/Quotient/quotient_info.ML
src/HOL/Tools/Quotient/quotient_typ.ML
--- a/src/HOL/Tools/Quotient/quotient_info.ML	Thu Oct 27 16:28:34 2011 +0200
+++ b/src/HOL/Tools/Quotient/quotient_info.ML	Thu Oct 27 19:41:08 2011 +0200
@@ -10,7 +10,6 @@
 sig
 
   type maps_info = {mapfun: string, relmap: string}
-  val maps_defined: theory -> string -> bool
   val maps_lookup: theory -> string -> maps_info option
   val maps_update_thy: string -> maps_info -> theory -> theory
   val maps_update: string -> maps_info -> Proof.context -> Proof.context
@@ -62,8 +61,6 @@
   fun merge data = Symtab.merge (K true) data
 )
 
-fun maps_defined thy s = Symtab.defined (MapsData.get thy) s
-
 fun maps_lookup thy s = Symtab.lookup (MapsData.get thy) s
 
 fun maps_update_thy k minfo = MapsData.map (Symtab.update (k, minfo))
--- a/src/HOL/Tools/Quotient/quotient_typ.ML	Thu Oct 27 16:28:34 2011 +0200
+++ b/src/HOL/Tools/Quotient/quotient_typ.ML	Thu Oct 27 19:41:08 2011 +0200
@@ -225,7 +225,7 @@
     fun map_check_aux rty warns =
       case rty of
         Type (_, []) => warns
-      | Type (s, _) => if Quotient_Info.maps_defined thy s then warns else s::warns
+      | Type (s, _) => if is_some (Quotient_Info.maps_lookup thy s) then warns else s :: warns
       | _ => warns
 
     val warns = map_check_aux rty []