more general Typedef.bindings;
authorwenzelm
Thu, 03 Sep 2015 21:50:39 +0200
changeset 61110 6b7c2ecc6aea
parent 61109 1c98bfc5d743
child 61111 2618e7e3b5ea
more general Typedef.bindings; tuned signature;
src/HOL/HOLCF/Tools/cpodef.ML
src/HOL/HOLCF/Tools/domaindef.ML
src/HOL/Import/import_rule.ML
src/HOL/Nominal/nominal_datatype.ML
src/HOL/Tools/BNF/bnf_util.ML
src/HOL/Tools/Old_Datatype/old_datatype.ML
src/HOL/Tools/Quotient/quotient_type.ML
src/HOL/Tools/record.ML
src/HOL/Tools/typedef.ML
--- a/src/HOL/HOLCF/Tools/cpodef.ML	Thu Sep 03 19:27:45 2015 +0200
+++ b/src/HOL/HOLCF/Tools/cpodef.ML	Thu Sep 03 21:50:39 2015 +0200
@@ -15,27 +15,27 @@
       Rep_bottom_iff: thm, Abs_bottom_iff: thm }
 
   val add_podef: binding * (string * sort) list * mixfix ->
-    term -> (binding * binding) option -> (Proof.context -> tactic) -> theory ->
+    term -> Typedef.bindings option -> (Proof.context -> tactic) -> theory ->
     (Typedef.info * thm) * theory
   val add_cpodef: binding * (string * sort) list * mixfix ->
-    term -> (binding * binding) option -> (Proof.context -> tactic) * (Proof.context -> tactic) ->
+    term -> Typedef.bindings option -> (Proof.context -> tactic) * (Proof.context -> tactic) ->
     theory -> (Typedef.info * cpo_info) * theory
   val add_pcpodef: binding * (string * sort) list * mixfix ->
-    term -> (binding * binding) option -> (Proof.context -> tactic) * (Proof.context -> tactic) ->
+    term -> Typedef.bindings option -> (Proof.context -> tactic) * (Proof.context -> tactic) ->
     theory -> (Typedef.info * cpo_info * pcpo_info) * theory
 
   val cpodef_proof:
     (binding * (string * sort) list * mixfix) * term
-    * (binding * binding) option -> theory -> Proof.state
+    * Typedef.bindings option -> theory -> Proof.state
   val cpodef_proof_cmd:
     (binding * (string * string option) list * mixfix) * string
-    * (binding * binding) option -> theory -> Proof.state
+    * Typedef.bindings option -> theory -> Proof.state
   val pcpodef_proof:
     (binding * (string * sort) list * mixfix) * term
-    * (binding * binding) option -> theory -> Proof.state
+    * Typedef.bindings option -> theory -> Proof.state
   val pcpodef_proof_cmd:
     (binding * (string * string option) list * mixfix) * string
-    * (binding * binding) option -> theory -> Proof.state
+    * Typedef.bindings option -> theory -> Proof.state
 end
 
 structure Cpodef : CPODEF =
@@ -63,13 +63,14 @@
 fun prove_cpo
       (name: binding)
       (newT: typ)
-      (Rep_name: binding, Abs_name: binding)
+      opt_bindings
       (type_definition: thm)  (* type_definition Rep Abs A *)
       (below_def: thm)        (* op << == %x y. Rep x << Rep y *)
       (admissible: thm)       (* adm (%x. x : set) *)
       (thy: theory)
     =
   let
+    val {Rep_name, Abs_name, ...} = Typedef.make_bindings name opt_bindings;
     val cpo_thms = map (Thm.transfer thy) [type_definition, below_def, admissible]
     val (full_tname, Ts) = dest_Type newT
     val lhs_sorts = map (snd o dest_TFree) Ts
@@ -102,13 +103,14 @@
 fun prove_pcpo
       (name: binding)
       (newT: typ)
