Quot/quotient_info.ML
changeset 835 c4fa87dd0208
parent 805 d193e2111811
child 836 c2501b2b262a
--- a/Quot/quotient_info.ML	Mon Jan 11 01:03:34 2010 +0100
+++ b/Quot/quotient_info.ML	Mon Jan 11 11:51:19 2010 +0100
@@ -12,6 +12,7 @@
   type quotdata_info = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm}
   val transform_quotdata: morphism -> quotdata_info -> quotdata_info
   val quotdata_lookup: theory -> string -> quotdata_info     (* raises NotFound *)
+  val quotdata_lookup_type: theory -> typ -> quotdata_info   (* raises NotFound *)
   val quotdata_update_thy: string -> quotdata_info -> theory -> theory
   val quotdata_update_gen: string -> quotdata_info -> Context.generic -> Context.generic
   val print_quotinfo: Proof.context -> unit
@@ -158,8 +159,15 @@
 fun quotdata_update_thy str qinfo = QuotData.map (Symtab.update (str, qinfo))
 fun quotdata_update_gen str qinfo = Context.mapping (quotdata_update_thy str qinfo) I
 
-fun quotdata_dest thy =
-    map snd (Symtab.dest (QuotData.get thy))
+fun quotdata_lookup_type thy qty =
+  let
+    val smt = Symtab.dest (QuotData.get thy);
+    fun matches (_, x) = Sign.typ_instance thy (qty, (#qtyp x))
+  in
+    case (find_first matches smt) of
+      SOME (_, x) => x
+    | _ => raise NotFound
+  end
 
 fun print_quotinfo ctxt =
 let