tuning -- moved ML files to subdirectory
authorblanchet
Mon, 09 Dec 2013 09:44:57 +0100
changeset 54701 4ed7454aebde
parent 54700 64177ce0a7bd
child 54702 3daeba5130f0
tuning -- moved ML files to subdirectory
src/HOL/Ctr_Sugar.thy
src/HOL/Tools/Ctr_Sugar/case_translation.ML
src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML
src/HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML
src/HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML
src/HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML
src/HOL/Tools/case_translation.ML
src/HOL/Tools/ctr_sugar.ML
src/HOL/Tools/ctr_sugar_code.ML
src/HOL/Tools/ctr_sugar_tactics.ML
src/HOL/Tools/ctr_sugar_util.ML
--- a/src/HOL/Ctr_Sugar.thy	Mon Dec 09 06:33:46 2013 +0100
+++ b/src/HOL/Ctr_Sugar.thy	Mon Dec 09 09:44:57 2013 +0100
@@ -25,7 +25,7 @@
 declare [[coercion_args case_abs -]]
 declare [[coercion_args case_elem - +]]
 
-ML_file "Tools/case_translation.ML"
+ML_file "Tools/Ctr_Sugar/case_translation.ML"
 setup Case_Translation.setup
 
 lemma iffI_np: "\<lbrakk>x \<Longrightarrow> \<not> y; \<not> x \<Longrightarrow> y\<rbrakk> \<Longrightarrow> \<not> x \<longleftrightarrow> y"
@@ -36,9 +36,9 @@
 "\<not> Q \<Longrightarrow> P \<longleftrightarrow> Q \<Longrightarrow> P \<Longrightarrow> R"
 by blast+
 
