src/HOL/Nominal/nominal.ML
changeset 31936 904f93c76650
parent 31784 bd3486c57ba3
equal deleted inserted replaced
31923:d6cd15601d8a 31936:904f93c76650
     1 (*  Title:      HOL/Nominal/nominal.ML
       
     2     Author:     Stefan Berghofer and Christian Urban, TU Muenchen
       
     3 
       
     4 Nominal datatype package for Isabelle/HOL.
       
     5 *)
       
     6 
       
     7 signature NOMINAL =
       
     8 sig
       
     9   val add_nominal_datatype : Datatype.config -> string list ->
       
    10     (string list * bstring * mixfix *
       
    11       (bstring * string list * mixfix) list) list -> theory -> theory
       
    12   type descr
       
    13   type nominal_datatype_info
       
    14   val get_nominal_datatypes : theory -> nominal_datatype_info Symtab.table
       
    15   val get_nominal_datatype : theory -> string -> nominal_datatype_info option
       
    16   val mk_perm: typ list -> term -> term -> term
       
    17   val perm_of_pair: term * term -> term
       
    18   val mk_not_sym: thm list -> thm list
       
    19   val perm_simproc: simproc
       
    20   val fresh_const: typ -> typ -> term
       
    21   val fresh_star_const: typ -> typ -> term
       
    22 end
       
    23 
       
    24 structure Nominal : NOMINAL =
       
    25 struct
       
    26 
       
    27 val finite_emptyI = thm "finite.emptyI";
       
    28 val finite_Diff = thm "finite_Diff";
       
    29 val finite_Un = thm "finite_Un";
       
    30 val Un_iff = thm "Un_iff";
       
    31 val In0_eq = thm "In0_eq";
       
    32 val In1_eq = thm "In1_eq";
       
    33 val In0_not_In1 = thm "In0_not_In1";
       
    34 val In1_not_In0 = thm "In1_not_In0";
       
    35 val Un_assoc = thm "Un_assoc";
       
    36 val Collect_disj_eq = thm "Collect_disj_eq";
       
    37 val empty_def = thm "empty_def";
       
    38 val empty_iff = thm "empty_iff";
       
    39 
       
    40 open DatatypeAux;
       
    41 open NominalAtoms;
       
    42 
       
    43 (** FIXME: Datatype should export this function **)
       
    44 
       
    45 local
       
    46 
       
    47 fun dt_recs (DtTFree _) = []
       
    48   | dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts)
       
    49   | dt_recs (DtRec i) = [i];
       
    50 
       
    51 fun dt_cases (descr: descr) (_, args, constrs) =
       
    52   let
       
    53     fun the_bname i = Long_Name.base_name (#1 (valOf (AList.lookup (op =) descr i)));
       
    54     val bnames = map the_bname (distinct op = (List.concat (map dt_recs args)));
       
    55   in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
       
    56 
       
    57 
       
    58 fun induct_cases descr =
       
    59   DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr)));
       
    60 
       
    61 fun exhaust_cases descr i = dt_cases descr (valOf (AList.lookup (op =) descr i));
       
    62 
       
    63 in
       
    64 
       
    65 fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
       
    66 
       
    67 fun mk_case_names_exhausts descr new =
       
    68   map (RuleCases.case_names o exhaust_cases descr o #1)
       
    69     (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr);
       
    70 
       
    71 end;
       
    72 
       
    73 (* theory data *)
       
    74 
       
    75 type descr = (int * (string * dtyp list * (string * (dtyp list * dtyp) list) list)) list;
       
    76 
       
    77 type nominal_datatype_info =
       
    78   {index : int,
       
    79    descr : descr,
       
    80    sorts : (string * sort) list,
       
    81    rec_names : string list,
       
    82    rec_rewrites : thm list,
       
    83    induction : thm,
       
    84    distinct : thm list,
       
    85    inject : thm list};
       
    86 
       
    87 structure NominalDatatypesData = TheoryDataFun
       
    88 (
       
    89   type T = nominal_datatype_info Symtab.table;
       
    90   val empty = Symtab.empty;
       
    91   val copy = I;
       
    92   val extend = I;
       
    93   fun merge _ tabs : T = Symtab.merge (K true) tabs;
       
    94 );
       
    95 
       
    96 val get_nominal_datatypes = NominalDatatypesData.get;
       
    97 val put_nominal_datatypes = NominalDatatypesData.put;
       
    98 val map_nominal_datatypes = NominalDatatypesData.map;
       
    99 val get_nominal_datatype = Symtab.lookup o get_nominal_datatypes;
       
   100 
       
   101 
       
   102 (**** make datatype info ****)
       
   103 
       
   104 fun make_dt_info descr sorts induct reccomb_names rec_thms
       
   105     (((i, (_, (tname, _, _))), distinct), inject) =
       
   106   (tname,
       
   107    {index = i,
       
   108     descr = descr,
       
   109     sorts = sorts,
       
   110     rec_names = reccomb_names,
       
   111     rec_rewrites = rec_thms,
       
   112     induction = induct,
       
   113     distinct = distinct,
       
   114     inject = inject});
       
   115 
       
   116 (*******************************)
       
   117 
       
   118 val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
       
   119 
       
   120 
       
   121 (** simplification procedure for sorting permutations **)
       
   122 
       
   123 val dj_cp = thm "dj_cp";
       
   124 
       
   125 fun dest_permT (Type ("fun", [Type ("List.list", [Type ("*", [T, _])]),
       
   126       Type ("fun", [_, U])])) = (T, U);
       
   127 
       
   128 fun permTs_of (Const ("Nominal.perm", T) $ t $ u) = fst (dest_permT T) :: permTs_of u
       
   129   | permTs_of _ = [];
       
   130 
       
   131 fun perm_simproc' thy ss (Const ("Nominal.perm", T) $ t $ (u as Const ("Nominal.perm", U) $ r $ s)) =
       
   132       let
       
   133         val (aT as Type (a, []), S) = dest_permT T;
       
   134         val (bT as Type (b, []), _) = dest_permT U
       
   135       in if aT mem permTs_of u andalso aT <> bT then
       
   136           let
       
   137             val cp = cp_inst_of thy a b;
       
   138             val dj = dj_thm_of thy b a;
       
   139             val dj_cp' = [cp, dj] MRS dj_cp;
       
   140             val cert = SOME o cterm_of thy
       
   141           in
       
   142             SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)]
       
   143               [cert t, cert r, cert s] dj_cp'))
       
   144           end
       
   145         else NONE
       
   146       end
       
   147   | perm_simproc' thy ss _ = NONE;
       
   148 
       
   149 val perm_simproc =
       
   150   Simplifier.simproc (the_context ()) "perm_simp" ["pi1 \<bullet> (pi2 \<bullet> x)"] perm_simproc';
       
   151 
       
   152 val meta_spec = thm "meta_spec";
       
   153 
       
   154 fun projections rule =
       
   155   ProjectRule.projections (ProofContext.init (Thm.theory_of_thm rule)) rule
       
   156   |> map (standard #> RuleCases.save rule);
       
   157 
       
   158 val supp_prod = thm "supp_prod";
       
   159 val fresh_prod = thm "fresh_prod";
       
   160 val supports_fresh = thm "supports_fresh";
       
   161 val supports_def = thm "Nominal.supports_def";
       
   162 val fresh_def = thm "fresh_def";
       
   163 val supp_def = thm "supp_def";
       
   164 val rev_simps = thms "rev.simps";
       
   165 val app_simps = thms "append.simps";
       
   166 val at_fin_set_supp = thm "at_fin_set_supp";
       
   167 val at_fin_set_fresh = thm "at_fin_set_fresh";
       
   168 val abs_fun_eq1 = thm "abs_fun_eq1";
       
   169 
       
   170 val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq];
       
   171 
       
   172 fun mk_perm Ts t u =
       
   173   let
       
   174     val T = fastype_of1 (Ts, t);
       
   175     val U = fastype_of1 (Ts, u)
       
   176   in Const ("Nominal.perm", T --> U --> U) $ t $ u end;
       
   177 
       
   178 fun perm_of_pair (x, y) =
       
   179   let
       
   180     val T = fastype_of x;
       
   181     val pT = mk_permT T
       
   182   in Const ("List.list.Cons", HOLogic.mk_prodT (T, T) --> pT --> pT) $
       
   183     HOLogic.mk_prod (x, y) $ Const ("List.list.Nil", pT)
       
   184   end;
       
   185 
       
   186 fun mk_not_sym ths = maps (fn th => case prop_of th of
       
   187     _ $ (Const ("Not", _) $ (Const ("op =", _) $ _ $ _)) => [th, th RS not_sym]
       
   188   | _ => [th]) ths;
       
   189 
       
   190 fun fresh_const T U = Const ("Nominal.fresh", T --> U --> HOLogic.boolT);
       
   191 fun fresh_star_const T U =
       
   192   Const ("Nominal.fresh_star", HOLogic.mk_setT T --> U --> HOLogic.boolT);
       
   193 
       
   194 fun gen_add_nominal_datatype prep_typ config new_type_names dts thy =
       
   195   let
       
   196     (* this theory is used just for parsing *)
       
   197 
       
   198     val tmp_thy = thy |>
       
   199       Theory.copy |>
       
   200       Sign.add_types (map (fn (tvs, tname, mx, _) =>
       
   201         (Binding.name tname, length tvs, mx)) dts);
       
   202 
       
   203     val atoms = atoms_of thy;
       
   204 
       
   205     fun prep_constr ((constrs, sorts), (cname, cargs, mx)) =
       
   206       let val (cargs', sorts') = Library.foldl (prep_typ tmp_thy) (([], sorts), cargs)
       
   207       in (constrs @ [(cname, cargs', mx)], sorts') end
       
   208 
       
   209     fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) =
       
   210       let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs)
       
   211       in (dts @ [(tvs, tname, mx, constrs')], sorts') end
       
   212 
       
   213     val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts);
       
   214     val tyvars = map (map (fn s =>
       
   215       (s, the (AList.lookup (op =) sorts s))) o #1) dts';
       
   216 
       
   217     fun inter_sort thy S S' = Type.inter_sort (Sign.tsig_of thy) (S, S');
       
   218     fun augment_sort_typ thy S =
       
   219       let val S = Sign.certify_sort thy S
       
   220       in map_type_tfree (fn (s, S') => TFree (s,
       
   221         if member (op = o apsnd fst) sorts s then inter_sort thy S S' else S'))
       
   222       end;
       
   223     fun augment_sort thy S = map_types (augment_sort_typ thy S);
       
   224 
       
   225     val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts';
       
   226     val constr_syntax = map (fn (tvs, tname, mx, constrs) =>
       
   227       map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts';
       
   228 
       
   229     val ps = map (fn (_, n, _, _) =>
       
   230       (Sign.full_bname tmp_thy n, Sign.full_bname tmp_thy (n ^ "_Rep"))) dts;
       
   231     val rps = map Library.swap ps;
       
   232 
       
   233     fun replace_types (Type ("Nominal.ABS", [T, U])) =
       
   234           Type ("fun", [T, Type ("Nominal.noption", [replace_types U])])
       
   235       | replace_types (Type (s, Ts)) =
       
   236           Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts)
       
   237       | replace_types T = T;
       
   238 
       
   239     val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, Binding.name (tname ^ "_Rep"), NoSyn,
       
   240       map (fn (cname, cargs, mx) => (Binding.name (cname ^ "_Rep"),
       
   241         map replace_types cargs, NoSyn)) constrs)) dts';
       
   242 
       
   243     val new_type_names' = map (fn n => n ^ "_Rep") new_type_names;
       
   244 
       
   245     val (full_new_type_names',thy1) =
       
   246       Datatype.add_datatype config new_type_names' dts'' thy;
       
   247 
       
   248     val {descr, induction, ...} =
       
   249       Datatype.the_info thy1 (hd full_new_type_names');
       
   250     fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
       
   251 
       
   252     val big_name = space_implode "_" new_type_names;
       
   253 
       
   254 
       
   255     (**** define permutation functions ****)
       
   256 
       
   257     val permT = mk_permT (TFree ("'x", HOLogic.typeS));
       
   258     val pi = Free ("pi", permT);
       
   259     val perm_types = map (fn (i, _) =>
       
   260       let val T = nth_dtyp i
       
   261       in permT --> T --> T end) descr;
       
   262     val perm_names' = DatatypeProp.indexify_names (map (fn (i, _) =>
       
   263       "perm_" ^ name_of_typ (nth_dtyp i)) descr);
       
   264     val perm_names = replicate (length new_type_names) "Nominal.perm" @
       
   265       map (Sign.full_bname thy1) (List.drop (perm_names', length new_type_names));
       
   266     val perm_names_types = perm_names ~~ perm_types;
       
   267     val perm_names_types' = perm_names' ~~ perm_types;
       
   268 
       
   269     val perm_eqs = maps (fn (i, (_, _, constrs)) =>
       
   270       let val T = nth_dtyp i
       
   271       in map (fn (cname, dts) =>
       
   272         let
       
   273           val Ts = map (typ_of_dtyp descr sorts) dts;
       
   274           val names = Name.variant_list ["pi"] (DatatypeProp.make_tnames Ts);
       
   275           val args = map Free (names ~~ Ts);
       
   276           val c = Const (cname, Ts ---> T);
       
   277           fun perm_arg (dt, x) =
       
   278             let val T = type_of x
       
   279             in if is_rec_type dt then
       
   280                 let val (Us, _) = strip_type T
       
   281                 in list_abs (map (pair "x") Us,
       
   282                   Free (nth perm_names_types' (body_index dt)) $ pi $
       
   283                     list_comb (x, map (fn (i, U) =>
       
   284                       Const ("Nominal.perm", permT --> U --> U) $
       
   285                         (Const ("List.rev", permT --> permT) $ pi) $
       
   286                         Bound i) ((length Us - 1 downto 0) ~~ Us)))
       
   287                 end
       
   288               else Const ("Nominal.perm", permT --> T --> T) $ pi $ x
       
   289             end;
       
   290         in
       
   291           (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq
       
   292             (Free (nth perm_names_types' i) $
       
   293                Free ("pi", mk_permT (TFree ("'x", HOLogic.typeS))) $
       
   294                list_comb (c, args),
       
   295              list_comb (c, map perm_arg (dts ~~ args)))))
       
   296         end) constrs
       
   297       end) descr;
       
   298 
       
   299     val (perm_simps, thy2) =
       
   300       Primrec.add_primrec_overloaded
       
   301         (map (fn (s, sT) => (s, sT, false))
       
   302            (List.take (perm_names' ~~ perm_names_types, length new_type_names)))
       
   303         (map (fn s => (Binding.name s, NONE, NoSyn)) perm_names') perm_eqs thy1;
       
   304 
       
   305     (**** prove that permutation functions introduced by unfolding are ****)
       
   306     (**** equivalent to already existing permutation functions         ****)
       
   307 
       
   308     val _ = warning ("length descr: " ^ string_of_int (length descr));
       
   309     val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names));
       
   310 
       
   311     val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types);
       
   312     val perm_fun_def = PureThy.get_thm thy2 "perm_fun_def";
       
   313 
       
   314     val unfolded_perm_eq_thms =
       
   315       if length descr = length new_type_names then []
       
   316       else map standard (List.drop (split_conj_thm
       
   317         (Goal.prove_global thy2 [] []
       
   318           (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
       
   319             (map (fn (c as (s, T), x) =>
       
   320                let val [T1, T2] = binder_types T
       
   321                in HOLogic.mk_eq (Const c $ pi $ Free (x, T2),
       
   322                  Const ("Nominal.perm", T) $ pi $ Free (x, T2))
       
   323                end)
       
   324              (perm_names_types ~~ perm_indnames))))
       
   325           (fn _ => EVERY [indtac induction perm_indnames 1,
       
   326             ALLGOALS (asm_full_simp_tac
       
   327               (simpset_of thy2 addsimps [perm_fun_def]))])),
       
   328         length new_type_names));
       
   329 
       
   330     (**** prove [] \<bullet> t = t ****)
       
   331 
       
   332     val _ = warning "perm_empty_thms";
       
   333 
       
   334     val perm_empty_thms = List.concat (map (fn a =>
       
   335       let val permT = mk_permT (Type (a, []))
       
   336       in map standard (List.take (split_conj_thm
       
   337         (Goal.prove_global thy2 [] []
       
   338           (augment_sort thy2 [pt_class_of thy2 a]
       
   339             (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
       
   340               (map (fn ((s, T), x) => HOLogic.mk_eq
       
   341                   (Const (s, permT --> T --> T) $
       
   342                      Const ("List.list.Nil", permT) $ Free (x, T),
       
   343                    Free (x, T)))
       
   344                (perm_names ~~
       
   345                 map body_type perm_types ~~ perm_indnames)))))
       
   346           (fn _ => EVERY [indtac induction perm_indnames 1,
       
   347             ALLGOALS (asm_full_simp_tac (simpset_of thy2))])),
       
   348         length new_type_names))
       
   349       end)
       
   350       atoms);
       
   351 
       
   352     (**** prove (pi1 @ pi2) \<bullet> t = pi1 \<bullet> (pi2 \<bullet> t) ****)
       
   353 
       
   354     val _ = warning "perm_append_thms";
       
   355 
       
   356     (*FIXME: these should be looked up statically*)
       
   357     val at_pt_inst = PureThy.get_thm thy2 "at_pt_inst";
       
   358     val pt2 = PureThy.get_thm thy2 "pt2";
       
   359 
       
   360     val perm_append_thms = List.concat (map (fn a =>
       
   361       let
       
   362         val permT = mk_permT (Type (a, []));
       
   363         val pi1 = Free ("pi1", permT);
       
   364         val pi2 = Free ("pi2", permT);
       
   365         val pt_inst = pt_inst_of thy2 a;
       
   366         val pt2' = pt_inst RS pt2;
       
   367         val pt2_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "2") a);
       
   368       in List.take (map standard (split_conj_thm
       
   369         (Goal.prove_global thy2 [] []
       
   370            (augment_sort thy2 [pt_class_of thy2 a]
       
   371              (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
       
   372                 (map (fn ((s, T), x) =>
       
   373                     let val perm = Const (s, permT --> T --> T)
       
   374                     in HOLogic.mk_eq
       
   375                       (perm $ (Const ("List.append", permT --> permT --> permT) $
       
   376                          pi1 $ pi2) $ Free (x, T),
       
   377                        perm $ pi1 $ (perm $ pi2 $ Free (x, T)))
       
   378                     end)
       
   379                   (perm_names ~~
       
   380                    map body_type perm_types ~~ perm_indnames)))))
       
   381            (fn _ => EVERY [indtac induction perm_indnames 1,
       
   382               ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))),
       
   383          length new_type_names)
       
   384       end) atoms);
       
   385 
       
   386     (**** prove pi1 ~ pi2 ==> pi1 \<bullet> t = pi2 \<bullet> t ****)
       
   387 
       
   388     val _ = warning "perm_eq_thms";
       
   389 
       
   390     val pt3 = PureThy.get_thm thy2 "pt3";
       
   391     val pt3_rev = PureThy.get_thm thy2 "pt3_rev";
       
   392 
       
   393     val perm_eq_thms = List.concat (map (fn a =>
       
   394       let
       
   395         val permT = mk_permT (Type (a, []));
       
   396         val pi1 = Free ("pi1", permT);
       
   397         val pi2 = Free ("pi2", permT);
       
   398         val at_inst = at_inst_of thy2 a;
       
   399         val pt_inst = pt_inst_of thy2 a;
       
   400         val pt3' = pt_inst RS pt3;
       
   401         val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
       
   402         val pt3_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "3") a);
       
   403       in List.take (map standard (split_conj_thm
       
   404         (Goal.prove_global thy2 [] []
       
   405           (augment_sort thy2 [pt_class_of thy2 a] (Logic.mk_implies
       
   406              (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq",
       
   407                 permT --> permT --> HOLogic.boolT) $ pi1 $ pi2),
       
   408               HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
       
   409                 (map (fn ((s, T), x) =>
       
   410                     let val perm = Const (s, permT --> T --> T)
       
   411                     in HOLogic.mk_eq
       
   412                       (perm $ pi1 $ Free (x, T),
       
   413                        perm $ pi2 $ Free (x, T))
       
   414                     end)
       
   415                   (perm_names ~~
       
   416                    map body_type perm_types ~~ perm_indnames))))))
       
   417            (fn _ => EVERY [indtac induction perm_indnames 1,
       
   418               ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))),
       
   419          length new_type_names)
       
   420       end) atoms);
       
   421 
       
   422     (**** prove pi1 \<bullet> (pi2 \<bullet> t) = (pi1 \<bullet> pi2) \<bullet> (pi1 \<bullet> t) ****)
       
   423 
       
   424     val cp1 = PureThy.get_thm thy2 "cp1";
       
   425     val dj_cp = PureThy.get_thm thy2 "dj_cp";
       
   426     val pt_perm_compose = PureThy.get_thm thy2 "pt_perm_compose";
       
   427     val pt_perm_compose_rev = PureThy.get_thm thy2 "pt_perm_compose_rev";
       
   428     val dj_perm_perm_forget = PureThy.get_thm thy2 "dj_perm_perm_forget";
       
   429 
       
   430     fun composition_instance name1 name2 thy =
       
   431       let
       
   432         val cp_class = cp_class_of thy name1 name2;
       
   433         val pt_class =
       
   434           if name1 = name2 then [pt_class_of thy name1]
       
   435           else [];
       
   436         val permT1 = mk_permT (Type (name1, []));
       
   437         val permT2 = mk_permT (Type (name2, []));
       
   438         val Ts = map body_type perm_types;
       
   439         val cp_inst = cp_inst_of thy name1 name2;
       
   440         val simps = simpset_of thy addsimps (perm_fun_def ::
       
   441           (if name1 <> name2 then
       
   442              let val dj = dj_thm_of thy name2 name1
       
   443              in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end
       
   444            else
       
   445              let
       
   446                val at_inst = at_inst_of thy name1;
       
   447                val pt_inst = pt_inst_of thy name1;
       
   448              in
       
   449                [cp_inst RS cp1 RS sym,
       
   450                 at_inst RS (pt_inst RS pt_perm_compose) RS sym,
       
   451                 at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym]
       
   452             end))
       
   453         val sort = Sign.certify_sort thy (cp_class :: pt_class);
       
   454         val thms = split_conj_thm (Goal.prove_global thy [] []
       
   455           (augment_sort thy sort
       
   456             (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
       
   457               (map (fn ((s, T), x) =>
       
   458                   let
       
   459                     val pi1 = Free ("pi1", permT1);
       
   460                     val pi2 = Free ("pi2", permT2);
       
   461                     val perm1 = Const (s, permT1 --> T --> T);
       
   462                     val perm2 = Const (s, permT2 --> T --> T);
       
   463                     val perm3 = Const ("Nominal.perm", permT1 --> permT2 --> permT2)
       
   464                   in HOLogic.mk_eq
       
   465                     (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)),
       
   466                      perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T)))
       
   467                   end)
       
   468                 (perm_names ~~ Ts ~~ perm_indnames)))))
       
   469           (fn _ => EVERY [indtac induction perm_indnames 1,
       
   470              ALLGOALS (asm_full_simp_tac simps)]))
       
   471       in
       
   472         fold (fn (s, tvs) => fn thy => AxClass.prove_arity
       
   473             (s, map (inter_sort thy sort o snd) tvs, [cp_class])
       
   474             (Class.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy)
       
   475           (full_new_type_names' ~~ tyvars) thy
       
   476       end;
       
   477 
       
   478     val (perm_thmss,thy3) = thy2 |>
       
   479       fold (fn name1 => fold (composition_instance name1) atoms) atoms |>
       
   480       fold (fn atom => fn thy =>
       
   481         let val pt_name = pt_class_of thy atom
       
   482         in
       
   483           fold (fn (s, tvs) => fn thy => AxClass.prove_arity
       
   484               (s, map (inter_sort thy [pt_name] o snd) tvs, [pt_name])
       
   485               (EVERY
       
   486                 [Class.intro_classes_tac [],
       
   487                  resolve_tac perm_empty_thms 1,
       
   488                  resolve_tac perm_append_thms 1,
       
   489                  resolve_tac perm_eq_thms 1, assume_tac 1]) thy)
       
   490             (full_new_type_names' ~~ tyvars) thy
       
   491         end) atoms |>
       
   492       PureThy.add_thmss
       
   493         [((Binding.name (space_implode "_" new_type_names ^ "_unfolded_perm_eq"),
       
   494           unfolded_perm_eq_thms), [Simplifier.simp_add]),
       
   495          ((Binding.name (space_implode "_" new_type_names ^ "_perm_empty"),
       
   496           perm_empty_thms), [Simplifier.simp_add]),
       
   497          ((Binding.name (space_implode "_" new_type_names ^ "_perm_append"),
       
   498           perm_append_thms), [Simplifier.simp_add]),
       
   499          ((Binding.name (space_implode "_" new_type_names ^ "_perm_eq"),
       
   500           perm_eq_thms), [Simplifier.simp_add])];
       
   501 
       
   502     (**** Define representing sets ****)
       
   503 
       
   504     val _ = warning "representing sets";
       
   505 
       
   506     val rep_set_names = DatatypeProp.indexify_names
       
   507       (map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr);
       
   508     val big_rep_name =
       
   509       space_implode "_" (DatatypeProp.indexify_names (List.mapPartial
       
   510         (fn (i, ("Nominal.noption", _, _)) => NONE
       
   511           | (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set";
       
   512     val _ = warning ("big_rep_name: " ^ big_rep_name);
       
   513 
       
   514     fun strip_option (dtf as DtType ("fun", [dt, DtRec i])) =
       
   515           (case AList.lookup op = descr i of
       
   516              SOME ("Nominal.noption", _, [(_, [dt']), _]) =>
       
   517                apfst (cons dt) (strip_option dt')
       
   518            | _ => ([], dtf))
       
   519       | strip_option (DtType ("fun", [dt, DtType ("Nominal.noption", [dt'])])) =
       
   520           apfst (cons dt) (strip_option dt')
       
   521       | strip_option dt = ([], dt);
       
   522 
       
   523     val dt_atomTs = distinct op = (map (typ_of_dtyp descr sorts)
       
   524       (List.concat (map (fn (_, (_, _, cs)) => List.concat
       
   525         (map (List.concat o map (fst o strip_option) o snd) cs)) descr)));
       
   526     val dt_atoms = map (fst o dest_Type) dt_atomTs;
       
   527 
       
   528     fun make_intr s T (cname, cargs) =
       
   529       let
       
   530         fun mk_prem (dt, (j, j', prems, ts)) =
       
   531           let
       
   532             val (dts, dt') = strip_option dt;
       
   533             val (dts', dt'') = strip_dtyp dt';
       
   534             val Ts = map (typ_of_dtyp descr sorts) dts;
       
   535             val Us = map (typ_of_dtyp descr sorts) dts';
       
   536             val T = typ_of_dtyp descr sorts dt'';
       
   537             val free = mk_Free "x" (Us ---> T) j;
       
   538             val free' = app_bnds free (length Us);
       
   539             fun mk_abs_fun (T, (i, t)) =
       
   540               let val U = fastype_of t
       
   541               in (i + 1, Const ("Nominal.abs_fun", [T, U, T] --->
       
   542                 Type ("Nominal.noption", [U])) $ mk_Free "y" T i $ t)
       
   543               end
       
   544           in (j + 1, j' + length Ts,
       
   545             case dt'' of
       
   546                 DtRec k => list_all (map (pair "x") Us,
       
   547                   HOLogic.mk_Trueprop (Free (List.nth (rep_set_names, k),
       
   548                     T --> HOLogic.boolT) $ free')) :: prems
       
   549               | _ => prems,
       
   550             snd (List.foldr mk_abs_fun (j', free) Ts) :: ts)
       
   551           end;
       
   552 
       
   553         val (_, _, prems, ts) = List.foldr mk_prem (1, 1, [], []) cargs;
       
   554         val concl = HOLogic.mk_Trueprop (Free (s, T --> HOLogic.boolT) $
       
   555           list_comb (Const (cname, map fastype_of ts ---> T), ts))
       
   556       in Logic.list_implies (prems, concl)
       
   557       end;
       
   558 
       
   559     val (intr_ts, (rep_set_names', recTs')) =
       
   560       apfst List.concat (apsnd ListPair.unzip (ListPair.unzip (List.mapPartial
       
   561         (fn ((_, ("Nominal.noption", _, _)), _) => NONE
       
   562           | ((i, (_, _, constrs)), rep_set_name) =>
       
   563               let val T = nth_dtyp i
       
   564               in SOME (map (make_intr rep_set_name T) constrs,
       
   565                 (rep_set_name, T))
       
   566               end)
       
   567                 (descr ~~ rep_set_names))));
       
   568     val rep_set_names'' = map (Sign.full_bname thy3) rep_set_names';
       
   569 
       
   570     val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy4) =
       
   571         Inductive.add_inductive_global (serial_string ())
       
   572           {quiet_mode = false, verbose = false, kind = Thm.internalK,
       
   573            alt_name = Binding.name big_rep_name, coind = false, no_elim = true, no_ind = false,
       
   574            skip_mono = true, fork_mono = false}
       
   575           (map (fn (s, T) => ((Binding.name s, T --> HOLogic.boolT), NoSyn))
       
   576              (rep_set_names' ~~ recTs'))
       
   577           [] (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy3;
       
   578 
       
   579     (**** Prove that representing set is closed under permutation ****)
       
   580 
       
   581     val _ = warning "proving closure under permutation...";
       
   582 
       
   583     val abs_perm = PureThy.get_thms thy4 "abs_perm";
       
   584 
       
   585     val perm_indnames' = List.mapPartial
       
   586       (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x)
       
   587       (perm_indnames ~~ descr);
       
   588 
       
   589     fun mk_perm_closed name = map (fn th => standard (th RS mp))
       
   590       (List.take (split_conj_thm (Goal.prove_global thy4 [] []
       
   591         (augment_sort thy4
       
   592           (pt_class_of thy4 name :: map (cp_class_of thy4 name) (dt_atoms \ name))
       
   593           (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map
       
   594             (fn ((s, T), x) =>
       
   595                let
       
   596                  val S = Const (s, T --> HOLogic.boolT);
       
   597                  val permT = mk_permT (Type (name, []))
       
   598                in HOLogic.mk_imp (S $ Free (x, T),
       
   599                  S $ (Const ("Nominal.perm", permT --> T --> T) $
       
   600                    Free ("pi", permT) $ Free (x, T)))
       
   601                end) (rep_set_names'' ~~ recTs' ~~ perm_indnames')))))
       
   602         (fn _ => EVERY
       
   603            [indtac rep_induct [] 1,
       
   604             ALLGOALS (simp_tac (simpset_of thy4 addsimps
       
   605               (symmetric perm_fun_def :: abs_perm))),
       
   606             ALLGOALS (resolve_tac rep_intrs THEN_ALL_NEW assume_tac)])),
       
   607         length new_type_names));
       
   608 
       
   609     val perm_closed_thmss = map mk_perm_closed atoms;
       
   610 
       
   611     (**** typedef ****)
       
   612 
       
   613     val _ = warning "defining type...";
       
   614 
       
   615     val (typedefs, thy6) =
       
   616       thy4
       
   617       |> fold_map (fn ((((name, mx), tvs), (cname, U)), name') => fn thy =>
       
   618           Typedef.add_typedef false (SOME (Binding.name name'))
       
   619             (Binding.name name, map fst tvs, mx)
       
   620             (Const ("Collect", (U --> HOLogic.boolT) --> HOLogic.mk_setT U) $
       
   621                Const (cname, U --> HOLogic.boolT)) NONE
       
   622             (rtac exI 1 THEN rtac CollectI 1 THEN
       
   623               QUIET_BREADTH_FIRST (has_fewer_prems 1)
       
   624               (resolve_tac rep_intrs 1)) thy |> (fn ((_, r), thy) =>
       
   625         let
       
   626           val permT = mk_permT
       
   627             (TFree (Name.variant (map fst tvs) "'a", HOLogic.typeS));
       
   628           val pi = Free ("pi", permT);
       
   629           val T = Type (Sign.intern_type thy name, map TFree tvs);
       
   630         in apfst (pair r o hd)
       
   631           (PureThy.add_defs_unchecked true [((Binding.name ("prm_" ^ name ^ "_def"), Logic.mk_equals
       
   632             (Const ("Nominal.perm", permT --> T --> T) $ pi $ Free ("x", T),
       
   633              Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $
       
   634                (Const ("Nominal.perm", permT --> U --> U) $ pi $
       
   635                  (Const (Sign.intern_const thy ("Rep_" ^ name), T --> U) $
       
   636                    Free ("x", T))))), [])] thy)
       
   637         end))
       
   638           (types_syntax ~~ tyvars ~~
       
   639             List.take (rep_set_names'' ~~ recTs', length new_type_names) ~~
       
   640             new_type_names);
       
   641 
       
   642     val perm_defs = map snd typedefs;
       
   643     val Abs_inverse_thms = map (collect_simp o #Abs_inverse o fst) typedefs;
       
   644     val Rep_inverse_thms = map (#Rep_inverse o fst) typedefs;
       
   645     val Rep_thms = map (collect_simp o #Rep o fst) typedefs;
       
   646 
       
   647 
       
   648     (** prove that new types are in class pt_<name> **)
       
   649 
       
   650     val _ = warning "prove that new types are in class pt_<name> ...";
       
   651 
       
   652     fun pt_instance (atom, perm_closed_thms) =
       
   653       fold (fn ((((((Abs_inverse, Rep_inverse), Rep),
       
   654         perm_def), name), tvs), perm_closed) => fn thy =>
       
   655           let
       
   656             val pt_class = pt_class_of thy atom;
       
   657             val sort = Sign.certify_sort thy
       
   658               (pt_class :: map (cp_class_of thy atom) (dt_atoms \ atom))
       
   659           in AxClass.prove_arity
       
   660             (Sign.intern_type thy name,
       
   661               map (inter_sort thy sort o snd) tvs, [pt_class])
       
   662             (EVERY [Class.intro_classes_tac [],
       
   663               rewrite_goals_tac [perm_def],
       
   664               asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1,
       
   665               asm_full_simp_tac (simpset_of thy addsimps
       
   666                 [Rep RS perm_closed RS Abs_inverse]) 1,
       
   667               asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy
       
   668                 ("pt_" ^ Long_Name.base_name atom ^ "3")]) 1]) thy
       
   669           end)
       
   670         (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~
       
   671            new_type_names ~~ tyvars ~~ perm_closed_thms);
       
   672 
       
   673 
       
   674     (** prove that new types are in class cp_<name1>_<name2> **)
       
   675 
       
   676     val _ = warning "prove that new types are in class cp_<name1>_<name2> ...";
       
   677 
       
   678     fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy =
       
   679       let
       
   680         val cp_class = cp_class_of thy atom1 atom2;
       
   681         val sort = Sign.certify_sort thy
       
   682           (pt_class_of thy atom1 :: map (cp_class_of thy atom1) (dt_atoms \ atom1) @
       
   683            (if atom1 = atom2 then [cp_class_of thy atom1 atom1] else
       
   684             pt_class_of thy atom2 :: map (cp_class_of thy atom2) (dt_atoms \ atom2)));
       
   685         val cp1' = cp_inst_of thy atom1 atom2 RS cp1
       
   686       in fold (fn ((((((Abs_inverse, Rep),
       
   687         perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy =>
       
   688           AxClass.prove_arity
       
   689             (Sign.intern_type thy name,
       
   690               map (inter_sort thy sort o snd) tvs, [cp_class])
       
   691             (EVERY [Class.intro_classes_tac [],
       
   692               rewrite_goals_tac [perm_def],
       
   693               asm_full_simp_tac (simpset_of thy addsimps
       
   694                 ((Rep RS perm_closed1 RS Abs_inverse) ::
       
   695                  (if atom1 = atom2 then []
       
   696                   else [Rep RS perm_closed2 RS Abs_inverse]))) 1,
       
   697               cong_tac 1,
       
   698               rtac refl 1,
       
   699               rtac cp1' 1]) thy)
       
   700         (Abs_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ new_type_names ~~
       
   701            tyvars ~~ perm_closed_thms1 ~~ perm_closed_thms2) thy
       
   702       end;
       
   703 
       
   704     val thy7 = fold (fn x => fn thy => thy |>
       
   705       pt_instance x |>
       
   706       fold (cp_instance x) (atoms ~~ perm_closed_thmss))
       
   707         (atoms ~~ perm_closed_thmss) thy6;
       
   708 
       
   709     (**** constructors ****)
       
   710 
       
   711     fun mk_abs_fun (x, t) =
       
   712       let
       
   713         val T = fastype_of x;
       
   714         val U = fastype_of t
       
   715       in
       
   716         Const ("Nominal.abs_fun", T --> U --> T -->
       
   717           Type ("Nominal.noption", [U])) $ x $ t
       
   718       end;
       
   719 
       
   720     val (ty_idxs, _) = List.foldl
       
   721       (fn ((i, ("Nominal.noption", _, _)), p) => p
       
   722         | ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr;
       
   723 
       
   724     fun reindex (DtType (s, dts)) = DtType (s, map reindex dts)
       
   725       | reindex (DtRec i) = DtRec (the (AList.lookup op = ty_idxs i))
       
   726       | reindex dt = dt;
       
   727 
       
   728     fun strip_suffix i s = implode (List.take (explode s, size s - i));
       
   729 
       
   730     (** strips the "_Rep" in type names *)
       
   731     fun strip_nth_name i s =
       
   732       let val xs = Long_Name.explode s;
       
   733       in Long_Name.implode (Library.nth_map (length xs - i) (strip_suffix 4) xs) end;
       
   734 
       
   735     val (descr'', ndescr) = ListPair.unzip (map_filter
       
   736       (fn (i, ("Nominal.noption", _, _)) => NONE
       
   737         | (i, (s, dts, constrs)) =>
       
   738              let
       
   739                val SOME index = AList.lookup op = ty_idxs i;
       
   740                val (constrs2, constrs1) =
       
   741                  map_split (fn (cname, cargs) =>
       
   742                    apsnd (pair (strip_nth_name 2 (strip_nth_name 1 cname)))
       
   743                    (fold_map (fn dt => fn dts =>
       
   744                      let val (dts', dt') = strip_option dt
       
   745                      in ((length dts, length dts'), dts @ dts' @ [reindex dt']) end)
       
   746                        cargs [])) constrs
       
   747              in SOME ((index, (strip_nth_name 1 s,  map reindex dts, constrs1)),
       
   748                (index, constrs2))
       
   749              end) descr);
       
   750 
       
   751     val (descr1, descr2) = chop (length new_type_names) descr'';
       
   752     val descr' = [descr1, descr2];
       
   753 
       
   754     fun partition_cargs idxs xs = map (fn (i, j) =>
       
   755       (List.take (List.drop (xs, i), j), List.nth (xs, i + j))) idxs;
       
   756 
       
   757     val pdescr = map (fn ((i, (s, dts, constrs)), (_, idxss)) => (i, (s, dts,
       
   758       map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs))
       
   759         (constrs ~~ idxss)))) (descr'' ~~ ndescr);
       
   760 
       
   761     fun nth_dtyp' i = typ_of_dtyp descr'' sorts (DtRec i);
       
   762 
       
   763     val rep_names = map (fn s =>
       
   764       Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names;
       
   765     val abs_names = map (fn s =>
       
   766       Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names;
       
   767 
       
   768     val recTs = get_rec_types descr'' sorts;
       
   769     val newTs' = Library.take (length new_type_names, recTs');
       
   770     val newTs = Library.take (length new_type_names, recTs);
       
   771 
       
   772     val full_new_type_names = map (Sign.full_bname thy) new_type_names;
       
   773 
       
   774     fun make_constr_def tname T T' ((thy, defs, eqns),
       
   775         (((cname_rep, _), (cname, cargs)), (cname', mx))) =
       
   776       let
       
   777         fun constr_arg ((dts, dt), (j, l_args, r_args)) =
       
   778           let
       
   779             val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp descr'' sorts dt) i)
       
   780               (dts ~~ (j upto j + length dts - 1))
       
   781             val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
       
   782           in
       
   783             (j + length dts + 1,
       
   784              xs @ x :: l_args,
       
   785              List.foldr mk_abs_fun
       
   786                (case dt of
       
   787                   DtRec k => if k < length new_type_names then
       
   788                       Const (List.nth (rep_names, k), typ_of_dtyp descr'' sorts dt -->
       
   789                         typ_of_dtyp descr sorts dt) $ x
       
   790                     else error "nested recursion not (yet) supported"
       
   791                 | _ => x) xs :: r_args)
       
   792           end
       
   793 
       
   794         val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) cargs;
       
   795         val abs_name = Sign.intern_const thy ("Abs_" ^ tname);
       
   796         val rep_name = Sign.intern_const thy ("Rep_" ^ tname);
       
   797         val constrT = map fastype_of l_args ---> T;
       
   798         val lhs = list_comb (Const (cname, constrT), l_args);
       
   799         val rhs = list_comb (Const (cname_rep, map fastype_of r_args ---> T'), r_args);
       
   800         val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs);
       
   801         val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
       
   802           (Const (rep_name, T --> T') $ lhs, rhs));
       
   803         val def_name = (Long_Name.base_name cname) ^ "_def";
       
   804         val ([def_thm], thy') = thy |>
       
   805           Sign.add_consts_i [(Binding.name cname', constrT, mx)] |>
       
   806           (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)]
       
   807       in (thy', defs @ [def_thm], eqns @ [eqn]) end;
       
   808 
       
   809     fun dt_constr_defs ((thy, defs, eqns, dist_lemmas), ((((((_, (_, _, constrs)),
       
   810         (_, (_, _, constrs'))), tname), T), T'), constr_syntax)) =
       
   811       let
       
   812         val rep_const = cterm_of thy
       
   813           (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T'));
       
   814         val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
       
   815         val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T')
       
   816           ((Sign.add_path tname thy, defs, []), constrs ~~ constrs' ~~ constr_syntax)
       
   817       in
       
   818         (parent_path (#flat_names config) thy', defs', eqns @ [eqns'], dist_lemmas @ [dist])
       
   819       end;
       
   820 
       
   821     val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs
       
   822       ((thy7, [], [], []), List.take (descr, length new_type_names) ~~
       
   823         List.take (pdescr, length new_type_names) ~~
       
   824         new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax);
       
   825 
       
   826     val abs_inject_thms = map (collect_simp o #Abs_inject o fst) typedefs
       
   827     val rep_inject_thms = map (#Rep_inject o fst) typedefs
       
   828 
       
   829     (* prove theorem  Rep_i (Constr_j ...) = Constr'_j ...  *)
       
   830 
       
   831     fun prove_constr_rep_thm eqn =
       
   832       let
       
   833         val inj_thms = map (fn r => r RS iffD1) abs_inject_thms;
       
   834         val rewrites = constr_defs @ map mk_meta_eq Rep_inverse_thms
       
   835       in Goal.prove_global thy8 [] [] eqn (fn _ => EVERY
       
   836         [resolve_tac inj_thms 1,
       
   837          rewrite_goals_tac rewrites,
       
   838          rtac refl 3,
       
   839          resolve_tac rep_intrs 2,
       
   840          REPEAT (resolve_tac Rep_thms 1)])
       
   841       end;
       
   842 
       
   843     val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns;
       
   844 
       
   845     (* prove theorem  pi \<bullet> Rep_i x = Rep_i (pi \<bullet> x) *)
       
   846 
       
   847     fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th =>
       
   848       let
       
   849         val _ $ (_ $ (Rep $ x)) = Logic.unvarify (prop_of th);
       
   850         val Type ("fun", [T, U]) = fastype_of Rep;
       
   851         val permT = mk_permT (Type (atom, []));
       
   852         val pi = Free ("pi", permT);
       
   853       in
       
   854         Goal.prove_global thy8 [] []
       
   855           (augment_sort thy8
       
   856             (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom))
       
   857             (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
   858               (Const ("Nominal.perm", permT --> U --> U) $ pi $ (Rep $ x),
       
   859                Rep $ (Const ("Nominal.perm", permT --> T --> T) $ pi $ x)))))
       
   860           (fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @
       
   861             perm_closed_thms @ Rep_thms)) 1)
       
   862       end) Rep_thms;
       
   863 
       
   864     val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm
       
   865       (atoms ~~ perm_closed_thmss));
       
   866 
       
   867     (* prove distinctness theorems *)
       
   868 
       
   869     val distinct_props = DatatypeProp.make_distincts descr' sorts;
       
   870     val dist_rewrites = map2 (fn rep_thms => fn dist_lemma =>
       
   871       dist_lemma :: rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])
       
   872         constr_rep_thmss dist_lemmas;
       
   873 
       
   874     fun prove_distinct_thms _ (_, []) = []
       
   875       | prove_distinct_thms (p as (rep_thms, dist_lemma)) (k, t :: ts) =
       
   876           let
       
   877             val dist_thm = Goal.prove_global thy8 [] [] t (fn _ =>
       
   878               simp_tac (simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1)
       
   879           in dist_thm :: standard (dist_thm RS not_sym) ::
       
   880             prove_distinct_thms p (k, ts)
       
   881           end;
       
   882 
       
   883     val distinct_thms = map2 prove_distinct_thms
       
   884       (constr_rep_thmss ~~ dist_lemmas) distinct_props;
       
   885 
       
   886     (** prove equations for permutation functions **)
       
   887 
       
   888     val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
       
   889       let val T = nth_dtyp' i
       
   890       in List.concat (map (fn (atom, perm_closed_thms) =>
       
   891           map (fn ((cname, dts), constr_rep_thm) =>
       
   892         let
       
   893           val cname = Sign.intern_const thy8
       
   894             (Long_Name.append tname (Long_Name.base_name cname));
       
   895           val permT = mk_permT (Type (atom, []));
       
   896           val pi = Free ("pi", permT);
       
   897 
       
   898           fun perm t =
       
   899             let val T = fastype_of t
       
   900             in Const ("Nominal.perm", permT --> T --> T) $ pi $ t end;
       
   901 
       
   902           fun constr_arg ((dts, dt), (j, l_args, r_args)) =
       
   903             let
       
   904               val Ts = map (typ_of_dtyp descr'' sorts) dts;
       
   905               val xs = map (fn (T, i) => mk_Free "x" T i)
       
   906                 (Ts ~~ (j upto j + length dts - 1))
       
   907               val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
       
   908             in
       
   909               (j + length dts + 1,
       
   910                xs @ x :: l_args,
       
   911                map perm (xs @ [x]) @ r_args)
       
   912             end
       
   913 
       
   914           val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) dts;
       
   915           val c = Const (cname, map fastype_of l_args ---> T)
       
   916         in
       
   917           Goal.prove_global thy8 [] []
       
   918             (augment_sort thy8
       
   919               (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom))
       
   920               (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
   921                 (perm (list_comb (c, l_args)), list_comb (c, r_args)))))
       
   922             (fn _ => EVERY
       
   923               [simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1,
       
   924                simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @
       
   925                  constr_defs @ perm_closed_thms)) 1,
       
   926                TRY (simp_tac (HOL_basic_ss addsimps
       
   927                  (symmetric perm_fun_def :: abs_perm)) 1),
       
   928                TRY (simp_tac (HOL_basic_ss addsimps
       
   929                  (perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @
       
   930                     perm_closed_thms)) 1)])
       
   931         end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss))
       
   932       end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
       
   933 
       
   934     (** prove injectivity of constructors **)
       
   935 
       
   936     val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms;
       
   937     val alpha = PureThy.get_thms thy8 "alpha";
       
   938     val abs_fresh = PureThy.get_thms thy8 "abs_fresh";
       
   939 
       
   940     val pt_cp_sort =
       
   941       map (pt_class_of thy8) dt_atoms @
       
   942       maps (fn s => map (cp_class_of thy8 s) (dt_atoms \ s)) dt_atoms;
       
   943 
       
   944     val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
       
   945       let val T = nth_dtyp' i
       
   946       in List.mapPartial (fn ((cname, dts), constr_rep_thm) =>
       
   947         if null dts then NONE else SOME
       
   948         let
       
   949           val cname = Sign.intern_const thy8
       
   950             (Long_Name.append tname (Long_Name.base_name cname));
       
   951 
       
   952           fun make_inj ((dts, dt), (j, args1, args2, eqs)) =
       
   953             let
       
   954               val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1);
       
   955               val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
       
   956               val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx;
       
   957               val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts);
       
   958               val y = mk_Free "y" (typ_of_dtyp descr'' sorts dt) (j + length dts)
       
   959             in
       
   960               (j + length dts + 1,
       
   961                xs @ (x :: args1), ys @ (y :: args2),
       
   962                HOLogic.mk_eq
       
   963                  (List.foldr mk_abs_fun x xs, List.foldr mk_abs_fun y ys) :: eqs)
       
   964             end;
       
   965 
       
   966           val (_, args1, args2, eqs) = List.foldr make_inj (1, [], [], []) dts;
       
   967           val Ts = map fastype_of args1;
       
   968           val c = Const (cname, Ts ---> T)
       
   969         in
       
   970           Goal.prove_global thy8 [] []
       
   971             (augment_sort thy8 pt_cp_sort
       
   972               (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
   973                 (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)),
       
   974                  foldr1 HOLogic.mk_conj eqs))))
       
   975             (fn _ => EVERY
       
   976                [asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm ::
       
   977                   rep_inject_thms')) 1,
       
   978                 TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def ::
       
   979                   alpha @ abs_perm @ abs_fresh @ rep_inject_thms @
       
   980                   perm_rep_perm_thms)) 1)])
       
   981         end) (constrs ~~ constr_rep_thms)
       
   982       end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
       
   983 
       
   984     (** equations for support and freshness **)
       
   985 
       
   986     val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip
       
   987       (map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') =>
       
   988       let val T = nth_dtyp' i
       
   989       in List.concat (map (fn (cname, dts) => map (fn atom =>
       
   990         let
       
   991           val cname = Sign.intern_const thy8
       
   992             (Long_Name.append tname (Long_Name.base_name cname));
       
   993           val atomT = Type (atom, []);
       
   994 
       
   995           fun process_constr ((dts, dt), (j, args1, args2)) =
       
   996             let
       
   997               val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1);
       
   998               val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
       
   999               val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
       
  1000             in
       
  1001               (j + length dts + 1,
       
  1002                xs @ (x :: args1), List.foldr mk_abs_fun x xs :: args2)
       
  1003             end;
       
  1004 
       
  1005           val (_, args1, args2) = List.foldr process_constr (1, [], []) dts;
       
  1006           val Ts = map fastype_of args1;
       
  1007           val c = list_comb (Const (cname, Ts ---> T), args1);
       
  1008           fun supp t =
       
  1009             Const ("Nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t;
       
  1010           fun fresh t = fresh_const atomT (fastype_of t) $ Free ("a", atomT) $ t;
       
  1011           val supp_thm = Goal.prove_global thy8 [] []
       
  1012             (augment_sort thy8 pt_cp_sort
       
  1013               (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
  1014                 (supp c,
       
  1015                  if null dts then HOLogic.mk_set atomT []
       
  1016                  else foldr1 (HOLogic.mk_binop @{const_name Un}) (map supp args2)))))
       
  1017             (fn _ =>
       
  1018               simp_tac (HOL_basic_ss addsimps (supp_def ::
       
  1019                  Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un ::
       
  1020                  symmetric empty_def :: finite_emptyI :: simp_thms @
       
  1021                  abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1)
       
  1022         in
       
  1023           (supp_thm,
       
  1024            Goal.prove_global thy8 [] [] (augment_sort thy8 pt_cp_sort
       
  1025              (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
  1026                (fresh c,
       
  1027                 if null dts then HOLogic.true_const
       
  1028                 else foldr1 HOLogic.mk_conj (map fresh args2)))))
       
  1029              (fn _ =>
       
  1030                simp_tac (HOL_ss addsimps [Un_iff, empty_iff, fresh_def, supp_thm]) 1))
       
  1031         end) atoms) constrs)
       
  1032       end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps')));
       
  1033 
       
  1034     (**** weak induction theorem ****)
       
  1035 
       
  1036     fun mk_indrule_lemma ((prems, concls), (((i, _), T), U)) =
       
  1037       let
       
  1038         val Rep_t = Const (List.nth (rep_names, i), T --> U) $
       
  1039           mk_Free "x" T i;
       
  1040 
       
  1041         val Abs_t =  Const (List.nth (abs_names, i), U --> T)
       
  1042 
       
  1043       in (prems @ [HOLogic.imp $
       
  1044             (Const (List.nth (rep_set_names'', i), U --> HOLogic.boolT) $ Rep_t) $
       
  1045               (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
       
  1046           concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
       
  1047       end;
       
  1048 
       
  1049     val (indrule_lemma_prems, indrule_lemma_concls) =
       
  1050       Library.foldl mk_indrule_lemma (([], []), (descr'' ~~ recTs ~~ recTs'));
       
  1051 
       
  1052     val indrule_lemma = Goal.prove_global thy8 [] []
       
  1053       (Logic.mk_implies
       
  1054         (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
       
  1055          HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY
       
  1056            [REPEAT (etac conjE 1),
       
  1057             REPEAT (EVERY
       
  1058               [TRY (rtac conjI 1), full_simp_tac (HOL_basic_ss addsimps Rep_inverse_thms) 1,
       
  1059                etac mp 1, resolve_tac Rep_thms 1])]);
       
  1060 
       
  1061     val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
       
  1062     val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
       
  1063       map (Free o apfst fst o dest_Var) Ps;
       
  1064     val indrule_lemma' = cterm_instantiate
       
  1065       (map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma;
       
  1066 
       
  1067     val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms;
       
  1068 
       
  1069     val dt_induct_prop = DatatypeProp.make_ind descr' sorts;
       
  1070     val dt_induct = Goal.prove_global thy8 []
       
  1071       (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
       
  1072       (fn {prems, ...} => EVERY
       
  1073         [rtac indrule_lemma' 1,
       
  1074          (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
       
  1075          EVERY (map (fn (prem, r) => (EVERY
       
  1076            [REPEAT (eresolve_tac Abs_inverse_thms' 1),
       
  1077             simp_tac (HOL_basic_ss addsimps [symmetric r]) 1,
       
  1078             DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
       
  1079                 (prems ~~ constr_defs))]);
       
  1080 
       
  1081     val case_names_induct = mk_case_names_induct descr'';
       
  1082 
       
  1083     (**** prove that new datatypes have finite support ****)
       
  1084 
       
  1085     val _ = warning "proving finite support for the new datatype";
       
  1086 
       
  1087     val indnames = DatatypeProp.make_tnames recTs;
       
  1088 
       
  1089     val abs_supp = PureThy.get_thms thy8 "abs_supp";
       
  1090     val supp_atm = PureThy.get_thms thy8 "supp_atm";
       
  1091 
       
  1092     val finite_supp_thms = map (fn atom =>
       
  1093       let val atomT = Type (atom, [])
       
  1094       in map standard (List.take
       
  1095         (split_conj_thm (Goal.prove_global thy8 [] []
       
  1096            (augment_sort thy8 (fs_class_of thy8 atom :: pt_cp_sort)
       
  1097              (HOLogic.mk_Trueprop
       
  1098                (foldr1 HOLogic.mk_conj (map (fn (s, T) =>
       
  1099                  Const ("Finite_Set.finite", HOLogic.mk_setT atomT --> HOLogic.boolT) $
       
  1100                    (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T)))
       
  1101                    (indnames ~~ recTs)))))
       
  1102            (fn _ => indtac dt_induct indnames 1 THEN
       
  1103             ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps
       
  1104               (abs_supp @ supp_atm @
       
  1105                PureThy.get_thms thy8 ("fs_" ^ Long_Name.base_name atom ^ "1") @
       
  1106                List.concat supp_thms))))),
       
  1107          length new_type_names))
       
  1108       end) atoms;
       
  1109 
       
  1110     val simp_atts = replicate (length new_type_names) [Simplifier.simp_add];
       
  1111 
       
  1112 	(* Function to add both the simp and eqvt attributes *)
       
  1113         (* These two attributes are duplicated on all the types in the mutual nominal datatypes *)
       
  1114 
       
  1115     val simp_eqvt_atts = replicate (length new_type_names) [Simplifier.simp_add, NominalThmDecls.eqvt_add];
       
  1116  
       
  1117     val (_, thy9) = thy8 |>
       
  1118       Sign.add_path big_name |>
       
  1119       PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] ||>>
       
  1120       PureThy.add_thmss [((Binding.name "inducts", projections dt_induct), [case_names_induct])] ||>
       
  1121       Sign.parent_path ||>>
       
  1122       DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>>
       
  1123       DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>>
       
  1124       DatatypeAux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>>
       
  1125       DatatypeAux.store_thmss "inject" new_type_names inject_thms ||>>
       
  1126       DatatypeAux.store_thmss "supp" new_type_names supp_thms ||>>
       
  1127       DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||>
       
  1128       fold (fn (atom, ths) => fn thy =>
       
  1129         let
       
  1130           val class = fs_class_of thy atom;
       
  1131           val sort = Sign.certify_sort thy (class :: pt_cp_sort)
       
  1132         in fold (fn Type (s, Ts) => AxClass.prove_arity
       
  1133           (s, map (inter_sort thy sort o snd o dest_TFree) Ts, [class])
       
  1134           (Class.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy
       
  1135         end) (atoms ~~ finite_supp_thms);
       
  1136 
       
  1137     (**** strong induction theorem ****)
       
  1138 
       
  1139     val pnames = if length descr'' = 1 then ["P"]
       
  1140       else map (fn i => "P" ^ string_of_int i) (1 upto length descr'');
       
  1141     val ind_sort = if null dt_atomTs then HOLogic.typeS
       
  1142       else Sign.certify_sort thy9 (map (fs_class_of thy9) dt_atoms);
       
  1143     val fsT = TFree ("'n", ind_sort);
       
  1144     val fsT' = TFree ("'n", HOLogic.typeS);
       
  1145 
       
  1146     val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T)))
       
  1147       (DatatypeProp.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs);
       
  1148 
       
  1149     fun make_pred fsT i T =
       
  1150       Free (List.nth (pnames, i), fsT --> T --> HOLogic.boolT);
       
  1151 
       
  1152     fun mk_fresh1 xs [] = []
       
  1153       | mk_fresh1 xs ((y as (_, T)) :: ys) = map (fn x => HOLogic.mk_Trueprop
       
  1154             (HOLogic.mk_not (HOLogic.mk_eq (Free y, Free x))))
       
  1155               (filter (fn (_, U) => T = U) (rev xs)) @
       
  1156           mk_fresh1 (y :: xs) ys;
       
  1157 
       
  1158     fun mk_fresh2 xss [] = []
       
  1159       | mk_fresh2 xss ((p as (ys, _)) :: yss) = List.concat (map (fn y as (_, T) =>
       
  1160             map (fn (_, x as (_, U)) => HOLogic.mk_Trueprop
       
  1161               (fresh_const T U $ Free y $ Free x)) (rev xss @ yss)) ys) @
       
  1162           mk_fresh2 (p :: xss) yss;
       
  1163 
       
  1164     fun make_ind_prem fsT f k T ((cname, cargs), idxs) =
       
  1165       let
       
  1166         val recs = List.filter is_rec_type cargs;
       
  1167         val Ts = map (typ_of_dtyp descr'' sorts) cargs;
       
  1168         val recTs' = map (typ_of_dtyp descr'' sorts) recs;
       
  1169         val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts);
       
  1170         val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
       
  1171         val frees = tnames ~~ Ts;
       
  1172         val frees' = partition_cargs idxs frees;
       
  1173         val z = (Name.variant tnames "z", fsT);
       
  1174 
       
  1175         fun mk_prem ((dt, s), T) =
       
  1176           let
       
  1177             val (Us, U) = strip_type T;
       
  1178             val l = length Us
       
  1179           in list_all (z :: map (pair "x") Us, HOLogic.mk_Trueprop
       
  1180             (make_pred fsT (body_index dt) U $ Bound l $ app_bnds (Free (s, T)) l))
       
  1181           end;
       
  1182 
       
  1183         val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
       
  1184         val prems' = map (fn p as (_, T) => HOLogic.mk_Trueprop
       
  1185             (f T (Free p) (Free z))) (List.concat (map fst frees')) @
       
  1186           mk_fresh1 [] (List.concat (map fst frees')) @
       
  1187           mk_fresh2 [] frees'
       
  1188 
       
  1189       in list_all_free (frees @ [z], Logic.list_implies (prems' @ prems,
       
  1190         HOLogic.mk_Trueprop (make_pred fsT k T $ Free z $
       
  1191           list_comb (Const (cname, Ts ---> T), map Free frees))))
       
  1192       end;
       
  1193 
       
  1194     val ind_prems = List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
       
  1195       map (make_ind_prem fsT (fn T => fn t => fn u =>
       
  1196         fresh_const T fsT $ t $ u) i T)
       
  1197           (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs));
       
  1198     val tnames = DatatypeProp.make_tnames recTs;
       
  1199     val zs = Name.variant_list tnames (replicate (length descr'') "z");
       
  1200     val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
       
  1201       (map (fn ((((i, _), T), tname), z) =>
       
  1202         make_pred fsT i T $ Free (z, fsT) $ Free (tname, T))
       
  1203         (descr'' ~~ recTs ~~ tnames ~~ zs)));
       
  1204     val induct = Logic.list_implies (ind_prems, ind_concl);
       
  1205 
       
  1206     val ind_prems' =
       
  1207       map (fn (_, f as Free (_, T)) => list_all_free ([("x", fsT')],
       
  1208         HOLogic.mk_Trueprop (Const ("Finite_Set.finite",
       
  1209           (snd (split_last (binder_types T)) --> HOLogic.boolT) -->
       
  1210             HOLogic.boolT) $ (f $ Free ("x", fsT'))))) fresh_fs @
       
  1211       List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
       
  1212         map (make_ind_prem fsT' (fn T => fn t => fn u => HOLogic.Not $
       
  1213           HOLogic.mk_mem (t, the (AList.lookup op = fresh_fs T) $ u)) i T)
       
  1214             (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs));
       
  1215     val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
       
  1216       (map (fn ((((i, _), T), tname), z) =>
       
  1217         make_pred fsT' i T $ Free (z, fsT') $ Free (tname, T))
       
  1218         (descr'' ~~ recTs ~~ tnames ~~ zs)));
       
  1219     val induct' = Logic.list_implies (ind_prems', ind_concl');
       
  1220 
       
  1221     val aux_ind_vars =
       
  1222       (DatatypeProp.indexify_names (replicate (length dt_atomTs) "pi") ~~
       
  1223        map mk_permT dt_atomTs) @ [("z", fsT')];
       
  1224     val aux_ind_Ts = rev (map snd aux_ind_vars);
       
  1225     val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
       
  1226       (map (fn (((i, _), T), tname) =>
       
  1227         HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $
       
  1228           fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1))
       
  1229             (Free (tname, T))))
       
  1230         (descr'' ~~ recTs ~~ tnames)));
       
  1231 
       
  1232     val fin_set_supp = map (fn s =>
       
  1233       at_inst_of thy9 s RS at_fin_set_supp) dt_atoms;
       
  1234     val fin_set_fresh = map (fn s =>
       
  1235       at_inst_of thy9 s RS at_fin_set_fresh) dt_atoms;
       
  1236     val pt1_atoms = map (fn Type (s, _) =>
       
  1237       PureThy.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "1")) dt_atomTs;
       
  1238     val pt2_atoms = map (fn Type (s, _) =>
       
  1239       PureThy.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "2") RS sym) dt_atomTs;
       
  1240     val exists_fresh' = PureThy.get_thms thy9 "exists_fresh'";
       
  1241     val fs_atoms = PureThy.get_thms thy9 "fin_supp";
       
  1242     val abs_supp = PureThy.get_thms thy9 "abs_supp";
       
  1243     val perm_fresh_fresh = PureThy.get_thms thy9 "perm_fresh_fresh";
       
  1244     val calc_atm = PureThy.get_thms thy9 "calc_atm";
       
  1245     val fresh_atm = PureThy.get_thms thy9 "fresh_atm";
       
  1246     val fresh_left = PureThy.get_thms thy9 "fresh_left";
       
  1247     val perm_swap = PureThy.get_thms thy9 "perm_swap";
       
  1248 
       
  1249     fun obtain_fresh_name' ths ts T (freshs1, freshs2, ctxt) =
       
  1250       let
       
  1251         val p = foldr1 HOLogic.mk_prod (ts @ freshs1);
       
  1252         val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
       
  1253             (HOLogic.exists_const T $ Abs ("x", T,
       
  1254               fresh_const T (fastype_of p) $
       
  1255                 Bound 0 $ p)))
       
  1256           (fn _ => EVERY
       
  1257             [resolve_tac exists_fresh' 1,
       
  1258              simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms @
       
  1259                fin_set_supp @ ths)) 1]);
       
  1260         val (([cx], ths), ctxt') = Obtain.result
       
  1261           (fn _ => EVERY
       
  1262             [etac exE 1,
       
  1263              full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
       
  1264              REPEAT (etac conjE 1)])
       
  1265           [ex] ctxt
       
  1266       in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
       
  1267 
       
  1268     fun fresh_fresh_inst thy a b =
       
  1269       let
       
  1270         val T = fastype_of a;
       
  1271         val SOME th = find_first (fn th => case prop_of th of
       
  1272             _ $ (_ $ (Const (_, Type (_, [U, _])) $ _ $ _)) $ _ => U = T
       
  1273           | _ => false) perm_fresh_fresh
       
  1274       in
       
  1275         Drule.instantiate' []
       
  1276           [SOME (cterm_of thy a), NONE, SOME (cterm_of thy b)] th
       
  1277       end;
       
  1278 
       
  1279     val fs_cp_sort =
       
  1280       map (fs_class_of thy9) dt_atoms @
       
  1281       maps (fn s => map (cp_class_of thy9 s) (dt_atoms \ s)) dt_atoms;
       
  1282 
       
  1283     (**********************************************************************
       
  1284       The subgoals occurring in the proof of induct_aux have the
       
  1285       following parameters:
       
  1286 
       
  1287         x_1 ... x_k p_1 ... p_m z
       
  1288 
       
  1289       where
       
  1290 
       
  1291         x_i : constructor arguments (introduced by weak induction rule)
       
  1292         p_i : permutations (one for each atom type in the data type)
       
  1293         z   : freshness context
       
  1294     ***********************************************************************)
       
  1295 
       
  1296     val _ = warning "proving strong induction theorem ...";
       
  1297 
       
  1298     val induct_aux = Goal.prove_global thy9 []
       
  1299         (map (augment_sort thy9 fs_cp_sort) ind_prems')
       
  1300         (augment_sort thy9 fs_cp_sort ind_concl') (fn {prems, context} =>
       
  1301       let
       
  1302         val (prems1, prems2) = chop (length dt_atomTs) prems;
       
  1303         val ind_ss2 = HOL_ss addsimps
       
  1304           finite_Diff :: abs_fresh @ abs_supp @ fs_atoms;
       
  1305         val ind_ss1 = ind_ss2 addsimps fresh_left @ calc_atm @
       
  1306           fresh_atm @ rev_simps @ app_simps;
       
  1307         val ind_ss3 = HOL_ss addsimps abs_fun_eq1 ::
       
  1308           abs_perm @ calc_atm @ perm_swap;
       
  1309         val ind_ss4 = HOL_basic_ss addsimps fresh_left @ prems1 @
       
  1310           fin_set_fresh @ calc_atm;
       
  1311         val ind_ss5 = HOL_basic_ss addsimps pt1_atoms;
       
  1312         val ind_ss6 = HOL_basic_ss addsimps flat perm_simps';
       
  1313         val th = Goal.prove context [] []
       
  1314           (augment_sort thy9 fs_cp_sort aux_ind_concl)
       
  1315           (fn {context = context1, ...} =>
       
  1316              EVERY (indtac dt_induct tnames 1 ::
       
  1317                maps (fn ((_, (_, _, constrs)), (_, constrs')) =>
       
  1318                  map (fn ((cname, cargs), is) =>
       
  1319                    REPEAT (rtac allI 1) THEN
       
  1320                    SUBPROOF (fn {prems = iprems, params, concl,
       
  1321                        context = context2, ...} =>
       
  1322                      let
       
  1323                        val concl' = term_of concl;
       
  1324                        val _ $ (_ $ _ $ u) = concl';
       
  1325                        val U = fastype_of u;
       
  1326                        val (xs, params') =
       
  1327                          chop (length cargs) (map term_of params);
       
  1328                        val Ts = map fastype_of xs;
       
  1329                        val cnstr = Const (cname, Ts ---> U);
       
  1330                        val (pis, z) = split_last params';
       
  1331                        val mk_pi = fold_rev (mk_perm []) pis;
       
  1332                        val xs' = partition_cargs is xs;
       
  1333                        val xs'' = map (fn (ts, u) => (map mk_pi ts, mk_pi u)) xs';
       
  1334                        val ts = maps (fn (ts, u) => ts @ [u]) xs'';
       
  1335                        val (freshs1, freshs2, context3) = fold (fn t =>
       
  1336                          let val T = fastype_of t
       
  1337                          in obtain_fresh_name' prems1
       
  1338                            (the (AList.lookup op = fresh_fs T) $ z :: ts) T
       
  1339                          end) (maps fst xs') ([], [], context2);
       
  1340                        val freshs1' = unflat (map fst xs') freshs1;
       
  1341                        val freshs2' = map (Simplifier.simplify ind_ss4)
       
  1342                          (mk_not_sym freshs2);
       
  1343                        val ind_ss1' = ind_ss1 addsimps freshs2';
       
  1344                        val ind_ss3' = ind_ss3 addsimps freshs2';
       
  1345                        val rename_eq =
       
  1346                          if forall (null o fst) xs' then []
       
  1347                          else [Goal.prove context3 [] []
       
  1348                            (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
  1349                              (list_comb (cnstr, ts),
       
  1350                               list_comb (cnstr, maps (fn ((bs, t), cs) =>
       
  1351                                 cs @ [fold_rev (mk_perm []) (map perm_of_pair
       
  1352                                   (bs ~~ cs)) t]) (xs'' ~~ freshs1')))))
       
  1353                            (fn _ => EVERY
       
  1354                               (simp_tac (HOL_ss addsimps flat inject_thms) 1 ::
       
  1355                                REPEAT (FIRSTGOAL (rtac conjI)) ::
       
  1356                                maps (fn ((bs, t), cs) =>
       
  1357                                  if null bs then []
       
  1358                                  else rtac sym 1 :: maps (fn (b, c) =>
       
  1359                                    [rtac trans 1, rtac sym 1,
       
  1360                                     rtac (fresh_fresh_inst thy9 b c) 1,
       
  1361                                     simp_tac ind_ss1' 1,
       
  1362                                     simp_tac ind_ss2 1,
       
  1363                                     simp_tac ind_ss3' 1]) (bs ~~ cs))
       
  1364                                  (xs'' ~~ freshs1')))];
       
  1365                        val th = Goal.prove context3 [] [] concl' (fn _ => EVERY
       
  1366                          [simp_tac (ind_ss6 addsimps rename_eq) 1,
       
  1367                           cut_facts_tac iprems 1,
       
  1368                           (resolve_tac prems THEN_ALL_NEW
       
  1369                             SUBGOAL (fn (t, i) => case Logic.strip_assums_concl t of
       
  1370                                 _ $ (Const ("Nominal.fresh", _) $ _ $ _) =>
       
  1371                                   simp_tac ind_ss1' i
       
  1372                               | _ $ (Const ("Not", _) $ _) =>
       
  1373                                   resolve_tac freshs2' i
       
  1374                               | _ => asm_simp_tac (HOL_basic_ss addsimps
       
  1375                                   pt2_atoms addsimprocs [perm_simproc]) i)) 1])
       
  1376                        val final = ProofContext.export context3 context2 [th]
       
  1377                      in
       
  1378                        resolve_tac final 1
       
  1379                      end) context1 1) (constrs ~~ constrs')) (descr'' ~~ ndescr)))
       
  1380       in
       
  1381         EVERY
       
  1382           [cut_facts_tac [th] 1,
       
  1383            REPEAT (eresolve_tac [conjE, @{thm allE_Nil}] 1),
       
  1384            REPEAT (etac allE 1),
       
  1385            REPEAT (TRY (rtac conjI 1) THEN asm_full_simp_tac ind_ss5 1)]
       
  1386       end);
       
  1387 
       
  1388     val induct_aux' = Thm.instantiate ([],
       
  1389       map (fn (s, v as Var (_, T)) =>
       
  1390         (cterm_of thy9 v, cterm_of thy9 (Free (s, T))))
       
  1391           (pnames ~~ map head_of (HOLogic.dest_conj
       
  1392              (HOLogic.dest_Trueprop (concl_of induct_aux)))) @
       
  1393       map (fn (_, f) =>
       
  1394         let val f' = Logic.varify f
       
  1395         in (cterm_of thy9 f',
       
  1396           cterm_of thy9 (Const ("Nominal.supp", fastype_of f')))
       
  1397         end) fresh_fs) induct_aux;
       
  1398 
       
  1399     val induct = Goal.prove_global thy9 []
       
  1400       (map (augment_sort thy9 fs_cp_sort) ind_prems)
       
  1401       (augment_sort thy9 fs_cp_sort ind_concl)
       
  1402       (fn {prems, ...} => EVERY
       
  1403          [rtac induct_aux' 1,
       
  1404           REPEAT (resolve_tac fs_atoms 1),
       
  1405           REPEAT ((resolve_tac prems THEN_ALL_NEW
       
  1406             (etac meta_spec ORELSE' full_simp_tac (HOL_basic_ss addsimps [fresh_def]))) 1)])
       
  1407 
       
  1408     val (_, thy10) = thy9 |>
       
  1409       Sign.add_path big_name |>
       
  1410       PureThy.add_thms [((Binding.name "strong_induct'", induct_aux), [])] ||>>
       
  1411       PureThy.add_thms [((Binding.name "strong_induct", induct), [case_names_induct])] ||>>
       
  1412       PureThy.add_thmss [((Binding.name "strong_inducts", projections induct), [case_names_induct])];
       
  1413 
       
  1414     (**** recursion combinator ****)
       
  1415 
       
  1416     val _ = warning "defining recursion combinator ...";
       
  1417 
       
  1418     val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
       
  1419 
       
  1420     val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts used;
       
  1421 
       
  1422     val rec_sort = if null dt_atomTs then HOLogic.typeS else
       
  1423       Sign.certify_sort thy10 pt_cp_sort;
       
  1424 
       
  1425     val rec_result_Ts = map (fn TFree (s, _) => TFree (s, rec_sort)) rec_result_Ts';
       
  1426     val rec_fn_Ts = map (typ_subst_atomic (rec_result_Ts' ~~ rec_result_Ts)) rec_fn_Ts';
       
  1427 
       
  1428     val rec_set_Ts = map (fn (T1, T2) =>
       
  1429       rec_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
       
  1430 
       
  1431     val big_rec_name = big_name ^ "_rec_set";
       
  1432     val rec_set_names' =
       
  1433       if length descr'' = 1 then [big_rec_name] else
       
  1434         map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
       
  1435           (1 upto (length descr''));
       
  1436     val rec_set_names =  map (Sign.full_bname thy10) rec_set_names';
       
  1437 
       
  1438     val rec_fns = map (uncurry (mk_Free "f"))
       
  1439       (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
       
  1440     val rec_sets' = map (fn c => list_comb (Free c, rec_fns))
       
  1441       (rec_set_names' ~~ rec_set_Ts);
       
  1442     val rec_sets = map (fn c => list_comb (Const c, rec_fns))
       
  1443       (rec_set_names ~~ rec_set_Ts);
       
  1444 
       
  1445     (* introduction rules for graph of recursion function *)
       
  1446 
       
  1447     val rec_preds = map (fn (a, T) =>
       
  1448       Free (a, T --> HOLogic.boolT)) (pnames ~~ rec_result_Ts);
       
  1449 
       
  1450     fun mk_fresh3 rs [] = []
       
  1451       | mk_fresh3 rs ((p as (ys, z)) :: yss) = List.concat (map (fn y as (_, T) =>
       
  1452             List.mapPartial (fn ((_, (_, x)), r as (_, U)) => if z = x then NONE
       
  1453               else SOME (HOLogic.mk_Trueprop
       
  1454                 (fresh_const T U $ Free y $ Free r))) rs) ys) @
       
  1455           mk_fresh3 rs yss;
       
  1456 
       
  1457     (* FIXME: avoid collisions with other variable names? *)
       
  1458     val rec_ctxt = Free ("z", fsT');
       
  1459 
       
  1460     fun make_rec_intr T p rec_set ((rec_intr_ts, rec_prems, rec_prems',
       
  1461           rec_eq_prems, l), ((cname, cargs), idxs)) =
       
  1462       let
       
  1463         val Ts = map (typ_of_dtyp descr'' sorts) cargs;
       
  1464         val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts;
       
  1465         val frees' = partition_cargs idxs frees;
       
  1466         val binders = List.concat (map fst frees');
       
  1467         val atomTs = distinct op = (maps (map snd o fst) frees');
       
  1468         val recs = List.mapPartial
       
  1469           (fn ((_, DtRec i), p) => SOME (i, p) | _ => NONE)
       
  1470           (partition_cargs idxs cargs ~~ frees');
       
  1471         val frees'' = map (fn i => "y" ^ string_of_int i) (1 upto length recs) ~~
       
  1472           map (fn (i, _) => List.nth (rec_result_Ts, i)) recs;
       
  1473         val prems1 = map (fn ((i, (_, x)), y) => HOLogic.mk_Trueprop
       
  1474           (List.nth (rec_sets', i) $ Free x $ Free y)) (recs ~~ frees'');
       
  1475         val prems2 =
       
  1476           map (fn f => map (fn p as (_, T) => HOLogic.mk_Trueprop
       
  1477             (fresh_const T (fastype_of f) $ Free p $ f)) binders) rec_fns;
       
  1478         val prems3 = mk_fresh1 [] binders @ mk_fresh2 [] frees';
       
  1479         val prems4 = map (fn ((i, _), y) =>
       
  1480           HOLogic.mk_Trueprop (List.nth (rec_preds, i) $ Free y)) (recs ~~ frees'');
       
  1481         val prems5 = mk_fresh3 (recs ~~ frees'') frees';
       
  1482         val prems6 = maps (fn aT => map (fn y as (_, T) => HOLogic.mk_Trueprop
       
  1483           (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
       
  1484              (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ Free y)))
       
  1485                frees'') atomTs;
       
  1486         val prems7 = map (fn x as (_, T) => HOLogic.mk_Trueprop
       
  1487           (fresh_const T fsT' $ Free x $ rec_ctxt)) binders;
       
  1488         val result = list_comb (List.nth (rec_fns, l), map Free (frees @ frees''));
       
  1489         val result_freshs = map (fn p as (_, T) =>
       
  1490           fresh_const T (fastype_of result) $ Free p $ result) binders;
       
  1491         val P = HOLogic.mk_Trueprop (p $ result)
       
  1492       in
       
  1493         (rec_intr_ts @ [Logic.list_implies (List.concat prems2 @ prems3 @ prems1,
       
  1494            HOLogic.mk_Trueprop (rec_set $
       
  1495              list_comb (Const (cname, Ts ---> T), map Free frees) $ result))],
       
  1496          rec_prems @ [list_all_free (frees @ frees'', Logic.list_implies (prems4, P))],
       
  1497          rec_prems' @ map (fn fr => list_all_free (frees @ frees'',
       
  1498            Logic.list_implies (List.nth (prems2, l) @ prems3 @ prems5 @ prems7 @ prems6 @ [P],
       
  1499              HOLogic.mk_Trueprop fr))) result_freshs,
       
  1500          rec_eq_prems @ [List.concat prems2 @ prems3],
       
  1501          l + 1)
       
  1502       end;
       
  1503 
       
  1504     val (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, _) =
       
  1505       Library.foldl (fn (x, ((((d, d'), T), p), rec_set)) =>
       
  1506         Library.foldl (make_rec_intr T p rec_set) (x, #3 (snd d) ~~ snd d'))
       
  1507           (([], [], [], [], 0), descr'' ~~ ndescr ~~ recTs ~~ rec_preds ~~ rec_sets');
       
  1508 
       
  1509     val ({intrs = rec_intrs, elims = rec_elims, raw_induct = rec_induct, ...}, thy11) =
       
  1510       thy10 |>
       
  1511         Inductive.add_inductive_global (serial_string ())
       
  1512           {quiet_mode = #quiet config, verbose = false, kind = Thm.internalK,
       
  1513            alt_name = Binding.name big_rec_name, coind = false, no_elim = false, no_ind = false,
       
  1514            skip_mono = true, fork_mono = false}
       
  1515           (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
       
  1516           (map dest_Free rec_fns)
       
  1517           (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] ||>
       
  1518       PureThy.hide_fact true (Long_Name.append (Sign.full_bname thy10 big_rec_name) "induct");
       
  1519 
       
  1520     (** equivariance **)
       
  1521 
       
  1522     val fresh_bij = PureThy.get_thms thy11 "fresh_bij";
       
  1523     val perm_bij = PureThy.get_thms thy11 "perm_bij";
       
  1524 
       
  1525     val (rec_equiv_thms, rec_equiv_thms') = ListPair.unzip (map (fn aT =>
       
  1526       let
       
  1527         val permT = mk_permT aT;
       
  1528         val pi = Free ("pi", permT);
       
  1529         val rec_fns_pi = map (mk_perm [] pi o uncurry (mk_Free "f"))
       
  1530           (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
       
  1531         val rec_sets_pi = map (fn c => list_comb (Const c, rec_fns_pi))
       
  1532           (rec_set_names ~~ rec_set_Ts);
       
  1533         val ps = map (fn ((((T, U), R), R'), i) =>
       
  1534           let
       
  1535             val x = Free ("x" ^ string_of_int i, T);
       
  1536             val y = Free ("y" ^ string_of_int i, U)
       
  1537           in
       
  1538             (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y)
       
  1539           end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs));
       
  1540         val ths = map (fn th => standard (th RS mp)) (split_conj_thm
       
  1541           (Goal.prove_global thy11 [] []
       
  1542             (augment_sort thy1 pt_cp_sort
       
  1543               (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps))))
       
  1544             (fn _ => rtac rec_induct 1 THEN REPEAT
       
  1545                (simp_tac (Simplifier.theory_context thy11 HOL_basic_ss
       
  1546                   addsimps flat perm_simps'
       
  1547                   addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
       
  1548                 (resolve_tac rec_intrs THEN_ALL_NEW
       
  1549                  asm_simp_tac (HOL_ss addsimps (fresh_bij @ perm_bij))) 1))))
       
  1550         val ths' = map (fn ((P, Q), th) =>
       
  1551           Goal.prove_global thy11 [] []
       
  1552             (augment_sort thy1 pt_cp_sort
       
  1553               (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P)))
       
  1554             (fn _ => dtac (Thm.instantiate ([],
       
  1555                  [(cterm_of thy11 (Var (("pi", 0), permT)),
       
  1556                    cterm_of thy11 (Const ("List.rev", permT --> permT) $ pi))]) th) 1 THEN
       
  1557                NominalPermeq.perm_simp_tac HOL_ss 1)) (ps ~~ ths)
       
  1558       in (ths, ths') end) dt_atomTs);
       
  1559 
       
  1560     (** finite support **)
       
  1561 
       
  1562     val rec_fin_supp_thms = map (fn aT =>
       
  1563       let
       
  1564         val name = Long_Name.base_name (fst (dest_Type aT));
       
  1565         val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1");
       
  1566         val aset = HOLogic.mk_setT aT;
       
  1567         val finite = Const ("Finite_Set.finite", aset --> HOLogic.boolT);
       
  1568         val fins = map (fn (f, T) => HOLogic.mk_Trueprop
       
  1569           (finite $ (Const ("Nominal.supp", T --> aset) $ f)))
       
  1570             (rec_fns ~~ rec_fn_Ts)
       
  1571       in
       
  1572         map (fn th => standard (th RS mp)) (split_conj_thm
       
  1573           (Goal.prove_global thy11 []
       
  1574             (map (augment_sort thy11 fs_cp_sort) fins)
       
  1575             (augment_sort thy11 fs_cp_sort
       
  1576               (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
       
  1577                 (map (fn (((T, U), R), i) =>
       
  1578                    let
       
  1579                      val x = Free ("x" ^ string_of_int i, T);
       
  1580                      val y = Free ("y" ^ string_of_int i, U)
       
  1581                    in
       
  1582                      HOLogic.mk_imp (R $ x $ y,
       
  1583                        finite $ (Const ("Nominal.supp", U --> aset) $ y))
       
  1584                    end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~
       
  1585                      (1 upto length recTs))))))
       
  1586             (fn {prems = fins, ...} =>
       
  1587               (rtac rec_induct THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT
       
  1588                (NominalPermeq.finite_guess_tac (HOL_ss addsimps [fs_name]) 1))))
       
  1589       end) dt_atomTs;
       
  1590 
       
  1591     (** freshness **)
       
  1592 
       
  1593     val finite_premss = map (fn aT =>
       
  1594       map (fn (f, T) => HOLogic.mk_Trueprop
       
  1595         (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
       
  1596            (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ f)))
       
  1597            (rec_fns ~~ rec_fn_Ts)) dt_atomTs;
       
  1598 
       
  1599     val rec_fns' = map (augment_sort thy11 fs_cp_sort) rec_fns;
       
  1600 
       
  1601     val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) =>
       
  1602       let
       
  1603         val name = Long_Name.base_name (fst (dest_Type aT));
       
  1604         val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1");
       
  1605         val a = Free ("a", aT);
       
  1606         val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop
       
  1607           (fresh_const aT fT $ a $ f)) (rec_fns ~~ rec_fn_Ts)
       
  1608       in
       
  1609         map (fn (((T, U), R), eqvt_th) =>
       
  1610           let
       
  1611             val x = Free ("x", augment_sort_typ thy11 fs_cp_sort T);
       
  1612             val y = Free ("y", U);
       
  1613             val y' = Free ("y'", U)
       
  1614           in
       
  1615             standard (Goal.prove (ProofContext.init thy11) []
       
  1616               (map (augment_sort thy11 fs_cp_sort)
       
  1617                 (finite_prems @
       
  1618                    [HOLogic.mk_Trueprop (R $ x $ y),
       
  1619                     HOLogic.mk_Trueprop (HOLogic.mk_all ("y'", U,
       
  1620                       HOLogic.mk_imp (R $ x $ y', HOLogic.mk_eq (y', y)))),
       
  1621                     HOLogic.mk_Trueprop (fresh_const aT T $ a $ x)] @
       
  1622                  freshs))
       
  1623               (HOLogic.mk_Trueprop (fresh_const aT U $ a $ y))
       
  1624               (fn {prems, context} =>
       
  1625                  let
       
  1626                    val (finite_prems, rec_prem :: unique_prem ::
       
  1627                      fresh_prems) = chop (length finite_prems) prems;
       
  1628                    val unique_prem' = unique_prem RS spec RS mp;
       
  1629                    val unique = [unique_prem', unique_prem' RS sym] MRS trans;
       
  1630                    val _ $ (_ $ (_ $ S $ _)) $ _ = prop_of supports_fresh;
       
  1631                    val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns')
       
  1632                  in EVERY
       
  1633                    [rtac (Drule.cterm_instantiate
       
  1634                       [(cterm_of thy11 S,
       
  1635                         cterm_of thy11 (Const ("Nominal.supp",
       
  1636                           fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))]
       
  1637                       supports_fresh) 1,
       
  1638                     simp_tac (HOL_basic_ss addsimps
       
  1639                       [supports_def, symmetric fresh_def, fresh_prod]) 1,
       
  1640                     REPEAT_DETERM (resolve_tac [allI, impI] 1),
       
  1641                     REPEAT_DETERM (etac conjE 1),
       
  1642                     rtac unique 1,
       
  1643                     SUBPROOF (fn {prems = prems', params = [a, b], ...} => EVERY
       
  1644                       [cut_facts_tac [rec_prem] 1,
       
  1645                        rtac (Thm.instantiate ([],
       
  1646                          [(cterm_of thy11 (Var (("pi", 0), mk_permT aT)),
       
  1647                            cterm_of thy11 (perm_of_pair (term_of a, term_of b)))]) eqvt_th) 1,
       
  1648                        asm_simp_tac (HOL_ss addsimps
       
  1649                          (prems' @ perm_swap @ perm_fresh_fresh)) 1]) context 1,
       
  1650                     rtac rec_prem 1,
       
  1651                     simp_tac (HOL_ss addsimps (fs_name ::
       
  1652                       supp_prod :: finite_Un :: finite_prems)) 1,
       
  1653                     simp_tac (HOL_ss addsimps (symmetric fresh_def ::
       
  1654                       fresh_prod :: fresh_prems)) 1]
       
  1655                  end))
       
  1656           end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ eqvt_ths)
       
  1657       end) (dt_atomTs ~~ rec_equiv_thms' ~~ finite_premss);
       
  1658 
       
  1659     (** uniqueness **)
       
  1660 
       
  1661     val fun_tuple = foldr1 HOLogic.mk_prod (rec_ctxt :: rec_fns);
       
  1662     val fun_tupleT = fastype_of fun_tuple;
       
  1663     val rec_unique_frees =
       
  1664       DatatypeProp.indexify_names (replicate (length recTs) "x") ~~ recTs;
       
  1665     val rec_unique_frees'' = map (fn (s, T) => (s ^ "'", T)) rec_unique_frees;
       
  1666     val rec_unique_frees' =
       
  1667       DatatypeProp.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts;
       
  1668     val rec_unique_concls = map (fn ((x, U), R) =>
       
  1669         Const ("Ex1", (U --> HOLogic.boolT) --> HOLogic.boolT) $
       
  1670           Abs ("y", U, R $ Free x $ Bound 0))
       
  1671       (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets);
       
  1672 
       
  1673     val induct_aux_rec = Drule.cterm_instantiate
       
  1674       (map (pairself (cterm_of thy11) o apsnd (augment_sort thy11 fs_cp_sort))
       
  1675          (map (fn (aT, f) => (Logic.varify f, Abs ("z", HOLogic.unitT,
       
  1676             Const ("Nominal.supp", fun_tupleT --> HOLogic.mk_setT aT) $ fun_tuple)))
       
  1677               fresh_fs @
       
  1678           map (fn (((P, T), (x, U)), Q) =>
       
  1679            (Var ((P, 0), Logic.varifyT (fsT' --> T --> HOLogic.boolT)),
       
  1680             Abs ("z", HOLogic.unitT, absfree (x, U, Q))))
       
  1681               (pnames ~~ recTs ~~ rec_unique_frees ~~ rec_unique_concls) @
       
  1682           map (fn (s, T) => (Var ((s, 0), Logic.varifyT T), Free (s, T)))
       
  1683             rec_unique_frees)) induct_aux;
       
  1684 
       
  1685     fun obtain_fresh_name vs ths rec_fin_supp T (freshs1, freshs2, ctxt) =
       
  1686       let
       
  1687         val p = foldr1 HOLogic.mk_prod (vs @ freshs1);
       
  1688         val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
       
  1689             (HOLogic.exists_const T $ Abs ("x", T,
       
  1690               fresh_const T (fastype_of p) $ Bound 0 $ p)))
       
  1691           (fn _ => EVERY
       
  1692             [cut_facts_tac ths 1,
       
  1693              REPEAT_DETERM (dresolve_tac (the (AList.lookup op = rec_fin_supp T)) 1),
       
  1694              resolve_tac exists_fresh' 1,
       
  1695              asm_simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]);
       
  1696         val (([cx], ths), ctxt') = Obtain.result
       
  1697           (fn _ => EVERY
       
  1698             [etac exE 1,
       
  1699              full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
       
  1700              REPEAT (etac conjE 1)])
       
  1701           [ex] ctxt
       
  1702       in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
       
  1703 
       
  1704     val finite_ctxt_prems = map (fn aT =>
       
  1705       HOLogic.mk_Trueprop
       
  1706         (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
       
  1707            (Const ("Nominal.supp", fsT' --> HOLogic.mk_setT aT) $ rec_ctxt))) dt_atomTs;
       
  1708 
       
  1709     val rec_unique_thms = split_conj_thm (Goal.prove
       
  1710       (ProofContext.init thy11) (map fst rec_unique_frees)
       
  1711       (map (augment_sort thy11 fs_cp_sort)
       
  1712         (List.concat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems'))
       
  1713       (augment_sort thy11 fs_cp_sort
       
  1714         (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj rec_unique_concls)))
       
  1715       (fn {prems, context} =>
       
  1716          let
       
  1717            val k = length rec_fns;
       
  1718            val (finite_thss, ths1) = fold_map (fn T => fn xs =>
       
  1719              apfst (pair T) (chop k xs)) dt_atomTs prems;
       
  1720            val (finite_ctxt_ths, ths2) = chop (length dt_atomTs) ths1;
       
  1721            val (P_ind_ths, fcbs) = chop k ths2;
       
  1722            val P_ths = map (fn th => th RS mp) (split_conj_thm
       
  1723              (Goal.prove context
       
  1724                (map fst (rec_unique_frees'' @ rec_unique_frees')) []
       
  1725                (augment_sort thy11 fs_cp_sort
       
  1726                  (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
       
  1727                     (map (fn (((x, y), S), P) => HOLogic.mk_imp
       
  1728                       (S $ Free x $ Free y, P $ (Free y)))
       
  1729                         (rec_unique_frees'' ~~ rec_unique_frees' ~~
       
  1730                            rec_sets ~~ rec_preds)))))
       
  1731                (fn _ =>
       
  1732                   rtac rec_induct 1 THEN
       
  1733                   REPEAT ((resolve_tac P_ind_ths THEN_ALL_NEW assume_tac) 1))));
       
  1734            val rec_fin_supp_thms' = map
       
  1735              (fn (ths, (T, fin_ths)) => (T, map (curry op MRS fin_ths) ths))
       
  1736              (rec_fin_supp_thms ~~ finite_thss);
       
  1737          in EVERY
       
  1738            ([rtac induct_aux_rec 1] @
       
  1739             maps (fn ((_, finite_ths), finite_th) =>
       
  1740               [cut_facts_tac (finite_th :: finite_ths) 1,
       
  1741                asm_simp_tac (HOL_ss addsimps [supp_prod, finite_Un]) 1])
       
  1742                 (finite_thss ~~ finite_ctxt_ths) @
       
  1743             maps (fn ((_, idxss), elim) => maps (fn idxs =>
       
  1744               [full_simp_tac (HOL_ss addsimps [symmetric fresh_def, supp_prod, Un_iff]) 1,
       
  1745                REPEAT_DETERM (eresolve_tac [conjE, ex1E] 1),
       
  1746                rtac ex1I 1,
       
  1747                (resolve_tac rec_intrs THEN_ALL_NEW atac) 1,
       
  1748                rotate_tac ~1 1,
       
  1749                ((DETERM o etac elim) THEN_ALL_NEW full_simp_tac
       
  1750                   (HOL_ss addsimps List.concat distinct_thms)) 1] @
       
  1751                (if null idxs then [] else [hyp_subst_tac 1,
       
  1752                 SUBPROOF (fn {asms, concl, prems = prems', params, context = context', ...} =>
       
  1753                   let
       
  1754                     val SOME prem = find_first (can (HOLogic.dest_eq o
       
  1755                       HOLogic.dest_Trueprop o prop_of)) prems';
       
  1756                     val _ $ (_ $ lhs $ rhs) = prop_of prem;
       
  1757                     val _ $ (_ $ lhs' $ rhs') = term_of concl;
       
  1758                     val rT = fastype_of lhs';
       
  1759                     val (c, cargsl) = strip_comb lhs;
       
  1760                     val cargsl' = partition_cargs idxs cargsl;
       
  1761                     val boundsl = List.concat (map fst cargsl');
       
  1762                     val (_, cargsr) = strip_comb rhs;
       
  1763                     val cargsr' = partition_cargs idxs cargsr;
       
  1764                     val boundsr = List.concat (map fst cargsr');
       
  1765                     val (params1, _ :: params2) =
       
  1766                       chop (length params div 2) (map term_of params);
       
  1767                     val params' = params1 @ params2;
       
  1768                     val rec_prems = filter (fn th => case prop_of th of
       
  1769                         _ $ p => (case head_of p of
       
  1770                           Const (s, _) => s mem rec_set_names
       
  1771                         | _ => false)
       
  1772                       | _ => false) prems';
       
  1773                     val fresh_prems = filter (fn th => case prop_of th of
       
  1774                         _ $ (Const ("Nominal.fresh", _) $ _ $ _) => true
       
  1775                       | _ $ (Const ("Not", _) $ _) => true
       
  1776                       | _ => false) prems';
       
  1777                     val Ts = map fastype_of boundsl;
       
  1778 
       
  1779                     val _ = warning "step 1: obtaining fresh names";
       
  1780                     val (freshs1, freshs2, context'') = fold
       
  1781                       (obtain_fresh_name (rec_ctxt :: rec_fns' @ params')
       
  1782                          (List.concat (map snd finite_thss) @
       
  1783                             finite_ctxt_ths @ rec_prems)
       
  1784                          rec_fin_supp_thms')
       
  1785                       Ts ([], [], context');
       
  1786                     val pi1 = map perm_of_pair (boundsl ~~ freshs1);
       
  1787                     val rpi1 = rev pi1;
       
  1788                     val pi2 = map perm_of_pair (boundsr ~~ freshs1);
       
  1789                     val rpi2 = rev pi2;
       
  1790 
       
  1791                     val fresh_prems' = mk_not_sym fresh_prems;
       
  1792                     val freshs2' = mk_not_sym freshs2;
       
  1793 
       
  1794                     (** as, bs, cs # K as ts, K bs us **)
       
  1795                     val _ = warning "step 2: as, bs, cs # K as ts, K bs us";
       
  1796                     val prove_fresh_ss = HOL_ss addsimps
       
  1797                       (finite_Diff :: List.concat fresh_thms @
       
  1798                        fs_atoms @ abs_fresh @ abs_supp @ fresh_atm);
       
  1799                     (* FIXME: avoid asm_full_simp_tac ? *)
       
  1800                     fun prove_fresh ths y x = Goal.prove context'' [] []
       
  1801                       (HOLogic.mk_Trueprop (fresh_const
       
  1802                          (fastype_of x) (fastype_of y) $ x $ y))
       
  1803                       (fn _ => cut_facts_tac ths 1 THEN asm_full_simp_tac prove_fresh_ss 1);
       
  1804                     val constr_fresh_thms =
       
  1805                       map (prove_fresh fresh_prems lhs) boundsl @
       
  1806                       map (prove_fresh fresh_prems rhs) boundsr @
       
  1807                       map (prove_fresh freshs2 lhs) freshs1 @
       
  1808                       map (prove_fresh freshs2 rhs) freshs1;
       
  1809 
       
  1810                     (** pi1 o (K as ts) = pi2 o (K bs us) **)
       
  1811                     val _ = warning "step 3: pi1 o (K as ts) = pi2 o (K bs us)";
       
  1812                     val pi1_pi2_eq = Goal.prove context'' [] []
       
  1813                       (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
  1814                         (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 rhs)))
       
  1815                       (fn _ => EVERY
       
  1816                          [cut_facts_tac constr_fresh_thms 1,
       
  1817                           asm_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh) 1,
       
  1818                           rtac prem 1]);
       
  1819 
       
  1820                     (** pi1 o ts = pi2 o us **)
       
  1821                     val _ = warning "step 4: pi1 o ts = pi2 o us";
       
  1822                     val pi1_pi2_eqs = map (fn (t, u) =>
       
  1823                       Goal.prove context'' [] []
       
  1824                         (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
  1825                           (fold_rev (mk_perm []) pi1 t, fold_rev (mk_perm []) pi2 u)))
       
  1826                         (fn _ => EVERY
       
  1827                            [cut_facts_tac [pi1_pi2_eq] 1,
       
  1828                             asm_full_simp_tac (HOL_ss addsimps
       
  1829                               (calc_atm @ List.concat perm_simps' @
       
  1830                                fresh_prems' @ freshs2' @ abs_perm @
       
  1831                                alpha @ List.concat inject_thms)) 1]))
       
  1832                         (map snd cargsl' ~~ map snd cargsr');
       
  1833 
       
  1834                     (** pi1^-1 o pi2 o us = ts **)
       
  1835                     val _ = warning "step 5: pi1^-1 o pi2 o us = ts";
       
  1836                     val rpi1_pi2_eqs = map (fn ((t, u), eq) =>
       
  1837                       Goal.prove context'' [] []
       
  1838                         (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
  1839                           (fold_rev (mk_perm []) (rpi1 @ pi2) u, t)))
       
  1840                         (fn _ => simp_tac (HOL_ss addsimps
       
  1841                            ((eq RS sym) :: perm_swap)) 1))
       
  1842                         (map snd cargsl' ~~ map snd cargsr' ~~ pi1_pi2_eqs);
       
  1843 
       
  1844                     val (rec_prems1, rec_prems2) =
       
  1845                       chop (length rec_prems div 2) rec_prems;
       
  1846 
       
  1847                     (** (ts, pi1^-1 o pi2 o vs) in rec_set **)
       
  1848                     val _ = warning "step 6: (ts, pi1^-1 o pi2 o vs) in rec_set";
       
  1849                     val rec_prems' = map (fn th =>
       
  1850                       let
       
  1851                         val _ $ (S $ x $ y) = prop_of th;
       
  1852                         val Const (s, _) = head_of S;
       
  1853                         val k = find_index (equal s) rec_set_names;
       
  1854                         val pi = rpi1 @ pi2;
       
  1855                         fun mk_pi z = fold_rev (mk_perm []) pi z;
       
  1856                         fun eqvt_tac p =
       
  1857                           let
       
  1858                             val U as Type (_, [Type (_, [T, _])]) = fastype_of p;
       
  1859                             val l = find_index (equal T) dt_atomTs;
       
  1860                             val th = List.nth (List.nth (rec_equiv_thms', l), k);
       
  1861                             val th' = Thm.instantiate ([],
       
  1862                               [(cterm_of thy11 (Var (("pi", 0), U)),
       
  1863                                 cterm_of thy11 p)]) th;
       
  1864                           in rtac th' 1 end;
       
  1865                         val th' = Goal.prove context'' [] []
       
  1866                           (HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y))
       
  1867                           (fn _ => EVERY
       
  1868                              (map eqvt_tac pi @
       
  1869                               [simp_tac (HOL_ss addsimps (fresh_prems' @ freshs2' @
       
  1870                                  perm_swap @ perm_fresh_fresh)) 1,
       
  1871                                rtac th 1]))
       
  1872                       in
       
  1873                         Simplifier.simplify
       
  1874                           (HOL_basic_ss addsimps rpi1_pi2_eqs) th'
       
  1875                       end) rec_prems2;
       
  1876 
       
  1877                     val ihs = filter (fn th => case prop_of th of
       
  1878                       _ $ (Const ("All", _) $ _) => true | _ => false) prems';
       
  1879 
       
  1880                     (** pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs **)
       
  1881                     val _ = warning "step 7: pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs";
       
  1882                     val rec_eqns = map (fn (th, ih) =>
       
  1883                       let
       
  1884                         val th' = th RS (ih RS spec RS mp) RS sym;
       
  1885                         val _ $ (_ $ lhs $ rhs) = prop_of th';
       
  1886                         fun strip_perm (_ $ _ $ t) = strip_perm t
       
  1887                           | strip_perm t = t;
       
  1888                       in
       
  1889                         Goal.prove context'' [] []
       
  1890                            (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
  1891                               (fold_rev (mk_perm []) pi1 lhs,
       
  1892                                fold_rev (mk_perm []) pi2 (strip_perm rhs))))
       
  1893                            (fn _ => simp_tac (HOL_basic_ss addsimps
       
  1894                               (th' :: perm_swap)) 1)
       
  1895                       end) (rec_prems' ~~ ihs);
       
  1896 
       
  1897                     (** as # rs **)
       
  1898                     val _ = warning "step 8: as # rs";
       
  1899                     val rec_freshs = List.concat
       
  1900                       (map (fn (rec_prem, ih) =>
       
  1901                         let
       
  1902                           val _ $ (S $ x $ (y as Free (_, T))) =
       
  1903                             prop_of rec_prem;
       
  1904                           val k = find_index (equal S) rec_sets;
       
  1905                           val atoms = List.concat (List.mapPartial (fn (bs, z) =>
       
  1906                             if z = x then NONE else SOME bs) cargsl')
       
  1907                         in
       
  1908                           map (fn a as Free (_, aT) =>
       
  1909                             let val l = find_index (equal aT) dt_atomTs;
       
  1910                             in
       
  1911                               Goal.prove context'' [] []
       
  1912                                 (HOLogic.mk_Trueprop (fresh_const aT T $ a $ y))
       
  1913                                 (fn _ => EVERY
       
  1914                                    (rtac (List.nth (List.nth (rec_fresh_thms, l), k)) 1 ::
       
  1915                                     map (fn th => rtac th 1)
       
  1916                                       (snd (List.nth (finite_thss, l))) @
       
  1917                                     [rtac rec_prem 1, rtac ih 1,
       
  1918                                      REPEAT_DETERM (resolve_tac fresh_prems 1)]))
       
  1919                             end) atoms
       
  1920                         end) (rec_prems1 ~~ ihs));
       
  1921 
       
  1922                     (** as # fK as ts rs , bs # fK bs us vs **)
       
  1923                     val _ = warning "step 9: as # fK as ts rs , bs # fK bs us vs";
       
  1924                     fun prove_fresh_result (a as Free (_, aT)) =
       
  1925                       Goal.prove context'' [] []
       
  1926                         (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ rhs'))
       
  1927                         (fn _ => EVERY
       
  1928                            [resolve_tac fcbs 1,
       
  1929                             REPEAT_DETERM (resolve_tac
       
  1930                               (fresh_prems @ rec_freshs) 1),
       
  1931                             REPEAT_DETERM (resolve_tac (maps snd rec_fin_supp_thms') 1
       
  1932                               THEN resolve_tac rec_prems 1),
       
  1933                             resolve_tac P_ind_ths 1,
       
  1934                             REPEAT_DETERM (resolve_tac (P_ths @ rec_prems) 1)]);
       
  1935 
       
  1936                     val fresh_results'' = map prove_fresh_result boundsl;
       
  1937 
       
  1938                     fun prove_fresh_result'' ((a as Free (_, aT), b), th) =
       
  1939                       let val th' = Goal.prove context'' [] []
       
  1940                         (HOLogic.mk_Trueprop (fresh_const aT rT $
       
  1941                             fold_rev (mk_perm []) (rpi2 @ pi1) a $
       
  1942                             fold_rev (mk_perm []) (rpi2 @ pi1) rhs'))
       
  1943                         (fn _ => simp_tac (HOL_ss addsimps fresh_bij) 1 THEN
       
  1944                            rtac th 1)
       
  1945                       in
       
  1946                         Goal.prove context'' [] []
       
  1947                           (HOLogic.mk_Trueprop (fresh_const aT rT $ b $ lhs'))
       
  1948                           (fn _ => EVERY
       
  1949                              [cut_facts_tac [th'] 1,
       
  1950                               full_simp_tac (Simplifier.theory_context thy11 HOL_ss
       
  1951                                 addsimps rec_eqns @ pi1_pi2_eqs @ perm_swap
       
  1952                                 addsimprocs [NominalPermeq.perm_simproc_app]) 1,
       
  1953                               full_simp_tac (HOL_ss addsimps (calc_atm @
       
  1954                                 fresh_prems' @ freshs2' @ perm_fresh_fresh)) 1])
       
  1955                       end;
       
  1956 
       
  1957                     val fresh_results = fresh_results'' @ map prove_fresh_result''
       
  1958                       (boundsl ~~ boundsr ~~ fresh_results'');
       
  1959 
       
  1960                     (** cs # fK as ts rs , cs # fK bs us vs **)
       
  1961                     val _ = warning "step 10: cs # fK as ts rs , cs # fK bs us vs";
       
  1962                     fun prove_fresh_result' recs t (a as Free (_, aT)) =
       
  1963                       Goal.prove context'' [] []
       
  1964                         (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ t))
       
  1965                         (fn _ => EVERY
       
  1966                           [cut_facts_tac recs 1,
       
  1967                            REPEAT_DETERM (dresolve_tac
       
  1968                              (the (AList.lookup op = rec_fin_supp_thms' aT)) 1),
       
  1969                            NominalPermeq.fresh_guess_tac
       
  1970                              (HOL_ss addsimps (freshs2 @
       
  1971                                 fs_atoms @ fresh_atm @
       
  1972                                 List.concat (map snd finite_thss))) 1]);
       
  1973 
       
  1974                     val fresh_results' =
       
  1975                       map (prove_fresh_result' rec_prems1 rhs') freshs1 @
       
  1976                       map (prove_fresh_result' rec_prems2 lhs') freshs1;
       
  1977 
       
  1978                     (** pi1 o (fK as ts rs) = pi2 o (fK bs us vs) **)
       
  1979                     val _ = warning "step 11: pi1 o (fK as ts rs) = pi2 o (fK bs us vs)";
       
  1980                     val pi1_pi2_result = Goal.prove context'' [] []
       
  1981                       (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
  1982                         (fold_rev (mk_perm []) pi1 rhs', fold_rev (mk_perm []) pi2 lhs')))
       
  1983                       (fn _ => simp_tac (Simplifier.context context'' HOL_ss
       
  1984                            addsimps pi1_pi2_eqs @ rec_eqns
       
  1985                            addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
       
  1986                          TRY (simp_tac (HOL_ss addsimps
       
  1987                            (fresh_prems' @ freshs2' @ calc_atm @ perm_fresh_fresh)) 1));
       
  1988 
       
  1989                     val _ = warning "final result";
       
  1990                     val final = Goal.prove context'' [] [] (term_of concl)
       
  1991                       (fn _ => cut_facts_tac [pi1_pi2_result RS sym] 1 THEN
       
  1992                         full_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh @
       
  1993                           fresh_results @ fresh_results') 1);
       
  1994                     val final' = ProofContext.export context'' context' [final];
       
  1995                     val _ = warning "finished!"
       
  1996                   in
       
  1997                     resolve_tac final' 1
       
  1998                   end) context 1])) idxss) (ndescr ~~ rec_elims))
       
  1999          end));
       
  2000 
       
  2001     val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
       
  2002 
       
  2003     (* define primrec combinators *)
       
  2004 
       
  2005     val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
       
  2006     val reccomb_names = map (Sign.full_bname thy11)
       
  2007       (if length descr'' = 1 then [big_reccomb_name] else
       
  2008         (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
       
  2009           (1 upto (length descr''))));
       
  2010     val reccombs = map (fn ((name, T), T') => list_comb
       
  2011       (Const (name, rec_fn_Ts @ [T] ---> T'), rec_fns))
       
  2012         (reccomb_names ~~ recTs ~~ rec_result_Ts);
       
  2013 
       
  2014     val (reccomb_defs, thy12) =
       
  2015       thy11
       
  2016       |> Sign.add_consts_i (map (fn ((name, T), T') =>
       
  2017           (Binding.name (Long_Name.base_name name), rec_fn_Ts @ [T] ---> T', NoSyn))
       
  2018           (reccomb_names ~~ recTs ~~ rec_result_Ts))
       
  2019       |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
       
  2020           (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T,
       
  2021            Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
       
  2022              set $ Free ("x", T) $ Free ("y", T'))))))
       
  2023                (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts));
       
  2024 
       
  2025     (* prove characteristic equations for primrec combinators *)
       
  2026 
       
  2027     val rec_thms = map (fn (prems, concl) =>
       
  2028       let
       
  2029         val _ $ (_ $ (_ $ x) $ _) = concl;
       
  2030         val (_, cargs) = strip_comb x;
       
  2031         val ps = map (fn (x as Free (_, T), i) =>
       
  2032           (Free ("x" ^ string_of_int i, T), x)) (cargs ~~ (1 upto length cargs));
       
  2033         val concl' = subst_atomic_types (rec_result_Ts' ~~ rec_result_Ts) concl;
       
  2034         val prems' = List.concat finite_premss @ finite_ctxt_prems @
       
  2035           rec_prems @ rec_prems' @ map (subst_atomic ps) prems;
       
  2036         fun solve rules prems = resolve_tac rules THEN_ALL_NEW
       
  2037           (resolve_tac prems THEN_ALL_NEW atac)
       
  2038       in
       
  2039         Goal.prove_global thy12 []
       
  2040           (map (augment_sort thy12 fs_cp_sort) prems')
       
  2041           (augment_sort thy12 fs_cp_sort concl')
       
  2042           (fn {prems, ...} => EVERY
       
  2043             [rewrite_goals_tac reccomb_defs,
       
  2044              rtac the1_equality 1,
       
  2045              solve rec_unique_thms prems 1,
       
  2046              resolve_tac rec_intrs 1,
       
  2047              REPEAT (solve (prems @ rec_total_thms) prems 1)])
       
  2048       end) (rec_eq_prems ~~
       
  2049         DatatypeProp.make_primrecs new_type_names descr' sorts thy12);
       
  2050 
       
  2051     val dt_infos = map (make_dt_info pdescr sorts induct reccomb_names rec_thms)
       
  2052       ((0 upto length descr1 - 1) ~~ descr1 ~~ distinct_thms ~~ inject_thms);
       
  2053 
       
  2054     (* FIXME: theorems are stored in database for testing only *)
       
  2055     val (_, thy13) = thy12 |>
       
  2056       PureThy.add_thmss
       
  2057         [((Binding.name "rec_equiv", List.concat rec_equiv_thms), []),
       
  2058          ((Binding.name "rec_equiv'", List.concat rec_equiv_thms'), []),
       
  2059          ((Binding.name "rec_fin_supp", List.concat rec_fin_supp_thms), []),
       
  2060          ((Binding.name "rec_fresh", List.concat rec_fresh_thms), []),
       
  2061          ((Binding.name "rec_unique", map standard rec_unique_thms), []),
       
  2062          ((Binding.name "recs", rec_thms), [])] ||>
       
  2063       Sign.parent_path ||>
       
  2064       map_nominal_datatypes (fold Symtab.update dt_infos);
       
  2065 
       
  2066   in
       
  2067     thy13
       
  2068   end;
       
  2069 
       
  2070 val add_nominal_datatype = gen_add_nominal_datatype Datatype.read_typ;
       
  2071 
       
  2072 
       
  2073 (* FIXME: The following stuff should be exported by Datatype *)
       
  2074 
       
  2075 local structure P = OuterParse and K = OuterKeyword in
       
  2076 
       
  2077 val datatype_decl =
       
  2078   Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix --
       
  2079     (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix));
       
  2080 
       
  2081 fun mk_datatype args =
       
  2082   let
       
  2083     val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args;
       
  2084     val specs = map (fn ((((_, vs), t), mx), cons) =>
       
  2085       (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
       
  2086   in add_nominal_datatype Datatype.default_config names specs end;
       
  2087 
       
  2088 val _ =
       
  2089   OuterSyntax.command "nominal_datatype" "define inductive datatypes" K.thy_decl
       
  2090     (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
       
  2091 
       
  2092 end;
       
  2093 
       
  2094 end