rename 'repdef' to 'domaindef'
authorhuffman
Tue, 16 Nov 2010 11:50:05 -0800
changeset 40575 b9a86f15e763
parent 40566 36d4f2757f4f
child 40576 fa5e0cacd5e7
rename 'repdef' to 'domaindef'
src/HOLCF/Domain.thy
src/HOLCF/IsaMakefile
src/HOLCF/Tools/Domain/domain_isomorphism.ML
src/HOLCF/Tools/domaindef.ML
src/HOLCF/Tools/repdef.ML
src/HOLCF/ex/Domain_Proofs.thy
--- a/src/HOLCF/Domain.thy	Tue Nov 16 15:29:01 2010 +0100
+++ b/src/HOLCF/Domain.thy	Tue Nov 16 11:50:05 2010 -0800
@@ -7,7 +7,7 @@
 theory Domain
 imports Bifinite Domain_Aux
 uses
-  ("Tools/repdef.ML")
+  ("Tools/domaindef.ML")
   ("Tools/Domain/domain_isomorphism.ML")
   ("Tools/Domain/domain_axioms.ML")
   ("Tools/Domain/domain.ML")
@@ -94,7 +94,7 @@
   , (@{const_name liftprj}, SOME @{typ "udom \<rightarrow> 'a::pcpo u"}) ]
 *}
 
-lemma typedef_rep_class:
+lemma typedef_liftdomain_class:
   fixes Rep :: "'a::pcpo \<Rightarrow> udom"
   fixes Abs :: "udom \<Rightarrow> 'a::pcpo"
   fixes t :: defl
@@ -156,7 +156,7 @@
   , (@{const_name liftprj}, SOME @{typ "udom \<rightarrow> 'a::predomain u"}) ]
 *}
 
-use "Tools/repdef.ML"
+use "Tools/domaindef.ML"
 
 subsection {* Isomorphic deflations *}
 
--- a/src/HOLCF/IsaMakefile	Tue Nov 16 15:29:01 2010 +0100
+++ b/src/HOLCF/IsaMakefile	Tue Nov 16 11:50:05 2010 -0800
@@ -77,9 +77,9 @@
   Tools/Domain/domain_induction.ML \
   Tools/Domain/domain_isomorphism.ML \
   Tools/Domain/domain_take_proofs.ML \
+  Tools/domaindef.ML \
   Tools/fixrec.ML \
   Tools/pcpodef.ML \
-  Tools/repdef.ML \
   document/root.tex
 	@$(ISABELLE_TOOL) usedir -b -g true -r $(OUT)/HOL HOLCF
 
--- a/src/HOLCF/Tools/Domain/domain_isomorphism.ML	Tue Nov 16 15:29:01 2010 +0100
+++ b/src/HOLCF/Tools/Domain/domain_isomorphism.ML	Tue Nov 16 11:50:05 2010 -0800
@@ -523,7 +523,7 @@
       let
         val spec = (tbind, map (rpair dummyS) vs, mx);
         val ((_, _, _, {DEFL, liftemb_def, liftprj_def, ...}), thy) =
