# HG changeset patch # User Cezary Kaliszyk # Date 1263452540 -3600 # Node ID 09d5b7f0e55da818e85a42144d09c08a9e1a50a5 # Parent 9e247b9505f0ede2bc1d6393b3377f9f383a5901 Undid changes from symtab to termtab, since we need to lookup specialized types. diff -r 9e247b9505f0 -r 09d5b7f0e55d Quot/quotient_def.ML --- a/Quot/quotient_def.ML Wed Jan 13 16:46:25 2010 +0100 +++ b/Quot/quotient_def.ML Thu Jan 14 08:02:20 2010 +0100 @@ -52,7 +52,7 @@ (* data storage *) fun qcinfo phi = transform_qconsts phi {qconst = trm, rconst = rhs, def = thm} val lthy'' = Local_Theory.declaration true - (fn phi => qconsts_update_gen trm (qcinfo phi)) lthy' + (fn phi => qconsts_update_gen lhs_str (qcinfo phi)) lthy' in ((trm, thm), lthy'') end diff -r 9e247b9505f0 -r 09d5b7f0e55d Quot/quotient_info.ML --- a/Quot/quotient_info.ML Wed Jan 13 16:46:25 2010 +0100 +++ b/Quot/quotient_info.ML Thu Jan 14 08:02:20 2010 +0100 @@ -19,8 +19,8 @@ type qconsts_info = {qconst: term, rconst: term, def: thm} val transform_qconsts: morphism -> qconsts_info -> qconsts_info val qconsts_lookup: theory -> term -> qconsts_info (* raises NotFound *) - val qconsts_update_thy: term -> qconsts_info -> theory -> theory - val qconsts_update_gen: term -> qconsts_info -> Context.generic -> Context.generic + val qconsts_update_thy: string -> qconsts_info -> theory -> theory + val qconsts_update_gen: string -> qconsts_info -> Context.generic -> Context.generic val qconsts_dest: theory -> qconsts_info list val print_qconstinfo: Proof.context -> unit @@ -179,28 +179,31 @@ (* info about quotient constants *) type qconsts_info = {qconst: term, rconst: term, def: thm} +(* We need to be able to lookup instances of lifted constants, + for example given "nat fset" we need to find "'a fset"; + but overloaded constants share the same name *) structure QConstsData = Theory_Data - (type T = qconsts_info Termtab.table - val empty = Termtab.empty + (type T = qconsts_info Symtab.table + val empty = Symtab.empty val extend = I - val merge = Termtab.merge (K true)) + val merge = Symtab.merge (K true)) fun transform_qconsts phi {qconst, rconst, def} = {qconst = Morphism.term phi qconst, rconst = Morphism.term phi rconst, def = Morphism.thm phi def} -fun qconsts_update_thy const qcinfo = QConstsData.map (Termtab.update (const, qcinfo)) -fun qconsts_update_gen const qcinfo = Context.mapping (qconsts_update_thy const qcinfo) I +fun qconsts_update_thy name qcinfo = QConstsData.map (Symtab.update (name, qcinfo)) +fun qconsts_update_gen name qcinfo = Context.mapping (qconsts_update_thy name qcinfo) I fun qconsts_dest thy = - map snd (Termtab.dest (QConstsData.get thy)) + map snd (Symtab.dest (QConstsData.get thy)) (* FIXME / TODO : better implementation of the lookup datastructure *) (* for example symtabs to alist; or tables with string type key *) fun qconsts_lookup thy t = let - val smt = Termtab.dest (QConstsData.get thy); + val smt = Symtab.dest (QConstsData.get thy); val (name, qty) = dest_Const t fun matches (_, x) = let @@ -225,7 +228,7 @@ Syntax.pretty_term ctxt (prop_of def)]) in QConstsData.get (ProofContext.theory_of ctxt) - |> Termtab.dest + |> Symtab.dest |> map (prt_qconst o snd) |> Pretty.big_list "quotient constants:" |> Pretty.writeln