Quot/quotient_info.ML
changeset 799 0755f8fd56b3
parent 798 a422a51bb0eb
child 805 d193e2111811
--- a/Quot/quotient_info.ML	Wed Dec 30 12:10:57 2009 +0000
+++ b/Quot/quotient_info.ML	Thu Dec 31 23:53:10 2009 +0100
@@ -9,14 +9,14 @@
   val print_mapsinfo: Proof.context -> unit
 
   type quotdata_info = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm}
-  val quotdata_transfer: morphism -> quotdata_info -> quotdata_info
+  val transform_quotdata: morphism -> quotdata_info -> quotdata_info
   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 print_quotinfo: Proof.context -> unit
 
   type qconsts_info = {qconst: term, rconst: term, def: thm}
-  val qconsts_transfer: morphism -> qconsts_info -> qconsts_info
+  val transform_qconsts: morphism -> qconsts_info -> qconsts_info
   val qconsts_lookup: theory -> term -> qconsts_info     (* raises NotFound *)
   val qconsts_update_thy: string -> qconsts_info -> theory -> theory
   val qconsts_update_gen: string -> qconsts_info -> Context.generic -> Context.generic
@@ -138,7 +138,7 @@
    val extend = I
    val merge = Symtab.merge (K true)) 
 
-fun quotdata_transfer phi {qtyp, rtyp, equiv_rel, equiv_thm} =
+fun transform_quotdata phi {qtyp, rtyp, equiv_rel, equiv_thm} =
     {qtyp = Morphism.typ phi qtyp,
      rtyp = Morphism.typ phi rtyp,
      equiv_rel = Morphism.term phi equiv_rel,
@@ -185,7 +185,7 @@
    val extend = I
    val merge = Symtab.merge (K true))
 
-fun qconsts_transfer phi {qconst, rconst, def} =
+fun transform_qconsts phi {qconst, rconst, def} =
     {qconst = Morphism.term phi qconst,
      rconst = Morphism.term phi rconst,
      def = Morphism.thm phi def}