-      (Rep_name: binding, Abs_name: binding)
+      opt_bindings
       (type_definition: thm)  (* type_definition Rep Abs A *)
       (below_def: thm)        (* op << == %x y. Rep x << Rep y *)
       (bottom_mem: thm)       (* bottom : set *)
       (thy: theory)
     =
   let
+    val {Rep_name, Abs_name, ...} = Typedef.make_bindings name opt_bindings;
     val pcpo_thms = map (Thm.transfer thy) [type_definition, below_def, bottom_mem]
     val (full_tname, Ts) = dest_Type newT
     val lhs_sorts = map (snd o dest_TFree) Ts
@@ -138,7 +140,7 @@
 
 (* prepare_cpodef *)
 
-fun prepare prep_term name (tname, raw_args, _) raw_set opt_morphs thy =
+fun prepare prep_term name (tname, raw_args, _) raw_set thy =
   let
     (*rhs*)
     val tmp_ctxt =
@@ -155,18 +157,15 @@
     val lhs_tfrees = map (Proof_Context.check_tfree tmp_ctxt') raw_args
     val full_tname = Sign.full_name thy tname
     val newT = Type (full_tname, map TFree lhs_tfrees)
-
-    val morphs = opt_morphs
-      |> the_default (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name)
   in
-    (newT, oldT, set, morphs)
+    (newT, oldT, set)
   end
 
-fun add_podef typ set opt_morphs tac thy =
+fun add_podef typ set opt_bindings tac thy =
   let
     val name = #1 typ
     val ((full_tname, info as ({Rep_name, ...}, {type_definition, ...})), thy) = thy
-      |> Typedef.add_typedef_global false typ set opt_morphs tac
+      |> Typedef.add_typedef_global typ set opt_bindings tac
     val oldT = #rep_type (#1 info)
     val newT = #abs_type (#1 info)
     val lhs_tfrees = map dest_TFree (snd (dest_Type newT))
@@ -189,13 +188,12 @@
       (prep_term: Proof.context -> 'a -> term)
       (typ: binding * (string * sort) list * mixfix)
       (raw_set: 'a)
-      (opt_morphs: (binding * binding) option)
+      opt_bindings
       (thy: theory)
     : term * term * (thm -> thm -> theory -> (Typedef.info * cpo_info) * theory) =
   let
     val name = #1 typ
-    val (newT, oldT, set, morphs) =
-      prepare prep_term name typ raw_set opt_morphs thy
+    val (newT, oldT, set) = prepare prep_term name typ raw_set thy
 
     val goal_nonempty =
       HOLogic.mk_Trueprop (HOLogic.mk_exists ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set)))
@@ -205,9 +203,9 @@
     fun cpodef_result nonempty admissible thy =
       let
         val ((info as (_, {type_definition, ...}), below_def), thy) = thy
-          |> add_podef typ set opt_morphs (fn ctxt => resolve_tac ctxt [nonempty] 1)
+          |> add_podef typ set opt_bindings (fn ctxt => resolve_tac ctxt [nonempty] 1)
         val (cpo_info, thy) = thy
-          |> prove_cpo name newT morphs type_definition below_def admissible
+          |> prove_cpo name newT opt_bindings type_definition below_def admissible
       in
         ((info, cpo_info), thy)
       end
@@ -221,13 +219,12 @@
       (prep_term: Proof.context -> 'a -> term)
       (typ: binding * (string * sort) list * mixfix)
       (raw_set: 'a)
-      (opt_morphs: (binding * binding) option)
+      opt_bindings
       (thy: theory)
     : term * term * (thm -> thm -> theory -> (Typedef.info * cpo_info * pcpo_info) * theory) =
   let
     val name = #1 typ
-    val (newT, oldT, set, morphs) =
-      prepare prep_term name typ raw_set opt_morphs thy
+    val (newT, oldT, set) = prepare prep_term name typ raw_set thy
 
     val goal_bottom_mem =
       HOLogic.mk_Trueprop (HOLogic.mk_mem (Const (@{const_name bottom}, oldT), set))
@@ -239,11 +236,11 @@
       let
         fun tac ctxt = resolve_tac ctxt [exI] 1 THEN resolve_tac ctxt [bottom_mem] 1
         val ((info as (_, {type_definition, ...}), below_def), thy) = thy
-          |> add_podef typ set opt_morphs tac
+          |> add_podef typ set opt_bindings tac
         val (cpo_info, thy) = thy
-          |> prove_cpo name newT morphs type_definition below_def admissible
+          |> prove_cpo name newT opt_bindings type_definition below_def admissible
         val (pcpo_info, thy) = thy
-          |> prove_pcpo name newT morphs type_definition below_def bottom_mem
+          |> prove_pcpo name newT opt_bindings type_definition below_def bottom_mem
       in
         ((info, cpo_info, pcpo_info), thy)
       end
@@ -256,10 +253,10 @@
 
 (* tactic interface *)
 
-fun add_cpodef typ set opt_morphs (tac1, tac2) thy =
+fun add_cpodef typ set opt_bindings (tac1, tac2) thy =
   let
     val (goal1, goal2, cpodef_result) =
-      prepare_cpodef Syntax.check_term typ set opt_morphs thy
+      prepare_cpodef Syntax.check_term typ set opt_bindings thy
     val thm1 = Goal.prove_global thy [] [] goal1 (tac1 o #context)
       handle ERROR msg => cat_error msg
         ("Failed to prove non-emptiness of " ^ quote (Syntax.string_of_term_global thy set))
@@ -268,10 +265,10 @@
         ("Failed to prove admissibility of " ^ quote (Syntax.string_of_term_global thy set))
   in cpodef_result thm1 thm2 thy end
 
-fun add_pcpodef typ set opt_morphs (tac1, tac2) thy =
+fun add_pcpodef typ set opt_bindings (tac1, tac2) thy =
   let
     val (goal1, goal2, pcpodef_result) =
-      prepare_pcpodef Syntax.check_term typ set opt_morphs thy
+      prepare_pcpodef Syntax.check_term typ set opt_bindings thy
     val thm1 = Goal.prove_global thy [] [] goal1 (tac1 o #context)
       handle ERROR msg => cat_error msg
         ("Failed to prove non-emptiness of " ^ quote (Syntax.string_of_term_global thy set))
@@ -286,23 +283,23 @@
 local
 
 fun gen_cpodef_proof prep_term prep_constraint
-    ((b, raw_args, mx), set, opt_morphs) thy =
+    ((b, raw_args, mx), set, opt_bindings) thy =
   let
     val ctxt = Proof_Context.init_global thy
     val args = map (apsnd (prep_constraint ctxt)) raw_args
     val (goal1, goal2, make_result) =
-      prepare_cpodef prep_term (b, args, mx) set opt_morphs thy
+      prepare_cpodef prep_term (b, args, mx) set opt_bindings thy
     fun after_qed [[th1, th2]] = Proof_Context.background_theory (snd o make_result th1 th2)
       | after_qed _ = raise Fail "cpodef_proof"
   in Proof.theorem NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end
 
 fun gen_pcpodef_proof prep_term prep_constraint
-    ((b, raw_args, mx), set, opt_morphs) thy =
+    ((b, raw_args, mx), set, opt_bindings) thy =
   let
     val ctxt = Proof_Context.init_global thy
     val args = map (apsnd (prep_constraint ctxt)) raw_args
     val (goal1, goal2, make_result) =
-      prepare_pcpodef prep_term (b, args, mx) set opt_morphs thy
+      prepare_pcpodef prep_term (b, args, mx) set opt_bindings thy
     fun after_qed [[th1, th2]] = Proof_Context.background_theory (snd o make_result th1 th2)
       | after_qed _ = raise Fail "pcpodef_proof"
   in Proof.theorem NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end
@@ -321,24 +318,30 @@
 
 (** outer syntax **)
 
-val typedef_proof_decl =
+local
+
+fun cpodef pcpo =
   (Parse.type_args_constrained -- Parse.binding) -- Parse.opt_mixfix --
   (@{keyword "="} |-- Parse.term) --
   Scan.option
     (@{keyword "morphisms"} |-- Parse.!!! (Parse.binding -- Parse.binding))
+  >> (fn ((((args, t), mx), A), morphs) =>
+      Toplevel.theory_to_proof
+        ((if pcpo then pcpodef_proof_cmd else cpodef_proof_cmd)
+          ((t, args, mx), A, SOME (Typedef.make_morphisms t morphs))))
 
-fun mk_pcpodef_proof pcpo ((((args, t), mx), A), morphs) =
-  (if pcpo then pcpodef_proof_cmd else cpodef_proof_cmd)
-    ((t, args, mx), A, morphs)
+in
 
 val _ =
   Outer_Syntax.command @{command_keyword pcpodef}
     "HOLCF type definition (requires admissibility proof)"
-    (typedef_proof_decl >> (Toplevel.theory_to_proof o mk_pcpodef_proof true))
+    (cpodef true)
 
 val _ =
   Outer_Syntax.command @{command_keyword cpodef}
     "HOLCF type definition (requires admissibility proof)"
-    (typedef_proof_decl >> (Toplevel.theory_to_proof o mk_pcpodef_proof false))
+    (cpodef false)
 
 end
+
+end
--- a/src/HOL/HOLCF/Tools/domaindef.ML	Thu Sep 03 19:27:45 2015 +0200
+++ b/src/HOL/HOLCF/Tools/domaindef.ML	Thu Sep 03 21:50:39 2015 +0200
@@ -18,11 +18,11 @@
     }
 
   val add_domaindef: binding * (string * sort) list * mixfix ->
-    term -> (binding * binding) option -> theory ->
+    term -> Typedef.bindings option -> theory ->
     (Typedef.info * Cpodef.cpo_info * Cpodef.pcpo_info * rep_info) * theory
 
   val domaindef_cmd: (binding * (string * string option) list * mixfix) * string
-    * (binding * binding) option -> theory -> theory
+    * Typedef.bindings option -> theory -> theory
 end
 
 structure Domaindef : DOMAINDEF =
@@ -80,7 +80,7 @@
       (prep_term: Proof.context -> 'a -> term)
       (typ as (tname, raw_args, _) : binding * (string * sort) list * mixfix)
       (raw_defl: 'a)
-      (opt_morphs: (binding * binding) option)
+      (opt_bindings: Typedef.bindings option)
       (thy: theory)
     : (Typedef.info * Cpodef.cpo_info * Cpodef.pcpo_info * rep_info) * theory =
   let
@@ -100,10 +100,6 @@
     val full_tname = Sign.full_name thy tname
     val newT = Type (full_tname, map TFree lhs_tfrees)
 
-    (*morphisms*)
-    val morphs = opt_morphs
-      |> the_default (Binding.prefix_name "Rep_" tname, Binding.prefix_name "Abs_" tname)
-
     (*set*)
     val set = @{term "defl_set :: udom defl => udom set"} $ defl
 
@@ -111,7 +107,7 @@
     fun tac1 ctxt = resolve_tac ctxt @{thms defl_set_bottom} 1
     fun tac2 ctxt = resolve_tac ctxt @{thms adm_defl_set} 1
     val ((info, cpo_info, pcpo_info), thy) = thy
-      |> Cpodef.add_pcpodef typ set (SOME morphs) (tac1, tac2)
+      |> Cpodef.add_pcpodef typ set opt_bindings (tac1, tac2)
 
     (*definitions*)
     val Rep_const = Const (#Rep_name (#1 info), newT --> udomT)
@@ -187,8 +183,8 @@
   handle ERROR msg =>
     cat_error msg ("The error(s) above occurred in domaindef " ^ Binding.print tname)
 
-fun add_domaindef typ defl opt_morphs thy =
-  gen_add_domaindef Syntax.check_term typ defl opt_morphs thy
+fun add_domaindef typ defl opt_bindings thy =
+  gen_add_domaindef Syntax.check_term typ defl opt_bindings thy
 
 fun domaindef_cmd ((b, raw_args, mx), A, morphs) thy =
   let
@@ -199,17 +195,13 @@
 
 (** outer syntax **)
 
-val domaindef_decl =
-  (Parse.type_args_constrained -- Parse.binding) --
-  Parse.opt_mixfix -- (@{keyword "="} |-- Parse.term) --
-  Scan.option
-    (@{keyword "morphisms"} |-- Parse.!!! (Parse.binding -- Parse.binding))
-
-fun mk_domaindef (((((args, t)), mx), A), morphs) =
-  domaindef_cmd ((t, args, mx), A, morphs)
-
 val _ =
   Outer_Syntax.command @{command_keyword domaindef} "HOLCF definition of domains from deflations"
-    (domaindef_decl >> (Toplevel.theory o mk_domaindef))
+    ((Parse.type_args_constrained -- Parse.binding) --
+      Parse.opt_mixfix -- (@{keyword "="} |-- Parse.term) --
+      Scan.option
+        (@{keyword "morphisms"} |-- Parse.!!! (Parse.binding -- Parse.binding)) >>
+     (fn (((((args, t)), mx), A), morphs) =>
+      Toplevel.theory (domaindef_cmd ((t, args, mx), A, SOME (Typedef.make_morphisms t morphs)))));
 
 end
--- a/src/HOL/Import/import_rule.ML	Thu Sep 03 19:27:45 2015 +0200
+++ b/src/HOL/Import/import_rule.ML	Thu Sep 03 21:50:39 2015 +0200
@@ -218,9 +218,12 @@
       | _ => error "type_introduction: bad type definition theorem"
     val tfrees = Term.add_tfrees c []
     val tnames = sort_strings (map fst tfrees)
+    val typedef_bindings =
+      Typedef.make_morphisms (Binding.name tycname)
+        (SOME (Binding.name rep_name, Binding.name abs_name))
     val ((_, typedef_info), thy') =
-     Typedef.add_typedef_global false (Binding.name tycname, map (rpair dummyS) tnames, NoSyn) c
-       (SOME (Binding.name rep_name, Binding.name abs_name)) (fn ctxt => resolve_tac ctxt [th2] 1) thy
+     Typedef.add_typedef_global (Binding.name tycname, map (rpair dummyS) tnames, NoSyn) c
+       (SOME typedef_bindings) (fn ctxt => resolve_tac ctxt [th2] 1) thy
     val aty = #abs_type (#1 typedef_info)
     val th = freezeT thy' (#type_definition (#2 typedef_info))
     val (th_s, _) = Thm.dest_comb (Thm.dest_arg (Thm.cprop_of th))
--- a/src/HOL/Nominal/nominal_datatype.ML	Thu Sep 03 19:27:45 2015 +0200
+++ b/src/HOL/Nominal/nominal_datatype.ML	Thu Sep 03 21:50:39 2015 +0200
@@ -582,7 +582,7 @@
     val (typedefs, thy6) =
       thy4
       |> fold_map (fn (((name, mx), tvs), (cname, U)) => fn thy =>
-          Typedef.add_typedef_global false
+          Typedef.add_typedef_global
             (name, map (fn (v, _) => (v, dummyS)) tvs, mx)  (* FIXME keep constraints!? *)
             (Const (@{const_name Collect}, (U --> HOLogic.boolT) --> HOLogic.mk_setT U) $
                Const (cname, U --> HOLogic.boolT)) NONE
--- a/src/HOL/Tools/BNF/bnf_util.ML	Thu Sep 03 19:27:45 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_util.ML	Thu Sep 03 21:50:39 2015 +0200
@@ -143,10 +143,20 @@
   let
     (*Work around loss of qualification in "typedef" axioms by replicating it in the name*)
     val b' = fold_rev Binding.prefix_name (map (suffix "_" o fst) (Binding.path_of b)) b;
+
+    val default_bindings = Typedef.default_bindings (Binding.concealed b');
+    val bindings =
+      (case opt_morphs of
+        NONE => default_bindings
+      | SOME (Rep_name, Abs_name) =>
+         {Rep_name = Binding.concealed Rep_name,
+          Abs_name = Binding.concealed Abs_name,
+          type_definition_name = #type_definition_name default_bindings});
+
     val ((name, info), (lthy, lthy_old)) =
       lthy
       |> Local_Theory.open_target |> snd
-      |> Typedef.add_typedef true (b', Ts, mx) set opt_morphs tac
+      |> Typedef.add_typedef (b', Ts, mx) set (SOME bindings) tac
       ||> `Local_Theory.close_target;
     val phi = Proof_Context.export_morphism lthy_old lthy;
   in
--- a/src/HOL/Tools/Old_Datatype/old_datatype.ML	Thu Sep 03 19:27:45 2015 +0200
+++ b/src/HOL/Tools/Old_Datatype/old_datatype.ML	Thu Sep 03 21:50:39 2015 +0200
@@ -186,7 +186,7 @@
       |> Sign.parent_path
       |> fold_map
         (fn (((name, mx), tvs), c) =>
-          Typedef.add_typedef_global false (name, tvs, mx)
+          Typedef.add_typedef_global (name, tvs, mx)
             (Collect $ Const (c, UnivT')) NONE
             (fn ctxt =>
               resolve_tac ctxt [exI] 1 THEN
--- a/src/HOL/Tools/Quotient/quotient_type.ML	Thu Sep 03 19:27:45 2015 +0200
+++ b/src/HOL/Tools/Quotient/quotient_type.ML	Thu Sep 03 21:50:39 2015 +0200
@@ -45,7 +45,7 @@
     fun typedef_tac ctxt =
       EVERY1 (map (resolve_tac ctxt o single) [@{thm part_equivp_typedef}, equiv_thm])
   in
-    Typedef.add_typedef false (qty_name, map (rpair dummyS) vs, mx)
+    Typedef.add_typedef (qty_name, map (rpair dummyS) vs, mx)
       (typedef_term rel rty lthy) NONE typedef_tac lthy
   end
 
--- a/src/HOL/Tools/record.ML	Thu Sep 03 19:27:45 2015 +0200
+++ b/src/HOL/Tools/record.ML	Thu Sep 03 21:50:39 2015 +0200
@@ -99,7 +99,7 @@
     val vs = map (Proof_Context.check_tfree ctxt) raw_vs;
   in
     thy
-    |> Typedef.add_typedef_global false (raw_tyco, vs, NoSyn)
+    |> Typedef.add_typedef_global (raw_tyco, vs, NoSyn)
         (HOLogic.mk_UNIV repT) NONE (fn ctxt' => resolve_tac ctxt' [UNIV_witness] 1)
     |-> (fn (tyco, info) => get_typedef_info tyco vs info)
   end;
--- a/src/HOL/Tools/typedef.ML	Thu Sep 03 19:27:45 2015 +0200
+++ b/src/HOL/Tools/typedef.ML	Thu Sep 03 21:50:39 2015 +0200
@@ -16,16 +16,20 @@
   val get_info: Proof.context -> string -> info list
   val get_info_global: theory -> string -> info list
   val interpretation: (string -> local_theory -> local_theory) -> theory -> theory
-  val add_typedef: bool -> binding * (string * sort) list * mixfix ->
-    term -> (binding * binding) option -> (Proof.context -> tactic) -> local_theory ->
+  type bindings = {Rep_name: binding, Abs_name: binding, type_definition_name: binding}
+  val default_bindings: binding -> bindings
+  val make_bindings: binding -> bindings option -> bindings
+  val make_morphisms: binding -> (binding * binding) option -> bindings
+  val add_typedef: binding * (string * sort) list * mixfix ->
+    term -> bindings option -> (Proof.context -> tactic) -> local_theory ->
     (string * info) * local_theory
-  val add_typedef_global: bool -> binding * (string * sort) list * mixfix ->
-    term -> (binding * binding) option -> (Proof.context -> tactic) -> theory ->
+  val add_typedef_global: binding * (string * sort) list * mixfix ->
+    term -> bindings option -> (Proof.context -> tactic) -> theory ->
     (string * info) * theory
-  val typedef: (binding * (string * sort) list * mixfix) * term *
-    (binding * binding) option -> local_theory -> Proof.state
-  val typedef_cmd: (binding * (string * string option) list * mixfix) * string *
-    (binding * binding) option -> local_theory -> Proof.state
+  val typedef: binding * (string * sort) list * mixfix -> term -> bindings option ->
+    local_theory -> Proof.state
+  val typedef_cmd: binding * (string * string option) list * mixfix -> string -> bindings option ->
+    local_theory -> Proof.state
 end;
 
 structure Typedef: TYPEDEF =
@@ -105,7 +109,7 @@
         (newT --> oldT) --> (oldT --> newT) --> HOLogic.mk_setT oldT --> HOLogic.boolT);
   in Logic.mk_implies (mk_inhabited A, HOLogic.mk_Trueprop (typedefC $ RepC $ AbsC $ A)) end;
 
-fun primitive_typedef typedef_name newT oldT Rep_name Abs_name A lthy =
+fun primitive_typedef type_definition_name newT oldT Rep_name Abs_name A lthy =
   let
     (* errors *)
 
@@ -133,18 +137,35 @@
 
     val ((axiom_name, axiom), axiom_lthy) = consts_lthy
       |> Local_Theory.background_theory_result
-        (Thm.add_axiom consts_lthy (typedef_name, mk_typedef newT oldT RepC AbsC A) ##>
+        (Thm.add_axiom consts_lthy (type_definition_name, mk_typedef newT oldT RepC AbsC A) ##>
           Theory.add_deps consts_lthy "" (dest_Const RepC) typedef_deps ##>
           Theory.add_deps consts_lthy "" (dest_Const AbsC) typedef_deps);
 
   in ((RepC, AbsC, axiom_name, axiom), axiom_lthy) end;
 
 
+(* derived bindings *)
+
+type bindings = {Rep_name: binding, Abs_name: binding, type_definition_name: binding};
+
+fun default_bindings name =
+ {Rep_name = Binding.prefix_name "Rep_" name,
+  Abs_name = Binding.prefix_name "Abs_" name,
+  type_definition_name = Binding.prefix_name "type_definition_" name};
+
+fun make_bindings name NONE = default_bindings name
+  | make_bindings _ (SOME bindings) = bindings;
+
+fun make_morphisms name NONE = default_bindings name
+  | make_morphisms name (SOME (Rep_name, Abs_name)) =
+      {Rep_name = Rep_name, Abs_name = Abs_name,
+        type_definition_name = #type_definition_name (default_bindings name)};
+
+
 (* prepare_typedef *)
 
-fun prepare_typedef concealed prep_term (name, raw_args, mx) raw_set opt_morphs lthy =
+fun prepare_typedef prep_term (name, raw_args, mx) raw_set opt_bindings lthy =
   let
-    val concealed_name = name |> concealed ? Binding.concealed;
     val bname = Binding.name_of name;
 
 
@@ -174,16 +195,10 @@
 
     (* axiomatization *)
 
-    val (Rep_name, Abs_name) =
-      (case opt_morphs of
-        NONE => (Binding.prefix_name "Rep_" concealed_name,
-          Binding.prefix_name "Abs_" concealed_name)
-      | SOME morphs => morphs);
-
-    val typedef_name = Binding.prefix_name "type_definition_" concealed_name;
+    val {Rep_name, Abs_name, type_definition_name} = make_bindings name opt_bindings;
 
     val ((RepC, AbsC, axiom_name, typedef), typedef_lthy) = typedecl_lthy
-      |> primitive_typedef typedef_name newT oldT Rep_name Abs_name set;
+      |> primitive_typedef type_definition_name newT oldT Rep_name Abs_name set;
 
     val alias_lthy = typedef_lthy
       |> Local_Theory.const_alias Rep_name (#1 (Term.dest_Const RepC))
@@ -202,7 +217,7 @@
         fun make th = Goal.norm_result lthy1 (typedef' RS th);
         val (((((((((((_, [type_definition]), Rep), Rep_inverse), Abs_inverse), Rep_inject),
             Abs_inject), Rep_cases), Abs_cases), Rep_induct), Abs_induct), lthy2) = lthy1
-          |> Local_Theory.note ((typedef_name, []), [typedef'])
+          |> Local_Theory.note ((type_definition_name, []), [typedef'])
           ||>> note_qualify ((Rep_name, []), make @{thm type_definition.Rep})
           ||>> note_qualify ((Binding.suffix_name "_inverse" Rep_name, []),
               make @{thm type_definition.Rep_inverse})
@@ -247,18 +262,18 @@
 
 (* add_typedef: tactic interface *)
 
-fun add_typedef concealed typ set opt_morphs tac lthy =
+fun add_typedef typ set opt_bindings tac lthy =
   let
     val ((goal, _, typedef_result), lthy') =
-      prepare_typedef concealed Syntax.check_term typ set opt_morphs lthy;
+      prepare_typedef Syntax.check_term typ set opt_bindings lthy;
     val inhabited =
       Goal.prove lthy' [] [] goal (tac o #context)
       |> Goal.norm_result lthy' |> Thm.close_derivation;
   in typedef_result inhabited lthy' end;
 
-fun add_typedef_global concealed typ set opt_morphs tac =
+fun add_typedef_global typ set opt_bindings tac =
   Named_Target.theory_init
-  #> add_typedef concealed typ set opt_morphs tac
+  #> add_typedef typ set opt_bindings tac
   #> Local_Theory.exit_result_global (apsnd o transform_info);
 
 
@@ -266,11 +281,11 @@
 
 local
 
-fun gen_typedef prep_term prep_constraint ((b, raw_args, mx), set, opt_morphs) lthy =
+fun gen_typedef prep_term prep_constraint (b, raw_args, mx) set opt_bindings lthy =
   let
     val args = map (apsnd (prep_constraint lthy)) raw_args;
     val ((goal, goal_pat, typedef_result), lthy') =
-      prepare_typedef false prep_term (b, args, mx) set opt_morphs lthy;
+      prepare_typedef prep_term (b, args, mx) set opt_bindings lthy;
     fun after_qed [[th]] = snd o typedef_result th;
   in Proof.theorem NONE after_qed [[(goal, [goal_pat])]] lthy' end;
 
@@ -291,6 +306,7 @@
     (Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix --
       (@{keyword "="} |-- Parse.term) --
       Scan.option (@{keyword "morphisms"} |-- Parse.!!! (Parse.binding -- Parse.binding))
-    >> (fn ((((vs, t), mx), A), morphs) => fn lthy => typedef_cmd ((t, vs, mx), A, morphs) lthy));
+    >> (fn ((((vs, t), mx), A), opt_morphs) => fn lthy =>
+        typedef_cmd (t, vs, mx) A (SOME (make_morphisms t opt_morphs)) lthy));
 
 end;