-          Repdef.add_repdef false NONE spec defl NONE thy;
+          Domaindef.add_domaindef false NONE spec defl NONE thy;
         (* declare domain_defl_simps rules *)
         val thy = Context.theory_map (RepData.add_thm DEFL) thy;
       in
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Tools/domaindef.ML	Tue Nov 16 11:50:05 2010 -0800
@@ -0,0 +1,236 @@
+(*  Title:      HOLCF/Tools/repdef.ML
+    Author:     Brian Huffman
+
+Defining representable domains using algebraic deflations.
+*)
+
+signature DOMAINDEF =
+sig
+  type rep_info =
+    {
+      emb_def : thm,
+      prj_def : thm,
+      defl_def : thm,
+      liftemb_def : thm,
+      liftprj_def : thm,
+      liftdefl_def : thm,
+      DEFL : thm
+    }
+
+  val add_domaindef: bool -> binding option -> binding * (string * sort) list * mixfix ->
+    term -> (binding * binding) option -> theory ->
+    (Typedef.info * Pcpodef.cpo_info * Pcpodef.pcpo_info * rep_info) * theory
+
+  val domaindef_cmd: (bool * binding) * (binding * (string * string option) list * mixfix) * string
+    * (binding * binding) option -> theory -> theory
+end;
+
+structure Domaindef :> DOMAINDEF =
+struct
+
+open HOLCF_Library;
+
+infixr 6 ->>;
+infix -->>;
+
+(** type definitions **)
+
+type rep_info =
+  {
+    emb_def : thm,
+    prj_def : thm,
+    defl_def : thm,
+    liftemb_def : thm,
+    liftprj_def : thm,
+    liftdefl_def : thm,
+    DEFL : thm
+  };
+
+(* building types and terms *)
+
+val udomT = @{typ udom};
+val deflT = @{typ defl};
+fun emb_const T = Const (@{const_name emb}, T ->> udomT);
+fun prj_const T = Const (@{const_name prj}, udomT ->> T);
+fun defl_const T = Const (@{const_name defl}, Term.itselfT T --> deflT);
+fun liftemb_const T = Const (@{const_name liftemb}, mk_upT T ->> udomT);
+fun liftprj_const T = Const (@{const_name liftprj}, udomT ->> mk_upT T);
+fun liftdefl_const T = Const (@{const_name liftdefl}, Term.itselfT T --> deflT);
+
+fun mk_u_map t =
+  let
+    val (T, U) = dest_cfunT (fastype_of t);
+    val u_map_type = (T ->> U) ->> (mk_upT T ->> mk_upT U);
+    val u_map_const = Const (@{const_name u_map}, u_map_type);
+  in
+    mk_capply (u_map_const, t)
+  end;
+
+fun mk_cast (t, x) =
+  capply_const (udomT, udomT)
+  $ (capply_const (deflT, udomT ->> udomT) $ @{const cast} $ t)
+  $ x;
+
+(* manipulating theorems *)
+
+(* proving class instances *)
+
+fun declare_type_name a =
+  Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS)));
+
+fun gen_add_domaindef
+      (prep_term: Proof.context -> 'a -> term)
+      (def: bool)
+      (name: binding)
+      (typ as (tname, raw_args, mx) : binding * (string * sort) list * mixfix)
+      (raw_defl: 'a)
+      (opt_morphs: (binding * binding) option)
+      (thy: theory)
+    : (Typedef.info * Pcpodef.cpo_info * Pcpodef.pcpo_info * rep_info) * theory =
+  let
+    val _ = Theory.requires thy "Domain" "domaindefs";
+
+    (*rhs*)
+    val tmp_ctxt =
+      ProofContext.init_global thy
+      |> fold (Variable.declare_typ o TFree) raw_args;
+    val defl = prep_term tmp_ctxt raw_defl;
+    val tmp_ctxt = tmp_ctxt |> Variable.declare_constraints defl;
+
+    val deflT = Term.fastype_of defl;
+    val _ = if deflT = @{typ "defl"} then ()
+            else error ("Not type defl: " ^ quote (Syntax.string_of_typ tmp_ctxt deflT));
+
+    (*lhs*)
+    val lhs_tfrees = map (ProofContext.check_tfree tmp_ctxt) raw_args;
+    val lhs_sorts = map snd lhs_tfrees;
+    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_" name, Binding.prefix_name "Abs_" name);
+
+    (*set*)
+    val set = @{const defl_set} $ defl;
+
+    (*pcpodef*)
+    val tac1 = rtac @{thm defl_set_bottom} 1;
+    val tac2 = rtac @{thm adm_defl_set} 1;
+    val ((info, cpo_info, pcpo_info), thy) = thy
+      |> Pcpodef.add_pcpodef def (SOME name) typ set (SOME morphs) (tac1, tac2);
+
+    (*definitions*)
+    val Rep_const = Const (#Rep_name (#1 info), newT --> udomT);
+    val Abs_const = Const (#Abs_name (#1 info), udomT --> newT);
+    val emb_eqn = Logic.mk_equals (emb_const newT, cabs_const (newT, udomT) $ Rep_const);
+    val prj_eqn = Logic.mk_equals (prj_const newT, cabs_const (udomT, newT) $
+      Abs ("x", udomT, Abs_const $ mk_cast (defl, Bound 0)));
+    val defl_eqn = Logic.mk_equals (defl_const newT,
+      Abs ("x", Term.itselfT newT, defl));
+    val liftemb_eqn =
+      Logic.mk_equals (liftemb_const newT,
+      mk_cfcomp (@{term "udom_emb u_approx"}, mk_u_map (emb_const newT)));
+    val liftprj_eqn =
+      Logic.mk_equals (liftprj_const newT,
+      mk_cfcomp (mk_u_map (prj_const newT), @{term "udom_prj u_approx"}));
+    val liftdefl_eqn =
+      Logic.mk_equals (liftdefl_const newT,
+        Abs ("t", Term.itselfT newT,
+          mk_capply (@{const u_defl}, defl_const newT $ Logic.mk_type newT)));
+
+    val name_def = Binding.suffix_name "_def" name;
+    val emb_bind = (Binding.prefix_name "emb_" name_def, []);
+    val prj_bind = (Binding.prefix_name "prj_" name_def, []);
+    val defl_bind = (Binding.prefix_name "defl_" name_def, []);
+    val liftemb_bind = (Binding.prefix_name "liftemb_" name_def, []);
+    val liftprj_bind = (Binding.prefix_name "liftprj_" name_def, []);
+    val liftdefl_bind = (Binding.prefix_name "liftdefl_" name_def, []);
+
+    (*instantiate class rep*)
+    val lthy = thy
+      |> Class.instantiation ([full_tname], lhs_tfrees, @{sort liftdomain});
+    val ((_, (_, emb_ldef)), lthy) =
+        Specification.definition (NONE, (emb_bind, emb_eqn)) lthy;
+    val ((_, (_, prj_ldef)), lthy) =
+        Specification.definition (NONE, (prj_bind, prj_eqn)) lthy;
+    val ((_, (_, defl_ldef)), lthy) =
+        Specification.definition (NONE, (defl_bind, defl_eqn)) lthy;
+    val ((_, (_, liftemb_ldef)), lthy) =
+        Specification.definition (NONE, (liftemb_bind, liftemb_eqn)) lthy;
+    val ((_, (_, liftprj_ldef)), lthy) =
+        Specification.definition (NONE, (liftprj_bind, liftprj_eqn)) lthy;
+    val ((_, (_, liftdefl_ldef)), lthy) =
+        Specification.definition (NONE, (liftdefl_bind, liftdefl_eqn)) lthy;
+    val ctxt_thy = ProofContext.init_global (ProofContext.theory_of lthy);
+    val emb_def = singleton (ProofContext.export lthy ctxt_thy) emb_ldef;
+    val prj_def = singleton (ProofContext.export lthy ctxt_thy) prj_ldef;
+    val defl_def = singleton (ProofContext.export lthy ctxt_thy) defl_ldef;
+    val liftemb_def = singleton (ProofContext.export lthy ctxt_thy) liftemb_ldef;
+    val liftprj_def = singleton (ProofContext.export lthy ctxt_thy) liftprj_ldef;
+    val liftdefl_def = singleton (ProofContext.export lthy ctxt_thy) liftdefl_ldef;
+    val type_definition_thm =
+      MetaSimplifier.rewrite_rule
+        (the_list (#set_def (#2 info)))
+        (#type_definition (#2 info));
+    val typedef_thms =
+      [type_definition_thm, #below_def cpo_info, emb_def, prj_def, defl_def,
+      liftemb_def, liftprj_def, liftdefl_def];
+    val thy = lthy
+      |> Class.prove_instantiation_instance
+          (K (Tactic.rtac (@{thm typedef_liftdomain_class} OF typedef_thms) 1))
+      |> Local_Theory.exit_global;
+
+    (*other theorems*)
+    val defl_thm' = Thm.transfer thy defl_def;
+    val (DEFL_thm, thy) = thy
+      |> Sign.add_path (Binding.name_of name)
+      |> Global_Theory.add_thm
+         ((Binding.prefix_name "DEFL_" name,
+          Drule.zero_var_indexes (@{thm typedef_DEFL} OF [defl_thm'])), [])
+      ||> Sign.restore_naming thy;
+
+    val rep_info =
+      { emb_def = emb_def, prj_def = prj_def, defl_def = defl_def,
+        liftemb_def = liftemb_def, liftprj_def = liftprj_def,
+        liftdefl_def = liftdefl_def, DEFL = DEFL_thm };
+  in
+    ((info, cpo_info, pcpo_info, rep_info), thy)
+  end
+  handle ERROR msg =>
+    cat_error msg ("The error(s) above occurred in domaindef " ^ quote (Binding.str_of name));
+
+fun add_domaindef def opt_name typ defl opt_morphs thy =
+  let
+    val name = the_default (#1 typ) opt_name;
+  in
+    gen_add_domaindef Syntax.check_term def name typ defl opt_morphs thy
+  end;
+
+fun domaindef_cmd ((def, name), (b, raw_args, mx), A, morphs) thy =
+  let
+    val ctxt = ProofContext.init_global thy;
+    val args = map (apsnd (Typedecl.read_constraint ctxt)) raw_args;
+  in snd (gen_add_domaindef Syntax.read_term def name (b, args, mx) A morphs thy) end;
+
+
+(** outer syntax **)
+
+val domaindef_decl =
+  Scan.optional (Parse.$$$ "(" |--
+      ((Parse.$$$ "open" >> K false) -- Scan.option Parse.binding ||
+        Parse.binding >> (fn s => (true, SOME s)))
+        --| Parse.$$$ ")") (true, NONE) --
+    (Parse.type_args_constrained -- Parse.binding) --
+    Parse.opt_mixfix -- (Parse.$$$ "=" |-- Parse.term) --
+    Scan.option (Parse.$$$ "morphisms" |-- Parse.!!! (Parse.binding -- Parse.binding));
+
+fun mk_domaindef ((((((def, opt_name), (args, t)), mx), A), morphs)) =
+  domaindef_cmd ((def, the_default t opt_name), (t, args, mx), A, morphs);
+
+val _ =
+  Outer_Syntax.command "domaindef" "HOLCF definition of domains from deflations" Keyword.thy_decl
+    (domaindef_decl >>
+      (Toplevel.print oo (Toplevel.theory o mk_domaindef)));
+
+end;
--- a/src/HOLCF/Tools/repdef.ML	Tue Nov 16 15:29:01 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,236 +0,0 @@
-(*  Title:      HOLCF/Tools/repdef.ML
-    Author:     Brian Huffman
-
-Defining representable domains using algebraic deflations.
-*)
-
-signature REPDEF =
-sig
-  type rep_info =
-    {
-      emb_def : thm,
-      prj_def : thm,
-      defl_def : thm,
-      liftemb_def : thm,
-      liftprj_def : thm,
-      liftdefl_def : thm,
-      DEFL : thm
-    }
-
-  val add_repdef: bool -> binding option -> binding * (string * sort) list * mixfix ->
-    term -> (binding * binding) option -> theory ->
-    (Typedef.info * Pcpodef.cpo_info * Pcpodef.pcpo_info * rep_info) * theory
-
-  val repdef_cmd: (bool * binding) * (binding * (string * string option) list * mixfix) * string
-    * (binding * binding) option -> theory -> theory
-end;
-
-structure Repdef :> REPDEF =
-struct
-
-open HOLCF_Library;
-
-infixr 6 ->>;
-infix -->>;
-
-(** type definitions **)
-
-type rep_info =
-  {
-    emb_def : thm,
-    prj_def : thm,
-    defl_def : thm,
-    liftemb_def : thm,
-    liftprj_def : thm,
-    liftdefl_def : thm,
-    DEFL : thm
-  };
-
-(* building types and terms *)
-
-val udomT = @{typ udom};
-val deflT = @{typ defl};
-fun emb_const T = Const (@{const_name emb}, T ->> udomT);
-fun prj_const T = Const (@{const_name prj}, udomT ->> T);
-fun defl_const T = Const (@{const_name defl}, Term.itselfT T --> deflT);
-fun liftemb_const T = Const (@{const_name liftemb}, mk_upT T ->> udomT);
-fun liftprj_const T = Const (@{const_name liftprj}, udomT ->> mk_upT T);
-fun liftdefl_const T = Const (@{const_name liftdefl}, Term.itselfT T --> deflT);
-
-fun mk_u_map t =
-  let
-    val (T, U) = dest_cfunT (fastype_of t);
-    val u_map_type = (T ->> U) ->> (mk_upT T ->> mk_upT U);
-    val u_map_const = Const (@{const_name u_map}, u_map_type);
-  in
-    mk_capply (u_map_const, t)
-  end;
-
-fun mk_cast (t, x) =
-  capply_const (udomT, udomT)
-  $ (capply_const (deflT, udomT ->> udomT) $ @{const cast} $ t)
-  $ x;
-
-(* manipulating theorems *)
-
-(* proving class instances *)
-
-fun declare_type_name a =
-  Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS)));
-
-fun gen_add_repdef
-      (prep_term: Proof.context -> 'a -> term)
-      (def: bool)
-      (name: binding)
-      (typ as (tname, raw_args, mx) : binding * (string * sort) list * mixfix)
-      (raw_defl: 'a)
-      (opt_morphs: (binding * binding) option)
-      (thy: theory)
-    : (Typedef.info * Pcpodef.cpo_info * Pcpodef.pcpo_info * rep_info) * theory =
-  let
-    val _ = Theory.requires thy "Domain" "repdefs";
-
-    (*rhs*)
-    val tmp_ctxt =
-      ProofContext.init_global thy
-      |> fold (Variable.declare_typ o TFree) raw_args;
-    val defl = prep_term tmp_ctxt raw_defl;
-    val tmp_ctxt = tmp_ctxt |> Variable.declare_constraints defl;
-
-    val deflT = Term.fastype_of defl;
-    val _ = if deflT = @{typ "defl"} then ()
-            else error ("Not type defl: " ^ quote (Syntax.string_of_typ tmp_ctxt deflT));
-
-    (*lhs*)
-    val lhs_tfrees = map (ProofContext.check_tfree tmp_ctxt) raw_args;
-    val lhs_sorts = map snd lhs_tfrees;
-    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_" name, Binding.prefix_name "Abs_" name);
-
-    (*set*)
-    val set = @{const defl_set} $ defl;
-
-    (*pcpodef*)
-    val tac1 = rtac @{thm defl_set_bottom} 1;
-    val tac2 = rtac @{thm adm_defl_set} 1;
-    val ((info, cpo_info, pcpo_info), thy) = thy
-      |> Pcpodef.add_pcpodef def (SOME name) typ set (SOME morphs) (tac1, tac2);
-
-    (*definitions*)
-    val Rep_const = Const (#Rep_name (#1 info), newT --> udomT);
-    val Abs_const = Const (#Abs_name (#1 info), udomT --> newT);
-    val emb_eqn = Logic.mk_equals (emb_const newT, cabs_const (newT, udomT) $ Rep_const);
-    val prj_eqn = Logic.mk_equals (prj_const newT, cabs_const (udomT, newT) $
-      Abs ("x", udomT, Abs_const $ mk_cast (defl, Bound 0)));
-    val defl_eqn = Logic.mk_equals (defl_const newT,
-      Abs ("x", Term.itselfT newT, defl));
-    val liftemb_eqn =
-      Logic.mk_equals (liftemb_const newT,
-      mk_cfcomp (@{term "udom_emb u_approx"}, mk_u_map (emb_const newT)));
-    val liftprj_eqn =
-      Logic.mk_equals (liftprj_const newT,
-      mk_cfcomp (mk_u_map (prj_const newT), @{term "udom_prj u_approx"}));
-    val liftdefl_eqn =
-      Logic.mk_equals (liftdefl_const newT,
-        Abs ("t", Term.itselfT newT,
-          mk_capply (@{const u_defl}, defl_const newT $ Logic.mk_type newT)));
-
-    val name_def = Binding.suffix_name "_def" name;
-    val emb_bind = (Binding.prefix_name "emb_" name_def, []);
-    val prj_bind = (Binding.prefix_name "prj_" name_def, []);
-    val defl_bind = (Binding.prefix_name "defl_" name_def, []);
-    val liftemb_bind = (Binding.prefix_name "liftemb_" name_def, []);
-    val liftprj_bind = (Binding.prefix_name "liftprj_" name_def, []);
-    val liftdefl_bind = (Binding.prefix_name "liftdefl_" name_def, []);
-
-    (*instantiate class rep*)
-    val lthy = thy
-      |> Class.instantiation ([full_tname], lhs_tfrees, @{sort liftdomain});
-    val ((_, (_, emb_ldef)), lthy) =
-        Specification.definition (NONE, (emb_bind, emb_eqn)) lthy;
-    val ((_, (_, prj_ldef)), lthy) =
-        Specification.definition (NONE, (prj_bind, prj_eqn)) lthy;
-    val ((_, (_, defl_ldef)), lthy) =
-        Specification.definition (NONE, (defl_bind, defl_eqn)) lthy;
-    val ((_, (_, liftemb_ldef)), lthy) =
-        Specification.definition (NONE, (liftemb_bind, liftemb_eqn)) lthy;
-    val ((_, (_, liftprj_ldef)), lthy) =
-        Specification.definition (NONE, (liftprj_bind, liftprj_eqn)) lthy;
-    val ((_, (_, liftdefl_ldef)), lthy) =
-        Specification.definition (NONE, (liftdefl_bind, liftdefl_eqn)) lthy;
-    val ctxt_thy = ProofContext.init_global (ProofContext.theory_of lthy);
-    val emb_def = singleton (ProofContext.export lthy ctxt_thy) emb_ldef;
-    val prj_def = singleton (ProofContext.export lthy ctxt_thy) prj_ldef;
-    val defl_def = singleton (ProofContext.export lthy ctxt_thy) defl_ldef;
-    val liftemb_def = singleton (ProofContext.export lthy ctxt_thy) liftemb_ldef;
-    val liftprj_def = singleton (ProofContext.export lthy ctxt_thy) liftprj_ldef;
-    val liftdefl_def = singleton (ProofContext.export lthy ctxt_thy) liftdefl_ldef;
-    val type_definition_thm =
-      MetaSimplifier.rewrite_rule
-        (the_list (#set_def (#2 info)))
-        (#type_definition (#2 info));
-    val typedef_thms =
-      [type_definition_thm, #below_def cpo_info, emb_def, prj_def, defl_def,
-      liftemb_def, liftprj_def, liftdefl_def];
-    val thy = lthy
-      |> Class.prove_instantiation_instance
-          (K (Tactic.rtac (@{thm typedef_rep_class} OF typedef_thms) 1))
-      |> Local_Theory.exit_global;
-
-    (*other theorems*)
-    val defl_thm' = Thm.transfer thy defl_def;
-    val (DEFL_thm, thy) = thy
-      |> Sign.add_path (Binding.name_of name)
-      |> Global_Theory.add_thm
-         ((Binding.prefix_name "DEFL_" name,
-          Drule.zero_var_indexes (@{thm typedef_DEFL} OF [defl_thm'])), [])
-      ||> Sign.restore_naming thy;
-
-    val rep_info =
-      { emb_def = emb_def, prj_def = prj_def, defl_def = defl_def,
-        liftemb_def = liftemb_def, liftprj_def = liftprj_def,
-        liftdefl_def = liftdefl_def, DEFL = DEFL_thm };
-  in
-    ((info, cpo_info, pcpo_info, rep_info), thy)
-  end
-  handle ERROR msg =>
-    cat_error msg ("The error(s) above occurred in repdef " ^ quote (Binding.str_of name));
-
-fun add_repdef def opt_name typ defl opt_morphs thy =
-  let
-    val name = the_default (#1 typ) opt_name;
-  in
-    gen_add_repdef Syntax.check_term def name typ defl opt_morphs thy
-  end;
-
-fun repdef_cmd ((def, name), (b, raw_args, mx), A, morphs) thy =
-  let
-    val ctxt = ProofContext.init_global thy;
-    val args = map (apsnd (Typedecl.read_constraint ctxt)) raw_args;
-  in snd (gen_add_repdef Syntax.read_term def name (b, args, mx) A morphs thy) end;
-
-
-(** outer syntax **)
-
-val repdef_decl =
-  Scan.optional (Parse.$$$ "(" |--
-      ((Parse.$$$ "open" >> K false) -- Scan.option Parse.binding ||
-        Parse.binding >> (fn s => (true, SOME s)))
-        --| Parse.$$$ ")") (true, NONE) --
-    (Parse.type_args_constrained -- Parse.binding) --
-    Parse.opt_mixfix -- (Parse.$$$ "=" |-- Parse.term) --
-    Scan.option (Parse.$$$ "morphisms" |-- Parse.!!! (Parse.binding -- Parse.binding));
-
-fun mk_repdef ((((((def, opt_name), (args, t)), mx), A), morphs)) =
-  repdef_cmd ((def, the_default t opt_name), (t, args, mx), A, morphs);
-
-val _ =
-  Outer_Syntax.command "repdef" "HOLCF definition of representable domains" Keyword.thy_decl
-    (repdef_decl >>
-      (Toplevel.print oo (Toplevel.theory o mk_repdef)));
-
-end;
--- a/src/HOLCF/ex/Domain_Proofs.thy	Tue Nov 16 15:29:01 2010 +0100
+++ b/src/HOLCF/ex/Domain_Proofs.thy	Tue Nov 16 11:50:05 2010 -0800
@@ -116,7 +116,7 @@
   "liftdefl \<equiv> \<lambda>(t::'a foo itself). u_defl\<cdot>DEFL('a foo)"
 
 instance
-apply (rule typedef_rep_class)
+apply (rule typedef_liftdomain_class)
 apply (rule type_definition_foo)
 apply (rule below_foo_def)
 apply (rule emb_foo_def)
@@ -151,7 +151,7 @@
   "liftdefl \<equiv> \<lambda>(t::'a bar itself). u_defl\<cdot>DEFL('a bar)"
 
 instance
-apply (rule typedef_rep_class)
+apply (rule typedef_liftdomain_class)
 apply (rule type_definition_bar)
 apply (rule below_bar_def)
 apply (rule emb_bar_def)
@@ -186,7 +186,7 @@
   "liftdefl \<equiv> \<lambda>(t::'a baz itself). u_defl\<cdot>DEFL('a baz)"
 
 instance
-apply (rule typedef_rep_class)
+apply (rule typedef_liftdomain_class)
 apply (rule type_definition_baz)
 apply (rule below_baz_def)
 apply (rule emb_baz_def)