17   val print_quotinfo: Proof.context -> unit  | 
    17   val print_quotinfo: Proof.context -> unit  | 
    18   | 
    18   | 
    19   type qconsts_info = {qconst: term, rconst: term, def: thm} | 
    19   type qconsts_info = {qconst: term, rconst: term, def: thm} | 
    20   val transform_qconsts: morphism -> qconsts_info -> qconsts_info  | 
    20   val transform_qconsts: morphism -> qconsts_info -> qconsts_info  | 
    21   val qconsts_lookup: theory -> term -> qconsts_info     (* raises NotFound *)  | 
    21   val qconsts_lookup: theory -> term -> qconsts_info     (* raises NotFound *)  | 
    22   val qconsts_update_thy: term -> qconsts_info -> theory -> theory  | 
    22   val qconsts_update_thy: string -> qconsts_info -> theory -> theory  | 
    23   val qconsts_update_gen: term -> qconsts_info -> Context.generic -> Context.generic  | 
    23   val qconsts_update_gen: string -> qconsts_info -> Context.generic -> Context.generic  | 
    24   val qconsts_dest: theory -> qconsts_info list  | 
    24   val qconsts_dest: Proof.context -> qconsts_info list  | 
    25   val print_qconstinfo: Proof.context -> unit  | 
    25   val print_qconstinfo: Proof.context -> unit  | 
    26   | 
    26   | 
    27   val equiv_rules_get: Proof.context -> thm list  | 
    27   val equiv_rules_get: Proof.context -> thm list  | 
    28   val equiv_rules_add: attribute  | 
    28   val equiv_rules_add: attribute  | 
    29   val rsp_rules_get: Proof.context -> thm list    | 
    29   val rsp_rules_get: Proof.context -> thm list    | 
   177   | 
   177   | 
   178   | 
   178   | 
   179 (* info about quotient constants *)  | 
   179 (* info about quotient constants *)  | 
   180 type qconsts_info = {qconst: term, rconst: term, def: thm} | 
   180 type qconsts_info = {qconst: term, rconst: term, def: thm} | 
   181   | 
   181   | 
         | 
   182 fun qconsts_info_eq (x : qconsts_info, y : qconsts_info) = #qconst x = #qconst y  | 
         | 
   183   | 
         | 
   184 (* We need to be able to lookup instances of lifted constants,  | 
         | 
   185    for example given "nat fset" we need to find "'a fset";  | 
         | 
   186    but overloaded constants share the same name *)  | 
   182 structure QConstsData = Theory_Data  | 
   187 structure QConstsData = Theory_Data  | 
   183   (type T = qconsts_info Termtab.table  | 
   188   (type T = (qconsts_info list) Symtab.table  | 
   184    val empty = Termtab.empty  | 
   189    val empty = Symtab.empty  | 
   185    val extend = I  | 
   190    val extend = I  | 
   186    val merge = Termtab.merge (K true))  | 
   191    val merge = Symtab.merge_list qconsts_info_eq)  | 
   187   | 
   192   | 
   188 fun transform_qconsts phi {qconst, rconst, def} = | 
   193 fun transform_qconsts phi {qconst, rconst, def} = | 
   189     {qconst = Morphism.term phi qconst, | 
   194     {qconst = Morphism.term phi qconst, | 
   190      rconst = Morphism.term phi rconst,  | 
   195      rconst = Morphism.term phi rconst,  | 
   191      def = Morphism.thm phi def}  | 
   196      def = Morphism.thm phi def}  | 
   192   | 
   197   | 
   193 fun qconsts_update_thy const qcinfo = QConstsData.map (Termtab.update (const, qcinfo))  | 
   198 fun qconsts_update_thy name qcinfo = QConstsData.map (Symtab.cons_list (name, qcinfo))  | 
   194 fun qconsts_update_gen const qcinfo = Context.mapping (qconsts_update_thy const qcinfo) I  | 
   199 fun qconsts_update_gen name qcinfo = Context.mapping (qconsts_update_thy name qcinfo) I  | 
   195   | 
   200   | 
   196 fun qconsts_dest thy =  | 
   201 fun qconsts_dest lthy =  | 
   197   map snd (Termtab.dest (QConstsData.get thy))  | 
   202   flat (map snd (Symtab.dest (QConstsData.get (ProofContext.theory_of lthy))))  | 
   198   | 
   203   | 
   199 (* FIXME / TODO : better implementation of the lookup datastructure *)  | 
         | 
   200 (* for example symtabs to alist; or tables with string type key     *)   | 
         | 
   201 fun qconsts_lookup thy t =  | 
   204 fun qconsts_lookup thy t =  | 
   202   let  | 
   205   let  | 
   203     val smt = Termtab.dest (QConstsData.get thy);  | 
         | 
   204     val (name, qty) = dest_Const t  | 
   206     val (name, qty) = dest_Const t  | 
   205     fun matches (_, x) =  | 
   207     fun matches x =  | 
   206       let  | 
   208       let  | 
   207         val (name', qty') = dest_Const (#qconst x);  | 
   209         val (name', qty') = dest_Const (#qconst x);  | 
   208       in  | 
   210       in  | 
   209         name = name' andalso Sign.typ_instance thy (qty, qty')  | 
   211         name = name' andalso Sign.typ_instance thy (qty, qty')  | 
   210       end  | 
   212       end  | 
   211   in  | 
   213   in  | 
   212     case (find_first matches smt) of  | 
   214     case Symtab.lookup (QConstsData.get thy) name of  | 
   213       SOME (_, x) => x  | 
   215       NONE => raise NotFound  | 
   214     | _ => raise NotFound  | 
   216     | SOME l =>  | 
         | 
   217       (case (find_first matches l) of  | 
         | 
   218         SOME x => x  | 
         | 
   219       | NONE => raise NotFound  | 
         | 
   220       )  | 
   215   end  | 
   221   end  | 
   216   | 
   222   | 
   217 fun print_qconstinfo ctxt =  | 
   223 fun print_qconstinfo ctxt =  | 
   218 let  | 
   224 let  | 
   219   fun prt_qconst {qconst, rconst, def} = | 
   225   fun prt_qconst {qconst, rconst, def} = |