-ML_file "Tools/ctr_sugar_util.ML"
-ML_file "Tools/ctr_sugar_tactics.ML"
-ML_file "Tools/ctr_sugar_code.ML"
-ML_file "Tools/ctr_sugar.ML"
+ML_file "Tools/Ctr_Sugar/ctr_sugar_util.ML"
+ML_file "Tools/Ctr_Sugar/ctr_sugar_tactics.ML"
+ML_file "Tools/Ctr_Sugar/ctr_sugar_code.ML"
+ML_file "Tools/Ctr_Sugar/ctr_sugar.ML"
 
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Ctr_Sugar/case_translation.ML	Mon Dec 09 09:44:57 2013 +0100
@@ -0,0 +1,646 @@
+(*  Title:      HOL/Tools/Ctr_Sugar/case_translation.ML
+    Author:     Konrad Slind, Cambridge University Computer Laboratory
+    Author:     Stefan Berghofer, TU Muenchen
+    Author:     Dmitriy Traytel, TU Muenchen
+
+Nested case expressions via a generic data slot for case combinators and constructors.
+*)
+
+signature CASE_TRANSLATION =
+sig
+  val indexify_names: string list -> string list
+  val make_tnames: typ list -> string list
+
+  datatype config = Error | Warning | Quiet
+  val case_tr: bool -> Proof.context -> term list -> term
+  val lookup_by_constr: Proof.context -> string * typ -> (term * term list) option
+  val lookup_by_constr_permissive: Proof.context -> string * typ -> (term * term list) option
+  val lookup_by_case: Proof.context -> string -> (term * term list) option
+  val make_case:  Proof.context -> config -> Name.context -> term -> (term * term) list -> term
+  val print_case_translations: Proof.context -> unit
+  val strip_case: Proof.context -> bool -> term -> (term * (term * term) list) option
+  val strip_case_full: Proof.context -> bool -> term -> term
+  val show_cases: bool Config.T
+  val setup: theory -> theory
+  val register: term -> term list -> Context.generic -> Context.generic
+end;
+
+structure Case_Translation: CASE_TRANSLATION =
+struct
+
+(** general utilities **)
+
+fun indexify_names names =
+  let
+    fun index (x :: xs) tab =
+        (case AList.lookup (op =) tab x of
+          NONE =>
+            if member (op =) xs x
+            then (x ^ "1") :: index xs ((x, 2) :: tab)
+            else x :: index xs tab
+        | SOME i => (x ^ string_of_int i) :: index xs ((x, i + 1) :: tab))
+      | index [] _ = [];
+  in index names [] end;
+
+fun make_tnames Ts =
+  let
+    fun type_name (TFree (name, _)) = unprefix "'" name
+      | type_name (Type (name, _)) =
+          let val name' = Long_Name.base_name name
+          in if Symbol_Pos.is_identifier name' then name' else "x" end;
+  in indexify_names (map type_name Ts) end;
+
+
+
+(** data management **)
+
+datatype data = Data of
+  {constrs: (string * (term * term list)) list Symtab.table,
+   cases: (term * term list) Symtab.table};
+
+fun make_data (constrs, cases) = Data {constrs = constrs, cases = cases};
+
+structure Data = Generic_Data
+(
+  type T = data;
+  val empty = make_data (Symtab.empty, Symtab.empty);
+  val extend = I;
+  fun merge
+    (Data {constrs = constrs1, cases = cases1},
+     Data {constrs = constrs2, cases = cases2}) =
+    make_data
+      (Symtab.join (K (AList.merge (op =) (K true))) (constrs1, constrs2),
+      Symtab.merge (K true) (cases1, cases2));
+);
+
+fun map_data f =
+  Data.map (fn Data {constrs, cases} => make_data (f (constrs, cases)));
+fun map_constrs f = map_data (fn (constrs, cases) => (f constrs, cases));
+fun map_cases f = map_data (fn (constrs, cases) => (constrs, f cases));
+
+val rep_data = (fn Data args => args) o Data.get o Context.Proof;
+
+fun T_of_data (comb, constrs : term list) =
+  fastype_of comb
+  |> funpow (length constrs) range_type
+  |> domain_type;
+
+val Tname_of_data = fst o dest_Type o T_of_data;
+
+val constrs_of = #constrs o rep_data;
+val cases_of = #cases o rep_data;
+
+fun lookup_by_constr ctxt (c, T) =
+  let
+    val tab = Symtab.lookup_list (constrs_of ctxt) c;
+  in
+    (case body_type T of
+      Type (tyco, _) => AList.lookup (op =) tab tyco
+    | _ => NONE)
+  end;
+
+fun lookup_by_constr_permissive ctxt (c, T) =
+  let
+    val tab = Symtab.lookup_list (constrs_of ctxt) c;
+    val hint = (case body_type T of Type (tyco, _) => SOME tyco | _ => NONE);
+    val default = if null tab then NONE else SOME (snd (List.last tab));
+    (*conservative wrt. overloaded constructors*)
+  in
+    (case hint of
+      NONE => default
+    | SOME tyco =>
+        (case AList.lookup (op =) tab tyco of
+          NONE => default (*permissive*)
+        | SOME info => SOME info))
+  end;
+
+val lookup_by_case = Symtab.lookup o cases_of;
+
+
+
+(** installation **)
+
+fun case_error s = error ("Error in case expression:\n" ^ s);
+
+val name_of = try (dest_Const #> fst);
+
+
+(* parse translation *)
+
+fun constrain_Abs tT t = Syntax.const @{syntax_const "_constrainAbs"} $ t $ tT;
+
+fun case_tr err ctxt [t, u] =
+      let
+        val consts = Proof_Context.consts_of ctxt;
+        fun is_const s = can (Consts.the_constraint consts) (Consts.intern consts s);
+
+        fun variant_free x used =
+          let val (x', used') = Name.variant x used
+          in if is_const x' then variant_free x' used' else (x', used') end;
+
+        fun abs p tTs t =
+          Syntax.const @{const_syntax case_abs} $
+            fold constrain_Abs tTs (absfree p t);
+
+        fun abs_pat (Const (@{syntax_const "_constrain"}, _) $ t $ tT) tTs =
+              abs_pat t (tT :: tTs)
+          | abs_pat (Free (p as (x, _))) tTs =
+              if is_const x then I else abs p tTs
+          | abs_pat (t $ u) _ = abs_pat u [] #> abs_pat t []
+          | abs_pat _ _ = I;
+
+        (* replace occurrences of dummy_pattern by distinct variables *)
+        fun replace_dummies (Const (@{const_syntax dummy_pattern}, T)) used =
+              let val (x, used') = variant_free "x" used
+              in (Free (x, T), used') end
+          | replace_dummies (t $ u) used =
+              let
+                val (t', used') = replace_dummies t used;
+                val (u', used'') = replace_dummies u used';
+              in (t' $ u', used'') end
+          | replace_dummies t used = (t, used);
+
+        fun dest_case1 (t as Const (@{syntax_const "_case1"}, _) $ l $ r) =
+              let val (l', _) = replace_dummies l (Term.declare_term_frees t Name.context) in
+                abs_pat l' []
+                  (Syntax.const @{const_syntax case_elem} $ Term_Position.strip_positions l' $ r)
+              end
+          | dest_case1 _ = case_error "dest_case1";
+
+        fun dest_case2 (Const (@{syntax_const "_case2"}, _) $ t $ u) = t :: dest_case2 u
+          | dest_case2 t = [t];
+
+        val errt = if err then @{term True} else @{term False};
+      in
+        Syntax.const @{const_syntax case_guard} $ errt $ t $
+          (fold_rev
+            (fn t => fn u => Syntax.const @{const_syntax case_cons} $ dest_case1 t $ u)
+            (dest_case2 u)
+            (Syntax.const @{const_syntax case_nil}))
+      end
+  | case_tr _ _ _ = case_error "case_tr";
+
+val trfun_setup = Sign.parse_translation [(@{syntax_const "_case_syntax"}, case_tr true)];
+
+
+(* print translation *)
+
+fun case_tr' (_ :: x :: t :: ts) =
+      let
+        fun mk_clause (Const (@{const_syntax case_abs}, _) $ Abs (s, T, t)) xs used =
+              let val (s', used') = Name.variant s used
+              in mk_clause t ((s', T) :: xs) used' end
+          | mk_clause (Const (@{const_syntax case_elem}, _) $ pat $ rhs) xs _ =
+              Syntax.const @{syntax_const "_case1"} $
+                subst_bounds (map Syntax_Trans.mark_bound_abs xs, pat) $
+                subst_bounds (map Syntax_Trans.mark_bound_body xs, rhs)
+          | mk_clause _ _ _ = raise Match;
+
+        fun mk_clauses (Const (@{const_syntax case_nil}, _)) = []
+          | mk_clauses (Const (@{const_syntax case_cons}, _) $ t $ u) =
+              mk_clause t [] (Term.declare_term_frees t Name.context) :: mk_clauses u
+          | mk_clauses _ = raise Match;
+      in
+        list_comb (Syntax.const @{syntax_const "_case_syntax"} $ x $
+          foldr1 (fn (t, u) => Syntax.const @{syntax_const "_case2"} $ t $ u)
+            (mk_clauses t), ts)
+      end
+  | case_tr' _ = raise Match;
+
+val trfun_setup' = Sign.print_translation [(@{const_syntax "case_guard"}, K case_tr')];
+
+
+(* declarations *)
+
+fun register raw_case_comb raw_constrs context =
+  let
+    val ctxt = Context.proof_of context;
+    val case_comb = singleton (Variable.polymorphic ctxt) raw_case_comb;
+    val constrs = Variable.polymorphic ctxt raw_constrs;
+    val case_key = case_comb |> dest_Const |> fst;
+    val constr_keys = map (fst o dest_Const) constrs;
+    val data = (case_comb, constrs);
+    val Tname = Tname_of_data data;
+    val update_constrs = fold (fn key => Symtab.cons_list (key, (Tname, data))) constr_keys;
+    val update_cases = Symtab.update (case_key, data);
+  in
+    context
+    |> map_constrs update_constrs
+    |> map_cases update_cases
+  end;
+
+
+(* (Un)check phases *)
+
+datatype config = Error | Warning | Quiet;
+
+exception CASE_ERROR of string * int;
+
+
+(*Each pattern carries with it a tag i, which denotes the clause it
+  came from. i = ~1 indicates that the clause was added by pattern
+  completion.*)
+
+fun add_row_used ((prfx, pats), (tm, _)) =
+  fold Term.declare_term_frees (tm :: pats @ map Free prfx);
+
+(*try to preserve names given by user*)
+fun default_name "" (Free (name', _)) = name'
+  | default_name name _ = name;
+
+
+(*Produce an instance of a constructor, plus fresh variables for its arguments.*)
+fun fresh_constr colty used c =
+  let
+    val (_, T) = dest_Const c;
+    val Ts = binder_types T;
+    val (names, _) =
+      fold_map Name.variant (make_tnames (map Logic.unvarifyT_global Ts)) used;
+    val ty = body_type T;
+    val ty_theta = Type.raw_match (ty, colty) Vartab.empty
+      handle Type.TYPE_MATCH => raise CASE_ERROR ("type mismatch", ~1);
+    val c' = Envir.subst_term_types ty_theta c;
+    val gvars = map (Envir.subst_term_types ty_theta o Free) (names ~~ Ts);
+  in (c', gvars) end;
+
+(*Go through a list of rows and pick out the ones beginning with a
+  pattern with constructor = name.*)
+fun mk_group (name, T) rows =
+  let val k = length (binder_types T) in
+    fold (fn (row as ((prfx, p :: ps), rhs as (_, i))) =>
+      fn ((in_group, not_in_group), names) =>
+        (case strip_comb p of
+          (Const (name', _), args) =>
+            if name = name' then
+              if length args = k then
+                ((((prfx, args @ ps), rhs) :: in_group, not_in_group),
+                 map2 default_name names args)
+              else raise CASE_ERROR ("Wrong number of arguments for constructor " ^ quote name, i)
+            else ((in_group, row :: not_in_group), names)
+        | _ => raise CASE_ERROR ("Not a constructor pattern", i)))
+    rows (([], []), replicate k "") |>> pairself rev
+  end;
+
+
+(* Partitioning *)
+
+fun partition _ _ _ _ [] = raise CASE_ERROR ("partition: no rows", ~1)
+  | partition used constructors colty res_ty (rows as (((prfx, _ :: ps), _) :: _)) =
+      let
+        fun part [] [] = []
+          | part [] ((_, (_, i)) :: _) = raise CASE_ERROR ("Not a constructor pattern", i)
+          | part (c :: cs) rows =
+              let
+                val ((in_group, not_in_group), names) = mk_group (dest_Const c) rows;
+                val used' = fold add_row_used in_group used;
+                val (c', gvars) = fresh_constr colty used' c;
+                val in_group' =
+                  if null in_group  (* Constructor not given *)
+                  then
+                    let
+                      val Ts = map fastype_of ps;
+                      val (xs, _) =
+                        fold_map Name.variant
+                          (replicate (length ps) "x")
+                          (fold Term.declare_term_frees gvars used');
+                    in
+                      [((prfx, gvars @ map Free (xs ~~ Ts)),
+                        (Const (@{const_name undefined}, res_ty), ~1))]
+                    end
+                  else in_group;
+              in
+                {constructor = c',
+                 new_formals = gvars,
+                 names = names,
+                 group = in_group'} :: part cs not_in_group
+              end;
+      in part constructors rows end;
+
+fun v_to_prfx (prfx, Free v :: pats) = (v :: prfx, pats)
+  | v_to_prfx _ = raise CASE_ERROR ("mk_case: v_to_prfx", ~1);
+
+
+(* Translation of pattern terms into nested case expressions. *)
+
+fun mk_case ctxt used range_ty =
+  let
+    val get_info = lookup_by_constr_permissive ctxt;
+
+    fun expand _ _ _ ((_, []), _) = raise CASE_ERROR ("mk_case: expand", ~1)
+      | expand constructors used ty (row as ((prfx, p :: ps), (rhs, tag))) =
+          if is_Free p then
+            let
+              val used' = add_row_used row used;
+              fun expnd c =
+                let val capp = list_comb (fresh_constr ty used' c)
+                in ((prfx, capp :: ps), (subst_free [(p, capp)] rhs, tag)) end;
+            in map expnd constructors end
+          else [row];
+
+    val (name, _) = Name.variant "a" used;
+
+    fun mk _ [] = raise CASE_ERROR ("no rows", ~1)
+      | mk [] (((_, []), (tm, tag)) :: _) = ([tag], tm) (* Done *)
+      | mk path ((row as ((_, [Free _]), _)) :: _ :: _) = mk path [row]
+      | mk (u :: us) (rows as ((_, _ :: _), _) :: _) =
+          let val col0 = map (fn ((_, p :: _), (_, i)) => (p, i)) rows in
+            (case Option.map (apfst head_of) (find_first (not o is_Free o fst) col0) of
+              NONE =>
+                let
+                  val rows' = map (fn ((v, _), row) => row ||>
+                    apfst (subst_free [(v, u)]) |>> v_to_prfx) (col0 ~~ rows);
+                in mk us rows' end
+            | SOME (Const (cname, cT), i) =>
+                (case get_info (cname, cT) of
+                  NONE => raise CASE_ERROR ("Not a datatype constructor: " ^ quote cname, i)
+                | SOME (case_comb, constructors) =>
+                    let
+                      val pty = body_type cT;
+                      val used' = fold Term.declare_term_frees us used;
+                      val nrows = maps (expand constructors used' pty) rows;
+                      val subproblems = partition used' constructors pty range_ty nrows;
+                      val (pat_rect, dtrees) =
+                        split_list (map (fn {new_formals, group, ...} =>
+                          mk (new_formals @ us) group) subproblems);
+                      val case_functions =
+                        map2 (fn {new_formals, names, ...} =>
+                          fold_rev (fn (x as Free (_, T), s) => fn t =>
+                            Abs (if s = "" then name else s, T, abstract_over (x, t)))
+                              (new_formals ~~ names))
+                        subproblems dtrees;
+                      val types = map fastype_of (case_functions @ [u]);
+                      val case_const = Const (name_of case_comb |> the, types ---> range_ty);
+                      val tree = list_comb (case_const, case_functions @ [u]);
+                    in (flat pat_rect, tree) end)
+            | SOME (t, i) =>
+                raise CASE_ERROR ("Not a datatype constructor: " ^ Syntax.string_of_term ctxt t, i))
+          end
+      | mk _ _ = raise CASE_ERROR ("Malformed row matrix", ~1)
+  in mk end;
+
+
+(*Repeated variable occurrences in a pattern are not allowed.*)
+fun no_repeat_vars ctxt pat = fold_aterms
+  (fn x as Free (s, _) =>
+      (fn xs =>
+        if member op aconv xs x then
+          case_error (quote s ^ " occurs repeatedly in the pattern " ^
+            quote (Syntax.string_of_term ctxt pat))
+        else x :: xs)
+    | _ => I) pat [];
+
+fun make_case ctxt config used x clauses =
+  let
+    fun string_of_clause (pat, rhs) =
+      Syntax.unparse_term ctxt
+        (Term.list_comb (Syntax.const @{syntax_const "_case1"},
+          Syntax.uncheck_terms ctxt [pat, rhs]))
+      |> Pretty.string_of;
+
+    val _ = map (no_repeat_vars ctxt o fst) clauses;
+    val rows = map_index (fn (i, (pat, rhs)) => (([], [pat]), (rhs, i))) clauses;
+    val rangeT =
+      (case distinct (op =) (map (fastype_of o snd) clauses) of
+        [] => case_error "no clauses given"
+      | [T] => T
+      | _ => case_error "all cases must have the same result type");
+    val used' = fold add_row_used rows used;
+    val (tags, case_tm) =
+      mk_case ctxt used' rangeT [x] rows
+        handle CASE_ERROR (msg, i) =>
+          case_error
+            (msg ^ (if i < 0 then "" else "\nIn clause\n" ^ string_of_clause (nth clauses i)));
+    val _ =
+      (case subtract (op =) tags (map (snd o snd) rows) of
+        [] => ()
+      | is =>
+          if config = Quiet then ()
+          else
+            (if config = Error then case_error else warning (*FIXME lack of syntactic context!?*))
+              ("The following clauses are redundant (covered by preceding clauses):\n" ^
+                cat_lines (map (string_of_clause o nth clauses) is)));
+  in
+    case_tm
+  end;
+
+
+(* term check *)
+
+fun decode_clause (Const (@{const_name case_abs}, _) $ Abs (s, T, t)) xs used =
+      let val (s', used') = Name.variant s used
+      in decode_clause t (Free (s', T) :: xs) used' end
+  | decode_clause (Const (@{const_name case_elem}, _) $ t $ u) xs _ =
+      (subst_bounds (xs, t), subst_bounds (xs, u))
+  | decode_clause _ _ _ = case_error "decode_clause";
+
+fun decode_cases (Const (@{const_name case_nil}, _)) = []
+  | decode_cases (Const (@{const_name case_cons}, _) $ t $ u) =
+      decode_clause t [] (Term.declare_term_frees t Name.context) ::
+      decode_cases u
+  | decode_cases _ = case_error "decode_cases";
+
+fun check_case ctxt =
+  let
+    fun decode_case (Const (@{const_name case_guard}, _) $ b $ u $ t) =
+          make_case ctxt (if b = @{term True} then Error else Warning)
+            Name.context (decode_case u) (decode_cases t)
+      | decode_case (t $ u) = decode_case t $ decode_case u
+      | decode_case (Abs (x, T, u)) =
+          let val (x', u') = Term.dest_abs (x, T, u);
+          in Term.absfree (x', T) (decode_case u') end
+      | decode_case t = t;
+  in
+    map decode_case
+  end;
+
+val term_check_setup =
+  Context.theory_map (Syntax_Phases.term_check 1 "case" check_case);
+
+
+(* Pretty printing of nested case expressions *)
+
+(* destruct one level of pattern matching *)
+
+fun dest_case ctxt d used t =
+  (case apfst name_of (strip_comb t) of
+    (SOME cname, ts as _ :: _) =>
+      let
+        val (fs, x) = split_last ts;
+        fun strip_abs i Us t =
+          let
+            val zs = strip_abs_vars t;
+            val j = length zs;
+            val (xs, ys) =
+              if j < i then (zs @ map (pair "x") (drop j Us), [])
+              else chop i zs;
+            val u = fold_rev Term.abs ys (strip_abs_body t);
+            val xs' = map Free
+              ((fold_map Name.variant (map fst xs)
+                  (Term.declare_term_names u used) |> fst) ~~
+               map snd xs);
+            val (xs1, xs2) = chop j xs'
+          in (xs', list_comb (subst_bounds (rev xs1, u), xs2)) end;
+        fun is_dependent i t =
+          let val k = length (strip_abs_vars t) - i
+          in k < 0 orelse exists (fn j => j >= k) (loose_bnos (strip_abs_body t)) end;
+        fun count_cases (_, _, true) = I
+          | count_cases (c, (_, body), false) = AList.map_default op aconv (body, []) (cons c);
+        val is_undefined = name_of #> equal (SOME @{const_name undefined});
+        fun mk_case (c, (xs, body), _) = (list_comb (c, xs), body);
+        val get_info = lookup_by_case ctxt;
+      in
+        (case get_info cname of
+          SOME (_, constructors) =>
+            if length fs = length constructors then
+              let
+                val cases = map (fn (Const (s, U), t) =>
+                  let
+                    val Us = binder_types U;
+                    val k = length Us;
+                    val p as (xs, _) = strip_abs k Us t;
+                  in
+                    (Const (s, map fastype_of xs ---> fastype_of x), p, is_dependent k t)
+                  end) (constructors ~~ fs);
+                val cases' =
+                  sort (int_ord o swap o pairself (length o snd))
+                    (fold_rev count_cases cases []);
+                val R = fastype_of t;
+                val dummy =
+                  if d then Term.dummy_pattern R
+                  else Free (Name.variant "x" used |> fst, R);
+              in
+                SOME (x,
+                  map mk_case
+                    (case find_first (is_undefined o fst) cases' of
+                      SOME (_, cs) =>
+                        if length cs = length constructors then [hd cases]
+                        else filter_out (fn (_, (_, body), _) => is_undefined body) cases
+                    | NONE =>
+                        (case cases' of
+                          [] => cases
+                        | (default, cs) :: _ =>
+                            if length cs = 1 then cases
+                            else if length cs = length constructors then
+                              [hd cases, (dummy, ([], default), false)]
+                            else
+                              filter_out (fn (c, _, _) => member op aconv cs c) cases @
+                                [(dummy, ([], default), false)])))
+              end
+            else NONE
+        | _ => NONE)
+      end
+  | _ => NONE);
+
+
+(* destruct nested patterns *)
+
+fun encode_clause recur S T (pat, rhs) =
+  fold (fn x as (_, U) => fn t =>
+    Const (@{const_name case_abs}, (U --> T) --> T) $ Term.absfree x t)
+      (Term.add_frees pat [])
+      (Const (@{const_name case_elem}, S --> T --> S --> T) $ pat $ recur rhs);
+
+fun encode_cases _ S T [] = Const (@{const_name case_nil}, S --> T)
+  | encode_cases recur S T (p :: ps) =
+      Const (@{const_name case_cons}, (S --> T) --> (S --> T) --> S --> T) $
+        encode_clause recur S T p $ encode_cases recur S T ps;
+
+fun encode_case recur (t, ps as (pat, rhs) :: _) =
+      let
+        val tT = fastype_of t;
+        val T = fastype_of rhs;
+      in
+        Const (@{const_name case_guard}, @{typ bool} --> tT --> (tT --> T) --> T) $
+          @{term True} $ t $ (encode_cases recur (fastype_of pat) (fastype_of rhs) ps)
+      end
+  | encode_case _ _ = case_error "encode_case";
+
+fun strip_case' ctxt d (pat, rhs) =
+  (case dest_case ctxt d (Term.declare_term_frees pat Name.context) rhs of
+    SOME (exp as Free _, clauses) =>
+      if Term.exists_subterm (curry (op aconv) exp) pat andalso
+        not (exists (fn (_, rhs') =>
+          Term.exists_subterm (curry (op aconv) exp) rhs') clauses)
+      then
+        maps (strip_case' ctxt d) (map (fn (pat', rhs') =>
+          (subst_free [(exp, pat')] pat, rhs')) clauses)
+      else [(pat, rhs)]
+  | _ => [(pat, rhs)]);
+
+fun strip_case ctxt d t =
+  (case dest_case ctxt d Name.context t of
+    SOME (x, clauses) => SOME (x, maps (strip_case' ctxt d) clauses)
+  | NONE => NONE);
+
+fun strip_case_full ctxt d t =
+  (case dest_case ctxt d Name.context t of
+    SOME (x, clauses) =>
+      encode_case (strip_case_full ctxt d)
+        (strip_case_full ctxt d x, maps (strip_case' ctxt d) clauses)
+  | NONE =>
+      (case t of
+        t $ u => strip_case_full ctxt d t $ strip_case_full ctxt d u
+      | Abs (x, T, u) =>
+          let val (x', u') = Term.dest_abs (x, T, u);
+          in Term.absfree (x', T) (strip_case_full ctxt d u') end
+      | _ => t));
+
+
+(* term uncheck *)
+
+val show_cases = Attrib.setup_config_bool @{binding show_cases} (K true);
+
+fun uncheck_case ctxt ts =
+  if Config.get ctxt show_cases
+  then map (fn t => if can Term.type_of t then strip_case_full ctxt true t else t) ts
+  else ts;
+
+val term_uncheck_setup =
+  Context.theory_map (Syntax_Phases.term_uncheck 1 "case" uncheck_case);
+
+
+(* theory setup *)
+
+val setup =
+  trfun_setup #>
+  trfun_setup' #>
+  term_check_setup #>
+  term_uncheck_setup #>
+  Attrib.setup @{binding case_translation}
+    (Args.term -- Scan.repeat1 Args.term >>
+      (fn (t, ts) => Thm.declaration_attribute (K (register t ts))))
+    "declaration of case combinators and constructors";
+
+
+(* outer syntax commands *)
+
+fun print_case_translations ctxt =
+  let
+    val cases = map snd (Symtab.dest (cases_of ctxt));
+    val type_space = Proof_Context.type_space ctxt;
+
+    val pretty_term = Syntax.pretty_term ctxt;
+
+    fun pretty_data (data as (comb, ctrs)) =
+      let
+        val name = Tname_of_data data;
+        val xname = Name_Space.extern ctxt type_space name;
+        val markup = Name_Space.markup type_space name;
+        val prt =
+          (Pretty.block o Pretty.fbreaks)
+           [Pretty.block [Pretty.mark_str (markup, xname), Pretty.str ":"],
+            Pretty.block [Pretty.str "combinator:", Pretty.brk 1, pretty_term comb],
+            Pretty.block (Pretty.str "constructors:" :: Pretty.brk 1 ::
+              Pretty.commas (map pretty_term ctrs))];
+      in (xname, prt) end;
+  in
+    Pretty.big_list "case translations:" (map #2 (sort_wrt #1 (map pretty_data cases)))
+    |> Pretty.writeln
+  end;
+
+val _ =
+  Outer_Syntax.improper_command @{command_spec "print_case_translations"}
+    "print registered case combinators and constructors"
+    (Scan.succeed (Toplevel.keep (print_case_translations o Toplevel.context_of)))
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML	Mon Dec 09 09:44:57 2013 +0100
@@ -0,0 +1,981 @@
+(*  Title:      HOL/Tools/Ctr_Sugar/ctr_sugar.ML
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012, 2013
+
+Wrapping existing freely generated type's constructors.
+*)
+
+signature CTR_SUGAR =
+sig
+  type ctr_sugar =
+    {ctrs: term list,
+     casex: term,
+     discs: term list,
+     selss: term list list,
+     exhaust: thm,
+     nchotomy: thm,
+     injects: thm list,
+     distincts: thm list,
+     case_thms: thm list,
+     case_cong: thm,
+     weak_case_cong: thm,
+     split: thm,
+     split_asm: thm,
+     disc_thmss: thm list list,
+     discIs: thm list,
+     sel_thmss: thm list list,
+     disc_exhausts: thm list,
+     sel_exhausts: thm list,
+     collapses: thm list,
+     expands: thm list,
+     sel_splits: thm list,
+     sel_split_asms: thm list,
+     case_eq_ifs: thm list};
+
+  val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar
+  val transfer_ctr_sugar: Proof.context -> ctr_sugar -> ctr_sugar
+  val ctr_sugar_of: Proof.context -> string -> ctr_sugar option
+  val ctr_sugars_of: Proof.context -> ctr_sugar list
+  val ctr_sugar_of_case: Proof.context -> string -> ctr_sugar option
+  val register_ctr_sugar: string -> ctr_sugar -> local_theory -> local_theory
+  val register_ctr_sugar_global: string -> ctr_sugar -> theory -> theory
+
+  val rep_compat_prefix: string
+
+  val mk_half_pairss: 'a list * 'a list -> ('a * 'a) list list
+  val join_halves: int -> 'a list list -> 'a list list -> 'a list * 'a list list list
+
+  val mk_ctr: typ list -> term -> term
+  val mk_case: typ list -> typ -> term -> term
+  val mk_disc_or_sel: typ list -> term -> term
+  val name_of_ctr: term -> string
+  val name_of_disc: term -> string
+  val dest_ctr: Proof.context -> string -> term -> term * term list
+  val dest_case: Proof.context -> string -> typ list -> term -> (term list * term list) option
+
+  val wrap_free_constructors: ({prems: thm list, context: Proof.context} -> tactic) list list ->
+    (((bool * (bool * bool)) * term list) * binding) *
+      (binding list * (binding list list * (binding * term) list list)) -> local_theory ->
+    ctr_sugar * local_theory
+  val parse_wrap_free_constructors_options: (bool * (bool * bool)) parser
+  val parse_bound_term: (binding * string) parser
+end;
+
+structure Ctr_Sugar : CTR_SUGAR =
+struct
+
+open Ctr_Sugar_Util
+open Ctr_Sugar_Tactics
+open Ctr_Sugar_Code
+
+type ctr_sugar =
+  {ctrs: term list,
+   casex: term,
+   discs: term list,
+   selss: term list list,
+   exhaust: thm,
+   nchotomy: thm,
+   injects: thm list,
+   distincts: thm list,
+   case_thms: thm list,
+   case_cong: thm,
+   weak_case_cong: thm,
+   split: thm,
+   split_asm: thm,
+   disc_thmss: thm list list,
+   discIs: thm list,
+   sel_thmss: thm list list,
+   disc_exhausts: thm list,
+   sel_exhausts: thm list,
+   collapses: thm list,
+   expands: thm list,
+   sel_splits: thm list,
+   sel_split_asms: thm list,
+   case_eq_ifs: thm list};
+
+fun eq_ctr_sugar ({ctrs = ctrs1, casex = case1, discs = discs1, selss = selss1, ...} : ctr_sugar,
+    {ctrs = ctrs2, casex = case2, discs = discs2, selss = selss2, ...} : ctr_sugar) =
+  ctrs1 = ctrs2 andalso case1 = case2 andalso discs1 = discs2 andalso selss1 = selss2;
+
+fun morph_ctr_sugar phi {ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts,
+    case_thms, case_cong, weak_case_cong, split, split_asm, disc_thmss, discIs, sel_thmss,
+    disc_exhausts, sel_exhausts, collapses, expands, sel_splits, sel_split_asms, case_eq_ifs} =
+  {ctrs = map (Morphism.term phi) ctrs,
+   casex = Morphism.term phi casex,
+   discs = map (Morphism.term phi) discs,
+   selss = map (map (Morphism.term phi)) selss,
+   exhaust = Morphism.thm phi exhaust,
+   nchotomy = Morphism.thm phi nchotomy,
+   injects = map (Morphism.thm phi) injects,
+   distincts = map (Morphism.thm phi) distincts,
+   case_thms = map (Morphism.thm phi) case_thms,
+   case_cong = Morphism.thm phi case_cong,
+   weak_case_cong = Morphism.thm phi weak_case_cong,
+   split = Morphism.thm phi split,
+   split_asm = Morphism.thm phi split_asm,
+   disc_thmss = map (map (Morphism.thm phi)) disc_thmss,
+   discIs = map (Morphism.thm phi) discIs,
+   sel_thmss = map (map (Morphism.thm phi)) sel_thmss,
+   disc_exhausts = map (Morphism.thm phi) disc_exhausts,
+   sel_exhausts = map (Morphism.thm phi) sel_exhausts,
+   collapses = map (Morphism.thm phi) collapses,
+   expands = map (Morphism.thm phi) expands,
+   sel_splits = map (Morphism.thm phi) sel_splits,
+   sel_split_asms = map (Morphism.thm phi) sel_split_asms,
+   case_eq_ifs = map (Morphism.thm phi) case_eq_ifs};
+
+val transfer_ctr_sugar =
+  morph_ctr_sugar o Morphism.thm_morphism o Thm.transfer o Proof_Context.theory_of;
+
+structure Data = Generic_Data
+(
+  type T = ctr_sugar Symtab.table;
+  val empty = Symtab.empty;
+  val extend = I;
+  val merge = Symtab.merge eq_ctr_sugar;
+);
+
+fun ctr_sugar_of ctxt =
+  Symtab.lookup (Data.get (Context.Proof ctxt))
+  #> Option.map (transfer_ctr_sugar ctxt);
+
+fun ctr_sugars_of ctxt =
+  Symtab.fold (cons o transfer_ctr_sugar ctxt o snd) (Data.get (Context.Proof ctxt)) [];
+
+fun ctr_sugar_of_case ctxt s =
+  find_first (fn {casex = Const (s', _), ...} => s' = s | _ => false) (ctr_sugars_of ctxt);
+
+fun register_ctr_sugar key ctr_sugar =
+  Local_Theory.declaration {syntax = false, pervasive = true}
+    (fn phi => Data.map (Symtab.default (key, morph_ctr_sugar phi ctr_sugar)));
+
+fun register_ctr_sugar_global key ctr_sugar =
+  Context.theory_map (Data.map (Symtab.default (key, ctr_sugar)));
+
+val rep_compat_prefix = "new";
+
+val isN = "is_";
+val unN = "un_";
+fun mk_unN 1 1 suf = unN ^ suf
+  | mk_unN _ l suf = unN ^ suf ^ string_of_int l;
+
+val caseN = "case";
+val case_congN = "case_cong";
+val case_eq_ifN = "case_eq_if";
+val collapseN = "collapse";
+val disc_excludeN = "disc_exclude";
+val disc_exhaustN = "disc_exhaust";
+val discN = "disc";
+val discIN = "discI";
+val distinctN = "distinct";
+val exhaustN = "exhaust";
+val expandN = "expand";
+val injectN = "inject";
+val nchotomyN = "nchotomy";
+val selN = "sel";
+val sel_exhaustN = "sel_exhaust";
+val sel_splitN = "sel_split";
+val sel_split_asmN = "sel_split_asm";
+val splitN = "split";
+val splitsN = "splits";
+val split_asmN = "split_asm";
+val weak_case_cong_thmsN = "weak_case_cong";
+
+val cong_attrs = @{attributes [cong]};
+val dest_attrs = @{attributes [dest]};
+val safe_elim_attrs = @{attributes [elim!]};
+val iff_attrs = @{attributes [iff]};
+val inductsimp_attrs = @{attributes [induct_simp]};
+val nitpicksimp_attrs = @{attributes [nitpick_simp]};
+val simp_attrs = @{attributes [simp]};
+val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
+val code_nitpicksimp_simp_attrs = code_nitpicksimp_attrs @ simp_attrs;
+
+fun unflat_lookup eq xs ys = map (fn xs' => permute_like eq xs xs' ys);
+
+fun mk_half_pairss' _ ([], []) = []
+  | mk_half_pairss' indent (x :: xs, _ :: ys) =
+    indent @ fold_rev (cons o single o pair x) ys (mk_half_pairss' ([] :: indent) (xs, ys));
+
+fun mk_half_pairss p = mk_half_pairss' [[]] p;
+
+fun join_halves n half_xss other_half_xss =
+  let
+    val xsss =
+      map2 (map2 append) (Library.chop_groups n half_xss)
+        (transpose (Library.chop_groups n other_half_xss))
+    val xs = splice (flat half_xss) (flat other_half_xss);
+  in (xs, xsss) end;
+
+fun mk_undefined T = Const (@{const_name undefined}, T);
+
+fun mk_ctr Ts t =
+  let val Type (_, Ts0) = body_type (fastype_of t) in
+    subst_nonatomic_types (Ts0 ~~ Ts) t
+  end;
+
+fun mk_case Ts T t =
+  let val (Type (_, Ts0), body) = strip_type (fastype_of t) |>> List.last in
+    subst_nonatomic_types ((body, T) :: (Ts0 ~~ Ts)) t
+  end;
+
+fun mk_disc_or_sel Ts t =
+  subst_nonatomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t;
+
+fun name_of_const what t =
+  (case head_of t of
+    Const (s, _) => s
+  | Free (s, _) => s
+  | _ => error ("Cannot extract name of " ^ what));
+
+val name_of_ctr = name_of_const "constructor";
+
+val notN = "not_";
+val eqN = "eq_";
+val neqN = "neq_";
+
+fun name_of_disc t =
+  (case head_of t of
+    Abs (_, _, @{const Not} $ (t' $ Bound 0)) =>
+    Long_Name.map_base_name (prefix notN) (name_of_disc t')
+  | Abs (_, _, Const (@{const_name HOL.eq}, _) $ Bound 0 $ t') =>
+    Long_Name.map_base_name (prefix eqN) (name_of_disc t')
+  | Abs (_, _, @{const Not} $ (Const (@{const_name HOL.eq}, _) $ Bound 0 $ t')) =>
+    Long_Name.map_base_name (prefix neqN) (name_of_disc t')
+  | t' => name_of_const "destructor" t');
+
+val base_name_of_ctr = Long_Name.base_name o name_of_ctr;
+
+fun dest_ctr ctxt s t =
+  let
+    val (f, args) = Term.strip_comb t;
+  in
+    (case ctr_sugar_of ctxt s of
+      SOME {ctrs, ...} =>
+      (case find_first (can (fo_match ctxt f)) ctrs of
+        SOME f' => (f', args)
+      | NONE => raise Fail "dest_ctr")
+    | NONE => raise Fail "dest_ctr")
+  end;
+
+fun dest_case ctxt s Ts t =
+  (case Term.strip_comb t of
+    (Const (c, _), args as _ :: _) =>
+    (case ctr_sugar_of ctxt s of
+      SOME {casex = Const (case_name, _), discs = discs0, selss = selss0, ...} =>
+      if case_name = c then
+        let val n = length discs0 in
+          if n < length args then
+            let
+              val (branches, obj :: leftovers) = chop n args;
+              val discs = map (mk_disc_or_sel Ts) discs0;
+              val selss = map (map (mk_disc_or_sel Ts)) selss0;
+              val conds = map (rapp obj) discs;
+              val branch_argss = map (fn sels => map (rapp obj) sels @ leftovers) selss;
+              val branches' = map2 (curry Term.betapplys) branches branch_argss;
+            in
+              SOME (conds, branches')
+            end
+          else
+            NONE
+        end
+      else
+        NONE
+    | _ => NONE)
+  | _ => NONE);
+
+fun eta_expand_arg xs f_xs = fold_rev Term.lambda xs f_xs;
+
+fun prepare_wrap_free_constructors prep_term ((((no_discs_sels, (no_code, rep_compat)), raw_ctrs),
+    raw_case_binding), (raw_disc_bindings, (raw_sel_bindingss, raw_sel_defaultss))) no_defs_lthy =
+  let
+    (* TODO: sanity checks on arguments *)
+
+    val n = length raw_ctrs;
+    val ks = 1 upto n;
+
+    val _ = if n > 0 then () else error "No constructors specified";
+
+    val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs;
+    val sel_defaultss =
+      pad_list [] n (map (map (apsnd (prep_term no_defs_lthy))) raw_sel_defaultss);
+
+    val Type (fcT_name, As0) = body_type (fastype_of (hd ctrs0));
+    val fc_b_name = Long_Name.base_name fcT_name;
+    val fc_b = Binding.name fc_b_name;
+
+    fun qualify mandatory =
+      Binding.qualify mandatory fc_b_name o (rep_compat ? Binding.qualify false rep_compat_prefix);
+
+    fun dest_TFree_or_TVar (TFree sS) = sS
+      | dest_TFree_or_TVar (TVar ((s, _), S)) = (s, S)
+      | dest_TFree_or_TVar _ = error "Invalid type argument";
+
+    val (unsorted_As, B) =
+      no_defs_lthy
+      |> variant_tfrees (map (fst o dest_TFree_or_TVar) As0)
+      ||> the_single o fst o mk_TFrees 1;
+
+    val As = map2 (resort_tfree o snd o dest_TFree_or_TVar) As0 unsorted_As;
+
+    val fcT = Type (fcT_name, As);
+    val ctrs = map (mk_ctr As) ctrs0;
+    val ctr_Tss = map (binder_types o fastype_of) ctrs;
+
+    val ms = map length ctr_Tss;
+
+    val raw_disc_bindings' = pad_list Binding.empty n raw_disc_bindings;
+
+    fun can_definitely_rely_on_disc k = not (Binding.is_empty (nth raw_disc_bindings' (k - 1)));
+    fun can_rely_on_disc k =
+      can_definitely_rely_on_disc k orelse (k = 1 andalso not (can_definitely_rely_on_disc 2));
+    fun should_omit_disc_binding k = n = 1 orelse (n = 2 andalso can_rely_on_disc (3 - k));
+
+    fun is_disc_binding_valid b =
+      not (Binding.is_empty b orelse Binding.eq_name (b, equal_binding));
+
+    val standard_disc_binding = Binding.name o prefix isN o base_name_of_ctr;
+
+    val disc_bindings =
+      raw_disc_bindings'
+      |> map4 (fn k => fn m => fn ctr => fn disc =>
+        qualify false
+          (if Binding.is_empty disc then
+             if should_omit_disc_binding k then disc else standard_disc_binding ctr
+           else if Binding.eq_name (disc, equal_binding) then
+             if m = 0 then disc
+             else error "Cannot use \"=\" syntax for discriminating nonnullary constructor"
+           else if Binding.eq_name (disc, standard_binding) then
+             standard_disc_binding ctr
+           else
+             disc)) ks ms ctrs0;
+
+    fun standard_sel_binding m l = Binding.name o mk_unN m l o base_name_of_ctr;
+
+    val sel_bindingss =
+      pad_list [] n raw_sel_bindingss
+      |> map3 (fn ctr => fn m => map2 (fn l => fn sel =>
+        qualify false
+          (if Binding.is_empty sel orelse Binding.eq_name (sel, standard_binding) then
+            standard_sel_binding m l ctr
+          else
+            sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms;
+
+    val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss;
+
+    val ((((((((xss, xss'), yss), fs), gs), [u', v']), [w]), (p, p')), names_lthy) = no_defs_lthy |>
+      mk_Freess' "x" ctr_Tss
+      ||>> mk_Freess "y" ctr_Tss
+      ||>> mk_Frees "f" case_Ts
+      ||>> mk_Frees "g" case_Ts
+      ||>> (apfst (map (rpair fcT)) oo Variable.variant_fixes) [fc_b_name, fc_b_name ^ "'"]
+      ||>> mk_Frees "z" [B]
+      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT;
+
+    val u = Free u';
+    val v = Free v';
+    val q = Free (fst p', mk_pred1T B);
+
+    val xctrs = map2 (curry Term.list_comb) ctrs xss;
+    val yctrs = map2 (curry Term.list_comb) ctrs yss;
+
+    val xfs = map2 (curry Term.list_comb) fs xss;
+    val xgs = map2 (curry Term.list_comb) gs xss;
+
+    (* TODO: Eta-expension is for compatibility with the old datatype package (but it also provides
+       nicer names). Consider removing. *)
+    val eta_fs = map2 eta_expand_arg xss xfs;
+    val eta_gs = map2 eta_expand_arg xss xgs;
+
+    val case_binding =
+      qualify false
+        (if Binding.is_empty raw_case_binding orelse
+            Binding.eq_name (raw_case_binding, standard_binding) then
+           Binding.prefix_name (caseN ^ "_") fc_b
+         else
+           raw_case_binding);
+
+    fun mk_case_disj xctr xf xs =
+      list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_eq (w, xf)));
+
+    val case_rhs = fold_rev (fold_rev Term.lambda) [fs, [u]]
+      (Const (@{const_name The}, (B --> HOLogic.boolT) --> B) $
+         Term.lambda w (Library.foldr1 HOLogic.mk_disj (map3 mk_case_disj xctrs xfs xss)));
+
+    val ((raw_case, (_, raw_case_def)), (lthy', lthy)) = no_defs_lthy
+      |> Local_Theory.define ((case_binding, NoSyn),
+        ((Binding.conceal (Thm.def_binding case_binding), []), case_rhs))
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy lthy';
+
+    val case_def = Morphism.thm phi raw_case_def;
+
+    val case0 = Morphism.term phi raw_case;
+    val casex = mk_case As B case0;
+
+    val fcase = Term.list_comb (casex, fs);
+
+    val ufcase = fcase $ u;
+    val vfcase = fcase $ v;
+
+    val eta_fcase = Term.list_comb (casex, eta_fs);
+    val eta_gcase = Term.list_comb (casex, eta_gs);
+
+    val eta_ufcase = eta_fcase $ u;
+    val eta_vgcase = eta_gcase $ v;
+
+    fun mk_uu_eq () = HOLogic.mk_eq (u, u);
+
+    val uv_eq = mk_Trueprop_eq (u, v);
+
+    val exist_xs_u_eq_ctrs =
+      map2 (fn xctr => fn xs => list_exists_free xs (HOLogic.mk_eq (u, xctr))) xctrs xss;
+
+    val unique_disc_no_def = TrueI; (*arbitrary marker*)
+    val alternate_disc_no_def = FalseE; (*arbitrary marker*)
+
+    fun alternate_disc_lhs get_udisc k =
+      HOLogic.mk_not
+        (let val b = nth disc_bindings (k - 1) in
+           if is_disc_binding_valid b then get_udisc b (k - 1) else nth exist_xs_u_eq_ctrs (k - 1)
+         end);
+
+    val (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') =
+      if no_discs_sels then
+        (true, [], [], [], [], [], lthy)
+      else
+        let
+          fun disc_free b = Free (Binding.name_of b, mk_pred1T fcT);
+
+          fun disc_spec b exist_xs_u_eq_ctr = mk_Trueprop_eq (disc_free b $ u, exist_xs_u_eq_ctr);
+
+          fun alternate_disc k =
+            Term.lambda u (alternate_disc_lhs (K o rapp u o disc_free) (3 - k));
+
+          fun mk_sel_case_args b proto_sels T =
+            map2 (fn Ts => fn k =>
+              (case AList.lookup (op =) proto_sels k of
+                NONE =>
+                (case AList.lookup Binding.eq_name (rev (nth sel_defaultss (k - 1))) b of
+                  NONE => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T)
+                | SOME t => t |> Type.constraint (Ts ---> T) |> Syntax.check_term lthy)
+              | SOME (xs, x) => fold_rev Term.lambda xs x)) ctr_Tss ks;
+
+          fun sel_spec b proto_sels =
+            let
+              val _ =
+                (case duplicates (op =) (map fst proto_sels) of
+                   k :: _ => error ("Duplicate selector name " ^ quote (Binding.name_of b) ^
+                     " for constructor " ^
+                     quote (Syntax.string_of_term lthy (nth ctrs (k - 1))))
+                 | [] => ())
+              val T =
+                (case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of
+                  [T] => T
+                | T :: T' :: _ => error ("Inconsistent range type for selector " ^
+                    quote (Binding.name_of b) ^ ": " ^ quote (Syntax.string_of_typ lthy T) ^ " vs. "
+                    ^ quote (Syntax.string_of_typ lthy T')));
+            in
+              mk_Trueprop_eq (Free (Binding.name_of b, fcT --> T) $ u,
+                Term.list_comb (mk_case As T case0, mk_sel_case_args b proto_sels T) $ u)
+            end;
+
+          val sel_bindings = flat sel_bindingss;
+          val uniq_sel_bindings = distinct Binding.eq_name sel_bindings;
+          val all_sels_distinct = (length uniq_sel_bindings = length sel_bindings);
+
+          val sel_binding_index =
+            if all_sels_distinct then 1 upto length sel_bindings
+            else map (fn b => find_index (curry Binding.eq_name b) uniq_sel_bindings) sel_bindings;
+
+          val proto_sels = flat (map3 (fn k => fn xs => map (fn x => (k, (xs, x)))) ks xss xss);
+          val sel_infos =
+            AList.group (op =) (sel_binding_index ~~ proto_sels)
+            |> sort (int_ord o pairself fst)
+            |> map snd |> curry (op ~~) uniq_sel_bindings;
+          val sel_bindings = map fst sel_infos;
+
+          fun unflat_selss xs = unflat_lookup Binding.eq_name sel_bindings xs sel_bindingss;
+
+          val (((raw_discs, raw_disc_defs), (raw_sels, raw_sel_defs)), (lthy', lthy)) =
+            lthy
+            |> apfst split_list o fold_map3 (fn k => fn exist_xs_u_eq_ctr => fn b =>
+                if Binding.is_empty b then
+                  if n = 1 then pair (Term.lambda u (mk_uu_eq ()), unique_disc_no_def)
+                  else pair (alternate_disc k, alternate_disc_no_def)
+                else if Binding.eq_name (b, equal_binding) then
+                  pair (Term.lambda u exist_xs_u_eq_ctr, refl)
+                else
+                  Specification.definition (SOME (b, NONE, NoSyn),
+                    ((Thm.def_binding b, []), disc_spec b exist_xs_u_eq_ctr)) #>> apsnd snd)
+              ks exist_xs_u_eq_ctrs disc_bindings
+            ||>> apfst split_list o fold_map (fn (b, proto_sels) =>
+              Specification.definition (SOME (b, NONE, NoSyn),
+                ((Thm.def_binding b, []), sel_spec b proto_sels)) #>> apsnd snd) sel_infos
+            ||> `Local_Theory.restore;
+
+          val phi = Proof_Context.export_morphism lthy lthy';
+
+          val disc_defs = map (Morphism.thm phi) raw_disc_defs;
+          val sel_defs = map (Morphism.thm phi) raw_sel_defs;
+          val sel_defss = unflat_selss sel_defs;
+
+          val discs0 = map (Morphism.term phi) raw_discs;
+          val selss0 = unflat_selss (map (Morphism.term phi) raw_sels);
+
+          val discs = map (mk_disc_or_sel As) discs0;
+          val selss = map (map (mk_disc_or_sel As)) selss0;
+        in
+          (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy')
+        end;
+
+    fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p);
+
+    val exhaust_goal =
+      let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (u, xctr)]) in
+        fold_rev Logic.all [p, u] (mk_imp_p (map2 mk_prem xctrs xss))
+      end;
+
+    val inject_goalss =
+      let
+        fun mk_goal _ _ [] [] = []
+          | mk_goal xctr yctr xs ys =
+            [fold_rev Logic.all (xs @ ys) (mk_Trueprop_eq (HOLogic.mk_eq (xctr, yctr),
+              Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))];
+      in
+        map4 mk_goal xctrs yctrs xss yss
+      end;
+
+    val half_distinct_goalss =
+      let
+        fun mk_goal ((xs, xc), (xs', xc')) =
+          fold_rev Logic.all (xs @ xs')
+            (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc'))));
+      in
+        map (map mk_goal) (mk_half_pairss (`I (xss ~~ xctrs)))
+      end;
+
+    val goalss = [exhaust_goal] :: inject_goalss @ half_distinct_goalss;
+
+    fun after_qed thmss lthy =
+      let
+        val ([exhaust_thm], (inject_thmss, half_distinct_thmss)) = (hd thmss, chop n (tl thmss));
+
+        val inject_thms = flat inject_thmss;
+
+        val rho_As = map (pairself (certifyT lthy)) (map Logic.varifyT_global As ~~ As);
+
+        fun inst_thm t thm =
+          Drule.instantiate' [] [SOME (certify lthy t)]
+            (Thm.instantiate (rho_As, []) (Drule.zero_var_indexes thm));
+
+        val uexhaust_thm = inst_thm u exhaust_thm;
+
+        val exhaust_cases = map base_name_of_ctr ctrs;
+
+        val other_half_distinct_thmss = map (map (fn thm => thm RS not_sym)) half_distinct_thmss;
+
+        val (distinct_thms, (distinct_thmsss', distinct_thmsss)) =
+          join_halves n half_distinct_thmss other_half_distinct_thmss ||> `transpose;
+
+        val nchotomy_thm =
+          let
+            val goal =
+              HOLogic.mk_Trueprop (HOLogic.mk_all (fst u', snd u',
+                Library.foldr1 HOLogic.mk_disj exist_xs_u_eq_ctrs));
+          in
+            Goal.prove_sorry lthy [] [] goal (fn _ => mk_nchotomy_tac n exhaust_thm)
+            |> Thm.close_derivation
+          end;
+
+        val case_thms =
+          let
+            val goals =
+              map3 (fn xctr => fn xf => fn xs =>
+                fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xctrs xfs xss;
+          in
+            map4 (fn k => fn goal => fn injects => fn distinctss =>
+                Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
+                  mk_case_tac ctxt n k case_def injects distinctss)
+                |> Thm.close_derivation)
+              ks goals inject_thmss distinct_thmsss
+          end;
+
+        val (case_cong_thm, weak_case_cong_thm) =
+          let
+            fun mk_prem xctr xs xf xg =
+              fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr),
+                mk_Trueprop_eq (xf, xg)));
+
+            val goal =
+              Logic.list_implies (uv_eq :: map4 mk_prem xctrs xss xfs xgs,
+                 mk_Trueprop_eq (eta_ufcase, eta_vgcase));
+            val weak_goal = Logic.mk_implies (uv_eq, mk_Trueprop_eq (ufcase, vfcase));
+          in
+            (Goal.prove_sorry lthy [] [] goal (fn _ => mk_case_cong_tac lthy uexhaust_thm case_thms),
+             Goal.prove_sorry lthy [] [] weak_goal (K (etac arg_cong 1)))
+            |> pairself (Thm.close_derivation #> singleton (Proof_Context.export names_lthy lthy))
+          end;
+
+        val split_lhs = q $ ufcase;
+
+        fun mk_split_conjunct xctr xs f_xs =
+          list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (u, xctr), q $ f_xs));
+        fun mk_split_disjunct xctr xs f_xs =
+          list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr),
+            HOLogic.mk_not (q $ f_xs)));
+
+        fun mk_split_goal xctrs xss xfs =
+          mk_Trueprop_eq (split_lhs, Library.foldr1 HOLogic.mk_conj
+            (map3 mk_split_conjunct xctrs xss xfs));
+        fun mk_split_asm_goal xctrs xss xfs =
+          mk_Trueprop_eq (split_lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj
+            (map3 mk_split_disjunct xctrs xss xfs)));
+
+        fun prove_split selss goal =
+          Goal.prove_sorry lthy [] [] goal (fn _ =>
+            mk_split_tac lthy uexhaust_thm case_thms selss inject_thmss distinct_thmsss)
+          |> Thm.close_derivation
+          |> singleton (Proof_Context.export names_lthy lthy);
+
+        fun prove_split_asm asm_goal split_thm =
+          Goal.prove_sorry lthy [] [] asm_goal (fn {context = ctxt, ...} =>
+            mk_split_asm_tac ctxt split_thm)
+          |> Thm.close_derivation
+          |> singleton (Proof_Context.export names_lthy lthy);
+
+        val (split_thm, split_asm_thm) =
+          let
+            val goal = mk_split_goal xctrs xss xfs;
+            val asm_goal = mk_split_asm_goal xctrs xss xfs;
+
+            val thm = prove_split (replicate n []) goal;
+            val asm_thm = prove_split_asm asm_goal thm;
+          in
+            (thm, asm_thm)
+          end;
+
+        val (all_sel_thms, sel_thmss, disc_thmss, nontriv_disc_thms, discI_thms, nontriv_discI_thms,
+             disc_exclude_thms, disc_exhaust_thms, sel_exhaust_thms, all_collapse_thms,
+             safe_collapse_thms, expand_thms, sel_split_thms, sel_split_asm_thms, case_eq_if_thms) =
+          if no_discs_sels then
+            ([], [], [], [], [], [], [], [], [], [], [], [], [], [], [])
+          else
+            let
+              val udiscs = map (rapp u) discs;
+              val uselss = map (map (rapp u)) selss;
+              val usel_ctrs = map2 (curry Term.list_comb) ctrs uselss;
+              val usel_fs = map2 (curry Term.list_comb) fs uselss;
+
+              val vdiscs = map (rapp v) discs;
+              val vselss = map (map (rapp v)) selss;
+
+              fun make_sel_thm xs' case_thm sel_def =
+                zero_var_indexes (Drule.gen_all (Drule.rename_bvars' (map (SOME o fst) xs')
+                    (Drule.forall_intr_vars (case_thm RS (sel_def RS trans)))));
+
+              val sel_thmss = map3 (map oo make_sel_thm) xss' case_thms sel_defss;
+
+              fun has_undefined_rhs thm =
+                (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of thm))) of
+                  Const (@{const_name undefined}, _) => true
+                | _ => false);
+
+              val all_sel_thms =
+                (if all_sels_distinct andalso forall null sel_defaultss then
+                   flat sel_thmss
+                 else
+                   map_product (fn s => fn (xs', c) => make_sel_thm xs' c s) sel_defs
+                     (xss' ~~ case_thms))
+                |> filter_out has_undefined_rhs;
+
+              fun mk_unique_disc_def () =
+                let
+                  val m = the_single ms;
+                  val goal = mk_Trueprop_eq (mk_uu_eq (), the_single exist_xs_u_eq_ctrs);
+                in
+                  Goal.prove_sorry lthy [] [] goal (fn _ => mk_unique_disc_def_tac m uexhaust_thm)
+                  |> Thm.close_derivation
+                  |> singleton (Proof_Context.export names_lthy lthy)
+                end;
+
+              fun mk_alternate_disc_def k =
+                let
+                  val goal =
+                    mk_Trueprop_eq (alternate_disc_lhs (K (nth udiscs)) (3 - k),
+                      nth exist_xs_u_eq_ctrs (k - 1));
+                in
+                  Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
+                    mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k))
+                      (nth distinct_thms (2 - k)) uexhaust_thm)
+                  |> Thm.close_derivation
+                  |> singleton (Proof_Context.export names_lthy lthy)
+                end;
+
+              val has_alternate_disc_def =
+                exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs;
+
+              val disc_defs' =
+                map2 (fn k => fn def =>
+                  if Thm.eq_thm_prop (def, unique_disc_no_def) then mk_unique_disc_def ()
+                  else if Thm.eq_thm_prop (def, alternate_disc_no_def) then mk_alternate_disc_def k
+                  else def) ks disc_defs;
+
+              val discD_thms = map (fn def => def RS iffD1) disc_defs';
+              val discI_thms =
+                map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms
+                  disc_defs';
+              val not_discI_thms =
+                map2 (fn m => fn def => funpow m (fn thm => allI RS thm)
+                    (unfold_thms lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]})))
+                  ms disc_defs';
+
+              val (disc_thmss', disc_thmss) =
+                let
+                  fun mk_thm discI _ [] = refl RS discI
+                    | mk_thm _ not_discI [distinct] = distinct RS not_discI;
+                  fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss;
+                in
+                  map3 mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose
+                end;
+
+              val nontriv_disc_thms =
+                flat (map2 (fn b => if is_disc_binding_valid b then I else K [])
+                  disc_bindings disc_thmss);
+
+              fun is_discI_boring b =
+                (n = 1 andalso Binding.is_empty b) orelse Binding.eq_name (b, equal_binding);
+
+              val nontriv_discI_thms =
+                flat (map2 (fn b => if is_discI_boring b then K [] else single) disc_bindings
+                  discI_thms);
+
+              val (disc_exclude_thms, (disc_exclude_thmsss', disc_exclude_thmsss)) =
+                let
+                  fun mk_goal [] = []
+                    | mk_goal [((_, udisc), (_, udisc'))] =
+                      [Logic.all u (Logic.mk_implies (HOLogic.mk_Trueprop udisc,
+                         HOLogic.mk_Trueprop (HOLogic.mk_not udisc')))];
+
+                  fun prove tac goal =
+                    Goal.prove_sorry lthy [] [] goal (K tac)
+                    |> Thm.close_derivation;
+
+                  val half_pairss = mk_half_pairss (`I (ms ~~ discD_thms ~~ udiscs));
+
+                  val half_goalss = map mk_goal half_pairss;
+                  val half_thmss =
+                    map3 (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] =>
+                        fn disc_thm => [prove (mk_half_disc_exclude_tac lthy m discD disc_thm) goal])
+                      half_goalss half_pairss (flat disc_thmss');
+
+                  val other_half_goalss = map (mk_goal o map swap) half_pairss;
+                  val other_half_thmss =
+                    map2 (map2 (prove o mk_other_half_disc_exclude_tac)) half_thmss
+                      other_half_goalss;
+                in
+                  join_halves n half_thmss other_half_thmss ||> `transpose
+                  |>> has_alternate_disc_def ? K []
+                end;
+
+              val disc_exhaust_thm =
+                let
+                  fun mk_prem udisc = mk_imp_p [HOLogic.mk_Trueprop udisc];
+                  val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem udiscs));
+                in
+                  Goal.prove_sorry lthy [] [] goal (fn _ =>
+                    mk_disc_exhaust_tac n exhaust_thm discI_thms)
+                  |> Thm.close_derivation
+                end;
+
+              val (safe_collapse_thms, all_collapse_thms) =
+                let
+                  fun mk_goal m udisc usel_ctr =
+                    let
+                      val prem = HOLogic.mk_Trueprop udisc;
+                      val concl = mk_Trueprop_eq ((usel_ctr, u) |> m = 0 ? swap);
+                    in
+                      (prem aconv concl, Logic.all u (Logic.mk_implies (prem, concl)))
+                    end;
+                  val (trivs, goals) = map3 mk_goal ms udiscs usel_ctrs |> split_list;
+                  val thms =
+                    map5 (fn m => fn discD => fn sel_thms => fn triv => fn goal =>
+                        Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
+                          mk_collapse_tac ctxt m discD sel_thms ORELSE HEADGOAL atac)
+                        |> Thm.close_derivation
+                        |> not triv ? perhaps (try (fn thm => refl RS thm)))
+                      ms discD_thms sel_thmss trivs goals;
+                in
+                  (map_filter (fn (true, _) => NONE | (false, thm) => SOME thm) (trivs ~~ thms),
+                   thms)
+                end;
+
+              val swapped_all_collapse_thms =
+                map2 (fn m => fn thm => if m = 0 then thm else thm RS sym) ms all_collapse_thms;
+
+              val sel_exhaust_thm =
+                let
+                  fun mk_prem usel_ctr = mk_imp_p [mk_Trueprop_eq (u, usel_ctr)];
+                  val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem usel_ctrs));
+                in
+                  Goal.prove_sorry lthy [] [] goal (fn _ =>
+                    mk_sel_exhaust_tac n disc_exhaust_thm swapped_all_collapse_thms)
+                  |> Thm.close_derivation
+                end;
+
+              val expand_thm =
+                let
+                  fun mk_prems k udisc usels vdisc vsels =
+                    (if k = n then [] else [mk_Trueprop_eq (udisc, vdisc)]) @
+                    (if null usels then
+                       []
+                     else
+                       [Logic.list_implies
+                          (if n = 1 then [] else map HOLogic.mk_Trueprop [udisc, vdisc],
+                             HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+                               (map2 (curry HOLogic.mk_eq) usels vsels)))]);
+
+                  val goal =
+                    Library.foldr Logic.list_implies
+                      (map5 mk_prems ks udiscs uselss vdiscs vselss, uv_eq);
+                  val uncollapse_thms =
+                    map2 (fn thm => fn [] => thm | _ => thm RS sym) all_collapse_thms uselss;
+                in
+                  Goal.prove_sorry lthy [] [] goal (fn _ =>
+                    mk_expand_tac lthy n ms (inst_thm u disc_exhaust_thm)
+                      (inst_thm v disc_exhaust_thm) uncollapse_thms disc_exclude_thmsss
+                      disc_exclude_thmsss')
+                  |> Thm.close_derivation
+                  |> singleton (Proof_Context.export names_lthy lthy)
+                end;
+
+              val (sel_split_thm, sel_split_asm_thm) =
+                let
+                  val zss = map (K []) xss;
+                  val goal = mk_split_goal usel_ctrs zss usel_fs;
+                  val asm_goal = mk_split_asm_goal usel_ctrs zss usel_fs;
+
+                  val thm = prove_split sel_thmss goal;
+                  val asm_thm = prove_split_asm asm_goal thm;
+                in
+                  (thm, asm_thm)
+                end;
+
+              val case_eq_if_thm =
+                let
+                  val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs usel_fs);
+                in
+                  Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
+                    mk_case_eq_if_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss)
+                  |> Thm.close_derivation
+                  |> singleton (Proof_Context.export names_lthy lthy)
+                end;
+            in
+              (all_sel_thms, sel_thmss, disc_thmss, nontriv_disc_thms, discI_thms,
+               nontriv_discI_thms, disc_exclude_thms, [disc_exhaust_thm], [sel_exhaust_thm],
+               all_collapse_thms, safe_collapse_thms, [expand_thm], [sel_split_thm],
+               [sel_split_asm_thm], [case_eq_if_thm])
+            end;
+
+        val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases));
+        val cases_type_attr = Attrib.internal (K (Induct.cases_type fcT_name));
+
+        val anonymous_notes =
+          [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs),
+           (map (fn th => th RS @{thm eq_False[THEN iffD2]}
+              handle THM _ => th RS @{thm eq_True[THEN iffD2]}) nontriv_disc_thms,
+            code_nitpicksimp_attrs)]
+          |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
+
+        val notes =
+          [(caseN, case_thms, code_nitpicksimp_simp_attrs),
+           (case_congN, [case_cong_thm], []),
+           (case_eq_ifN, case_eq_if_thms, []),
+           (collapseN, safe_collapse_thms, simp_attrs),
+           (discN, nontriv_disc_thms, simp_attrs),
+           (discIN, nontriv_discI_thms, []),
+           (disc_excludeN, disc_exclude_thms, dest_attrs),
+           (disc_exhaustN, disc_exhaust_thms, [exhaust_case_names_attr]),
+           (distinctN, distinct_thms, simp_attrs @ inductsimp_attrs),
+           (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]),
+           (expandN, expand_thms, []),
+           (injectN, inject_thms, iff_attrs @ inductsimp_attrs),
+           (nchotomyN, [nchotomy_thm], []),
+           (selN, all_sel_thms, code_nitpicksimp_simp_attrs),
+           (sel_exhaustN, sel_exhaust_thms, [exhaust_case_names_attr]),
+           (sel_splitN, sel_split_thms, []),
+           (sel_split_asmN, sel_split_asm_thms, []),
+           (splitN, [split_thm], []),
+           (split_asmN, [split_asm_thm], []),
+           (splitsN, [split_thm, split_asm_thm], []),
+           (weak_case_cong_thmsN, [weak_case_cong_thm], cong_attrs)]
+          |> filter_out (null o #2)
+          |> map (fn (thmN, thms, attrs) =>
+            ((qualify true (Binding.name thmN), attrs), [(thms, [])]));
+
+        val ctr_sugar =
+          {ctrs = ctrs, casex = casex, discs = discs, selss = selss, exhaust = exhaust_thm,
+           nchotomy = nchotomy_thm, injects = inject_thms, distincts = distinct_thms,
+           case_thms = case_thms, case_cong = case_cong_thm, weak_case_cong = weak_case_cong_thm,
+           split = split_thm, split_asm = split_asm_thm, disc_thmss = disc_thmss,
+           discIs = discI_thms, sel_thmss = sel_thmss, disc_exhausts = disc_exhaust_thms,
+           sel_exhausts = sel_exhaust_thms, collapses = all_collapse_thms, expands = expand_thms,
+           sel_splits = sel_split_thms, sel_split_asms = sel_split_asm_thms,
+           case_eq_ifs = case_eq_if_thms};
+      in
+        (ctr_sugar,
+         lthy
+         |> not rep_compat ?
+            Local_Theory.declaration {syntax = false, pervasive = true}
+              (fn phi => Case_Translation.register
+                 (Morphism.term phi casex) (map (Morphism.term phi) ctrs))
+         |> Local_Theory.background_theory (fold (fold Code.del_eqn) [disc_defs, sel_defs])
+         |> not no_code ?
+            Local_Theory.declaration {syntax = false, pervasive = false}
+              (fn phi => Context.mapping
+                (add_ctr_code fcT_name (map (Morphism.typ phi) As)
+                  (map (dest_Const o Morphism.term phi) ctrs) (Morphism.fact phi inject_thms)
+                  (Morphism.fact phi distinct_thms) (Morphism.fact phi case_thms))
+                I)
+         |> Local_Theory.notes (anonymous_notes @ notes) |> snd
+         |> register_ctr_sugar fcT_name ctr_sugar)
+      end;
+  in
+    (goalss, after_qed, lthy')
+  end;
+
+fun wrap_free_constructors tacss = (fn (goalss, after_qed, lthy) =>
+  map2 (map2 (Thm.close_derivation oo Goal.prove_sorry lthy [] [])) goalss tacss
+  |> (fn thms => after_qed thms lthy)) oo prepare_wrap_free_constructors (K I);
+
+val wrap_free_constructors_cmd = (fn (goalss, after_qed, lthy) =>
+  Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo
+  prepare_wrap_free_constructors Syntax.read_term;
+
+fun parse_bracket_list parser = @{keyword "["} |-- Parse.list parser --|  @{keyword "]"};
+
+val parse_bindings = parse_bracket_list parse_binding;
+val parse_bindingss = parse_bracket_list parse_bindings;
+
+val parse_bound_term = (parse_binding --| @{keyword ":"}) -- Parse.term;
+val parse_bound_terms = parse_bracket_list parse_bound_term;
+val parse_bound_termss = parse_bracket_list parse_bound_terms;
+
+val parse_wrap_free_constructors_options =
+  Scan.optional (@{keyword "("} |-- Parse.list1
+        (Parse.reserved "no_discs_sels" >> K 0 || Parse.reserved "no_code" >> K 1 ||
+         Parse.reserved "rep_compat" >> K 2) --| @{keyword ")"}
+      >> (fn js => (member (op =) js 0, (member (op =) js 1, member (op =) js 2))))
+    (false, (false, false));
+
+val _ =
+  Outer_Syntax.local_theory_to_proof @{command_spec "wrap_free_constructors"}
+    "wrap an existing freely generated type's constructors"
+    ((parse_wrap_free_constructors_options -- (@{keyword "["} |-- Parse.list Parse.term --|
+        @{keyword "]"}) --
+      parse_binding -- Scan.optional (parse_bindings -- Scan.optional (parse_bindingss --
+        Scan.optional parse_bound_termss []) ([], [])) ([], ([], [])))
+     >> wrap_free_constructors_cmd);
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML	Mon Dec 09 09:44:57 2013 +0100
@@ -0,0 +1,131 @@
+(*  Title:      HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML
+    Author:     Jasmin Blanchette, TU Muenchen
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Stefan Berghofer, TU Muenchen
+    Author:     Florian Haftmann, TU Muenchen
+    Copyright   2001-2013
+
+Code generation for freely generated types.
+*)
+
+signature CTR_SUGAR_CODE =
+sig
+  val add_ctr_code: string -> typ list -> (string * typ) list -> thm list -> thm list -> thm list ->
+    theory -> theory
+end;
+
+structure Ctr_Sugar_Code : CTR_SUGAR_CODE =
+struct
+
+open Ctr_Sugar_Util
+
+val eqN = "eq"
+val reflN = "refl"
+val simpsN = "simps"
+
+fun mk_case_certificate thy raw_thms =
+  let
+    val thms as thm1 :: _ = raw_thms
+      |> Conjunction.intr_balanced
+      |> Thm.unvarify_global
+      |> Conjunction.elim_balanced (length raw_thms)
+      |> map Simpdata.mk_meta_eq
+      |> map Drule.zero_var_indexes;
+    val params = Term.add_free_names (Thm.prop_of thm1) [];
+    val rhs = thm1
+      |> Thm.prop_of |> Logic.dest_equals |> fst |> Term.strip_comb
+      ||> fst o split_last |> list_comb;
+    val lhs = Free (singleton (Name.variant_list params) "case", Term.fastype_of rhs);
+    val assum = Thm.cterm_of thy (Logic.mk_equals (lhs, rhs));
+  in
+    thms
+    |> Conjunction.intr_balanced
+    |> rewrite_rule [Thm.symmetric (Thm.assume assum)]
+    |> Thm.implies_intr assum
+    |> Thm.generalize ([], params) 0
+    |> Axclass.unoverload thy
+    |> Thm.varifyT_global
+  end;
+
+fun mk_free_ctr_equations fcT ctrs inject_thms distinct_thms thy =
+  let
+    fun mk_fcT_eq (t, u) = Const (@{const_name HOL.equal}, fcT --> fcT --> HOLogic.boolT) $ t $ u;
+    fun true_eq tu = HOLogic.mk_eq (mk_fcT_eq tu, @{term True});
+    fun false_eq tu = HOLogic.mk_eq (mk_fcT_eq tu, @{term False});
+
+    val monomorphic_prop_of = prop_of o Thm.unvarify_global o Drule.zero_var_indexes;
+
+    fun massage_inject (tp $ (eqv $ (_ $ t $ u) $ rhs)) = tp $ (eqv $ mk_fcT_eq (t, u) $ rhs);
+    fun massage_distinct (tp $ (_ $ (_ $ t $ u))) = [tp $ false_eq (t, u), tp $ false_eq (u, t)];
+
+    val triv_inject_goals =
+      map_filter (fn c as (_, T) =>
+          if T = fcT then SOME (HOLogic.mk_Trueprop (true_eq (Const c, Const c))) else NONE)
+        ctrs;
+    val inject_goals = map (massage_inject o monomorphic_prop_of) inject_thms;
+    val distinct_goals = maps (massage_distinct o monomorphic_prop_of) distinct_thms;
+    val refl_goal = HOLogic.mk_Trueprop (true_eq (Free ("x", fcT), Free ("x", fcT)));
+
+    val simp_ctxt =
+      Simplifier.global_context thy HOL_basic_ss
+        addsimps (map Simpdata.mk_eq (@{thms equal eq_True} @ inject_thms @ distinct_thms));
+
+    fun prove goal =
+      Goal.prove_sorry_global thy [] [] goal (K (ALLGOALS (simp_tac simp_ctxt)))
+      |> Simpdata.mk_eq;
+  in
+    (map prove (triv_inject_goals @ inject_goals @ distinct_goals), prove refl_goal)
+  end;
+
+fun add_equality fcT fcT_name As ctrs inject_thms distinct_thms =
+  let
+    fun add_def lthy =
+      let
+        fun mk_side const_name =
+          Const (const_name, fcT --> fcT --> HOLogic.boolT) $ Free ("x", fcT) $ Free ("y", fcT);
+        val spec =
+          mk_Trueprop_eq (mk_side @{const_name HOL.equal}, mk_side @{const_name HOL.eq})
+          |> Syntax.check_term lthy;
+        val ((_, (_, raw_def)), lthy') =
+          Specification.definition (NONE, (Attrib.empty_binding, spec)) lthy;
+        val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of lthy); (* FIXME? *)
+        val def = singleton (Proof_Context.export lthy' ctxt_thy) raw_def;
+      in
+        (def, lthy')
+      end;
+
+    fun tac thms = Class.intro_classes_tac [] THEN ALLGOALS (Proof_Context.fact_tac thms);
+
+    val qualify =
+      Binding.qualify true (Long_Name.base_name fcT_name) o Binding.qualify true eqN o Binding.name;
+  in
+    Class.instantiation ([fcT_name], map dest_TFree As, [HOLogic.class_equal])
+    #> add_def
+    #-> Class.prove_instantiation_exit_result (map o Morphism.thm) (K tac) o single
+    #-> fold Code.del_eqn
+    #> `(mk_free_ctr_equations fcT ctrs inject_thms distinct_thms)
+    #-> (fn (thms, thm) => Global_Theory.note_thmss Thm.lemmaK
+      [((qualify reflN, [Code.add_nbe_default_eqn_attribute]), [([thm], [])]),
+        ((qualify simpsN, [Code.add_default_eqn_attribute]), [(rev thms, [])])])
+    #> snd
+  end;
+
+fun add_ctr_code fcT_name raw_As raw_ctrs inject_thms distinct_thms case_thms thy =
+  let
+    val As = map (perhaps (try Logic.unvarifyT_global)) raw_As;
+    val ctrs = map (apsnd (perhaps (try Logic.unvarifyT_global))) raw_ctrs;
+    val fcT = Type (fcT_name, As);
+    val unover_ctrs = map (fn ctr as (_, fcT) => (Axclass.unoverload_const thy ctr, fcT)) ctrs;
+  in
+    if can (Code.constrset_of_consts thy) unover_ctrs then
+      thy
+      |> Code.add_datatype ctrs
+      |> fold_rev Code.add_default_eqn case_thms
+      |> Code.add_case (mk_case_certificate thy case_thms)
+      |> not (Sorts.has_instance (Sign.classes_of thy) fcT_name [HOLogic.class_equal])
+        ? add_equality fcT fcT_name As ctrs inject_thms distinct_thms
+    else
+      thy
+  end;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML	Mon Dec 09 09:44:57 2013 +0100
@@ -0,0 +1,172 @@
+(*  Title:      HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012, 2013
+
+Tactics for wrapping existing freely generated type's constructors.
+*)
+
+signature CTR_SUGAR_GENERAL_TACTICS =
+sig
+  val select_prem_tac: int -> (int -> tactic) -> int -> int -> tactic
+  val unfold_thms_tac: Proof.context -> thm list -> tactic
+end;
+
+signature CTR_SUGAR_TACTICS =
+sig
+  include CTR_SUGAR_GENERAL_TACTICS
+
+  val mk_alternate_disc_def_tac: Proof.context -> int -> thm -> thm -> thm -> tactic
+  val mk_case_tac: Proof.context -> int -> int -> thm -> thm list -> thm list list -> tactic
+  val mk_case_cong_tac: Proof.context -> thm -> thm list -> tactic
+  val mk_case_eq_if_tac: Proof.context -> int -> thm -> thm list -> thm list list ->
+    thm list list -> tactic
+  val mk_collapse_tac: Proof.context -> int -> thm -> thm list -> tactic
+  val mk_disc_exhaust_tac: int -> thm -> thm list -> tactic
+  val mk_expand_tac: Proof.context -> int -> int list -> thm -> thm -> thm list ->
+    thm list list list -> thm list list list -> tactic
+  val mk_half_disc_exclude_tac: Proof.context -> int -> thm -> thm -> tactic
+  val mk_nchotomy_tac: int -> thm -> tactic
+  val mk_other_half_disc_exclude_tac: thm -> tactic
+  val mk_sel_exhaust_tac: int -> thm -> thm list -> tactic
+  val mk_split_tac: Proof.context -> thm -> thm list -> thm list list -> thm list list ->
+    thm list list list -> tactic
+  val mk_split_asm_tac: Proof.context -> thm -> tactic
+  val mk_unique_disc_def_tac: int -> thm -> tactic
+end;
+
+structure Ctr_Sugar_Tactics : CTR_SUGAR_TACTICS =
+struct
+
+open Ctr_Sugar_Util
+
+val meta_mp = @{thm meta_mp};
+
+fun select_prem_tac n tac k = DETERM o (EVERY' [REPEAT_DETERM_N (k - 1) o etac thin_rl,
+  tac, REPEAT_DETERM_N (n - k) o etac thin_rl]);
+
+fun unfold_thms_tac _ [] = all_tac
+  | unfold_thms_tac ctxt thms = Local_Defs.unfold_tac ctxt (distinct Thm.eq_thm_prop thms);
+
+fun if_P_or_not_P_OF pos thm = thm RS (if pos then @{thm if_P} else @{thm if_not_P});
+
+fun mk_nchotomy_tac n exhaust =
+  HEADGOAL (rtac allI THEN' rtac exhaust THEN'
+   EVERY' (maps (fn k => [rtac (mk_disjIN n k), REPEAT_DETERM o rtac exI, atac]) (1 upto n)));
+
+fun mk_unique_disc_def_tac m uexhaust =
+  HEADGOAL (EVERY' [rtac iffI, rtac uexhaust, REPEAT_DETERM_N m o rtac exI, atac, rtac refl]);
+
+fun mk_alternate_disc_def_tac ctxt k other_disc_def distinct uexhaust =
+  HEADGOAL (EVERY' ([rtac (other_disc_def RS @{thm arg_cong[of _ _ Not]} RS trans),
+    rtac @{thm iffI_np}, REPEAT_DETERM o etac exE,
+    hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt [not_ex]), REPEAT_DETERM o rtac allI,
+    rtac distinct, rtac uexhaust] @
+    (([etac notE, REPEAT_DETERM o rtac exI, atac], [REPEAT_DETERM o rtac exI, atac])
+     |> k = 1 ? swap |> op @)));
+
+fun mk_half_disc_exclude_tac ctxt m discD disc' =
+  HEADGOAL (dtac discD THEN' REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac ctxt THEN'
+    rtac disc');
+
+fun mk_other_half_disc_exclude_tac half = HEADGOAL (etac @{thm contrapos_pn} THEN' etac half);
+
+fun mk_disc_or_sel_exhaust_tac n exhaust destIs =
+  HEADGOAL (rtac exhaust THEN'
+    EVERY' (map2 (fn k => fn destI => dtac destI THEN'
+      select_prem_tac n (etac meta_mp) k THEN' atac) (1 upto n) destIs));
+
+val mk_disc_exhaust_tac = mk_disc_or_sel_exhaust_tac;
+
+fun mk_sel_exhaust_tac n disc_exhaust collapses =
+  mk_disc_or_sel_exhaust_tac n disc_exhaust collapses ORELSE
+  HEADGOAL (etac meta_mp THEN' resolve_tac collapses);
+
+fun mk_collapse_tac ctxt m discD sels =
+  HEADGOAL (dtac discD THEN'
+    (if m = 0 then
+       atac
+     else
+       REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac ctxt THEN'
+       SELECT_GOAL (unfold_thms_tac ctxt sels) THEN' rtac refl));
+
+fun mk_expand_tac ctxt n ms udisc_exhaust vdisc_exhaust uncollapses disc_excludesss
+    disc_excludesss' =
+  if ms = [0] then
+    HEADGOAL (rtac (@{thm trans_sym} OF (replicate 2 (the_single uncollapses))) THEN'
+      TRY o EVERY' [rtac udisc_exhaust, atac, rtac vdisc_exhaust, atac])
+  else
+    let val ks = 1 upto n in
+      HEADGOAL (rtac udisc_exhaust THEN'
+        EVERY' (map5 (fn k => fn m => fn disc_excludess => fn disc_excludess' =>
+            fn uuncollapse =>
+          EVERY' [rtac (uuncollapse RS trans) THEN' TRY o atac,
+            rtac sym, rtac vdisc_exhaust,
+            EVERY' (map4 (fn k' => fn disc_excludes => fn disc_excludes' => fn vuncollapse =>
+              EVERY'
+                (if k' = k then
+                   [rtac (vuncollapse RS trans), TRY o atac] @
+                   (if m = 0 then
+                      [rtac refl]
+                    else
+                      [if n = 1 then K all_tac else EVERY' [dtac meta_mp, atac, dtac meta_mp, atac],
+                       REPEAT_DETERM_N (Int.max (0, m - 1)) o etac conjE,
+                       asm_simp_tac (ss_only [] ctxt)])
+                 else
+                   [dtac (the_single (if k = n then disc_excludes else disc_excludes')),
+                    etac (if k = n then @{thm iff_contradict(1)} else @{thm iff_contradict(2)}),
+                    atac, atac]))
+              ks disc_excludess disc_excludess' uncollapses)])
+          ks ms disc_excludesss disc_excludesss' uncollapses))
+    end;
+
+fun mk_case_same_ctr_tac ctxt injects =
+  REPEAT_DETERM o etac exE THEN' etac conjE THEN'
+    (case injects of
+      [] => atac
+    | [inject] => dtac (inject RS iffD1) THEN' REPEAT_DETERM o etac conjE THEN'
+        hyp_subst_tac ctxt THEN' rtac refl);
+
+fun mk_case_distinct_ctrs_tac ctxt distincts =
+  REPEAT_DETERM o etac exE THEN' etac conjE THEN' full_simp_tac (ss_only distincts ctxt);
+
+fun mk_case_tac ctxt n k case_def injects distinctss =
+  let
+    val case_def' = mk_unabs_def (n + 1) (case_def RS meta_eq_to_obj_eq);
+    val ks = 1 upto n;
+  in
+    HEADGOAL (rtac (case_def' RS trans) THEN' rtac @{thm the_equality} THEN'
+      rtac (mk_disjIN n k) THEN' REPEAT_DETERM o rtac exI THEN' rtac conjI THEN' rtac refl THEN'
+      rtac refl THEN'
+      EVERY' (map2 (fn k' => fn distincts =>
+        (if k' < n then etac disjE else K all_tac) THEN'
+        (if k' = k then mk_case_same_ctr_tac ctxt injects
+         else mk_case_distinct_ctrs_tac ctxt distincts)) ks distinctss))
+  end;
+
+fun mk_case_cong_tac ctxt uexhaust cases =
+  HEADGOAL (rtac uexhaust THEN'
+    EVERY' (maps (fn casex => [dtac sym, asm_simp_tac (ss_only [casex] ctxt)]) cases));
+
+fun mk_case_eq_if_tac ctxt n uexhaust cases discss' selss =
+  HEADGOAL (rtac uexhaust THEN'
+    EVERY' (map3 (fn casex => fn if_discs => fn sels =>
+        EVERY' [hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt (if_discs @ sels)),
+          rtac casex])
+      cases (map2 (seq_conds if_P_or_not_P_OF n) (1 upto n) discss') selss));
+
+fun mk_split_tac ctxt uexhaust cases selss injectss distinctsss =
+  HEADGOAL (rtac uexhaust) THEN
+  ALLGOALS (fn k => (hyp_subst_tac ctxt THEN'
+     simp_tac (ss_only (@{thms simp_thms} @ cases @ nth selss (k - 1) @ nth injectss (k - 1) @
+       flat (nth distinctsss (k - 1))) ctxt)) k) THEN
+  ALLGOALS (blast_tac (put_claset (claset_of @{theory_context HOL}) ctxt));
+
+val split_asm_thms = @{thms imp_conv_disj de_Morgan_conj de_Morgan_disj not_not not_ex};
+
+fun mk_split_asm_tac ctxt split =
+  HEADGOAL (rtac (split RS trans)) THEN unfold_thms_tac ctxt split_asm_thms THEN
+  HEADGOAL (rtac refl);
+
+end;
+
+structure Ctr_Sugar_General_Tactics : CTR_SUGAR_GENERAL_TACTICS = Ctr_Sugar_Tactics;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML	Mon Dec 09 09:44:57 2013 +0100
@@ -0,0 +1,244 @@
+(*  Title:      HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012, 2013
+
+Library for wrapping existing freely generated type's constructors.
+*)
+
+signature CTR_SUGAR_UTIL =
+sig
+  val map3: ('a -> 'b -> 'c -> 'd) -> 'a list -> 'b list -> 'c list -> 'd list
+  val map4: ('a -> 'b -> 'c -> 'd -> 'e) -> 'a list -> 'b list -> 'c list -> 'd list -> 'e list
+  val map5: ('a -> 'b -> 'c -> 'd -> 'e -> 'f) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list
+  val fold_map2: ('a -> 'b -> 'c -> 'd * 'c) -> 'a list -> 'b list -> 'c -> 'd list * 'c
+  val fold_map3: ('a -> 'b -> 'c -> 'd -> 'e * 'd) ->
+    'a list -> 'b list -> 'c list -> 'd -> 'e list * 'd
+  val seq_conds: (bool -> 'a -> 'b) -> int -> int -> 'a list -> 'b list
+  val transpose: 'a list list -> 'a list list
+  val pad_list: 'a -> int -> 'a list -> 'a list
+  val splice: 'a list -> 'a list -> 'a list
+  val permute_like: ('a * 'b -> bool) -> 'a list -> 'b list -> 'c list -> 'c list
+
+  val mk_names: int -> string -> string list
+  val mk_fresh_names: Proof.context -> int -> string -> string list * Proof.context
+  val mk_TFrees': sort list -> Proof.context -> typ list * Proof.context
+  val mk_TFrees: int -> Proof.context -> typ list * Proof.context
+  val mk_Frees': string -> typ list -> Proof.context ->
+    (term list * (string * typ) list) * Proof.context
+  val mk_Freess': string -> typ list list -> Proof.context ->
+    (term list list * (string * typ) list list) * Proof.context
+  val mk_Frees: string -> typ list -> Proof.context -> term list * Proof.context
+  val mk_Freess: string -> typ list list -> Proof.context -> term list list * Proof.context
+  val resort_tfree: sort -> typ -> typ
+  val variant_types: string list -> sort list -> Proof.context ->
+    (string * sort) list * Proof.context
+  val variant_tfrees: string list -> Proof.context -> typ list * Proof.context
+
+  val typ_subst_nonatomic: (typ * typ) list -> typ -> typ
+  val subst_nonatomic_types: (typ * typ) list -> term -> term
+
+  val mk_predT: typ list -> typ
+  val mk_pred1T: typ -> typ
+
+  val mk_disjIN: int -> int -> thm
+
+  val mk_unabs_def: int -> thm -> thm
+
+  val mk_IfN: typ -> term list -> term list -> term
+  val mk_Trueprop_eq: term * term -> term
+
+  val rapp: term -> term -> term
+
+  val list_all_free: term list -> term -> term
+  val list_exists_free: term list -> term -> term
+
+  val fo_match: Proof.context -> term -> term -> Type.tyenv * Envir.tenv
+
+  val cterm_instantiate_pos: cterm option list -> thm -> thm
+  val unfold_thms: Proof.context -> thm list -> thm -> thm
+
+  val certifyT: Proof.context -> typ -> ctyp
+  val certify: Proof.context -> term -> cterm
+
+  val standard_binding: binding
+  val equal_binding: binding
+  val parse_binding: binding parser
+
+  val ss_only: thm list -> Proof.context -> Proof.context
+
+  val WRAP: ('a -> tactic) -> ('a -> tactic) -> 'a list -> tactic -> tactic
+  val WRAP': ('a -> int -> tactic) -> ('a -> int -> tactic) -> 'a list -> (int -> tactic) -> int ->
+    tactic
+  val CONJ_WRAP_GEN: tactic -> ('a -> tactic) -> 'a list -> tactic
+  val CONJ_WRAP_GEN': (int -> tactic) -> ('a -> int -> tactic) -> 'a list -> int -> tactic
+  val CONJ_WRAP: ('a -> tactic) -> 'a list -> tactic
+  val CONJ_WRAP': ('a -> int -> tactic) -> 'a list -> int -> tactic
+end;
+
+structure Ctr_Sugar_Util : CTR_SUGAR_UTIL =
+struct
+
+fun map3 _ [] [] [] = []
+  | map3 f (x1::x1s) (x2::x2s) (x3::x3s) = f x1 x2 x3 :: map3 f x1s x2s x3s
+  | map3 _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map4 _ [] [] [] [] = []
+  | map4 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) = f x1 x2 x3 x4 :: map4 f x1s x2s x3s x4s
+  | map4 _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map5 _ [] [] [] [] [] = []
+  | map5 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) =
+    f x1 x2 x3 x4 x5 :: map5 f x1s x2s x3s x4s x5s
+  | map5 _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun fold_map2 _ [] [] acc = ([], acc)
+  | fold_map2 f (x1::x1s) (x2::x2s) acc =
+    let
+      val (x, acc') = f x1 x2 acc;
+      val (xs, acc'') = fold_map2 f x1s x2s acc';
+    in (x :: xs, acc'') end
+  | fold_map2 _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun fold_map3 _ [] [] [] acc = ([], acc)
+  | fold_map3 f (x1::x1s) (x2::x2s) (x3::x3s) acc =
+    let
+      val (x, acc') = f x1 x2 x3 acc;
+      val (xs, acc'') = fold_map3 f x1s x2s x3s acc';
+    in (x :: xs, acc'') end
+  | fold_map3 _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun seq_conds f n k xs =
+  if k = n then
+    map (f false) (take (k - 1) xs)
+  else
+    let val (negs, pos) = split_last (take k xs) in
+      map (f false) negs @ [f true pos]
+    end;
+
+fun transpose [] = []
+  | transpose ([] :: xss) = transpose xss
+  | transpose xss = map hd xss :: transpose (map tl xss);
+
+fun pad_list x n xs = xs @ replicate (n - length xs) x;
+
+fun splice xs ys = flat (map2 (fn x => fn y => [x, y]) xs ys);
+
+fun permute_like eq xs xs' ys = map (nth ys o (fn y => find_index (fn x => eq (x, y)) xs)) xs';
+
+fun mk_names n x = if n = 1 then [x] else map (fn i => x ^ string_of_int i) (1 upto n);
+fun mk_fresh_names ctxt = (fn xs => Variable.variant_fixes xs ctxt) oo mk_names;
+
+val mk_TFrees' = apfst (map TFree) oo Variable.invent_types;
+
+fun mk_TFrees n = mk_TFrees' (replicate n HOLogic.typeS);
+
+fun mk_Frees' x Ts ctxt = mk_fresh_names ctxt (length Ts) x |>> (fn xs => `(map Free) (xs ~~ Ts));
+fun mk_Freess' x Tss = fold_map2 mk_Frees' (mk_names (length Tss) x) Tss #>> split_list;
+
+fun mk_Frees x Ts ctxt = mk_fresh_names ctxt (length Ts) x |>> (fn xs => map2 (curry Free) xs Ts);
+fun mk_Freess x Tss = fold_map2 mk_Frees (mk_names (length Tss) x) Tss;
+
+fun resort_tfree S (TFree (s, _)) = TFree (s, S);
+
+fun ensure_prefix pre s = s |> not (String.isPrefix pre s) ? prefix pre;
+
+fun variant_types ss Ss ctxt =
+  let
+    val (tfrees, _) =
+      fold_map2 (fn s => fn S => Name.variant s #> apfst (rpair S)) ss Ss (Variable.names_of ctxt);
+    val ctxt' = fold (Variable.declare_constraints o Logic.mk_type o TFree) tfrees ctxt;
+  in (tfrees, ctxt') end;
+
+fun variant_tfrees ss =
+  apfst (map TFree) o
+    variant_types (map (ensure_prefix "'") ss) (replicate (length ss) HOLogic.typeS);
+
+(*Replace each Ti by Ui (starting from the leaves); inst = [(T1, U1), ..., (Tn, Un)].*)
+fun typ_subst_nonatomic [] = I
+  | typ_subst_nonatomic inst =
+    let
+      fun subst (Type (s, Ts)) = perhaps (AList.lookup (op =) inst) (Type (s, map subst Ts))
+        | subst T = perhaps (AList.lookup (op =) inst) T;
+    in subst end;
+
+fun subst_nonatomic_types [] = I
+  | subst_nonatomic_types inst = map_types (typ_subst_nonatomic inst);
+
+fun mk_predT Ts = Ts ---> HOLogic.boolT;
+fun mk_pred1T T = mk_predT [T];
+
+fun mk_disjIN 1 1 = @{thm TrueE[OF TrueI]}
+  | mk_disjIN _ 1 = disjI1
+  | mk_disjIN 2 2 = disjI2
+  | mk_disjIN n m = (mk_disjIN (n - 1) (m - 1)) RS disjI2;
+
+fun mk_unabs_def n = funpow n (fn thm => thm RS fun_cong);
+
+fun mk_IfN _ _ [t] = t
+  | mk_IfN T (c :: cs) (t :: ts) =
+    Const (@{const_name If}, HOLogic.boolT --> T --> T --> T) $ c $ t $ mk_IfN T cs ts;
+
+val mk_Trueprop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq;
+
+fun rapp u t = betapply (t, u);
+
+fun list_quant_free quant_const =
+  fold_rev (fn Free (xT as (_, T)) => fn P => quant_const T $ Term.absfree xT P);
+
+val list_all_free = list_quant_free HOLogic.all_const;
+val list_exists_free = list_quant_free HOLogic.exists_const;
+
+fun fo_match ctxt t pat =
+  let val thy = Proof_Context.theory_of ctxt in
+    Pattern.first_order_match thy (pat, t) (Vartab.empty, Vartab.empty)
+  end;
+
+fun cterm_instantiate_pos cts thm =
+  let
+    val cert = Thm.cterm_of (Thm.theory_of_thm thm);
+    val vars = Term.add_vars (prop_of thm) [];
+    val vars' = rev (drop (length vars - length cts) vars);
+    val ps = map_filter (fn (_, NONE) => NONE
+      | (var, SOME ct) => SOME (cert (Var var), ct)) (vars' ~~ cts);
+  in
+    Drule.cterm_instantiate ps thm
+  end;
+
+fun unfold_thms ctxt thms = Local_Defs.unfold ctxt (distinct Thm.eq_thm_prop thms);
+
+(*stolen from ~~/src/HOL/Tools/SMT/smt_utils.ML*)
+fun certifyT ctxt = Thm.ctyp_of (Proof_Context.theory_of ctxt);
+fun certify ctxt = Thm.cterm_of (Proof_Context.theory_of ctxt);
+
+(* The standard binding stands for a name generated following the canonical convention (e.g.,
+   "is_Nil" from "Nil"). In contrast, the empty binding is either the standard binding or no
+   binding at all, depending on the context. *)
+val standard_binding = @{binding _};
+val equal_binding = @{binding "="};
+
+val parse_binding = Parse.binding || @{keyword "="} >> K equal_binding;
+
+fun ss_only thms ctxt = clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps thms;
+
+(*Tactical WRAP surrounds a static given tactic (core) with two deterministic chains of tactics*)
+fun WRAP gen_before gen_after xs core_tac =
+  fold_rev (fn x => fn tac => gen_before x THEN tac THEN gen_after x) xs core_tac;
+
+fun WRAP' gen_before gen_after xs core_tac =
+  fold_rev (fn x => fn tac => gen_before x THEN' tac THEN' gen_after x) xs core_tac;
+
+fun CONJ_WRAP_GEN conj_tac gen_tac xs =
+  let val (butlast, last) = split_last xs;
+  in WRAP (fn thm => conj_tac THEN gen_tac thm) (K all_tac) butlast (gen_tac last) end;
+
+fun CONJ_WRAP_GEN' conj_tac gen_tac xs =
+  let val (butlast, last) = split_last xs;
+  in WRAP' (fn thm => conj_tac THEN' gen_tac thm) (K (K all_tac)) butlast (gen_tac last) end;
+
+(*not eta-converted because of monotype restriction*)
+fun CONJ_WRAP gen_tac = CONJ_WRAP_GEN (rtac conjI 1) gen_tac;
+fun CONJ_WRAP' gen_tac = CONJ_WRAP_GEN' (rtac conjI) gen_tac;
+
+end;
--- a/src/HOL/Tools/case_translation.ML	Mon Dec 09 06:33:46 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,646 +0,0 @@
-(*  Title:      HOL/Tools/case_translation.ML
-    Author:     Konrad Slind, Cambridge University Computer Laboratory
-    Author:     Stefan Berghofer, TU Muenchen
-    Author:     Dmitriy Traytel, TU Muenchen
-
-Nested case expressions via a generic data slot for case combinators and constructors.
-*)
-
-signature CASE_TRANSLATION =
-sig
-  val indexify_names: string list -> string list
-  val make_tnames: typ list -> string list
-
-  datatype config = Error | Warning | Quiet
-  val case_tr: bool -> Proof.context -> term list -> term
-  val lookup_by_constr: Proof.context -> string * typ -> (term * term list) option
-  val lookup_by_constr_permissive: Proof.context -> string * typ -> (term * term list) option
-  val lookup_by_case: Proof.context -> string -> (term * term list) option
-  val make_case:  Proof.context -> config -> Name.context -> term -> (term * term) list -> term
-  val print_case_translations: Proof.context -> unit
-  val strip_case: Proof.context -> bool -> term -> (term * (term * term) list) option
-  val strip_case_full: Proof.context -> bool -> term -> term
-  val show_cases: bool Config.T
-  val setup: theory -> theory
-  val register: term -> term list -> Context.generic -> Context.generic
-end;
-
-structure Case_Translation: CASE_TRANSLATION =
-struct
-
-(** general utilities **)
-
-fun indexify_names names =
-  let
-    fun index (x :: xs) tab =
-        (case AList.lookup (op =) tab x of
-          NONE =>
-            if member (op =) xs x
-            then (x ^ "1") :: index xs ((x, 2) :: tab)
-            else x :: index xs tab
-        | SOME i => (x ^ string_of_int i) :: index xs ((x, i + 1) :: tab))
-      | index [] _ = [];
-  in index names [] end;
-
-fun make_tnames Ts =
-  let
-    fun type_name (TFree (name, _)) = unprefix "'" name
-      | type_name (Type (name, _)) =
-          let val name' = Long_Name.base_name name
-          in if Symbol_Pos.is_identifier name' then name' else "x" end;
-  in indexify_names (map type_name Ts) end;
-
-
-
-(** data management **)
-
-datatype data = Data of
-  {constrs: (string * (term * term list)) list Symtab.table,
-   cases: (term * term list) Symtab.table};
-
-fun make_data (constrs, cases) = Data {constrs = constrs, cases = cases};
-
-structure Data = Generic_Data
-(
-  type T = data;
-  val empty = make_data (Symtab.empty, Symtab.empty);
-  val extend = I;
-  fun merge
-    (Data {constrs = constrs1, cases = cases1},
-     Data {constrs = constrs2, cases = cases2}) =
-    make_data
-      (Symtab.join (K (AList.merge (op =) (K true))) (constrs1, constrs2),
-      Symtab.merge (K true) (cases1, cases2));
-);
-
-fun map_data f =
-  Data.map (fn Data {constrs, cases} => make_data (f (constrs, cases)));
-fun map_constrs f = map_data (fn (constrs, cases) => (f constrs, cases));
-fun map_cases f = map_data (fn (constrs, cases) => (constrs, f cases));
-
-val rep_data = (fn Data args => args) o Data.get o Context.Proof;
-
-fun T_of_data (comb, constrs : term list) =
-  fastype_of comb
-  |> funpow (length constrs) range_type
-  |> domain_type;
-
-val Tname_of_data = fst o dest_Type o T_of_data;
-
-val constrs_of = #constrs o rep_data;
-val cases_of = #cases o rep_data;
-
-fun lookup_by_constr ctxt (c, T) =
-  let
-    val tab = Symtab.lookup_list (constrs_of ctxt) c;
-  in
-    (case body_type T of
-      Type (tyco, _) => AList.lookup (op =) tab tyco
-    | _ => NONE)
-  end;
-
-fun lookup_by_constr_permissive ctxt (c, T) =
-  let
-    val tab = Symtab.lookup_list (constrs_of ctxt) c;
-    val hint = (case body_type T of Type (tyco, _) => SOME tyco | _ => NONE);
-    val default = if null tab then NONE else SOME (snd (List.last tab));
-    (*conservative wrt. overloaded constructors*)
-  in
-    (case hint of
-      NONE => default
-    | SOME tyco =>
-        (case AList.lookup (op =) tab tyco of
-          NONE => default (*permissive*)
-        | SOME info => SOME info))
-  end;
-
-val lookup_by_case = Symtab.lookup o cases_of;
-
-
-
-(** installation **)
-
-fun case_error s = error ("Error in case expression:\n" ^ s);
-
-val name_of = try (dest_Const #> fst);
-
-
-(* parse translation *)
-
-fun constrain_Abs tT t = Syntax.const @{syntax_const "_constrainAbs"} $ t $ tT;
-
-fun case_tr err ctxt [t, u] =
-      let
-        val consts = Proof_Context.consts_of ctxt;
-        fun is_const s = can (Consts.the_constraint consts) (Consts.intern consts s);
-
-        fun variant_free x used =
-          let val (x', used') = Name.variant x used
-          in if is_const x' then variant_free x' used' else (x', used') end;
-
-        fun abs p tTs t =
-          Syntax.const @{const_syntax case_abs} $
-            fold constrain_Abs tTs (absfree p t);
-
-        fun abs_pat (Const (@{syntax_const "_constrain"}, _) $ t $ tT) tTs =
-              abs_pat t (tT :: tTs)
-          | abs_pat (Free (p as (x, _))) tTs =
-              if is_const x then I else abs p tTs
-          | abs_pat (t $ u) _ = abs_pat u [] #> abs_pat t []
-          | abs_pat _ _ = I;
-
-        (* replace occurrences of dummy_pattern by distinct variables *)
-        fun replace_dummies (Const (@{const_syntax dummy_pattern}, T)) used =
-              let val (x, used') = variant_free "x" used
-              in (Free (x, T), used') end
-          | replace_dummies (t $ u) used =
-              let
-                val (t', used') = replace_dummies t used;
-                val (u', used'') = replace_dummies u used';
-              in (t' $ u', used'') end
-          | replace_dummies t used = (t, used);
-
-        fun dest_case1 (t as Const (@{syntax_const "_case1"}, _) $ l $ r) =
-              let val (l', _) = replace_dummies l (Term.declare_term_frees t Name.context) in
-                abs_pat l' []
-                  (Syntax.const @{const_syntax case_elem} $ Term_Position.strip_positions l' $ r)
-              end
-          | dest_case1 _ = case_error "dest_case1";
-
-        fun dest_case2 (Const (@{syntax_const "_case2"}, _) $ t $ u) = t :: dest_case2 u
-          | dest_case2 t = [t];
-
-        val errt = if err then @{term True} else @{term False};
-      in
-        Syntax.const @{const_syntax case_guard} $ errt $ t $
-          (fold_rev
-            (fn t => fn u => Syntax.const @{const_syntax case_cons} $ dest_case1 t $ u)
-            (dest_case2 u)
-            (Syntax.const @{const_syntax case_nil}))
-      end
-  | case_tr _ _ _ = case_error "case_tr";
-
-val trfun_setup = Sign.parse_translation [(@{syntax_const "_case_syntax"}, case_tr true)];
-
-
-(* print translation *)
-
-fun case_tr' (_ :: x :: t :: ts) =
-      let
-        fun mk_clause (Const (@{const_syntax case_abs}, _) $ Abs (s, T, t)) xs used =
-              let val (s', used') = Name.variant s used
-              in mk_clause t ((s', T) :: xs) used' end
-          | mk_clause (Const (@{const_syntax case_elem}, _) $ pat $ rhs) xs _ =
-              Syntax.const @{syntax_const "_case1"} $
-                subst_bounds (map Syntax_Trans.mark_bound_abs xs, pat) $
-                subst_bounds (map Syntax_Trans.mark_bound_body xs, rhs)
-          | mk_clause _ _ _ = raise Match;
-
-        fun mk_clauses (Const (@{const_syntax case_nil}, _)) = []
-          | mk_clauses (Const (@{const_syntax case_cons}, _) $ t $ u) =
-              mk_clause t [] (Term.declare_term_frees t Name.context) :: mk_clauses u
-          | mk_clauses _ = raise Match;
-      in
-        list_comb (Syntax.const @{syntax_const "_case_syntax"} $ x $
-          foldr1 (fn (t, u) => Syntax.const @{syntax_const "_case2"} $ t $ u)
-            (mk_clauses t), ts)
-      end
-  | case_tr' _ = raise Match;
-
-val trfun_setup' = Sign.print_translation [(@{const_syntax "case_guard"}, K case_tr')];
-
-
-(* declarations *)
-
-fun register raw_case_comb raw_constrs context =
-  let
-    val ctxt = Context.proof_of context;
-    val case_comb = singleton (Variable.polymorphic ctxt) raw_case_comb;
-    val constrs = Variable.polymorphic ctxt raw_constrs;
-    val case_key = case_comb |> dest_Const |> fst;
-    val constr_keys = map (fst o dest_Const) constrs;
-    val data = (case_comb, constrs);
-    val Tname = Tname_of_data data;
-    val update_constrs = fold (fn key => Symtab.cons_list (key, (Tname, data))) constr_keys;
-    val update_cases = Symtab.update (case_key, data);
-  in
-    context
-    |> map_constrs update_constrs
-    |> map_cases update_cases
-  end;
-
-
-(* (Un)check phases *)
-
-datatype config = Error | Warning | Quiet;
-
-exception CASE_ERROR of string * int;
-
-
-(*Each pattern carries with it a tag i, which denotes the clause it
-  came from. i = ~1 indicates that the clause was added by pattern
-  completion.*)
-
-fun add_row_used ((prfx, pats), (tm, _)) =
-  fold Term.declare_term_frees (tm :: pats @ map Free prfx);
-
-(*try to preserve names given by user*)
-fun default_name "" (Free (name', _)) = name'
-  | default_name name _ = name;
-
-
-(*Produce an instance of a constructor, plus fresh variables for its arguments.*)
-fun fresh_constr colty used c =
-  let
-    val (_, T) = dest_Const c;
-    val Ts = binder_types T;
-    val (names, _) =
-      fold_map Name.variant (make_tnames (map Logic.unvarifyT_global Ts)) used;
-    val ty = body_type T;
-    val ty_theta = Type.raw_match (ty, colty) Vartab.empty
-      handle Type.TYPE_MATCH => raise CASE_ERROR ("type mismatch", ~1);
-    val c' = Envir.subst_term_types ty_theta c;
-    val gvars = map (Envir.subst_term_types ty_theta o Free) (names ~~ Ts);
-  in (c', gvars) end;
-
-(*Go through a list of rows and pick out the ones beginning with a
-  pattern with constructor = name.*)
-fun mk_group (name, T) rows =
-  let val k = length (binder_types T) in
-    fold (fn (row as ((prfx, p :: ps), rhs as (_, i))) =>
-      fn ((in_group, not_in_group), names) =>
-        (case strip_comb p of
-          (Const (name', _), args) =>
-            if name = name' then
-              if length args = k then
-                ((((prfx, args @ ps), rhs) :: in_group, not_in_group),
-                 map2 default_name names args)
-              else raise CASE_ERROR ("Wrong number of arguments for constructor " ^ quote name, i)
-            else ((in_group, row :: not_in_group), names)
-        | _ => raise CASE_ERROR ("Not a constructor pattern", i)))
-    rows (([], []), replicate k "") |>> pairself rev
-  end;
-
-
-(* Partitioning *)
-
-fun partition _ _ _ _ [] = raise CASE_ERROR ("partition: no rows", ~1)
-  | partition used constructors colty res_ty (rows as (((prfx, _ :: ps), _) :: _)) =
-      let
-        fun part [] [] = []
-          | part [] ((_, (_, i)) :: _) = raise CASE_ERROR ("Not a constructor pattern", i)
-          | part (c :: cs) rows =
-              let
-                val ((in_group, not_in_group), names) = mk_group (dest_Const c) rows;
-                val used' = fold add_row_used in_group used;
-                val (c', gvars) = fresh_constr colty used' c;
-                val in_group' =
-                  if null in_group  (* Constructor not given *)
-                  then
-                    let
-                      val Ts = map fastype_of ps;
-                      val (xs, _) =
-                        fold_map Name.variant
-                          (replicate (length ps) "x")
-                          (fold Term.declare_term_frees gvars used');
-                    in
-                      [((prfx, gvars @ map Free (xs ~~ Ts)),
-                        (Const (@{const_name undefined}, res_ty), ~1))]
-                    end
-                  else in_group;
-              in
-                {constructor = c',
-                 new_formals = gvars,
-                 names = names,
-                 group = in_group'} :: part cs not_in_group
-              end;
-      in part constructors rows end;
-
-fun v_to_prfx (prfx, Free v :: pats) = (v :: prfx, pats)
-  | v_to_prfx _ = raise CASE_ERROR ("mk_case: v_to_prfx", ~1);
-
-
-(* Translation of pattern terms into nested case expressions. *)
-
-fun mk_case ctxt used range_ty =
-  let
-    val get_info = lookup_by_constr_permissive ctxt;
-
-    fun expand _ _ _ ((_, []), _) = raise CASE_ERROR ("mk_case: expand", ~1)
-      | expand constructors used ty (row as ((prfx, p :: ps), (rhs, tag))) =
-          if is_Free p then
-            let
-              val used' = add_row_used row used;
-              fun expnd c =
-                let val capp = list_comb (fresh_constr ty used' c)
-                in ((prfx, capp :: ps), (subst_free [(p, capp)] rhs, tag)) end;
-            in map expnd constructors end
-          else [row];
-
-    val (name, _) = Name.variant "a" used;
-
-    fun mk _ [] = raise CASE_ERROR ("no rows", ~1)
-      | mk [] (((_, []), (tm, tag)) :: _) = ([tag], tm) (* Done *)
-      | mk path ((row as ((_, [Free _]), _)) :: _ :: _) = mk path [row]
-      | mk (u :: us) (rows as ((_, _ :: _), _) :: _) =
-          let val col0 = map (fn ((_, p :: _), (_, i)) => (p, i)) rows in
-            (case Option.map (apfst head_of) (find_first (not o is_Free o fst) col0) of
-              NONE =>
-                let
-                  val rows' = map (fn ((v, _), row) => row ||>
-                    apfst (subst_free [(v, u)]) |>> v_to_prfx) (col0 ~~ rows);
-                in mk us rows' end
-            | SOME (Const (cname, cT), i) =>
-                (case get_info (cname, cT) of
-                  NONE => raise CASE_ERROR ("Not a datatype constructor: " ^ quote cname, i)
-                | SOME (case_comb, constructors) =>
-                    let
-                      val pty = body_type cT;
-                      val used' = fold Term.declare_term_frees us used;
-                      val nrows = maps (expand constructors used' pty) rows;
-                      val subproblems = partition used' constructors pty range_ty nrows;
-                      val (pat_rect, dtrees) =
-                        split_list (map (fn {new_formals, group, ...} =>
-                          mk (new_formals @ us) group) subproblems);
-                      val case_functions =
-                        map2 (fn {new_formals, names, ...} =>
-                          fold_rev (fn (x as Free (_, T), s) => fn t =>
-                            Abs (if s = "" then name else s, T, abstract_over (x, t)))
-                              (new_formals ~~ names))
-                        subproblems dtrees;
-                      val types = map fastype_of (case_functions @ [u]);
-                      val case_const = Const (name_of case_comb |> the, types ---> range_ty);
-                      val tree = list_comb (case_const, case_functions @ [u]);
-                    in (flat pat_rect, tree) end)
-            | SOME (t, i) =>
-                raise CASE_ERROR ("Not a datatype constructor: " ^ Syntax.string_of_term ctxt t, i))
-          end
-      | mk _ _ = raise CASE_ERROR ("Malformed row matrix", ~1)
-  in mk end;
-
-
-(*Repeated variable occurrences in a pattern are not allowed.*)
-fun no_repeat_vars ctxt pat = fold_aterms
-  (fn x as Free (s, _) =>
-      (fn xs =>
-        if member op aconv xs x then
-          case_error (quote s ^ " occurs repeatedly in the pattern " ^
-            quote (Syntax.string_of_term ctxt pat))
-        else x :: xs)
-    | _ => I) pat [];
-
-fun make_case ctxt config used x clauses =
-  let
-    fun string_of_clause (pat, rhs) =
-      Syntax.unparse_term ctxt
-        (Term.list_comb (Syntax.const @{syntax_const "_case1"},
-          Syntax.uncheck_terms ctxt [pat, rhs]))
-      |> Pretty.string_of;
-
-    val _ = map (no_repeat_vars ctxt o fst) clauses;
-    val rows = map_index (fn (i, (pat, rhs)) => (([], [pat]), (rhs, i))) clauses;
-    val rangeT =
-      (case distinct (op =) (map (fastype_of o snd) clauses) of
-        [] => case_error "no clauses given"
-      | [T] => T
-      | _ => case_error "all cases must have the same result type");
-    val used' = fold add_row_used rows used;
-    val (tags, case_tm) =
-      mk_case ctxt used' rangeT [x] rows
-        handle CASE_ERROR (msg, i) =>
-          case_error
-            (msg ^ (if i < 0 then "" else "\nIn clause\n" ^ string_of_clause (nth clauses i)));
-    val _ =
-      (case subtract (op =) tags (map (snd o snd) rows) of
-        [] => ()
-      | is =>
-          if config = Quiet then ()
-          else
-            (if config = Error then case_error else warning (*FIXME lack of syntactic context!?*))
-              ("The following clauses are redundant (covered by preceding clauses):\n" ^
-                cat_lines (map (string_of_clause o nth clauses) is)));
-  in
-    case_tm
-  end;
-
-
-(* term check *)
-
-fun decode_clause (Const (@{const_name case_abs}, _) $ Abs (s, T, t)) xs used =
-      let val (s', used') = Name.variant s used
-      in decode_clause t (Free (s', T) :: xs) used' end
-  | decode_clause (Const (@{const_name case_elem}, _) $ t $ u) xs _ =
-      (subst_bounds (xs, t), subst_bounds (xs, u))
-  | decode_clause _ _ _ = case_error "decode_clause";
-
-fun decode_cases (Const (@{const_name case_nil}, _)) = []
-  | decode_cases (Const (@{const_name case_cons}, _) $ t $ u) =
-      decode_clause t [] (Term.declare_term_frees t Name.context) ::
-      decode_cases u
-  | decode_cases _ = case_error "decode_cases";
-
-fun check_case ctxt =
-  let
-    fun decode_case (Const (@{const_name case_guard}, _) $ b $ u $ t) =
-          make_case ctxt (if b = @{term True} then Error else Warning)
-            Name.context (decode_case u) (decode_cases t)
-      | decode_case (t $ u) = decode_case t $ decode_case u
-      | decode_case (Abs (x, T, u)) =
-          let val (x', u') = Term.dest_abs (x, T, u);
-          in Term.absfree (x', T) (decode_case u') end
-      | decode_case t = t;
-  in
-    map decode_case
-  end;
-
-val term_check_setup =
-  Context.theory_map (Syntax_Phases.term_check 1 "case" check_case);
-
-
-(* Pretty printing of nested case expressions *)
-
-(* destruct one level of pattern matching *)
-
-fun dest_case ctxt d used t =
-  (case apfst name_of (strip_comb t) of
-    (SOME cname, ts as _ :: _) =>
-      let
-        val (fs, x) = split_last ts;
-        fun strip_abs i Us t =
-          let
-            val zs = strip_abs_vars t;
-            val j = length zs;
-            val (xs, ys) =
-              if j < i then (zs @ map (pair "x") (drop j Us), [])
-              else chop i zs;
-            val u = fold_rev Term.abs ys (strip_abs_body t);
-            val xs' = map Free
-              ((fold_map Name.variant (map fst xs)
-                  (Term.declare_term_names u used) |> fst) ~~
-               map snd xs);
-            val (xs1, xs2) = chop j xs'
-          in (xs', list_comb (subst_bounds (rev xs1, u), xs2)) end;
-        fun is_dependent i t =
-          let val k = length (strip_abs_vars t) - i
-          in k < 0 orelse exists (fn j => j >= k) (loose_bnos (strip_abs_body t)) end;
-        fun count_cases (_, _, true) = I
-          | count_cases (c, (_, body), false) = AList.map_default op aconv (body, []) (cons c);
-        val is_undefined = name_of #> equal (SOME @{const_name undefined});
-        fun mk_case (c, (xs, body), _) = (list_comb (c, xs), body);
-        val get_info = lookup_by_case ctxt;
-      in
-        (case get_info cname of
-          SOME (_, constructors) =>
-            if length fs = length constructors then
-              let
-                val cases = map (fn (Const (s, U), t) =>
-                  let
-                    val Us = binder_types U;
-                    val k = length Us;
-                    val p as (xs, _) = strip_abs k Us t;
-                  in
-                    (Const (s, map fastype_of xs ---> fastype_of x), p, is_dependent k t)
-                  end) (constructors ~~ fs);
-                val cases' =
-                  sort (int_ord o swap o pairself (length o snd))
-                    (fold_rev count_cases cases []);
-                val R = fastype_of t;
-                val dummy =
-                  if d then Term.dummy_pattern R
-                  else Free (Name.variant "x" used |> fst, R);
-              in
-                SOME (x,
-                  map mk_case
-                    (case find_first (is_undefined o fst) cases' of
-                      SOME (_, cs) =>
-                        if length cs = length constructors then [hd cases]
-                        else filter_out (fn (_, (_, body), _) => is_undefined body) cases
-                    | NONE =>
-                        (case cases' of
-                          [] => cases
-                        | (default, cs) :: _ =>
-                            if length cs = 1 then cases
-                            else if length cs = length constructors then
-                              [hd cases, (dummy, ([], default), false)]
-                            else
-                              filter_out (fn (c, _, _) => member op aconv cs c) cases @
-                                [(dummy, ([], default), false)])))
-              end
-            else NONE
-        | _ => NONE)
-      end
-  | _ => NONE);
-
-
-(* destruct nested patterns *)
-
-fun encode_clause recur S T (pat, rhs) =
-  fold (fn x as (_, U) => fn t =>
-    Const (@{const_name case_abs}, (U --> T) --> T) $ Term.absfree x t)
-      (Term.add_frees pat [])
-      (Const (@{const_name case_elem}, S --> T --> S --> T) $ pat $ recur rhs);
-
-fun encode_cases _ S T [] = Const (@{const_name case_nil}, S --> T)
-  | encode_cases recur S T (p :: ps) =
-      Const (@{const_name case_cons}, (S --> T) --> (S --> T) --> S --> T) $
-        encode_clause recur S T p $ encode_cases recur S T ps;
-
-fun encode_case recur (t, ps as (pat, rhs) :: _) =
-      let
-        val tT = fastype_of t;
-        val T = fastype_of rhs;
-      in
-        Const (@{const_name case_guard}, @{typ bool} --> tT --> (tT --> T) --> T) $
-          @{term True} $ t $ (encode_cases recur (fastype_of pat) (fastype_of rhs) ps)
-      end
-  | encode_case _ _ = case_error "encode_case";
-
-fun strip_case' ctxt d (pat, rhs) =
-  (case dest_case ctxt d (Term.declare_term_frees pat Name.context) rhs of
-    SOME (exp as Free _, clauses) =>
-      if Term.exists_subterm (curry (op aconv) exp) pat andalso
-        not (exists (fn (_, rhs') =>
-          Term.exists_subterm (curry (op aconv) exp) rhs') clauses)
-      then
-        maps (strip_case' ctxt d) (map (fn (pat', rhs') =>
-          (subst_free [(exp, pat')] pat, rhs')) clauses)
-      else [(pat, rhs)]
-  | _ => [(pat, rhs)]);
-
-fun strip_case ctxt d t =
-  (case dest_case ctxt d Name.context t of
-    SOME (x, clauses) => SOME (x, maps (strip_case' ctxt d) clauses)
-  | NONE => NONE);
-
-fun strip_case_full ctxt d t =
-  (case dest_case ctxt d Name.context t of
-    SOME (x, clauses) =>
-      encode_case (strip_case_full ctxt d)
-        (strip_case_full ctxt d x, maps (strip_case' ctxt d) clauses)
-  | NONE =>
-      (case t of
-        t $ u => strip_case_full ctxt d t $ strip_case_full ctxt d u
-      | Abs (x, T, u) =>
-          let val (x', u') = Term.dest_abs (x, T, u);
-          in Term.absfree (x', T) (strip_case_full ctxt d u') end
-      | _ => t));
-
-
-(* term uncheck *)
-
-val show_cases = Attrib.setup_config_bool @{binding show_cases} (K true);
-
-fun uncheck_case ctxt ts =
-  if Config.get ctxt show_cases
-  then map (fn t => if can Term.type_of t then strip_case_full ctxt true t else t) ts
-  else ts;
-
-val term_uncheck_setup =
-  Context.theory_map (Syntax_Phases.term_uncheck 1 "case" uncheck_case);
-
-
-(* theory setup *)
-
-val setup =
-  trfun_setup #>
-  trfun_setup' #>
-  term_check_setup #>
-  term_uncheck_setup #>
-  Attrib.setup @{binding case_translation}
-    (Args.term -- Scan.repeat1 Args.term >>
-      (fn (t, ts) => Thm.declaration_attribute (K (register t ts))))
-    "declaration of case combinators and constructors";
-
-
-(* outer syntax commands *)
-
-fun print_case_translations ctxt =
-  let
-    val cases = map snd (Symtab.dest (cases_of ctxt));
-    val type_space = Proof_Context.type_space ctxt;
-
-    val pretty_term = Syntax.pretty_term ctxt;
-
-    fun pretty_data (data as (comb, ctrs)) =
-      let
-        val name = Tname_of_data data;
-        val xname = Name_Space.extern ctxt type_space name;
-        val markup = Name_Space.markup type_space name;
-        val prt =
-          (Pretty.block o Pretty.fbreaks)
-           [Pretty.block [Pretty.mark_str (markup, xname), Pretty.str ":"],
-            Pretty.block [Pretty.str "combinator:", Pretty.brk 1, pretty_term comb],
-            Pretty.block (Pretty.str "constructors:" :: Pretty.brk 1 ::
-              Pretty.commas (map pretty_term ctrs))];
-      in (xname, prt) end;
-  in
-    Pretty.big_list "case translations:" (map #2 (sort_wrt #1 (map pretty_data cases)))
-    |> Pretty.writeln
-  end;
-
-val _ =
-  Outer_Syntax.improper_command @{command_spec "print_case_translations"}
-    "print registered case combinators and constructors"
-    (Scan.succeed (Toplevel.keep (print_case_translations o Toplevel.context_of)))
-
-end;
--- a/src/HOL/Tools/ctr_sugar.ML	Mon Dec 09 06:33:46 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,981 +0,0 @@
-(*  Title:      HOL/Tools/ctr_sugar.ML
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012, 2013
-
-Wrapping existing freely generated type's constructors.
-*)
-
-signature CTR_SUGAR =
-sig
-  type ctr_sugar =
-    {ctrs: term list,
-     casex: term,
-     discs: term list,
-     selss: term list list,
-     exhaust: thm,
-     nchotomy: thm,
-     injects: thm list,
-     distincts: thm list,
-     case_thms: thm list,
-     case_cong: thm,
-     weak_case_cong: thm,
-     split: thm,
-     split_asm: thm,
-     disc_thmss: thm list list,
-     discIs: thm list,
-     sel_thmss: thm list list,
-     disc_exhausts: thm list,
-     sel_exhausts: thm list,
-     collapses: thm list,
-     expands: thm list,
-     sel_splits: thm list,
-     sel_split_asms: thm list,
-     case_eq_ifs: thm list};
-
-  val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar
-  val transfer_ctr_sugar: Proof.context -> ctr_sugar -> ctr_sugar
-  val ctr_sugar_of: Proof.context -> string -> ctr_sugar option
-  val ctr_sugars_of: Proof.context -> ctr_sugar list
-  val ctr_sugar_of_case: Proof.context -> string -> ctr_sugar option
-  val register_ctr_sugar: string -> ctr_sugar -> local_theory -> local_theory
-  val register_ctr_sugar_global: string -> ctr_sugar -> theory -> theory
-
-  val rep_compat_prefix: string
-
-  val mk_half_pairss: 'a list * 'a list -> ('a * 'a) list list
-  val join_halves: int -> 'a list list -> 'a list list -> 'a list * 'a list list list
-
-  val mk_ctr: typ list -> term -> term
-  val mk_case: typ list -> typ -> term -> term
-  val mk_disc_or_sel: typ list -> term -> term
-  val name_of_ctr: term -> string
-  val name_of_disc: term -> string
-  val dest_ctr: Proof.context -> string -> term -> term * term list
-  val dest_case: Proof.context -> string -> typ list -> term -> (term list * term list) option
-
-  val wrap_free_constructors: ({prems: thm list, context: Proof.context} -> tactic) list list ->
-    (((bool * (bool * bool)) * term list) * binding) *
-      (binding list * (binding list list * (binding * term) list list)) -> local_theory ->
-    ctr_sugar * local_theory
-  val parse_wrap_free_constructors_options: (bool * (bool * bool)) parser
-  val parse_bound_term: (binding * string) parser
-end;
-
-structure Ctr_Sugar : CTR_SUGAR =
-struct
-
-open Ctr_Sugar_Util
-open Ctr_Sugar_Tactics
-open Ctr_Sugar_Code
-
-type ctr_sugar =
-  {ctrs: term list,
-   casex: term,
-   discs: term list,
-   selss: term list list,
-   exhaust: thm,
-   nchotomy: thm,
-   injects: thm list,
-   distincts: thm list,
-   case_thms: thm list,
-   case_cong: thm,
-   weak_case_cong: thm,
-   split: thm,
-   split_asm: thm,
-   disc_thmss: thm list list,
-   discIs: thm list,
-   sel_thmss: thm list list,
-   disc_exhausts: thm list,
-   sel_exhausts: thm list,
-   collapses: thm list,
-   expands: thm list,
-   sel_splits: thm list,
-   sel_split_asms: thm list,
-   case_eq_ifs: thm list};
-
-fun eq_ctr_sugar ({ctrs = ctrs1, casex = case1, discs = discs1, selss = selss1, ...} : ctr_sugar,
-    {ctrs = ctrs2, casex = case2, discs = discs2, selss = selss2, ...} : ctr_sugar) =
-  ctrs1 = ctrs2 andalso case1 = case2 andalso discs1 = discs2 andalso selss1 = selss2;
-
-fun morph_ctr_sugar phi {ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts,
-    case_thms, case_cong, weak_case_cong, split, split_asm, disc_thmss, discIs, sel_thmss,
-    disc_exhausts, sel_exhausts, collapses, expands, sel_splits, sel_split_asms, case_eq_ifs} =
-  {ctrs = map (Morphism.term phi) ctrs,
-   casex = Morphism.term phi casex,
-   discs = map (Morphism.term phi) discs,
-   selss = map (map (Morphism.term phi)) selss,
-   exhaust = Morphism.thm phi exhaust,
-   nchotomy = Morphism.thm phi nchotomy,
-   injects = map (Morphism.thm phi) injects,
-   distincts = map (Morphism.thm phi) distincts,
-   case_thms = map (Morphism.thm phi) case_thms,
-   case_cong = Morphism.thm phi case_cong,
-   weak_case_cong = Morphism.thm phi weak_case_cong,
-   split = Morphism.thm phi split,
-   split_asm = Morphism.thm phi split_asm,
-   disc_thmss = map (map (Morphism.thm phi)) disc_thmss,
-   discIs = map (Morphism.thm phi) discIs,
-   sel_thmss = map (map (Morphism.thm phi)) sel_thmss,
-   disc_exhausts = map (Morphism.thm phi) disc_exhausts,
-   sel_exhausts = map (Morphism.thm phi) sel_exhausts,
-   collapses = map (Morphism.thm phi) collapses,
-   expands = map (Morphism.thm phi) expands,
-   sel_splits = map (Morphism.thm phi) sel_splits,
-   sel_split_asms = map (Morphism.thm phi) sel_split_asms,
-   case_eq_ifs = map (Morphism.thm phi) case_eq_ifs};
-
-val transfer_ctr_sugar =
-  morph_ctr_sugar o Morphism.thm_morphism o Thm.transfer o Proof_Context.theory_of;
-
-structure Data = Generic_Data
-(
-  type T = ctr_sugar Symtab.table;
-  val empty = Symtab.empty;
-  val extend = I;
-  val merge = Symtab.merge eq_ctr_sugar;
-);
-
-fun ctr_sugar_of ctxt =
-  Symtab.lookup (Data.get (Context.Proof ctxt))
-  #> Option.map (transfer_ctr_sugar ctxt);
-
-fun ctr_sugars_of ctxt =
-  Symtab.fold (cons o transfer_ctr_sugar ctxt o snd) (Data.get (Context.Proof ctxt)) [];
-
-fun ctr_sugar_of_case ctxt s =
-  find_first (fn {casex = Const (s', _), ...} => s' = s | _ => false) (ctr_sugars_of ctxt);
-
-fun register_ctr_sugar key ctr_sugar =
-  Local_Theory.declaration {syntax = false, pervasive = true}
-    (fn phi => Data.map (Symtab.default (key, morph_ctr_sugar phi ctr_sugar)));
-
-fun register_ctr_sugar_global key ctr_sugar =
-  Context.theory_map (Data.map (Symtab.default (key, ctr_sugar)));
-
-val rep_compat_prefix = "new";
-
-val isN = "is_";
-val unN = "un_";
-fun mk_unN 1 1 suf = unN ^ suf
-  | mk_unN _ l suf = unN ^ suf ^ string_of_int l;
-
-val caseN = "case";
-val case_congN = "case_cong";
-val case_eq_ifN = "case_eq_if";
-val collapseN = "collapse";
-val disc_excludeN = "disc_exclude";
-val disc_exhaustN = "disc_exhaust";
-val discN = "disc";
-val discIN = "discI";
-val distinctN = "distinct";
-val exhaustN = "exhaust";
-val expandN = "expand";
-val injectN = "inject";
-val nchotomyN = "nchotomy";
-val selN = "sel";
-val sel_exhaustN = "sel_exhaust";
-val sel_splitN = "sel_split";
-val sel_split_asmN = "sel_split_asm";
-val splitN = "split";
-val splitsN = "splits";
-val split_asmN = "split_asm";
-val weak_case_cong_thmsN = "weak_case_cong";
-
-val cong_attrs = @{attributes [cong]};
-val dest_attrs = @{attributes [dest]};
-val safe_elim_attrs = @{attributes [elim!]};
-val iff_attrs = @{attributes [iff]};
-val inductsimp_attrs = @{attributes [induct_simp]};
-val nitpicksimp_attrs = @{attributes [nitpick_simp]};
-val simp_attrs = @{attributes [simp]};
-val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
-val code_nitpicksimp_simp_attrs = code_nitpicksimp_attrs @ simp_attrs;
-
-fun unflat_lookup eq xs ys = map (fn xs' => permute_like eq xs xs' ys);
-
-fun mk_half_pairss' _ ([], []) = []
-  | mk_half_pairss' indent (x :: xs, _ :: ys) =
-    indent @ fold_rev (cons o single o pair x) ys (mk_half_pairss' ([] :: indent) (xs, ys));
-
-fun mk_half_pairss p = mk_half_pairss' [[]] p;
-
-fun join_halves n half_xss other_half_xss =
-  let
-    val xsss =
-      map2 (map2 append) (Library.chop_groups n half_xss)
-        (transpose (Library.chop_groups n other_half_xss))
-    val xs = splice (flat half_xss) (flat other_half_xss);
-  in (xs, xsss) end;
-
-fun mk_undefined T = Const (@{const_name undefined}, T);
-
-fun mk_ctr Ts t =
-  let val Type (_, Ts0) = body_type (fastype_of t) in
-    subst_nonatomic_types (Ts0 ~~ Ts) t
-  end;
-
-fun mk_case Ts T t =
-  let val (Type (_, Ts0), body) = strip_type (fastype_of t) |>> List.last in
-    subst_nonatomic_types ((body, T) :: (Ts0 ~~ Ts)) t
-  end;
-
-fun mk_disc_or_sel Ts t =
-  subst_nonatomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t;
-
-fun name_of_const what t =
-  (case head_of t of
-    Const (s, _) => s
-  | Free (s, _) => s
-  | _ => error ("Cannot extract name of " ^ what));
-
-val name_of_ctr = name_of_const "constructor";
-
-val notN = "not_";
-val eqN = "eq_";
-val neqN = "neq_";
-
-fun name_of_disc t =
-  (case head_of t of
-    Abs (_, _, @{const Not} $ (t' $ Bound 0)) =>
-    Long_Name.map_base_name (prefix notN) (name_of_disc t')
-  | Abs (_, _, Const (@{const_name HOL.eq}, _) $ Bound 0 $ t') =>
-    Long_Name.map_base_name (prefix eqN) (name_of_disc t')
-  | Abs (_, _, @{const Not} $ (Const (@{const_name HOL.eq}, _) $ Bound 0 $ t')) =>
-    Long_Name.map_base_name (prefix neqN) (name_of_disc t')
-  | t' => name_of_const "destructor" t');
-
-val base_name_of_ctr = Long_Name.base_name o name_of_ctr;
-
-fun dest_ctr ctxt s t =
-  let
-    val (f, args) = Term.strip_comb t;
-  in
-    (case ctr_sugar_of ctxt s of
-      SOME {ctrs, ...} =>
-      (case find_first (can (fo_match ctxt f)) ctrs of
-        SOME f' => (f', args)
-      | NONE => raise Fail "dest_ctr")
-    | NONE => raise Fail "dest_ctr")
-  end;
-
-fun dest_case ctxt s Ts t =
-  (case Term.strip_comb t of
-    (Const (c, _), args as _ :: _) =>
-    (case ctr_sugar_of ctxt s of
-      SOME {casex = Const (case_name, _), discs = discs0, selss = selss0, ...} =>
-      if case_name = c then
-        let val n = length discs0 in
-          if n < length args then
-            let
-              val (branches, obj :: leftovers) = chop n args;
-              val discs = map (mk_disc_or_sel Ts) discs0;
-              val selss = map (map (mk_disc_or_sel Ts)) selss0;
-              val conds = map (rapp obj) discs;
-              val branch_argss = map (fn sels => map (rapp obj) sels @ leftovers) selss;
-              val branches' = map2 (curry Term.betapplys) branches branch_argss;
-            in
-              SOME (conds, branches')
-            end
-          else
-            NONE
-        end
-      else
-        NONE
-    | _ => NONE)
-  | _ => NONE);
-
-fun eta_expand_arg xs f_xs = fold_rev Term.lambda xs f_xs;
-
-fun prepare_wrap_free_constructors prep_term ((((no_discs_sels, (no_code, rep_compat)), raw_ctrs),
-    raw_case_binding), (raw_disc_bindings, (raw_sel_bindingss, raw_sel_defaultss))) no_defs_lthy =
-  let
-    (* TODO: sanity checks on arguments *)
-
-    val n = length raw_ctrs;
-    val ks = 1 upto n;
-
-    val _ = if n > 0 then () else error "No constructors specified";
-
-    val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs;
-    val sel_defaultss =
-      pad_list [] n (map (map (apsnd (prep_term no_defs_lthy))) raw_sel_defaultss);
-
-    val Type (fcT_name, As0) = body_type (fastype_of (hd ctrs0));
-    val fc_b_name = Long_Name.base_name fcT_name;
-    val fc_b = Binding.name fc_b_name;
-
-    fun qualify mandatory =
-      Binding.qualify mandatory fc_b_name o (rep_compat ? Binding.qualify false rep_compat_prefix);
-
-    fun dest_TFree_or_TVar (TFree sS) = sS
-      | dest_TFree_or_TVar (TVar ((s, _), S)) = (s, S)
-      | dest_TFree_or_TVar _ = error "Invalid type argument";
-
-    val (unsorted_As, B) =
-      no_defs_lthy
-      |> variant_tfrees (map (fst o dest_TFree_or_TVar) As0)
-      ||> the_single o fst o mk_TFrees 1;
-
-    val As = map2 (resort_tfree o snd o dest_TFree_or_TVar) As0 unsorted_As;
-
-    val fcT = Type (fcT_name, As);
-    val ctrs = map (mk_ctr As) ctrs0;
-    val ctr_Tss = map (binder_types o fastype_of) ctrs;
-
-    val ms = map length ctr_Tss;
-
-    val raw_disc_bindings' = pad_list Binding.empty n raw_disc_bindings;
-
-    fun can_definitely_rely_on_disc k = not (Binding.is_empty (nth raw_disc_bindings' (k - 1)));
-    fun can_rely_on_disc k =
-      can_definitely_rely_on_disc k orelse (k = 1 andalso not (can_definitely_rely_on_disc 2));
-    fun should_omit_disc_binding k = n = 1 orelse (n = 2 andalso can_rely_on_disc (3 - k));
-
-    fun is_disc_binding_valid b =
-      not (Binding.is_empty b orelse Binding.eq_name (b, equal_binding));
-
-    val standard_disc_binding = Binding.name o prefix isN o base_name_of_ctr;
-
-    val disc_bindings =
-      raw_disc_bindings'
-      |> map4 (fn k => fn m => fn ctr => fn disc =>
-        qualify false
-          (if Binding.is_empty disc then
-             if should_omit_disc_binding k then disc else standard_disc_binding ctr
-           else if Binding.eq_name (disc, equal_binding) then
-             if m = 0 then disc
-             else error "Cannot use \"=\" syntax for discriminating nonnullary constructor"
-           else if Binding.eq_name (disc, standard_binding) then
-             standard_disc_binding ctr
-           else
-             disc)) ks ms ctrs0;
-
-    fun standard_sel_binding m l = Binding.name o mk_unN m l o base_name_of_ctr;
-
-    val sel_bindingss =
-      pad_list [] n raw_sel_bindingss
-      |> map3 (fn ctr => fn m => map2 (fn l => fn sel =>
-        qualify false
-          (if Binding.is_empty sel orelse Binding.eq_name (sel, standard_binding) then
-            standard_sel_binding m l ctr
-          else
-            sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms;
-
-    val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss;
-
-    val ((((((((xss, xss'), yss), fs), gs), [u', v']), [w]), (p, p')), names_lthy) = no_defs_lthy |>
-      mk_Freess' "x" ctr_Tss
-      ||>> mk_Freess "y" ctr_Tss
-      ||>> mk_Frees "f" case_Ts
-      ||>> mk_Frees "g" case_Ts
-      ||>> (apfst (map (rpair fcT)) oo Variable.variant_fixes) [fc_b_name, fc_b_name ^ "'"]
-      ||>> mk_Frees "z" [B]
-      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT;
-
-    val u = Free u';
-    val v = Free v';
-    val q = Free (fst p', mk_pred1T B);
-
-    val xctrs = map2 (curry Term.list_comb) ctrs xss;
-    val yctrs = map2 (curry Term.list_comb) ctrs yss;
-
-    val xfs = map2 (curry Term.list_comb) fs xss;
-    val xgs = map2 (curry Term.list_comb) gs xss;
-
-    (* TODO: Eta-expension is for compatibility with the old datatype package (but it also provides
-       nicer names). Consider removing. *)
-    val eta_fs = map2 eta_expand_arg xss xfs;
-    val eta_gs = map2 eta_expand_arg xss xgs;
-
-    val case_binding =
-      qualify false
-        (if Binding.is_empty raw_case_binding orelse
-            Binding.eq_name (raw_case_binding, standard_binding) then
-           Binding.prefix_name (caseN ^ "_") fc_b
-         else
-           raw_case_binding);
-
-    fun mk_case_disj xctr xf xs =
-      list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_eq (w, xf)));
-
-    val case_rhs = fold_rev (fold_rev Term.lambda) [fs, [u]]
-      (Const (@{const_name The}, (B --> HOLogic.boolT) --> B) $
-         Term.lambda w (Library.foldr1 HOLogic.mk_disj (map3 mk_case_disj xctrs xfs xss)));
-
-    val ((raw_case, (_, raw_case_def)), (lthy', lthy)) = no_defs_lthy
-      |> Local_Theory.define ((case_binding, NoSyn),
-        ((Binding.conceal (Thm.def_binding case_binding), []), case_rhs))
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy lthy';
-
-    val case_def = Morphism.thm phi raw_case_def;
-
-    val case0 = Morphism.term phi raw_case;
-    val casex = mk_case As B case0;
-
-    val fcase = Term.list_comb (casex, fs);
-
-    val ufcase = fcase $ u;
-    val vfcase = fcase $ v;
-
-    val eta_fcase = Term.list_comb (casex, eta_fs);
-    val eta_gcase = Term.list_comb (casex, eta_gs);
-
-    val eta_ufcase = eta_fcase $ u;
-    val eta_vgcase = eta_gcase $ v;
-
-    fun mk_uu_eq () = HOLogic.mk_eq (u, u);
-
-    val uv_eq = mk_Trueprop_eq (u, v);
-
-    val exist_xs_u_eq_ctrs =
-      map2 (fn xctr => fn xs => list_exists_free xs (HOLogic.mk_eq (u, xctr))) xctrs xss;
-
-    val unique_disc_no_def = TrueI; (*arbitrary marker*)
-    val alternate_disc_no_def = FalseE; (*arbitrary marker*)
-
-    fun alternate_disc_lhs get_udisc k =
-      HOLogic.mk_not
-        (let val b = nth disc_bindings (k - 1) in
-           if is_disc_binding_valid b then get_udisc b (k - 1) else nth exist_xs_u_eq_ctrs (k - 1)
-         end);
-
-    val (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') =
-      if no_discs_sels then
-        (true, [], [], [], [], [], lthy)
-      else
-        let
-          fun disc_free b = Free (Binding.name_of b, mk_pred1T fcT);
-
-          fun disc_spec b exist_xs_u_eq_ctr = mk_Trueprop_eq (disc_free b $ u, exist_xs_u_eq_ctr);
-
-          fun alternate_disc k =
-            Term.lambda u (alternate_disc_lhs (K o rapp u o disc_free) (3 - k));
-
-          fun mk_sel_case_args b proto_sels T =
-            map2 (fn Ts => fn k =>
-              (case AList.lookup (op =) proto_sels k of
-                NONE =>
-                (case AList.lookup Binding.eq_name (rev (nth sel_defaultss (k - 1))) b of
-                  NONE => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T)
-                | SOME t => t |> Type.constraint (Ts ---> T) |> Syntax.check_term lthy)
-              | SOME (xs, x) => fold_rev Term.lambda xs x)) ctr_Tss ks;
-
-          fun sel_spec b proto_sels =
-            let
-              val _ =
-                (case duplicates (op =) (map fst proto_sels) of
-                   k :: _ => error ("Duplicate selector name " ^ quote (Binding.name_of b) ^
-                     " for constructor " ^
-                     quote (Syntax.string_of_term lthy (nth ctrs (k - 1))))
-                 | [] => ())
-              val T =
-                (case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of
-                  [T] => T
-                | T :: T' :: _ => error ("Inconsistent range type for selector " ^
-                    quote (Binding.name_of b) ^ ": " ^ quote (Syntax.string_of_typ lthy T) ^ " vs. "
-                    ^ quote (Syntax.string_of_typ lthy T')));
-            in
-              mk_Trueprop_eq (Free (Binding.name_of b, fcT --> T) $ u,
-                Term.list_comb (mk_case As T case0, mk_sel_case_args b proto_sels T) $ u)
-            end;
-
-          val sel_bindings = flat sel_bindingss;
-          val uniq_sel_bindings = distinct Binding.eq_name sel_bindings;
-          val all_sels_distinct = (length uniq_sel_bindings = length sel_bindings);
-
-          val sel_binding_index =
-            if all_sels_distinct then 1 upto length sel_bindings
-            else map (fn b => find_index (curry Binding.eq_name b) uniq_sel_bindings) sel_bindings;
-
-          val proto_sels = flat (map3 (fn k => fn xs => map (fn x => (k, (xs, x)))) ks xss xss);
-          val sel_infos =
-            AList.group (op =) (sel_binding_index ~~ proto_sels)
-            |> sort (int_ord o pairself fst)
-            |> map snd |> curry (op ~~) uniq_sel_bindings;
-          val sel_bindings = map fst sel_infos;
-
-          fun unflat_selss xs = unflat_lookup Binding.eq_name sel_bindings xs sel_bindingss;
-
-          val (((raw_discs, raw_disc_defs), (raw_sels, raw_sel_defs)), (lthy', lthy)) =
-            lthy
-            |> apfst split_list o fold_map3 (fn k => fn exist_xs_u_eq_ctr => fn b =>
-                if Binding.is_empty b then
-                  if n = 1 then pair (Term.lambda u (mk_uu_eq ()), unique_disc_no_def)
-                  else pair (alternate_disc k, alternate_disc_no_def)
-                else if Binding.eq_name (b, equal_binding) then
-                  pair (Term.lambda u exist_xs_u_eq_ctr, refl)
-                else
-                  Specification.definition (SOME (b, NONE, NoSyn),
-                    ((Thm.def_binding b, []), disc_spec b exist_xs_u_eq_ctr)) #>> apsnd snd)
-              ks exist_xs_u_eq_ctrs disc_bindings
-            ||>> apfst split_list o fold_map (fn (b, proto_sels) =>
-              Specification.definition (SOME (b, NONE, NoSyn),
-                ((Thm.def_binding b, []), sel_spec b proto_sels)) #>> apsnd snd) sel_infos
-            ||> `Local_Theory.restore;
-
-          val phi = Proof_Context.export_morphism lthy lthy';
-
-          val disc_defs = map (Morphism.thm phi) raw_disc_defs;
-          val sel_defs = map (Morphism.thm phi) raw_sel_defs;
-          val sel_defss = unflat_selss sel_defs;
-
-          val discs0 = map (Morphism.term phi) raw_discs;
-          val selss0 = unflat_selss (map (Morphism.term phi) raw_sels);
-
-          val discs = map (mk_disc_or_sel As) discs0;
-          val selss = map (map (mk_disc_or_sel As)) selss0;
-        in
-          (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy')
-        end;
-
-    fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p);
-
-    val exhaust_goal =
-      let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (u, xctr)]) in
-        fold_rev Logic.all [p, u] (mk_imp_p (map2 mk_prem xctrs xss))
-      end;
-
-    val inject_goalss =
-      let
-        fun mk_goal _ _ [] [] = []
-          | mk_goal xctr yctr xs ys =
-            [fold_rev Logic.all (xs @ ys) (mk_Trueprop_eq (HOLogic.mk_eq (xctr, yctr),
-              Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))];
-      in
-        map4 mk_goal xctrs yctrs xss yss
-      end;
-
-    val half_distinct_goalss =
-      let
-        fun mk_goal ((xs, xc), (xs', xc')) =
-          fold_rev Logic.all (xs @ xs')
-            (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc'))));
-      in
-        map (map mk_goal) (mk_half_pairss (`I (xss ~~ xctrs)))
-      end;
-
-    val goalss = [exhaust_goal] :: inject_goalss @ half_distinct_goalss;
-
-    fun after_qed thmss lthy =
-      let
-        val ([exhaust_thm], (inject_thmss, half_distinct_thmss)) = (hd thmss, chop n (tl thmss));
-
-        val inject_thms = flat inject_thmss;
-
-        val rho_As = map (pairself (certifyT lthy)) (map Logic.varifyT_global As ~~ As);
-
-        fun inst_thm t thm =
-          Drule.instantiate' [] [SOME (certify lthy t)]
-            (Thm.instantiate (rho_As, []) (Drule.zero_var_indexes thm));
-
-        val uexhaust_thm = inst_thm u exhaust_thm;
-
-        val exhaust_cases = map base_name_of_ctr ctrs;
-
-        val other_half_distinct_thmss = map (map (fn thm => thm RS not_sym)) half_distinct_thmss;
-
-        val (distinct_thms, (distinct_thmsss', distinct_thmsss)) =
-          join_halves n half_distinct_thmss other_half_distinct_thmss ||> `transpose;
-
-        val nchotomy_thm =
-          let
-            val goal =
-              HOLogic.mk_Trueprop (HOLogic.mk_all (fst u', snd u',
-                Library.foldr1 HOLogic.mk_disj exist_xs_u_eq_ctrs));
-          in
-            Goal.prove_sorry lthy [] [] goal (fn _ => mk_nchotomy_tac n exhaust_thm)
-            |> Thm.close_derivation
-          end;
-
-        val case_thms =
-          let
-            val goals =
-              map3 (fn xctr => fn xf => fn xs =>
-                fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xctrs xfs xss;
-          in
-            map4 (fn k => fn goal => fn injects => fn distinctss =>
-                Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
-                  mk_case_tac ctxt n k case_def injects distinctss)
-                |> Thm.close_derivation)
-              ks goals inject_thmss distinct_thmsss
-          end;
-
-        val (case_cong_thm, weak_case_cong_thm) =
-          let
-            fun mk_prem xctr xs xf xg =
-              fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr),
-                mk_Trueprop_eq (xf, xg)));
-
-            val goal =
-              Logic.list_implies (uv_eq :: map4 mk_prem xctrs xss xfs xgs,
-                 mk_Trueprop_eq (eta_ufcase, eta_vgcase));
-            val weak_goal = Logic.mk_implies (uv_eq, mk_Trueprop_eq (ufcase, vfcase));
-          in
-            (Goal.prove_sorry lthy [] [] goal (fn _ => mk_case_cong_tac lthy uexhaust_thm case_thms),
-             Goal.prove_sorry lthy [] [] weak_goal (K (etac arg_cong 1)))
-            |> pairself (Thm.close_derivation #> singleton (Proof_Context.export names_lthy lthy))
-          end;
-
-        val split_lhs = q $ ufcase;
-
-        fun mk_split_conjunct xctr xs f_xs =
-          list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (u, xctr), q $ f_xs));
-        fun mk_split_disjunct xctr xs f_xs =
-          list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr),
-            HOLogic.mk_not (q $ f_xs)));
-
-        fun mk_split_goal xctrs xss xfs =
-          mk_Trueprop_eq (split_lhs, Library.foldr1 HOLogic.mk_conj
-            (map3 mk_split_conjunct xctrs xss xfs));
-        fun mk_split_asm_goal xctrs xss xfs =
-          mk_Trueprop_eq (split_lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj
-            (map3 mk_split_disjunct xctrs xss xfs)));
-
-        fun prove_split selss goal =
-          Goal.prove_sorry lthy [] [] goal (fn _ =>
-            mk_split_tac lthy uexhaust_thm case_thms selss inject_thmss distinct_thmsss)
-          |> Thm.close_derivation
-          |> singleton (Proof_Context.export names_lthy lthy);
-
-        fun prove_split_asm asm_goal split_thm =
-          Goal.prove_sorry lthy [] [] asm_goal (fn {context = ctxt, ...} =>
-            mk_split_asm_tac ctxt split_thm)
-          |> Thm.close_derivation
-          |> singleton (Proof_Context.export names_lthy lthy);
-
-        val (split_thm, split_asm_thm) =
-          let
-            val goal = mk_split_goal xctrs xss xfs;
-            val asm_goal = mk_split_asm_goal xctrs xss xfs;
-
-            val thm = prove_split (replicate n []) goal;
-            val asm_thm = prove_split_asm asm_goal thm;
-          in
-            (thm, asm_thm)
-          end;
-
-        val (all_sel_thms, sel_thmss, disc_thmss, nontriv_disc_thms, discI_thms, nontriv_discI_thms,
-             disc_exclude_thms, disc_exhaust_thms, sel_exhaust_thms, all_collapse_thms,
-             safe_collapse_thms, expand_thms, sel_split_thms, sel_split_asm_thms, case_eq_if_thms) =
-          if no_discs_sels then
-            ([], [], [], [], [], [], [], [], [], [], [], [], [], [], [])
-          else
-            let
-              val udiscs = map (rapp u) discs;
-              val uselss = map (map (rapp u)) selss;
-              val usel_ctrs = map2 (curry Term.list_comb) ctrs uselss;
-              val usel_fs = map2 (curry Term.list_comb) fs uselss;
-
-              val vdiscs = map (rapp v) discs;
-              val vselss = map (map (rapp v)) selss;
-
-              fun make_sel_thm xs' case_thm sel_def =
-                zero_var_indexes (Drule.gen_all (Drule.rename_bvars' (map (SOME o fst) xs')
-                    (Drule.forall_intr_vars (case_thm RS (sel_def RS trans)))));
-
-              val sel_thmss = map3 (map oo make_sel_thm) xss' case_thms sel_defss;
-
-              fun has_undefined_rhs thm =
-                (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of thm))) of
-                  Const (@{const_name undefined}, _) => true
-                | _ => false);
-
-              val all_sel_thms =
-                (if all_sels_distinct andalso forall null sel_defaultss then
-                   flat sel_thmss
-                 else
-                   map_product (fn s => fn (xs', c) => make_sel_thm xs' c s) sel_defs
-                     (xss' ~~ case_thms))
-                |> filter_out has_undefined_rhs;
-
-              fun mk_unique_disc_def () =
-                let
-                  val m = the_single ms;
-                  val goal = mk_Trueprop_eq (mk_uu_eq (), the_single exist_xs_u_eq_ctrs);
-                in
-                  Goal.prove_sorry lthy [] [] goal (fn _ => mk_unique_disc_def_tac m uexhaust_thm)
-                  |> Thm.close_derivation
-                  |> singleton (Proof_Context.export names_lthy lthy)
-                end;
-
-              fun mk_alternate_disc_def k =
-                let
-                  val goal =
-                    mk_Trueprop_eq (alternate_disc_lhs (K (nth udiscs)) (3 - k),
-                      nth exist_xs_u_eq_ctrs (k - 1));
-                in
-                  Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
-                    mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k))
-                      (nth distinct_thms (2 - k)) uexhaust_thm)
-                  |> Thm.close_derivation
-                  |> singleton (Proof_Context.export names_lthy lthy)
-                end;
-
-              val has_alternate_disc_def =
-                exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs;
-
-              val disc_defs' =
-                map2 (fn k => fn def =>
-                  if Thm.eq_thm_prop (def, unique_disc_no_def) then mk_unique_disc_def ()
-                  else if Thm.eq_thm_prop (def, alternate_disc_no_def) then mk_alternate_disc_def k
-                  else def) ks disc_defs;
-
-              val discD_thms = map (fn def => def RS iffD1) disc_defs';
-              val discI_thms =
-                map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms
-                  disc_defs';
-              val not_discI_thms =
-                map2 (fn m => fn def => funpow m (fn thm => allI RS thm)
-                    (unfold_thms lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]})))
-                  ms disc_defs';
-
-              val (disc_thmss', disc_thmss) =
-                let
-                  fun mk_thm discI _ [] = refl RS discI
-                    | mk_thm _ not_discI [distinct] = distinct RS not_discI;
-                  fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss;
-                in
-                  map3 mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose
-                end;
-
-              val nontriv_disc_thms =
-                flat (map2 (fn b => if is_disc_binding_valid b then I else K [])
-                  disc_bindings disc_thmss);
-
-              fun is_discI_boring b =
-                (n = 1 andalso Binding.is_empty b) orelse Binding.eq_name (b, equal_binding);
-
-              val nontriv_discI_thms =
-                flat (map2 (fn b => if is_discI_boring b then K [] else single) disc_bindings
-                  discI_thms);
-
-              val (disc_exclude_thms, (disc_exclude_thmsss', disc_exclude_thmsss)) =
-                let
-                  fun mk_goal [] = []
-                    | mk_goal [((_, udisc), (_, udisc'))] =
-                      [Logic.all u (Logic.mk_implies (HOLogic.mk_Trueprop udisc,
-                         HOLogic.mk_Trueprop (HOLogic.mk_not udisc')))];
-
-                  fun prove tac goal =
-                    Goal.prove_sorry lthy [] [] goal (K tac)
-                    |> Thm.close_derivation;
-
-                  val half_pairss = mk_half_pairss (`I (ms ~~ discD_thms ~~ udiscs));
-
-                  val half_goalss = map mk_goal half_pairss;
-                  val half_thmss =
-                    map3 (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] =>
-                        fn disc_thm => [prove (mk_half_disc_exclude_tac lthy m discD disc_thm) goal])
-                      half_goalss half_pairss (flat disc_thmss');
-
-                  val other_half_goalss = map (mk_goal o map swap) half_pairss;
-                  val other_half_thmss =
-                    map2 (map2 (prove o mk_other_half_disc_exclude_tac)) half_thmss
-                      other_half_goalss;
-                in
-                  join_halves n half_thmss other_half_thmss ||> `transpose
-                  |>> has_alternate_disc_def ? K []
-                end;
-
-              val disc_exhaust_thm =
-                let
-                  fun mk_prem udisc = mk_imp_p [HOLogic.mk_Trueprop udisc];
-                  val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem udiscs));
-                in
-                  Goal.prove_sorry lthy [] [] goal (fn _ =>
-                    mk_disc_exhaust_tac n exhaust_thm discI_thms)
-                  |> Thm.close_derivation
-                end;
-
-              val (safe_collapse_thms, all_collapse_thms) =
-                let
-                  fun mk_goal m udisc usel_ctr =
-                    let
-                      val prem = HOLogic.mk_Trueprop udisc;
-                      val concl = mk_Trueprop_eq ((usel_ctr, u) |> m = 0 ? swap);
-                    in
-                      (prem aconv concl, Logic.all u (Logic.mk_implies (prem, concl)))
-                    end;
-                  val (trivs, goals) = map3 mk_goal ms udiscs usel_ctrs |> split_list;
-                  val thms =
-                    map5 (fn m => fn discD => fn sel_thms => fn triv => fn goal =>
-                        Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
-                          mk_collapse_tac ctxt m discD sel_thms ORELSE HEADGOAL atac)
-                        |> Thm.close_derivation
-                        |> not triv ? perhaps (try (fn thm => refl RS thm)))
-                      ms discD_thms sel_thmss trivs goals;
-                in
-                  (map_filter (fn (true, _) => NONE | (false, thm) => SOME thm) (trivs ~~ thms),
-                   thms)
-                end;
-
-              val swapped_all_collapse_thms =
-                map2 (fn m => fn thm => if m = 0 then thm else thm RS sym) ms all_collapse_thms;
-
-              val sel_exhaust_thm =
-                let
-                  fun mk_prem usel_ctr = mk_imp_p [mk_Trueprop_eq (u, usel_ctr)];
-                  val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem usel_ctrs));
-                in
-                  Goal.prove_sorry lthy [] [] goal (fn _ =>
-                    mk_sel_exhaust_tac n disc_exhaust_thm swapped_all_collapse_thms)
-                  |> Thm.close_derivation
-                end;
-
-              val expand_thm =
-                let
-                  fun mk_prems k udisc usels vdisc vsels =
-                    (if k = n then [] else [mk_Trueprop_eq (udisc, vdisc)]) @
-                    (if null usels then
-                       []
-                     else
-                       [Logic.list_implies
-                          (if n = 1 then [] else map HOLogic.mk_Trueprop [udisc, vdisc],
-                             HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-                               (map2 (curry HOLogic.mk_eq) usels vsels)))]);
-
-                  val goal =
-                    Library.foldr Logic.list_implies
-                      (map5 mk_prems ks udiscs uselss vdiscs vselss, uv_eq);
-                  val uncollapse_thms =
-                    map2 (fn thm => fn [] => thm | _ => thm RS sym) all_collapse_thms uselss;
-                in
-                  Goal.prove_sorry lthy [] [] goal (fn _ =>
-                    mk_expand_tac lthy n ms (inst_thm u disc_exhaust_thm)
-                      (inst_thm v disc_exhaust_thm) uncollapse_thms disc_exclude_thmsss
-                      disc_exclude_thmsss')
-                  |> Thm.close_derivation
-                  |> singleton (Proof_Context.export names_lthy lthy)
-                end;
-
-              val (sel_split_thm, sel_split_asm_thm) =
-                let
-                  val zss = map (K []) xss;
-                  val goal = mk_split_goal usel_ctrs zss usel_fs;
-                  val asm_goal = mk_split_asm_goal usel_ctrs zss usel_fs;
-
-                  val thm = prove_split sel_thmss goal;
-                  val asm_thm = prove_split_asm asm_goal thm;
-                in
-                  (thm, asm_thm)
-                end;
-
-              val case_eq_if_thm =
-                let
-                  val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs usel_fs);
-                in
-                  Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
-                    mk_case_eq_if_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss)
-                  |> Thm.close_derivation
-                  |> singleton (Proof_Context.export names_lthy lthy)
-                end;
-            in
-              (all_sel_thms, sel_thmss, disc_thmss, nontriv_disc_thms, discI_thms,
-               nontriv_discI_thms, disc_exclude_thms, [disc_exhaust_thm], [sel_exhaust_thm],
-               all_collapse_thms, safe_collapse_thms, [expand_thm], [sel_split_thm],
-               [sel_split_asm_thm], [case_eq_if_thm])
-            end;
-
-        val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases));
-        val cases_type_attr = Attrib.internal (K (Induct.cases_type fcT_name));
-
-        val anonymous_notes =
-          [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs),
-           (map (fn th => th RS @{thm eq_False[THEN iffD2]}
-              handle THM _ => th RS @{thm eq_True[THEN iffD2]}) nontriv_disc_thms,
-            code_nitpicksimp_attrs)]
-          |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
-
-        val notes =
-          [(caseN, case_thms, code_nitpicksimp_simp_attrs),
-           (case_congN, [case_cong_thm], []),
-           (case_eq_ifN, case_eq_if_thms, []),
-           (collapseN, safe_collapse_thms, simp_attrs),
-           (discN, nontriv_disc_thms, simp_attrs),
-           (discIN, nontriv_discI_thms, []),
-           (disc_excludeN, disc_exclude_thms, dest_attrs),
-           (disc_exhaustN, disc_exhaust_thms, [exhaust_case_names_attr]),
-           (distinctN, distinct_thms, simp_attrs @ inductsimp_attrs),
-           (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]),
-           (expandN, expand_thms, []),
-           (injectN, inject_thms, iff_attrs @ inductsimp_attrs),
-           (nchotomyN, [nchotomy_thm], []),
-           (selN, all_sel_thms, code_nitpicksimp_simp_attrs),
-           (sel_exhaustN, sel_exhaust_thms, [exhaust_case_names_attr]),
-           (sel_splitN, sel_split_thms, []),
-           (sel_split_asmN, sel_split_asm_thms, []),
-           (splitN, [split_thm], []),
-           (split_asmN, [split_asm_thm], []),
-           (splitsN, [split_thm, split_asm_thm], []),
-           (weak_case_cong_thmsN, [weak_case_cong_thm], cong_attrs)]
-          |> filter_out (null o #2)
-          |> map (fn (thmN, thms, attrs) =>
-            ((qualify true (Binding.name thmN), attrs), [(thms, [])]));
-
-        val ctr_sugar =
-          {ctrs = ctrs, casex = casex, discs = discs, selss = selss, exhaust = exhaust_thm,
-           nchotomy = nchotomy_thm, injects = inject_thms, distincts = distinct_thms,
-           case_thms = case_thms, case_cong = case_cong_thm, weak_case_cong = weak_case_cong_thm,
-           split = split_thm, split_asm = split_asm_thm, disc_thmss = disc_thmss,
-           discIs = discI_thms, sel_thmss = sel_thmss, disc_exhausts = disc_exhaust_thms,
-           sel_exhausts = sel_exhaust_thms, collapses = all_collapse_thms, expands = expand_thms,
-           sel_splits = sel_split_thms, sel_split_asms = sel_split_asm_thms,
-           case_eq_ifs = case_eq_if_thms};
-      in
-        (ctr_sugar,
-         lthy
-         |> not rep_compat ?
-            Local_Theory.declaration {syntax = false, pervasive = true}
-              (fn phi => Case_Translation.register
-                 (Morphism.term phi casex) (map (Morphism.term phi) ctrs))
-         |> Local_Theory.background_theory (fold (fold Code.del_eqn) [disc_defs, sel_defs])
-         |> not no_code ?
-            Local_Theory.declaration {syntax = false, pervasive = false}
-              (fn phi => Context.mapping
-                (add_ctr_code fcT_name (map (Morphism.typ phi) As)
-                  (map (dest_Const o Morphism.term phi) ctrs) (Morphism.fact phi inject_thms)
-                  (Morphism.fact phi distinct_thms) (Morphism.fact phi case_thms))
-                I)
-         |> Local_Theory.notes (anonymous_notes @ notes) |> snd
-         |> register_ctr_sugar fcT_name ctr_sugar)
-      end;
-  in
-    (goalss, after_qed, lthy')
-  end;
-
-fun wrap_free_constructors tacss = (fn (goalss, after_qed, lthy) =>
-  map2 (map2 (Thm.close_derivation oo Goal.prove_sorry lthy [] [])) goalss tacss
-  |> (fn thms => after_qed thms lthy)) oo prepare_wrap_free_constructors (K I);
-
-val wrap_free_constructors_cmd = (fn (goalss, after_qed, lthy) =>
-  Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo
-  prepare_wrap_free_constructors Syntax.read_term;
-
-fun parse_bracket_list parser = @{keyword "["} |-- Parse.list parser --|  @{keyword "]"};
-
-val parse_bindings = parse_bracket_list parse_binding;
-val parse_bindingss = parse_bracket_list parse_bindings;
-
-val parse_bound_term = (parse_binding --| @{keyword ":"}) -- Parse.term;
-val parse_bound_terms = parse_bracket_list parse_bound_term;
-val parse_bound_termss = parse_bracket_list parse_bound_terms;
-
-val parse_wrap_free_constructors_options =
-  Scan.optional (@{keyword "("} |-- Parse.list1
-        (Parse.reserved "no_discs_sels" >> K 0 || Parse.reserved "no_code" >> K 1 ||
-         Parse.reserved "rep_compat" >> K 2) --| @{keyword ")"}
-      >> (fn js => (member (op =) js 0, (member (op =) js 1, member (op =) js 2))))
-    (false, (false, false));
-
-val _ =
-  Outer_Syntax.local_theory_to_proof @{command_spec "wrap_free_constructors"}
-    "wrap an existing freely generated type's constructors"
-    ((parse_wrap_free_constructors_options -- (@{keyword "["} |-- Parse.list Parse.term --|
-        @{keyword "]"}) --
-      parse_binding -- Scan.optional (parse_bindings -- Scan.optional (parse_bindingss --
-        Scan.optional parse_bound_termss []) ([], [])) ([], ([], [])))
-     >> wrap_free_constructors_cmd);
-
-end;
--- a/src/HOL/Tools/ctr_sugar_code.ML	Mon Dec 09 06:33:46 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,131 +0,0 @@
-(*  Title:      HOL/Tools/ctr_sugar_code.ML
-    Author:     Jasmin Blanchette, TU Muenchen
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Stefan Berghofer, TU Muenchen
-    Author:     Florian Haftmann, TU Muenchen
-    Copyright   2001-2013
-
-Code generation for freely generated types.
-*)
-
-signature CTR_SUGAR_CODE =
-sig
-  val add_ctr_code: string -> typ list -> (string * typ) list -> thm list -> thm list -> thm list ->
-    theory -> theory
-end;
-
-structure Ctr_Sugar_Code : CTR_SUGAR_CODE =
-struct
-
-open Ctr_Sugar_Util
-
-val eqN = "eq"
-val reflN = "refl"
-val simpsN = "simps"
-
-fun mk_case_certificate thy raw_thms =
-  let
-    val thms as thm1 :: _ = raw_thms
-      |> Conjunction.intr_balanced
-      |> Thm.unvarify_global
-      |> Conjunction.elim_balanced (length raw_thms)
-      |> map Simpdata.mk_meta_eq
-      |> map Drule.zero_var_indexes;
-    val params = Term.add_free_names (Thm.prop_of thm1) [];
-    val rhs = thm1
-      |> Thm.prop_of |> Logic.dest_equals |> fst |> Term.strip_comb
-      ||> fst o split_last |> list_comb;
-    val lhs = Free (singleton (Name.variant_list params) "case", Term.fastype_of rhs);
-    val assum = Thm.cterm_of thy (Logic.mk_equals (lhs, rhs));
-  in
-    thms
-    |> Conjunction.intr_balanced
-    |> rewrite_rule [Thm.symmetric (Thm.assume assum)]
-    |> Thm.implies_intr assum
-    |> Thm.generalize ([], params) 0
-    |> Axclass.unoverload thy
-    |> Thm.varifyT_global
-  end;
-
-fun mk_free_ctr_equations fcT ctrs inject_thms distinct_thms thy =
-  let
-    fun mk_fcT_eq (t, u) = Const (@{const_name HOL.equal}, fcT --> fcT --> HOLogic.boolT) $ t $ u;
-    fun true_eq tu = HOLogic.mk_eq (mk_fcT_eq tu, @{term True});
-    fun false_eq tu = HOLogic.mk_eq (mk_fcT_eq tu, @{term False});
-
-    val monomorphic_prop_of = prop_of o Thm.unvarify_global o Drule.zero_var_indexes;
-
-    fun massage_inject (tp $ (eqv $ (_ $ t $ u) $ rhs)) = tp $ (eqv $ mk_fcT_eq (t, u) $ rhs);
-    fun massage_distinct (tp $ (_ $ (_ $ t $ u))) = [tp $ false_eq (t, u), tp $ false_eq (u, t)];
-
-    val triv_inject_goals =
-      map_filter (fn c as (_, T) =>
-          if T = fcT then SOME (HOLogic.mk_Trueprop (true_eq (Const c, Const c))) else NONE)
-        ctrs;
-    val inject_goals = map (massage_inject o monomorphic_prop_of) inject_thms;
-    val distinct_goals = maps (massage_distinct o monomorphic_prop_of) distinct_thms;
-    val refl_goal = HOLogic.mk_Trueprop (true_eq (Free ("x", fcT), Free ("x", fcT)));
-
-    val simp_ctxt =
-      Simplifier.global_context thy HOL_basic_ss
-        addsimps (map Simpdata.mk_eq (@{thms equal eq_True} @ inject_thms @ distinct_thms));
-
-    fun prove goal =
-      Goal.prove_sorry_global thy [] [] goal (K (ALLGOALS (simp_tac simp_ctxt)))
-      |> Simpdata.mk_eq;
-  in
-    (map prove (triv_inject_goals @ inject_goals @ distinct_goals), prove refl_goal)
-  end;
-
-fun add_equality fcT fcT_name As ctrs inject_thms distinct_thms =
-  let
-    fun add_def lthy =
-      let
-        fun mk_side const_name =
-          Const (const_name, fcT --> fcT --> HOLogic.boolT) $ Free ("x", fcT) $ Free ("y", fcT);
-        val spec =
-          mk_Trueprop_eq (mk_side @{const_name HOL.equal}, mk_side @{const_name HOL.eq})
-          |> Syntax.check_term lthy;
-        val ((_, (_, raw_def)), lthy') =
-          Specification.definition (NONE, (Attrib.empty_binding, spec)) lthy;
-        val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of lthy); (* FIXME? *)
-        val def = singleton (Proof_Context.export lthy' ctxt_thy) raw_def;
-      in
-        (def, lthy')
-      end;
-
-    fun tac thms = Class.intro_classes_tac [] THEN ALLGOALS (Proof_Context.fact_tac thms);
-
-    val qualify =
-      Binding.qualify true (Long_Name.base_name fcT_name) o Binding.qualify true eqN o Binding.name;
-  in
-    Class.instantiation ([fcT_name], map dest_TFree As, [HOLogic.class_equal])
-    #> add_def
-    #-> Class.prove_instantiation_exit_result (map o Morphism.thm) (K tac) o single
-    #-> fold Code.del_eqn
-    #> `(mk_free_ctr_equations fcT ctrs inject_thms distinct_thms)
-    #-> (fn (thms, thm) => Global_Theory.note_thmss Thm.lemmaK
-      [((qualify reflN, [Code.add_nbe_default_eqn_attribute]), [([thm], [])]),
-        ((qualify simpsN, [Code.add_default_eqn_attribute]), [(rev thms, [])])])
-    #> snd
-  end;
-
-fun add_ctr_code fcT_name raw_As raw_ctrs inject_thms distinct_thms case_thms thy =
-  let
-    val As = map (perhaps (try Logic.unvarifyT_global)) raw_As;
-    val ctrs = map (apsnd (perhaps (try Logic.unvarifyT_global))) raw_ctrs;
-    val fcT = Type (fcT_name, As);
-    val unover_ctrs = map (fn ctr as (_, fcT) => (Axclass.unoverload_const thy ctr, fcT)) ctrs;
-  in
-    if can (Code.constrset_of_consts thy) unover_ctrs then
-      thy
-      |> Code.add_datatype ctrs
-      |> fold_rev Code.add_default_eqn case_thms
-      |> Code.add_case (mk_case_certificate thy case_thms)
-      |> not (Sorts.has_instance (Sign.classes_of thy) fcT_name [HOLogic.class_equal])
-        ? add_equality fcT fcT_name As ctrs inject_thms distinct_thms
-    else
-      thy
-  end;
-
-end;
--- a/src/HOL/Tools/ctr_sugar_tactics.ML	Mon Dec 09 06:33:46 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,172 +0,0 @@
-(*  Title:      HOL/Tools/ctr_sugar_tactics.ML
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012, 2013
-
-Tactics for wrapping existing freely generated type's constructors.
-*)
-
-signature CTR_SUGAR_GENERAL_TACTICS =
-sig
-  val select_prem_tac: int -> (int -> tactic) -> int -> int -> tactic
-  val unfold_thms_tac: Proof.context -> thm list -> tactic
-end;
-
-signature CTR_SUGAR_TACTICS =
-sig
-  include CTR_SUGAR_GENERAL_TACTICS
-
-  val mk_alternate_disc_def_tac: Proof.context -> int -> thm -> thm -> thm -> tactic
-  val mk_case_tac: Proof.context -> int -> int -> thm -> thm list -> thm list list -> tactic
-  val mk_case_cong_tac: Proof.context -> thm -> thm list -> tactic
-  val mk_case_eq_if_tac: Proof.context -> int -> thm -> thm list -> thm list list ->
-    thm list list -> tactic
-  val mk_collapse_tac: Proof.context -> int -> thm -> thm list -> tactic
-  val mk_disc_exhaust_tac: int -> thm -> thm list -> tactic
-  val mk_expand_tac: Proof.context -> int -> int list -> thm -> thm -> thm list ->
-    thm list list list -> thm list list list -> tactic
-  val mk_half_disc_exclude_tac: Proof.context -> int -> thm -> thm -> tactic
-  val mk_nchotomy_tac: int -> thm -> tactic
-  val mk_other_half_disc_exclude_tac: thm -> tactic
-  val mk_sel_exhaust_tac: int -> thm -> thm list -> tactic
-  val mk_split_tac: Proof.context -> thm -> thm list -> thm list list -> thm list list ->
-    thm list list list -> tactic
-  val mk_split_asm_tac: Proof.context -> thm -> tactic
-  val mk_unique_disc_def_tac: int -> thm -> tactic
-end;
-
-structure Ctr_Sugar_Tactics : CTR_SUGAR_TACTICS =
-struct
-
-open Ctr_Sugar_Util
-
-val meta_mp = @{thm meta_mp};
-
-fun select_prem_tac n tac k = DETERM o (EVERY' [REPEAT_DETERM_N (k - 1) o etac thin_rl,
-  tac, REPEAT_DETERM_N (n - k) o etac thin_rl]);
-
-fun unfold_thms_tac _ [] = all_tac
-  | unfold_thms_tac ctxt thms = Local_Defs.unfold_tac ctxt (distinct Thm.eq_thm_prop thms);
-
-fun if_P_or_not_P_OF pos thm = thm RS (if pos then @{thm if_P} else @{thm if_not_P});
-
-fun mk_nchotomy_tac n exhaust =
-  HEADGOAL (rtac allI THEN' rtac exhaust THEN'
-   EVERY' (maps (fn k => [rtac (mk_disjIN n k), REPEAT_DETERM o rtac exI, atac]) (1 upto n)));
-
-fun mk_unique_disc_def_tac m uexhaust =
-  HEADGOAL (EVERY' [rtac iffI, rtac uexhaust, REPEAT_DETERM_N m o rtac exI, atac, rtac refl]);
-
-fun mk_alternate_disc_def_tac ctxt k other_disc_def distinct uexhaust =
-  HEADGOAL (EVERY' ([rtac (other_disc_def RS @{thm arg_cong[of _ _ Not]} RS trans),
-    rtac @{thm iffI_np}, REPEAT_DETERM o etac exE,
-    hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt [not_ex]), REPEAT_DETERM o rtac allI,
-    rtac distinct, rtac uexhaust] @
-    (([etac notE, REPEAT_DETERM o rtac exI, atac], [REPEAT_DETERM o rtac exI, atac])
-     |> k = 1 ? swap |> op @)));
-
-fun mk_half_disc_exclude_tac ctxt m discD disc' =
-  HEADGOAL (dtac discD THEN' REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac ctxt THEN'
-    rtac disc');
-
-fun mk_other_half_disc_exclude_tac half = HEADGOAL (etac @{thm contrapos_pn} THEN' etac half);
-
-fun mk_disc_or_sel_exhaust_tac n exhaust destIs =
-  HEADGOAL (rtac exhaust THEN'
-    EVERY' (map2 (fn k => fn destI => dtac destI THEN'
-      select_prem_tac n (etac meta_mp) k THEN' atac) (1 upto n) destIs));
-
-val mk_disc_exhaust_tac = mk_disc_or_sel_exhaust_tac;
-
-fun mk_sel_exhaust_tac n disc_exhaust collapses =
-  mk_disc_or_sel_exhaust_tac n disc_exhaust collapses ORELSE
-  HEADGOAL (etac meta_mp THEN' resolve_tac collapses);
-
-fun mk_collapse_tac ctxt m discD sels =
-  HEADGOAL (dtac discD THEN'
-    (if m = 0 then
-       atac
-     else
-       REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac ctxt THEN'
-       SELECT_GOAL (unfold_thms_tac ctxt sels) THEN' rtac refl));
-
-fun mk_expand_tac ctxt n ms udisc_exhaust vdisc_exhaust uncollapses disc_excludesss
-    disc_excludesss' =
-  if ms = [0] then
-    HEADGOAL (rtac (@{thm trans_sym} OF (replicate 2 (the_single uncollapses))) THEN'
-      TRY o EVERY' [rtac udisc_exhaust, atac, rtac vdisc_exhaust, atac])
-  else
-    let val ks = 1 upto n in
-      HEADGOAL (rtac udisc_exhaust THEN'
-        EVERY' (map5 (fn k => fn m => fn disc_excludess => fn disc_excludess' =>
-            fn uuncollapse =>
-          EVERY' [rtac (uuncollapse RS trans) THEN' TRY o atac,
-            rtac sym, rtac vdisc_exhaust,
-            EVERY' (map4 (fn k' => fn disc_excludes => fn disc_excludes' => fn vuncollapse =>
-              EVERY'
-                (if k' = k then
-                   [rtac (vuncollapse RS trans), TRY o atac] @
-                   (if m = 0 then
-                      [rtac refl]
-                    else
-                      [if n = 1 then K all_tac else EVERY' [dtac meta_mp, atac, dtac meta_mp, atac],
-                       REPEAT_DETERM_N (Int.max (0, m - 1)) o etac conjE,
-                       asm_simp_tac (ss_only [] ctxt)])
-                 else
-                   [dtac (the_single (if k = n then disc_excludes else disc_excludes')),
-                    etac (if k = n then @{thm iff_contradict(1)} else @{thm iff_contradict(2)}),
-                    atac, atac]))
-              ks disc_excludess disc_excludess' uncollapses)])
-          ks ms disc_excludesss disc_excludesss' uncollapses))
-    end;
-
-fun mk_case_same_ctr_tac ctxt injects =
-  REPEAT_DETERM o etac exE THEN' etac conjE THEN'
-    (case injects of
-      [] => atac
-    | [inject] => dtac (inject RS iffD1) THEN' REPEAT_DETERM o etac conjE THEN'
-        hyp_subst_tac ctxt THEN' rtac refl);
-
-fun mk_case_distinct_ctrs_tac ctxt distincts =
-  REPEAT_DETERM o etac exE THEN' etac conjE THEN' full_simp_tac (ss_only distincts ctxt);
-
-fun mk_case_tac ctxt n k case_def injects distinctss =
-  let
-    val case_def' = mk_unabs_def (n + 1) (case_def RS meta_eq_to_obj_eq);
-    val ks = 1 upto n;
-  in
-    HEADGOAL (rtac (case_def' RS trans) THEN' rtac @{thm the_equality} THEN'
-      rtac (mk_disjIN n k) THEN' REPEAT_DETERM o rtac exI THEN' rtac conjI THEN' rtac refl THEN'
-      rtac refl THEN'
-      EVERY' (map2 (fn k' => fn distincts =>
-        (if k' < n then etac disjE else K all_tac) THEN'
-        (if k' = k then mk_case_same_ctr_tac ctxt injects
-         else mk_case_distinct_ctrs_tac ctxt distincts)) ks distinctss))
-  end;
-
-fun mk_case_cong_tac ctxt uexhaust cases =
-  HEADGOAL (rtac uexhaust THEN'
-    EVERY' (maps (fn casex => [dtac sym, asm_simp_tac (ss_only [casex] ctxt)]) cases));
-
-fun mk_case_eq_if_tac ctxt n uexhaust cases discss' selss =
-  HEADGOAL (rtac uexhaust THEN'
-    EVERY' (map3 (fn casex => fn if_discs => fn sels =>
-        EVERY' [hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt (if_discs @ sels)),
-          rtac casex])
-      cases (map2 (seq_conds if_P_or_not_P_OF n) (1 upto n) discss') selss));
-
-fun mk_split_tac ctxt uexhaust cases selss injectss distinctsss =
-  HEADGOAL (rtac uexhaust) THEN
-  ALLGOALS (fn k => (hyp_subst_tac ctxt THEN'
-     simp_tac (ss_only (@{thms simp_thms} @ cases @ nth selss (k - 1) @ nth injectss (k - 1) @
-       flat (nth distinctsss (k - 1))) ctxt)) k) THEN
-  ALLGOALS (blast_tac (put_claset (claset_of @{theory_context HOL}) ctxt));
-
-val split_asm_thms = @{thms imp_conv_disj de_Morgan_conj de_Morgan_disj not_not not_ex};
-
-fun mk_split_asm_tac ctxt split =
-  HEADGOAL (rtac (split RS trans)) THEN unfold_thms_tac ctxt split_asm_thms THEN
-  HEADGOAL (rtac refl);
-
-end;
-
-structure Ctr_Sugar_General_Tactics : CTR_SUGAR_GENERAL_TACTICS = Ctr_Sugar_Tactics;
--- a/src/HOL/Tools/ctr_sugar_util.ML	Mon Dec 09 06:33:46 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,244 +0,0 @@
-(*  Title:      HOL/Tools/ctr_sugar_util.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012, 2013
-
-Library for wrapping existing freely generated type's constructors.
-*)
-
-signature CTR_SUGAR_UTIL =
-sig
-  val map3: ('a -> 'b -> 'c -> 'd) -> 'a list -> 'b list -> 'c list -> 'd list
-  val map4: ('a -> 'b -> 'c -> 'd -> 'e) -> 'a list -> 'b list -> 'c list -> 'd list -> 'e list
-  val map5: ('a -> 'b -> 'c -> 'd -> 'e -> 'f) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list
-  val fold_map2: ('a -> 'b -> 'c -> 'd * 'c) -> 'a list -> 'b list -> 'c -> 'd list * 'c
-  val fold_map3: ('a -> 'b -> 'c -> 'd -> 'e * 'd) ->
-    'a list -> 'b list -> 'c list -> 'd -> 'e list * 'd
-  val seq_conds: (bool -> 'a -> 'b) -> int -> int -> 'a list -> 'b list
-  val transpose: 'a list list -> 'a list list
-  val pad_list: 'a -> int -> 'a list -> 'a list
-  val splice: 'a list -> 'a list -> 'a list
-  val permute_like: ('a * 'b -> bool) -> 'a list -> 'b list -> 'c list -> 'c list
-
-  val mk_names: int -> string -> string list
-  val mk_fresh_names: Proof.context -> int -> string -> string list * Proof.context
-  val mk_TFrees': sort list -> Proof.context -> typ list * Proof.context
-  val mk_TFrees: int -> Proof.context -> typ list * Proof.context
-  val mk_Frees': string -> typ list -> Proof.context ->
-    (term list * (string * typ) list) * Proof.context
-  val mk_Freess': string -> typ list list -> Proof.context ->
-    (term list list * (string * typ) list list) * Proof.context
-  val mk_Frees: string -> typ list -> Proof.context -> term list * Proof.context
-  val mk_Freess: string -> typ list list -> Proof.context -> term list list * Proof.context
-  val resort_tfree: sort -> typ -> typ
-  val variant_types: string list -> sort list -> Proof.context ->
-    (string * sort) list * Proof.context
-  val variant_tfrees: string list -> Proof.context -> typ list * Proof.context
-
-  val typ_subst_nonatomic: (typ * typ) list -> typ -> typ
-  val subst_nonatomic_types: (typ * typ) list -> term -> term
-
-  val mk_predT: typ list -> typ
-  val mk_pred1T: typ -> typ
-
-  val mk_disjIN: int -> int -> thm
-
-  val mk_unabs_def: int -> thm -> thm
-
-  val mk_IfN: typ -> term list -> term list -> term
-  val mk_Trueprop_eq: term * term -> term
-
-  val rapp: term -> term -> term
-
-  val list_all_free: term list -> term -> term
-  val list_exists_free: term list -> term -> term
-
-  val fo_match: Proof.context -> term -> term -> Type.tyenv * Envir.tenv
-
-  val cterm_instantiate_pos: cterm option list -> thm -> thm
-  val unfold_thms: Proof.context -> thm list -> thm -> thm
-
-  val certifyT: Proof.context -> typ -> ctyp
-  val certify: Proof.context -> term -> cterm
-
-  val standard_binding: binding
-  val equal_binding: binding
-  val parse_binding: binding parser
-
-  val ss_only: thm list -> Proof.context -> Proof.context
-
-  val WRAP: ('a -> tactic) -> ('a -> tactic) -> 'a list -> tactic -> tactic
-  val WRAP': ('a -> int -> tactic) -> ('a -> int -> tactic) -> 'a list -> (int -> tactic) -> int ->
-    tactic
-  val CONJ_WRAP_GEN: tactic -> ('a -> tactic) -> 'a list -> tactic
-  val CONJ_WRAP_GEN': (int -> tactic) -> ('a -> int -> tactic) -> 'a list -> int -> tactic
-  val CONJ_WRAP: ('a -> tactic) -> 'a list -> tactic
-  val CONJ_WRAP': ('a -> int -> tactic) -> 'a list -> int -> tactic
-end;
-
-structure Ctr_Sugar_Util : CTR_SUGAR_UTIL =
-struct
-
-fun map3 _ [] [] [] = []
-  | map3 f (x1::x1s) (x2::x2s) (x3::x3s) = f x1 x2 x3 :: map3 f x1s x2s x3s
-  | map3 _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map4 _ [] [] [] [] = []
-  | map4 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) = f x1 x2 x3 x4 :: map4 f x1s x2s x3s x4s
-  | map4 _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map5 _ [] [] [] [] [] = []
-  | map5 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) =
-    f x1 x2 x3 x4 x5 :: map5 f x1s x2s x3s x4s x5s
-  | map5 _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun fold_map2 _ [] [] acc = ([], acc)
-  | fold_map2 f (x1::x1s) (x2::x2s) acc =
-    let
-      val (x, acc') = f x1 x2 acc;
-      val (xs, acc'') = fold_map2 f x1s x2s acc';
-    in (x :: xs, acc'') end
-  | fold_map2 _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun fold_map3 _ [] [] [] acc = ([], acc)
-  | fold_map3 f (x1::x1s) (x2::x2s) (x3::x3s) acc =
-    let
-      val (x, acc') = f x1 x2 x3 acc;
-      val (xs, acc'') = fold_map3 f x1s x2s x3s acc';
-    in (x :: xs, acc'') end
-  | fold_map3 _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun seq_conds f n k xs =
-  if k = n then
-    map (f false) (take (k - 1) xs)
-  else
-    let val (negs, pos) = split_last (take k xs) in
-      map (f false) negs @ [f true pos]
-    end;
-
-fun transpose [] = []
-  | transpose ([] :: xss) = transpose xss
-  | transpose xss = map hd xss :: transpose (map tl xss);
-
-fun pad_list x n xs = xs @ replicate (n - length xs) x;
-
-fun splice xs ys = flat (map2 (fn x => fn y => [x, y]) xs ys);
-
-fun permute_like eq xs xs' ys = map (nth ys o (fn y => find_index (fn x => eq (x, y)) xs)) xs';
-
-fun mk_names n x = if n = 1 then [x] else map (fn i => x ^ string_of_int i) (1 upto n);
-fun mk_fresh_names ctxt = (fn xs => Variable.variant_fixes xs ctxt) oo mk_names;
-
-val mk_TFrees' = apfst (map TFree) oo Variable.invent_types;
-
-fun mk_TFrees n = mk_TFrees' (replicate n HOLogic.typeS);
-
-fun mk_Frees' x Ts ctxt = mk_fresh_names ctxt (length Ts) x |>> (fn xs => `(map Free) (xs ~~ Ts));
-fun mk_Freess' x Tss = fold_map2 mk_Frees' (mk_names (length Tss) x) Tss #>> split_list;
-
-fun mk_Frees x Ts ctxt = mk_fresh_names ctxt (length Ts) x |>> (fn xs => map2 (curry Free) xs Ts);
-fun mk_Freess x Tss = fold_map2 mk_Frees (mk_names (length Tss) x) Tss;
-
-fun resort_tfree S (TFree (s, _)) = TFree (s, S);
-
-fun ensure_prefix pre s = s |> not (String.isPrefix pre s) ? prefix pre;
-
-fun variant_types ss Ss ctxt =
-  let
-    val (tfrees, _) =
-      fold_map2 (fn s => fn S => Name.variant s #> apfst (rpair S)) ss Ss (Variable.names_of ctxt);
-    val ctxt' = fold (Variable.declare_constraints o Logic.mk_type o TFree) tfrees ctxt;
-  in (tfrees, ctxt') end;
-
-fun variant_tfrees ss =
-  apfst (map TFree) o
-    variant_types (map (ensure_prefix "'") ss) (replicate (length ss) HOLogic.typeS);
-
-(*Replace each Ti by Ui (starting from the leaves); inst = [(T1, U1), ..., (Tn, Un)].*)
-fun typ_subst_nonatomic [] = I
-  | typ_subst_nonatomic inst =
-    let
-      fun subst (Type (s, Ts)) = perhaps (AList.lookup (op =) inst) (Type (s, map subst Ts))
-        | subst T = perhaps (AList.lookup (op =) inst) T;
-    in subst end;
-
-fun subst_nonatomic_types [] = I
-  | subst_nonatomic_types inst = map_types (typ_subst_nonatomic inst);
-
-fun mk_predT Ts = Ts ---> HOLogic.boolT;
-fun mk_pred1T T = mk_predT [T];
-
-fun mk_disjIN 1 1 = @{thm TrueE[OF TrueI]}
-  | mk_disjIN _ 1 = disjI1
-  | mk_disjIN 2 2 = disjI2
-  | mk_disjIN n m = (mk_disjIN (n - 1) (m - 1)) RS disjI2;
-
-fun mk_unabs_def n = funpow n (fn thm => thm RS fun_cong);
-
-fun mk_IfN _ _ [t] = t
-  | mk_IfN T (c :: cs) (t :: ts) =
-    Const (@{const_name If}, HOLogic.boolT --> T --> T --> T) $ c $ t $ mk_IfN T cs ts;
-
-val mk_Trueprop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq;
-
-fun rapp u t = betapply (t, u);
-
-fun list_quant_free quant_const =
-  fold_rev (fn Free (xT as (_, T)) => fn P => quant_const T $ Term.absfree xT P);
-
-val list_all_free = list_quant_free HOLogic.all_const;
-val list_exists_free = list_quant_free HOLogic.exists_const;
-
-fun fo_match ctxt t pat =
-  let val thy = Proof_Context.theory_of ctxt in
-    Pattern.first_order_match thy (pat, t) (Vartab.empty, Vartab.empty)
-  end;
-
-fun cterm_instantiate_pos cts thm =
-  let
-    val cert = Thm.cterm_of (Thm.theory_of_thm thm);
-    val vars = Term.add_vars (prop_of thm) [];
-    val vars' = rev (drop (length vars - length cts) vars);
-    val ps = map_filter (fn (_, NONE) => NONE
-      | (var, SOME ct) => SOME (cert (Var var), ct)) (vars' ~~ cts);
-  in
-    Drule.cterm_instantiate ps thm
-  end;
-
-fun unfold_thms ctxt thms = Local_Defs.unfold ctxt (distinct Thm.eq_thm_prop thms);
-
-(*stolen from ~~/src/HOL/Tools/SMT/smt_utils.ML*)
-fun certifyT ctxt = Thm.ctyp_of (Proof_Context.theory_of ctxt);
-fun certify ctxt = Thm.cterm_of (Proof_Context.theory_of ctxt);
-
-(* The standard binding stands for a name generated following the canonical convention (e.g.,
-   "is_Nil" from "Nil"). In contrast, the empty binding is either the standard binding or no
-   binding at all, depending on the context. *)
-val standard_binding = @{binding _};
-val equal_binding = @{binding "="};
-
-val parse_binding = Parse.binding || @{keyword "="} >> K equal_binding;
-
-fun ss_only thms ctxt = clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps thms;
-
-(*Tactical WRAP surrounds a static given tactic (core) with two deterministic chains of tactics*)
-fun WRAP gen_before gen_after xs core_tac =
-  fold_rev (fn x => fn tac => gen_before x THEN tac THEN gen_after x) xs core_tac;
-
-fun WRAP' gen_before gen_after xs core_tac =
-  fold_rev (fn x => fn tac => gen_before x THEN' tac THEN' gen_after x) xs core_tac;
-
-fun CONJ_WRAP_GEN conj_tac gen_tac xs =
-  let val (butlast, last) = split_last xs;
-  in WRAP (fn thm => conj_tac THEN gen_tac thm) (K all_tac) butlast (gen_tac last) end;
-
-fun CONJ_WRAP_GEN' conj_tac gen_tac xs =
-  let val (butlast, last) = split_last xs;
-  in WRAP' (fn thm => conj_tac THEN' gen_tac thm) (K (K all_tac)) butlast (gen_tac last) end;
-
-(*not eta-converted because of monotype restriction*)
-fun CONJ_WRAP gen_tac = CONJ_WRAP_GEN (rtac conjI 1) gen_tac;
-fun CONJ_WRAP' gen_tac = CONJ_WRAP_GEN' (rtac conjI) gen_tac;
-
-end;