diff -r 96651cddeba9 -r 8d4d52f51e0d Quot/quotient_info.ML --- a/Quot/quotient_info.ML Fri Feb 05 09:06:27 2010 +0100 +++ b/Quot/quotient_info.ML Fri Feb 05 09:06:49 2010 +0100 @@ -23,6 +23,7 @@ val quotdata_lookup: theory -> string -> 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 quotdata_dest: Proof.context -> quotdata_info list val print_quotinfo: Proof.context -> unit type qconsts_info = {qconst: term, rconst: term, def: thm} @@ -166,6 +167,9 @@ 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 lthy = + map snd (Symtab.dest (QuotData.get (ProofContext.theory_of lthy))) + fun print_quotinfo ctxt = let fun prt_quot {qtyp, rtyp, equiv_rel, equiv_thm} =