src/HOL/Nominal/nominal_package.ML
changeset 28731 c60ac7923a06
parent 28662 64ab5bb68d4c
child 28736 b1fd60fee652
equal deleted inserted replaced
28730:71c946ce7eb9 28731:c60ac7923a06
   132       let
   132       let
   133         val (aT as Type (a, []), S) = dest_permT T;
   133         val (aT as Type (a, []), S) = dest_permT T;
   134         val (bT as Type (b, []), _) = dest_permT U
   134         val (bT as Type (b, []), _) = dest_permT U
   135       in if aT mem permTs_of u andalso aT <> bT then
   135       in if aT mem permTs_of u andalso aT <> bT then
   136           let
   136           let
   137             val a' = Sign.base_name a;
   137             val cp = cp_inst_of thy a b;
   138             val b' = Sign.base_name b;
   138             val dj = dj_thm_of thy b a;
   139             val cp = PureThy.get_thm thy ("cp_" ^ a' ^ "_" ^ b' ^ "_inst");
       
   140             val dj = PureThy.get_thm thy ("dj_" ^ b' ^ "_" ^ a');
       
   141             val dj_cp' = [cp, dj] MRS dj_cp;
   139             val dj_cp' = [cp, dj] MRS dj_cp;
   142             val cert = SOME o cterm_of thy
   140             val cert = SOME o cterm_of thy
   143           in
   141           in
   144             SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)]
   142             SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)]
   145               [cert t, cert r, cert s] dj_cp'))
   143               [cert t, cert r, cert s] dj_cp'))
   201       Theory.copy |>
   199       Theory.copy |>
   202       Sign.add_types (map (fn (tvs, tname, mx, _) =>
   200       Sign.add_types (map (fn (tvs, tname, mx, _) =>
   203         (tname, length tvs, mx)) dts);
   201         (tname, length tvs, mx)) dts);
   204 
   202 
   205     val atoms = atoms_of thy;
   203     val atoms = atoms_of thy;
   206     val classes = map (NameSpace.map_base (fn s => "pt_" ^ s)) atoms;
       
   207     val cp_classes = List.concat (map (fn atom1 => map (fn atom2 =>
       
   208       Sign.intern_class thy ("cp_" ^ Sign.base_name atom1 ^ "_" ^
       
   209         Sign.base_name atom2)) atoms) atoms);
       
   210     fun augment_sort S = S union classes;
       
   211     val augment_sort_typ = map_type_tfree (fn (s, S) => TFree (s, augment_sort S));
       
   212 
   204 
   213     fun prep_constr ((constrs, sorts), (cname, cargs, mx)) =
   205     fun prep_constr ((constrs, sorts), (cname, cargs, mx)) =
   214       let val (cargs', sorts') = Library.foldl (prep_typ tmp_thy) (([], sorts), cargs)
   206       let val (cargs', sorts') = Library.foldl (prep_typ tmp_thy) (([], sorts), cargs)
   215       in (constrs @ [(cname, cargs', mx)], sorts') end
   207       in (constrs @ [(cname, cargs', mx)], sorts') end
   216 
   208 
   217     fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) =
   209     fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) =
   218       let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs)
   210       let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs)
   219       in (dts @ [(tvs, tname, mx, constrs')], sorts') end
   211       in (dts @ [(tvs, tname, mx, constrs')], sorts') end
   220 
   212 
   221     val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts);
   213     val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts);
   222     val sorts' = map (apsnd augment_sort) sorts;
       
   223     val tyvars = map #1 dts';
   214     val tyvars = map #1 dts';
       
   215 
       
   216     fun inter_sort thy S S' = Type.inter_sort (Sign.tsig_of thy) (S, S');
       
   217     fun augment_sort_typ thy S =
       
   218       let val S = Sign.certify_sort thy S
       
   219       in map_type_tfree (fn (s, S') => TFree (s,
       
   220         if member (op = o apsnd fst) sorts s then inter_sort thy S S' else S'))
       
   221       end;
       
   222     fun augment_sort thy S = map_types (augment_sort_typ thy S);
   224 
   223 
   225     val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts';
   224     val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts';
   226     val constr_syntax = map (fn (tvs, tname, mx, constrs) =>
   225     val constr_syntax = map (fn (tvs, tname, mx, constrs) =>
   227       map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts';
   226       map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts';
   228 
   227 
   236           Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts)
   235           Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts)
   237       | replace_types T = T;
   236       | replace_types T = T;
   238 
   237 
   239     val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, tname ^ "_Rep", NoSyn,
   238     val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, tname ^ "_Rep", NoSyn,
   240       map (fn (cname, cargs, mx) => (cname ^ "_Rep",
   239       map (fn (cname, cargs, mx) => (cname ^ "_Rep",
   241         map (augment_sort_typ o replace_types) cargs, NoSyn)) constrs)) dts';
   240         map replace_types cargs, NoSyn)) constrs)) dts';
   242 
   241 
   243     val new_type_names' = map (fn n => n ^ "_Rep") new_type_names;
   242     val new_type_names' = map (fn n => n ^ "_Rep") new_type_names;
   244     val full_new_type_names' = map (Sign.full_name thy) new_type_names';
   243     val full_new_type_names' = map (Sign.full_name thy) new_type_names';
   245 
   244 
   246     val ({induction, ...},thy1) =
   245     val ({induction, ...},thy1) =
   247       DatatypePackage.add_datatype err flat_names new_type_names' dts'' thy;
   246       DatatypePackage.add_datatype err flat_names new_type_names' dts'' thy;
   248 
   247 
   249     val SOME {descr, ...} = Symtab.lookup
   248     val SOME {descr, ...} = Symtab.lookup
   250       (DatatypePackage.get_datatypes thy1) (hd full_new_type_names');
   249       (DatatypePackage.get_datatypes thy1) (hd full_new_type_names');
   251     fun nth_dtyp i = typ_of_dtyp descr sorts' (DtRec i);
   250     fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
   252 
   251 
   253     val big_name = space_implode "_" new_type_names;
   252     val big_name = space_implode "_" new_type_names;
   254 
   253 
   255 
   254 
   256     (**** define permutation functions ****)
   255     (**** define permutation functions ****)
   269 
   268 
   270     val perm_eqs = maps (fn (i, (_, _, constrs)) =>
   269     val perm_eqs = maps (fn (i, (_, _, constrs)) =>
   271       let val T = nth_dtyp i
   270       let val T = nth_dtyp i
   272       in map (fn (cname, dts) =>
   271       in map (fn (cname, dts) =>
   273         let
   272         let
   274           val Ts = map (typ_of_dtyp descr sorts') dts;
   273           val Ts = map (typ_of_dtyp descr sorts) dts;
   275           val names = Name.variant_list ["pi"] (DatatypeProp.make_tnames Ts);
   274           val names = Name.variant_list ["pi"] (DatatypeProp.make_tnames Ts);
   276           val args = map Free (names ~~ Ts);
   275           val args = map Free (names ~~ Ts);
   277           val c = Const (cname, Ts ---> T);
   276           val c = Const (cname, Ts ---> T);
   278           fun perm_arg (dt, x) =
   277           fun perm_arg (dt, x) =
   279             let val T = type_of x
   278             let val T = type_of x
   334 
   333 
   335     val perm_empty_thms = List.concat (map (fn a =>
   334     val perm_empty_thms = List.concat (map (fn a =>
   336       let val permT = mk_permT (Type (a, []))
   335       let val permT = mk_permT (Type (a, []))
   337       in map standard (List.take (split_conj_thm
   336       in map standard (List.take (split_conj_thm
   338         (Goal.prove_global thy2 [] []
   337         (Goal.prove_global thy2 [] []
   339           (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   338           (augment_sort thy2 [pt_class_of thy2 a]
   340             (map (fn ((s, T), x) => HOLogic.mk_eq
   339             (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   341                 (Const (s, permT --> T --> T) $
   340               (map (fn ((s, T), x) => HOLogic.mk_eq
   342                    Const ("List.list.Nil", permT) $ Free (x, T),
   341                   (Const (s, permT --> T --> T) $
   343                  Free (x, T)))
   342                      Const ("List.list.Nil", permT) $ Free (x, T),
   344              (perm_names ~~
   343                    Free (x, T)))
   345               map body_type perm_types ~~ perm_indnames))))
   344                (perm_names ~~
       
   345                 map body_type perm_types ~~ perm_indnames)))))
   346           (fn _ => EVERY [indtac induction perm_indnames 1,
   346           (fn _ => EVERY [indtac induction perm_indnames 1,
   347             ALLGOALS (asm_full_simp_tac (simpset_of thy2))])),
   347             ALLGOALS (asm_full_simp_tac (simpset_of thy2))])),
   348         length new_type_names))
   348         length new_type_names))
   349       end)
   349       end)
   350       atoms);
   350       atoms);
   360     val perm_append_thms = List.concat (map (fn a =>
   360     val perm_append_thms = List.concat (map (fn a =>
   361       let
   361       let
   362         val permT = mk_permT (Type (a, []));
   362         val permT = mk_permT (Type (a, []));
   363         val pi1 = Free ("pi1", permT);
   363         val pi1 = Free ("pi1", permT);
   364         val pi2 = Free ("pi2", permT);
   364         val pi2 = Free ("pi2", permT);
   365         val pt_inst = PureThy.get_thm thy2 ("pt_" ^ Sign.base_name a ^ "_inst");
   365         val pt_inst = pt_inst_of thy2 a;
   366         val pt2' = pt_inst RS pt2;
   366         val pt2' = pt_inst RS pt2;
   367         val pt2_ax = PureThy.get_thm thy2 (NameSpace.map_base (fn s => "pt_" ^ s ^ "2") a);
   367         val pt2_ax = PureThy.get_thm thy2 (NameSpace.map_base (fn s => "pt_" ^ s ^ "2") a);
   368       in List.take (map standard (split_conj_thm
   368       in List.take (map standard (split_conj_thm
   369         (Goal.prove_global thy2 [] []
   369         (Goal.prove_global thy2 [] []
       
   370            (augment_sort thy2 [pt_class_of thy2 a]
   370              (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   371              (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   371                 (map (fn ((s, T), x) =>
   372                 (map (fn ((s, T), x) =>
   372                     let val perm = Const (s, permT --> T --> T)
   373                     let val perm = Const (s, permT --> T --> T)
   373                     in HOLogic.mk_eq
   374                     in HOLogic.mk_eq
   374                       (perm $ (Const ("List.append", permT --> permT --> permT) $
   375                       (perm $ (Const ("List.append", permT --> permT --> permT) $
   375                          pi1 $ pi2) $ Free (x, T),
   376                          pi1 $ pi2) $ Free (x, T),
   376                        perm $ pi1 $ (perm $ pi2 $ Free (x, T)))
   377                        perm $ pi1 $ (perm $ pi2 $ Free (x, T)))
   377                     end)
   378                     end)
   378                   (perm_names ~~
   379                   (perm_names ~~
   379                    map body_type perm_types ~~ perm_indnames))))
   380                    map body_type perm_types ~~ perm_indnames)))))
   380            (fn _ => EVERY [indtac induction perm_indnames 1,
   381            (fn _ => EVERY [indtac induction perm_indnames 1,
   381               ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))),
   382               ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))),
   382          length new_type_names)
   383          length new_type_names)
   383       end) atoms);
   384       end) atoms);
   384 
   385 
   392     val perm_eq_thms = List.concat (map (fn a =>
   393     val perm_eq_thms = List.concat (map (fn a =>
   393       let
   394       let
   394         val permT = mk_permT (Type (a, []));
   395         val permT = mk_permT (Type (a, []));
   395         val pi1 = Free ("pi1", permT);
   396         val pi1 = Free ("pi1", permT);
   396         val pi2 = Free ("pi2", permT);
   397         val pi2 = Free ("pi2", permT);
   397         (*FIXME: not robust - better access these theorems using NominalData?*)
   398         val at_inst = at_inst_of thy2 a;
   398         val at_inst = PureThy.get_thm thy2 ("at_" ^ Sign.base_name a ^ "_inst");
   399         val pt_inst = pt_inst_of thy2 a;
   399         val pt_inst = PureThy.get_thm thy2 ("pt_" ^ Sign.base_name a ^ "_inst");
       
   400         val pt3' = pt_inst RS pt3;
   400         val pt3' = pt_inst RS pt3;
   401         val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
   401         val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
   402         val pt3_ax = PureThy.get_thm thy2 (NameSpace.map_base (fn s => "pt_" ^ s ^ "3") a);
   402         val pt3_ax = PureThy.get_thm thy2 (NameSpace.map_base (fn s => "pt_" ^ s ^ "3") a);
   403       in List.take (map standard (split_conj_thm
   403       in List.take (map standard (split_conj_thm
   404         (Goal.prove_global thy2 [] [] (Logic.mk_implies
   404         (Goal.prove_global thy2 [] []
       
   405           (augment_sort thy2 [pt_class_of thy2 a] (Logic.mk_implies
   405              (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq",
   406              (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq",
   406                 permT --> permT --> HOLogic.boolT) $ pi1 $ pi2),
   407                 permT --> permT --> HOLogic.boolT) $ pi1 $ pi2),
   407               HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   408               HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   408                 (map (fn ((s, T), x) =>
   409                 (map (fn ((s, T), x) =>
   409                     let val perm = Const (s, permT --> T --> T)
   410                     let val perm = Const (s, permT --> T --> T)
   410                     in HOLogic.mk_eq
   411                     in HOLogic.mk_eq
   411                       (perm $ pi1 $ Free (x, T),
   412                       (perm $ pi1 $ Free (x, T),
   412                        perm $ pi2 $ Free (x, T))
   413                        perm $ pi2 $ Free (x, T))
   413                     end)
   414                     end)
   414                   (perm_names ~~
   415                   (perm_names ~~
   415                    map body_type perm_types ~~ perm_indnames)))))
   416                    map body_type perm_types ~~ perm_indnames))))))
   416            (fn _ => EVERY [indtac induction perm_indnames 1,
   417            (fn _ => EVERY [indtac induction perm_indnames 1,
   417               ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))),
   418               ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))),
   418          length new_type_names)
   419          length new_type_names)
   419       end) atoms);
   420       end) atoms);
   420 
   421 
   426     val pt_perm_compose_rev = PureThy.get_thm thy2 "pt_perm_compose_rev";
   427     val pt_perm_compose_rev = PureThy.get_thm thy2 "pt_perm_compose_rev";
   427     val dj_perm_perm_forget = PureThy.get_thm thy2 "dj_perm_perm_forget";
   428     val dj_perm_perm_forget = PureThy.get_thm thy2 "dj_perm_perm_forget";
   428 
   429 
   429     fun composition_instance name1 name2 thy =
   430     fun composition_instance name1 name2 thy =
   430       let
   431       let
   431         val name1' = Sign.base_name name1;
   432         val cp_class = cp_class_of thy name1 name2;
   432         val name2' = Sign.base_name name2;
   433         val pt_class =
   433         val cp_class = Sign.intern_class thy ("cp_" ^ name1' ^ "_" ^ name2');
   434           if name1 = name2 then [pt_class_of thy name1]
       
   435           else [];
   434         val permT1 = mk_permT (Type (name1, []));
   436         val permT1 = mk_permT (Type (name1, []));
   435         val permT2 = mk_permT (Type (name2, []));
   437         val permT2 = mk_permT (Type (name2, []));
   436         val augment = map_type_tfree
   438         val Ts = map body_type perm_types;
   437           (fn (x, S) => TFree (x, cp_class :: S));
   439         val cp_inst = cp_inst_of thy name1 name2;
   438         val Ts = map (augment o body_type) perm_types;
       
   439         val cp_inst = PureThy.get_thm thy ("cp_" ^ name1' ^ "_" ^ name2' ^ "_inst");
       
   440         val simps = simpset_of thy addsimps (perm_fun_def ::
   440         val simps = simpset_of thy addsimps (perm_fun_def ::
   441           (if name1 <> name2 then
   441           (if name1 <> name2 then
   442              let val dj = PureThy.get_thm thy ("dj_" ^ name2' ^ "_" ^ name1')
   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
   443              in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end
   444            else
   444            else
   445              let
   445              let
   446                val at_inst = PureThy.get_thm thy ("at_" ^ name1' ^ "_inst");
   446                val at_inst = at_inst_of thy name1;
   447                val pt_inst = PureThy.get_thm thy ("pt_" ^ name1' ^ "_inst");
   447                val pt_inst = pt_inst_of thy name1;
   448              in
   448              in
   449                [cp_inst RS cp1 RS sym,
   449                [cp_inst RS cp1 RS sym,
   450                 at_inst RS (pt_inst RS pt_perm_compose) 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]
   451                 at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym]
   452             end))
   452             end))
   453         val thms = split_conj_thm (Goal.prove_global thy [] []
   453         val thms = split_conj_thm (Goal.prove_global thy [] []
       
   454           (augment_sort thy (cp_class :: pt_class)
   454             (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   455             (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
   455               (map (fn ((s, T), x) =>
   456               (map (fn ((s, T), x) =>
   456                   let
   457                   let
   457                     val pi1 = Free ("pi1", permT1);
   458                     val pi1 = Free ("pi1", permT1);
   458                     val pi2 = Free ("pi2", permT2);
   459                     val pi2 = Free ("pi2", permT2);
   461                     val perm3 = Const ("Nominal.perm", permT1 --> permT2 --> permT2)
   462                     val perm3 = Const ("Nominal.perm", permT1 --> permT2 --> permT2)
   462                   in HOLogic.mk_eq
   463                   in HOLogic.mk_eq
   463                     (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)),
   464                     (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)),
   464                      perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T)))
   465                      perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T)))
   465                   end)
   466                   end)
   466                 (perm_names ~~ Ts ~~ perm_indnames))))
   467                 (perm_names ~~ Ts ~~ perm_indnames)))))
   467           (fn _ => EVERY [indtac induction perm_indnames 1,
   468           (fn _ => EVERY [indtac induction perm_indnames 1,
   468              ALLGOALS (asm_full_simp_tac simps)]))
   469              ALLGOALS (asm_full_simp_tac simps)]))
   469       in
   470       in
   470         foldl (fn ((s, tvs), thy) => AxClass.prove_arity
   471         foldl (fn ((s, tvs), thy) => AxClass.prove_arity
   471             (s, replicate (length tvs) (cp_class :: classes), [cp_class])
   472             (s, replicate (length tvs) (cp_class :: pt_class), [cp_class])
   472             (Class.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy)
   473             (Class.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy)
   473           thy (full_new_type_names' ~~ tyvars)
   474           thy (full_new_type_names' ~~ tyvars)
   474       end;
   475       end;
   475 
   476 
   476     val (perm_thmss,thy3) = thy2 |>
   477     val (perm_thmss,thy3) = thy2 |>
   477       fold (fn name1 => fold (composition_instance name1) atoms) atoms |>
   478       fold (fn name1 => fold (composition_instance name1) atoms) atoms |>
   478       curry (Library.foldr (fn ((i, (tyname, args, _)), thy) =>
   479       fold (fn atom => fn thy =>
   479         AxClass.prove_arity (tyname, replicate (length args) classes, classes)
   480         let val pt_name = pt_class_of thy atom
   480         (Class.intro_classes_tac [] THEN REPEAT (EVERY
   481         in
   481            [resolve_tac perm_empty_thms 1,
   482           fold (fn (i, (tyname, args, _)) => AxClass.prove_arity
   482             resolve_tac perm_append_thms 1,
   483               (tyname, replicate (length args) [pt_name], [pt_name])
   483             resolve_tac perm_eq_thms 1, assume_tac 1])) thy))
   484               (EVERY
   484         (List.take (descr, length new_type_names)) |>
   485                 [Class.intro_classes_tac [],
       
   486                  resolve_tac perm_empty_thms 1,
       
   487                  resolve_tac perm_append_thms 1,
       
   488                  resolve_tac perm_eq_thms 1, assume_tac 1]))
       
   489             (List.take (descr, length new_type_names)) thy
       
   490         end) atoms |>
   485       PureThy.add_thmss
   491       PureThy.add_thmss
   486         [((space_implode "_" new_type_names ^ "_unfolded_perm_eq",
   492         [((space_implode "_" new_type_names ^ "_unfolded_perm_eq",
   487           unfolded_perm_eq_thms), [Simplifier.simp_add]),
   493           unfolded_perm_eq_thms), [Simplifier.simp_add]),
   488          ((space_implode "_" new_type_names ^ "_perm_empty",
   494          ((space_implode "_" new_type_names ^ "_perm_empty",
   489           perm_empty_thms), [Simplifier.simp_add]),
   495           perm_empty_thms), [Simplifier.simp_add]),
   511            | _ => ([], dtf))
   517            | _ => ([], dtf))
   512       | strip_option (DtType ("fun", [dt, DtType ("Nominal.noption", [dt'])])) =
   518       | strip_option (DtType ("fun", [dt, DtType ("Nominal.noption", [dt'])])) =
   513           apfst (cons dt) (strip_option dt')
   519           apfst (cons dt) (strip_option dt')
   514       | strip_option dt = ([], dt);
   520       | strip_option dt = ([], dt);
   515 
   521 
   516     val dt_atomTs = distinct op = (map (typ_of_dtyp descr sorts')
   522     val dt_atomTs = distinct op = (map (typ_of_dtyp descr sorts)
   517       (List.concat (map (fn (_, (_, _, cs)) => List.concat
   523       (List.concat (map (fn (_, (_, _, cs)) => List.concat
   518         (map (List.concat o map (fst o strip_option) o snd) cs)) descr)));
   524         (map (List.concat o map (fst o strip_option) o snd) cs)) descr)));
       
   525     val dt_atoms = map (fst o dest_Type) dt_atomTs;
   519 
   526 
   520     fun make_intr s T (cname, cargs) =
   527     fun make_intr s T (cname, cargs) =
   521       let
   528       let
   522         fun mk_prem (dt, (j, j', prems, ts)) =
   529         fun mk_prem (dt, (j, j', prems, ts)) =
   523           let
   530           let
   524             val (dts, dt') = strip_option dt;
   531             val (dts, dt') = strip_option dt;
   525             val (dts', dt'') = strip_dtyp dt';
   532             val (dts', dt'') = strip_dtyp dt';
   526             val Ts = map (typ_of_dtyp descr sorts') dts;
   533             val Ts = map (typ_of_dtyp descr sorts) dts;
   527             val Us = map (typ_of_dtyp descr sorts') dts';
   534             val Us = map (typ_of_dtyp descr sorts) dts';
   528             val T = typ_of_dtyp descr sorts' dt'';
   535             val T = typ_of_dtyp descr sorts dt'';
   529             val free = mk_Free "x" (Us ---> T) j;
   536             val free = mk_Free "x" (Us ---> T) j;
   530             val free' = app_bnds free (length Us);
   537             val free' = app_bnds free (length Us);
   531             fun mk_abs_fun (T, (i, t)) =
   538             fun mk_abs_fun (T, (i, t)) =
   532               let val U = fastype_of t
   539               let val U = fastype_of t
   533               in (i + 1, Const ("Nominal.abs_fun", [T, U, T] --->
   540               in (i + 1, Const ("Nominal.abs_fun", [T, U, T] --->
   578       (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x)
   585       (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x)
   579       (perm_indnames ~~ descr);
   586       (perm_indnames ~~ descr);
   580 
   587 
   581     fun mk_perm_closed name = map (fn th => standard (th RS mp))
   588     fun mk_perm_closed name = map (fn th => standard (th RS mp))
   582       (List.take (split_conj_thm (Goal.prove_global thy4 [] []
   589       (List.take (split_conj_thm (Goal.prove_global thy4 [] []
   583         (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map
   590         (augment_sort thy4
   584            (fn ((s, T), x) =>
   591           (pt_class_of thy4 name :: map (cp_class_of thy4 name) (dt_atoms \ name))
   585               let
   592           (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map
   586                 val T = map_type_tfree
   593             (fn ((s, T), x) =>
   587                   (fn (s, cs) => TFree (s, cs union cp_classes)) T;
   594                let
   588                 val S = Const (s, T --> HOLogic.boolT);
   595                  val S = Const (s, T --> HOLogic.boolT);
   589                 val permT = mk_permT (Type (name, []))
   596                  val permT = mk_permT (Type (name, []))
   590               in HOLogic.mk_imp (S $ Free (x, T),
   597                in HOLogic.mk_imp (S $ Free (x, T),
   591                 S $ (Const ("Nominal.perm", permT --> T --> T) $
   598                  S $ (Const ("Nominal.perm", permT --> T --> T) $
   592                   Free ("pi", permT) $ Free (x, T)))
   599                    Free ("pi", permT) $ Free (x, T)))
   593               end) (rep_set_names'' ~~ recTs' ~~ perm_indnames'))))
   600                end) (rep_set_names'' ~~ recTs' ~~ perm_indnames')))))
   594         (fn _ => EVERY (* CU: added perm_fun_def in the final tactic in order to deal with funs *)
   601         (fn _ => EVERY
   595            [indtac rep_induct [] 1,
   602            [indtac rep_induct [] 1,
   596             ALLGOALS (simp_tac (simpset_of thy4 addsimps
   603             ALLGOALS (simp_tac (simpset_of thy4 addsimps
   597               (symmetric perm_fun_def :: abs_perm))),
   604               (symmetric perm_fun_def :: abs_perm))),
   598             ALLGOALS (resolve_tac rep_intrs
   605             ALLGOALS (resolve_tac rep_intrs THEN_ALL_NEW assume_tac)])),
   599                THEN_ALL_NEW (asm_full_simp_tac (simpset_of thy4 addsimps [perm_fun_def])))])),
       
   600         length new_type_names));
   606         length new_type_names));
   601 
   607 
   602     val perm_closed_thmss = map mk_perm_closed atoms;
   608     val perm_closed_thmss = map mk_perm_closed atoms;
   603 
   609 
   604     (**** typedef ****)
   610     (**** typedef ****)
   615               QUIET_BREADTH_FIRST (has_fewer_prems 1)
   621               QUIET_BREADTH_FIRST (has_fewer_prems 1)
   616               (resolve_tac rep_intrs 1)) thy |> (fn ((_, r), thy) =>
   622               (resolve_tac rep_intrs 1)) thy |> (fn ((_, r), thy) =>
   617         let
   623         let
   618           val permT = mk_permT (TFree (Name.variant tvs "'a", HOLogic.typeS));
   624           val permT = mk_permT (TFree (Name.variant tvs "'a", HOLogic.typeS));
   619           val pi = Free ("pi", permT);
   625           val pi = Free ("pi", permT);
   620           val tvs' = map (fn s => TFree (s, the (AList.lookup op = sorts' s))) tvs;
   626           val tvs' = map (fn s => TFree (s, the (AList.lookup op = sorts s))) tvs;
   621           val T = Type (Sign.intern_type thy name, tvs');
   627           val T = Type (Sign.intern_type thy name, tvs');
   622         in apfst (pair r o hd)
   628         in apfst (pair r o hd)
   623           (PureThy.add_defs_unchecked true [(("prm_" ^ name ^ "_def", Logic.mk_equals
   629           (PureThy.add_defs_unchecked true [(("prm_" ^ name ^ "_def", Logic.mk_equals
   624             (Const ("Nominal.perm", permT --> T --> T) $ pi $ Free ("x", T),
   630             (Const ("Nominal.perm", permT --> T --> T) $ pi $ Free ("x", T),
   625              Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $
   631              Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $
   639 
   645 
   640     (** prove that new types are in class pt_<name> **)
   646     (** prove that new types are in class pt_<name> **)
   641 
   647 
   642     val _ = warning "prove that new types are in class pt_<name> ...";
   648     val _ = warning "prove that new types are in class pt_<name> ...";
   643 
   649 
   644     fun pt_instance ((class, atom), perm_closed_thms) =
   650     fun pt_instance (atom, perm_closed_thms) =
   645       fold (fn ((((((Abs_inverse, Rep_inverse), Rep),
   651       fold (fn ((((((Abs_inverse, Rep_inverse), Rep),
   646         perm_def), name), tvs), perm_closed) => fn thy =>
   652         perm_def), name), tvs), perm_closed) => fn thy =>
   647           AxClass.prove_arity
   653           let
       
   654             val pt_class = pt_class_of thy atom;
       
   655             val cp_sort = map (cp_class_of thy atom) (dt_atoms \ atom)
       
   656           in AxClass.prove_arity
   648             (Sign.intern_type thy name,
   657             (Sign.intern_type thy name,
   649               replicate (length tvs) (classes @ cp_classes), [class])
   658               replicate (length tvs) (pt_class :: cp_sort), [pt_class])
   650             (EVERY [Class.intro_classes_tac [],
   659             (EVERY [Class.intro_classes_tac [],
   651               rewrite_goals_tac [perm_def],
   660               rewrite_goals_tac [perm_def],
   652               asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1,
   661               asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1,
   653               asm_full_simp_tac (simpset_of thy addsimps
   662               asm_full_simp_tac (simpset_of thy addsimps
   654                 [Rep RS perm_closed RS Abs_inverse]) 1,
   663                 [Rep RS perm_closed RS Abs_inverse]) 1,
   655               asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy
   664               asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy
   656                 ("pt_" ^ Sign.base_name atom ^ "3")]) 1]) thy)
   665                 ("pt_" ^ Sign.base_name atom ^ "3")]) 1]) thy
       
   666           end)
   657         (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~
   667         (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~
   658            new_type_names ~~ tyvars ~~ perm_closed_thms);
   668            new_type_names ~~ tyvars ~~ perm_closed_thms);
   659 
   669 
   660 
   670 
   661     (** prove that new types are in class cp_<name1>_<name2> **)
   671     (** prove that new types are in class cp_<name1>_<name2> **)
   662 
   672 
   663     val _ = warning "prove that new types are in class cp_<name1>_<name2> ...";
   673     val _ = warning "prove that new types are in class cp_<name1>_<name2> ...";
   664 
   674 
   665     fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy =
   675     fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy =
   666       let
   676       let
   667         val name = "cp_" ^ Sign.base_name atom1 ^ "_" ^ Sign.base_name atom2;
   677         val cp_class = cp_class_of thy atom1 atom2;
   668         val class = Sign.intern_class thy name;
   678         val sort =
   669         val cp1' = PureThy.get_thm thy (name ^ "_inst") RS cp1
   679           pt_class_of thy atom1 :: map (cp_class_of thy atom1) (dt_atoms \ atom1) @
       
   680           (if atom1 = atom2 then [cp_class_of thy atom1 atom1] else
       
   681            pt_class_of thy atom2 :: map (cp_class_of thy atom2) (dt_atoms \ atom2));
       
   682         val cp1' = cp_inst_of thy atom1 atom2 RS cp1
   670       in fold (fn ((((((Abs_inverse, Rep),
   683       in fold (fn ((((((Abs_inverse, Rep),
   671         perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy =>
   684         perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy =>
   672           AxClass.prove_arity
   685           AxClass.prove_arity
   673             (Sign.intern_type thy name,
   686             (Sign.intern_type thy name,
   674               replicate (length tvs) (classes @ cp_classes), [class])
   687               replicate (length tvs) sort, [cp_class])
   675             (EVERY [Class.intro_classes_tac [],
   688             (EVERY [Class.intro_classes_tac [],
   676               rewrite_goals_tac [perm_def],
   689               rewrite_goals_tac [perm_def],
   677               asm_full_simp_tac (simpset_of thy addsimps
   690               asm_full_simp_tac (simpset_of thy addsimps
   678                 ((Rep RS perm_closed1 RS Abs_inverse) ::
   691                 ((Rep RS perm_closed1 RS Abs_inverse) ::
   679                  (if atom1 = atom2 then []
   692                  (if atom1 = atom2 then []
   685            tyvars ~~ perm_closed_thms1 ~~ perm_closed_thms2) thy
   698            tyvars ~~ perm_closed_thms1 ~~ perm_closed_thms2) thy
   686       end;
   699       end;
   687 
   700 
   688     val thy7 = fold (fn x => fn thy => thy |>
   701     val thy7 = fold (fn x => fn thy => thy |>
   689       pt_instance x |>
   702       pt_instance x |>
   690       fold (cp_instance (apfst snd x)) (atoms ~~ perm_closed_thmss))
   703       fold (cp_instance x) (atoms ~~ perm_closed_thmss))
   691         (classes ~~ atoms ~~ perm_closed_thmss) thy6;
   704         (atoms ~~ perm_closed_thmss) thy6;
   692 
   705 
   693     (**** constructors ****)
   706     (**** constructors ****)
   694 
   707 
   695     fun mk_abs_fun (x, t) =
   708     fun mk_abs_fun (x, t) =
   696       let
   709       let
   739 
   752 
   740     val pdescr = map (fn ((i, (s, dts, constrs)), (_, idxss)) => (i, (s, dts,
   753     val pdescr = map (fn ((i, (s, dts, constrs)), (_, idxss)) => (i, (s, dts,
   741       map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs))
   754       map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs))
   742         (constrs ~~ idxss)))) (descr'' ~~ ndescr);
   755         (constrs ~~ idxss)))) (descr'' ~~ ndescr);
   743 
   756 
   744     fun nth_dtyp' i = typ_of_dtyp descr'' sorts' (DtRec i);
   757     fun nth_dtyp' i = typ_of_dtyp descr'' sorts (DtRec i);
   745 
   758 
   746     val rep_names = map (fn s =>
   759     val rep_names = map (fn s =>
   747       Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names;
   760       Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names;
   748     val abs_names = map (fn s =>
   761     val abs_names = map (fn s =>
   749       Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names;
   762       Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names;
   750 
   763 
   751     val recTs = get_rec_types descr'' sorts';
   764     val recTs = get_rec_types descr'' sorts;
   752     val newTs' = Library.take (length new_type_names, recTs');
   765     val newTs' = Library.take (length new_type_names, recTs');
   753     val newTs = Library.take (length new_type_names, recTs);
   766     val newTs = Library.take (length new_type_names, recTs);
   754 
   767 
   755     val full_new_type_names = map (Sign.full_name thy) new_type_names;
   768     val full_new_type_names = map (Sign.full_name thy) new_type_names;
   756 
   769 
   757     fun make_constr_def tname T T' ((thy, defs, eqns),
   770     fun make_constr_def tname T T' ((thy, defs, eqns),
   758         (((cname_rep, _), (cname, cargs)), (cname', mx))) =
   771         (((cname_rep, _), (cname, cargs)), (cname', mx))) =
   759       let
   772       let
   760         fun constr_arg ((dts, dt), (j, l_args, r_args)) =
   773         fun constr_arg ((dts, dt), (j, l_args, r_args)) =
   761           let
   774           let
   762             val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp descr'' sorts' dt) i)
   775             val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp descr'' sorts dt) i)
   763               (dts ~~ (j upto j + length dts - 1))
   776               (dts ~~ (j upto j + length dts - 1))
   764             val x = mk_Free "x" (typ_of_dtyp descr'' sorts' dt) (j + length dts)
   777             val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
   765           in
   778           in
   766             (j + length dts + 1,
   779             (j + length dts + 1,
   767              xs @ x :: l_args,
   780              xs @ x :: l_args,
   768              foldr mk_abs_fun
   781              foldr mk_abs_fun
   769                (case dt of
   782                (case dt of
   770                   DtRec k => if k < length new_type_names then
   783                   DtRec k => if k < length new_type_names then
   771                       Const (List.nth (rep_names, k), typ_of_dtyp descr'' sorts' dt -->
   784                       Const (List.nth (rep_names, k), typ_of_dtyp descr'' sorts dt -->
   772                         typ_of_dtyp descr sorts' dt) $ x
   785                         typ_of_dtyp descr sorts dt) $ x
   773                     else error "nested recursion not (yet) supported"
   786                     else error "nested recursion not (yet) supported"
   774                 | _ => x) xs :: r_args)
   787                 | _ => x) xs :: r_args)
   775           end
   788           end
   776 
   789 
   777         val (_, l_args, r_args) = foldr constr_arg (1, [], []) cargs;
   790         val (_, l_args, r_args) = foldr constr_arg (1, [], []) cargs;
   832         val _ $ (_ $ (Rep $ x)) = Logic.unvarify (prop_of th);
   845         val _ $ (_ $ (Rep $ x)) = Logic.unvarify (prop_of th);
   833         val Type ("fun", [T, U]) = fastype_of Rep;
   846         val Type ("fun", [T, U]) = fastype_of Rep;
   834         val permT = mk_permT (Type (atom, []));
   847         val permT = mk_permT (Type (atom, []));
   835         val pi = Free ("pi", permT);
   848         val pi = Free ("pi", permT);
   836       in
   849       in
   837         Goal.prove_global thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq
   850         Goal.prove_global thy8 [] []
   838             (Const ("Nominal.perm", permT --> U --> U) $ pi $ (Rep $ x),
   851           (augment_sort thy8
   839              Rep $ (Const ("Nominal.perm", permT --> T --> T) $ pi $ x))))
   852             (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom))
       
   853             (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
   854               (Const ("Nominal.perm", permT --> U --> U) $ pi $ (Rep $ x),
       
   855                Rep $ (Const ("Nominal.perm", permT --> T --> T) $ pi $ x)))))
   840           (fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @
   856           (fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @
   841             perm_closed_thms @ Rep_thms)) 1)
   857             perm_closed_thms @ Rep_thms)) 1)
   842       end) Rep_thms;
   858       end) Rep_thms;
   843 
   859 
   844     val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm
   860     val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm
   845       (atoms ~~ perm_closed_thmss));
   861       (atoms ~~ perm_closed_thmss));
   846 
   862 
   847     (* prove distinctness theorems *)
   863     (* prove distinctness theorems *)
   848 
   864 
   849     val distinct_props = DatatypeProp.make_distincts descr' sorts';
   865     val distinct_props = DatatypeProp.make_distincts descr' sorts;
   850     val dist_rewrites = map2 (fn rep_thms => fn dist_lemma =>
   866     val dist_rewrites = map2 (fn rep_thms => fn dist_lemma =>
   851       dist_lemma :: rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])
   867       dist_lemma :: rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])
   852         constr_rep_thmss dist_lemmas;
   868         constr_rep_thmss dist_lemmas;
   853 
   869 
   854     fun prove_distinct_thms _ (_, []) = []
   870     fun prove_distinct_thms _ (_, []) = []
   879             let val T = fastype_of t
   895             let val T = fastype_of t
   880             in Const ("Nominal.perm", permT --> T --> T) $ pi $ t end;
   896             in Const ("Nominal.perm", permT --> T --> T) $ pi $ t end;
   881 
   897 
   882           fun constr_arg ((dts, dt), (j, l_args, r_args)) =
   898           fun constr_arg ((dts, dt), (j, l_args, r_args)) =
   883             let
   899             let
   884               val Ts = map (typ_of_dtyp descr'' sorts') dts;
   900               val Ts = map (typ_of_dtyp descr'' sorts) dts;
   885               val xs = map (fn (T, i) => mk_Free "x" T i)
   901               val xs = map (fn (T, i) => mk_Free "x" T i)
   886                 (Ts ~~ (j upto j + length dts - 1))
   902                 (Ts ~~ (j upto j + length dts - 1))
   887               val x = mk_Free "x" (typ_of_dtyp descr'' sorts' dt) (j + length dts)
   903               val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
   888             in
   904             in
   889               (j + length dts + 1,
   905               (j + length dts + 1,
   890                xs @ x :: l_args,
   906                xs @ x :: l_args,
   891                map perm (xs @ [x]) @ r_args)
   907                map perm (xs @ [x]) @ r_args)
   892             end
   908             end
   893 
   909 
   894           val (_, l_args, r_args) = foldr constr_arg (1, [], []) dts;
   910           val (_, l_args, r_args) = foldr constr_arg (1, [], []) dts;
   895           val c = Const (cname, map fastype_of l_args ---> T)
   911           val c = Const (cname, map fastype_of l_args ---> T)
   896         in
   912         in
   897           Goal.prove_global thy8 [] []
   913           Goal.prove_global thy8 [] []
   898             (HOLogic.mk_Trueprop (HOLogic.mk_eq
   914             (augment_sort thy8
   899               (perm (list_comb (c, l_args)), list_comb (c, r_args))))
   915               (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom))
       
   916               (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
   917                 (perm (list_comb (c, l_args)), list_comb (c, r_args)))))
   900             (fn _ => EVERY
   918             (fn _ => EVERY
   901               [simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1,
   919               [simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1,
   902                simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @
   920                simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @
   903                  constr_defs @ perm_closed_thms)) 1,
   921                  constr_defs @ perm_closed_thms)) 1,
   904                TRY (simp_tac (HOL_basic_ss addsimps
   922                TRY (simp_tac (HOL_basic_ss addsimps
   913 
   931 
   914     val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms;
   932     val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms;
   915     val alpha = PureThy.get_thms thy8 "alpha";
   933     val alpha = PureThy.get_thms thy8 "alpha";
   916     val abs_fresh = PureThy.get_thms thy8 "abs_fresh";
   934     val abs_fresh = PureThy.get_thms thy8 "abs_fresh";
   917 
   935 
       
   936     val pt_cp_sort =
       
   937       map (pt_class_of thy8) dt_atoms @
       
   938       maps (fn s => map (cp_class_of thy8 s) (dt_atoms \ s)) dt_atoms;
       
   939 
   918     val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
   940     val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
   919       let val T = nth_dtyp' i
   941       let val T = nth_dtyp' i
   920       in List.mapPartial (fn ((cname, dts), constr_rep_thm) =>
   942       in List.mapPartial (fn ((cname, dts), constr_rep_thm) =>
   921         if null dts then NONE else SOME
   943         if null dts then NONE else SOME
   922         let
   944         let
   923           val cname = Sign.intern_const thy8
   945           val cname = Sign.intern_const thy8
   924             (NameSpace.append tname (Sign.base_name cname));
   946             (NameSpace.append tname (Sign.base_name cname));
   925 
   947 
   926           fun make_inj ((dts, dt), (j, args1, args2, eqs)) =
   948           fun make_inj ((dts, dt), (j, args1, args2, eqs)) =
   927             let
   949             let
   928               val Ts_idx = map (typ_of_dtyp descr'' sorts') dts ~~ (j upto j + length dts - 1);
   950               val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1);
   929               val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
   951               val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
   930               val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx;
   952               val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx;
   931               val x = mk_Free "x" (typ_of_dtyp descr'' sorts' dt) (j + length dts);
   953               val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts);
   932               val y = mk_Free "y" (typ_of_dtyp descr'' sorts' dt) (j + length dts)
   954               val y = mk_Free "y" (typ_of_dtyp descr'' sorts dt) (j + length dts)
   933             in
   955             in
   934               (j + length dts + 1,
   956               (j + length dts + 1,
   935                xs @ (x :: args1), ys @ (y :: args2),
   957                xs @ (x :: args1), ys @ (y :: args2),
   936                HOLogic.mk_eq
   958                HOLogic.mk_eq
   937                  (foldr mk_abs_fun x xs, foldr mk_abs_fun y ys) :: eqs)
   959                  (foldr mk_abs_fun x xs, foldr mk_abs_fun y ys) :: eqs)
   939 
   961 
   940           val (_, args1, args2, eqs) = foldr make_inj (1, [], [], []) dts;
   962           val (_, args1, args2, eqs) = foldr make_inj (1, [], [], []) dts;
   941           val Ts = map fastype_of args1;
   963           val Ts = map fastype_of args1;
   942           val c = Const (cname, Ts ---> T)
   964           val c = Const (cname, Ts ---> T)
   943         in
   965         in
   944           Goal.prove_global thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq
   966           Goal.prove_global thy8 [] []
   945               (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)),
   967             (augment_sort thy8 pt_cp_sort
   946                foldr1 HOLogic.mk_conj eqs)))
   968               (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
   969                 (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)),
       
   970                  foldr1 HOLogic.mk_conj eqs))))
   947             (fn _ => EVERY
   971             (fn _ => EVERY
   948                [asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm ::
   972                [asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm ::
   949                   rep_inject_thms')) 1,
   973                   rep_inject_thms')) 1,
   950                 TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def ::
   974                 TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def ::
   951                   alpha @ abs_perm @ abs_fresh @ rep_inject_thms @
   975                   alpha @ abs_perm @ abs_fresh @ rep_inject_thms @
   952                   perm_rep_perm_thms)) 1),
   976                   perm_rep_perm_thms)) 1)])
   953                 TRY (asm_full_simp_tac (HOL_basic_ss addsimps (perm_fun_def ::
       
   954                   @{thm expand_fun_eq} :: rep_inject_thms @ perm_rep_perm_thms)) 1)])
       
   955         end) (constrs ~~ constr_rep_thms)
   977         end) (constrs ~~ constr_rep_thms)
   956       end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
   978       end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
   957 
   979 
   958     (** equations for support and freshness **)
   980     (** equations for support and freshness **)
   959 
   981 
   966             (NameSpace.append tname (Sign.base_name cname));
   988             (NameSpace.append tname (Sign.base_name cname));
   967           val atomT = Type (atom, []);
   989           val atomT = Type (atom, []);
   968 
   990 
   969           fun process_constr ((dts, dt), (j, args1, args2)) =
   991           fun process_constr ((dts, dt), (j, args1, args2)) =
   970             let
   992             let
   971               val Ts_idx = map (typ_of_dtyp descr'' sorts') dts ~~ (j upto j + length dts - 1);
   993               val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1);
   972               val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
   994               val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
   973               val x = mk_Free "x" (typ_of_dtyp descr'' sorts' dt) (j + length dts)
   995               val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
   974             in
   996             in
   975               (j + length dts + 1,
   997               (j + length dts + 1,
   976                xs @ (x :: args1), foldr mk_abs_fun x xs :: args2)
   998                xs @ (x :: args1), foldr mk_abs_fun x xs :: args2)
   977             end;
   999             end;
   978 
  1000 
   981           val c = list_comb (Const (cname, Ts ---> T), args1);
  1003           val c = list_comb (Const (cname, Ts ---> T), args1);
   982           fun supp t =
  1004           fun supp t =
   983             Const ("Nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t;
  1005             Const ("Nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t;
   984           fun fresh t = fresh_const atomT (fastype_of t) $ Free ("a", atomT) $ t;
  1006           fun fresh t = fresh_const atomT (fastype_of t) $ Free ("a", atomT) $ t;
   985           val supp_thm = Goal.prove_global thy8 [] []
  1007           val supp_thm = Goal.prove_global thy8 [] []
       
  1008             (augment_sort thy8 pt_cp_sort
   986               (HOLogic.mk_Trueprop (HOLogic.mk_eq
  1009               (HOLogic.mk_Trueprop (HOLogic.mk_eq
   987                 (supp c,
  1010                 (supp c,
   988                  if null dts then Const ("{}", HOLogic.mk_setT atomT)
  1011                  if null dts then Const ("{}", HOLogic.mk_setT atomT)
   989                  else foldr1 (HOLogic.mk_binop "op Un") (map supp args2))))
  1012                  else foldr1 (HOLogic.mk_binop "op Un") (map supp args2)))))
   990             (fn _ =>
  1013             (fn _ =>
   991               simp_tac (HOL_basic_ss addsimps (supp_def ::
  1014               simp_tac (HOL_basic_ss addsimps (supp_def ::
   992                  Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un ::
  1015                  Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un ::
   993                  symmetric empty_def :: finite_emptyI :: simp_thms @
  1016                  symmetric empty_def :: finite_emptyI :: simp_thms @
   994                  abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1)
  1017                  abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1)
   995         in
  1018         in
   996           (supp_thm,
  1019           (supp_thm,
   997            Goal.prove_global thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq
  1020            Goal.prove_global thy8 [] [] (augment_sort thy8 pt_cp_sort
   998               (fresh c,
  1021              (HOLogic.mk_Trueprop (HOLogic.mk_eq
   999                if null dts then HOLogic.true_const
  1022                (fresh c,
  1000                else foldr1 HOLogic.mk_conj (map fresh args2))))
  1023                 if null dts then HOLogic.true_const
       
  1024                 else foldr1 HOLogic.mk_conj (map fresh args2)))))
  1001              (fn _ =>
  1025              (fn _ =>
  1002                simp_tac (HOL_ss addsimps [Un_iff, empty_iff, fresh_def, supp_thm]) 1))
  1026                simp_tac (HOL_ss addsimps [Un_iff, empty_iff, fresh_def, supp_thm]) 1))
  1003         end) atoms) constrs)
  1027         end) atoms) constrs)
  1004       end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps')));
  1028       end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps')));
  1005 
  1029 
  1036     val indrule_lemma' = cterm_instantiate
  1060     val indrule_lemma' = cterm_instantiate
  1037       (map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma;
  1061       (map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma;
  1038 
  1062 
  1039     val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms;
  1063     val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms;
  1040 
  1064 
  1041     val dt_induct_prop = DatatypeProp.make_ind descr' sorts';
  1065     val dt_induct_prop = DatatypeProp.make_ind descr' sorts;
  1042     val dt_induct = Goal.prove_global thy8 []
  1066     val dt_induct = Goal.prove_global thy8 []
  1043       (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
  1067       (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
  1044       (fn {prems, ...} => EVERY
  1068       (fn {prems, ...} => EVERY
  1045         [rtac indrule_lemma' 1,
  1069         [rtac indrule_lemma' 1,
  1046          (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
  1070          (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
  1062     val supp_atm = PureThy.get_thms thy8 "supp_atm";
  1086     val supp_atm = PureThy.get_thms thy8 "supp_atm";
  1063 
  1087 
  1064     val finite_supp_thms = map (fn atom =>
  1088     val finite_supp_thms = map (fn atom =>
  1065       let val atomT = Type (atom, [])
  1089       let val atomT = Type (atom, [])
  1066       in map standard (List.take
  1090       in map standard (List.take
  1067         (split_conj_thm (Goal.prove_global thy8 [] [] (HOLogic.mk_Trueprop
  1091         (split_conj_thm (Goal.prove_global thy8 [] []
  1068            (foldr1 HOLogic.mk_conj (map (fn (s, T) =>
  1092            (augment_sort thy8 (fs_class_of thy8 atom :: pt_cp_sort)
  1069              Const ("Finite_Set.finite", HOLogic.mk_setT atomT --> HOLogic.boolT) $
  1093              (HOLogic.mk_Trueprop
  1070                (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T)))
  1094                (foldr1 HOLogic.mk_conj (map (fn (s, T) =>
  1071                (indnames ~~ recTs))))
  1095                  Const ("Finite_Set.finite", HOLogic.mk_setT atomT --> HOLogic.boolT) $
       
  1096                    (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T)))
       
  1097                    (indnames ~~ recTs)))))
  1072            (fn _ => indtac dt_induct indnames 1 THEN
  1098            (fn _ => indtac dt_induct indnames 1 THEN
  1073             ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps
  1099             ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps
  1074               (abs_supp @ supp_atm @
  1100               (abs_supp @ supp_atm @
  1075                PureThy.get_thms thy8 ("fs_" ^ Sign.base_name atom ^ "1") @
  1101                PureThy.get_thms thy8 ("fs_" ^ Sign.base_name atom ^ "1") @
  1076                List.concat supp_thms))))),
  1102                List.concat supp_thms))))),
  1094       DatatypeAux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>>
  1120       DatatypeAux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>>
  1095       DatatypeAux.store_thmss "inject" new_type_names inject_thms ||>>
  1121       DatatypeAux.store_thmss "inject" new_type_names inject_thms ||>>
  1096       DatatypeAux.store_thmss "supp" new_type_names supp_thms ||>>
  1122       DatatypeAux.store_thmss "supp" new_type_names supp_thms ||>>
  1097       DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||>
  1123       DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||>
  1098       fold (fn (atom, ths) => fn thy =>
  1124       fold (fn (atom, ths) => fn thy =>
  1099         let val class = Sign.intern_class thy ("fs_" ^ Sign.base_name atom)
  1125         let val class = fs_class_of thy atom
  1100         in fold (fn T => AxClass.prove_arity
  1126         in fold (fn T => AxClass.prove_arity
  1101             (fst (dest_Type T),
  1127             (fst (dest_Type T),
  1102               replicate (length sorts) [class], [class])
  1128               replicate (length sorts) (class :: pt_cp_sort), [class])
  1103             (Class.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy
  1129             (Class.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy
  1104         end) (atoms ~~ finite_supp_thms);
  1130         end) (atoms ~~ finite_supp_thms);
  1105 
  1131 
  1106     (**** strong induction theorem ****)
  1132     (**** strong induction theorem ****)
  1107 
  1133 
  1108     val pnames = if length descr'' = 1 then ["P"]
  1134     val pnames = if length descr'' = 1 then ["P"]
  1109       else map (fn i => "P" ^ string_of_int i) (1 upto length descr'');
  1135       else map (fn i => "P" ^ string_of_int i) (1 upto length descr'');
  1110     val ind_sort = if null dt_atomTs then HOLogic.typeS
  1136     val ind_sort = if null dt_atomTs then HOLogic.typeS
  1111       else Sign.certify_sort thy9 (map (fn T => Sign.intern_class thy9 ("fs_" ^
  1137       else Sign.certify_sort thy9 (map (fs_class_of thy9) dt_atoms);
  1112         Sign.base_name (fst (dest_Type T)))) dt_atomTs);
       
  1113     val fsT = TFree ("'n", ind_sort);
  1138     val fsT = TFree ("'n", ind_sort);
  1114     val fsT' = TFree ("'n", HOLogic.typeS);
  1139     val fsT' = TFree ("'n", HOLogic.typeS);
  1115 
  1140 
  1116     val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T)))
  1141     val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T)))
  1117       (DatatypeProp.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs);
  1142       (DatatypeProp.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs);
  1132           mk_fresh2 (p :: xss) yss;
  1157           mk_fresh2 (p :: xss) yss;
  1133 
  1158 
  1134     fun make_ind_prem fsT f k T ((cname, cargs), idxs) =
  1159     fun make_ind_prem fsT f k T ((cname, cargs), idxs) =
  1135       let
  1160       let
  1136         val recs = List.filter is_rec_type cargs;
  1161         val recs = List.filter is_rec_type cargs;
  1137         val Ts = map (typ_of_dtyp descr'' sorts') cargs;
  1162         val Ts = map (typ_of_dtyp descr'' sorts) cargs;
  1138         val recTs' = map (typ_of_dtyp descr'' sorts') recs;
  1163         val recTs' = map (typ_of_dtyp descr'' sorts) recs;
  1139         val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts);
  1164         val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts);
  1140         val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
  1165         val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
  1141         val frees = tnames ~~ Ts;
  1166         val frees = tnames ~~ Ts;
  1142         val frees' = partition_cargs idxs frees;
  1167         val frees' = partition_cargs idxs frees;
  1143         val z = (Name.variant tnames "z", fsT);
  1168         val z = (Name.variant tnames "z", fsT);
  1197         HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $
  1222         HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $
  1198           fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1))
  1223           fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1))
  1199             (Free (tname, T))))
  1224             (Free (tname, T))))
  1200         (descr'' ~~ recTs ~~ tnames)));
  1225         (descr'' ~~ recTs ~~ tnames)));
  1201 
  1226 
  1202     val fin_set_supp = map (fn Type (s, _) =>
  1227     val fin_set_supp = map (fn s =>
  1203       PureThy.get_thm thy9 ("at_" ^ Sign.base_name s ^ "_inst") RS
  1228       at_inst_of thy9 s RS at_fin_set_supp) dt_atoms;
  1204         at_fin_set_supp) dt_atomTs;
  1229     val fin_set_fresh = map (fn s =>
  1205     val fin_set_fresh = map (fn Type (s, _) =>
  1230       at_inst_of thy9 s RS at_fin_set_fresh) dt_atoms;
  1206       PureThy.get_thm thy9 ("at_" ^ Sign.base_name s ^ "_inst") RS
       
  1207         at_fin_set_fresh) dt_atomTs;
       
  1208     val pt1_atoms = map (fn Type (s, _) =>
  1231     val pt1_atoms = map (fn Type (s, _) =>
  1209       PureThy.get_thm thy9 ("pt_" ^ Sign.base_name s ^ "1")) dt_atomTs;
  1232       PureThy.get_thm thy9 ("pt_" ^ Sign.base_name s ^ "1")) dt_atomTs;
  1210     val pt2_atoms = map (fn Type (s, _) =>
  1233     val pt2_atoms = map (fn Type (s, _) =>
  1211       PureThy.get_thm thy9 ("pt_" ^ Sign.base_name s ^ "2") RS sym) dt_atomTs;
  1234       PureThy.get_thm thy9 ("pt_" ^ Sign.base_name s ^ "2") RS sym) dt_atomTs;
  1212     val exists_fresh' = PureThy.get_thms thy9 "exists_fresh'";
  1235     val exists_fresh' = PureThy.get_thms thy9 "exists_fresh'";
  1246       in
  1269       in
  1247         Drule.instantiate' []
  1270         Drule.instantiate' []
  1248           [SOME (cterm_of thy a), NONE, SOME (cterm_of thy b)] th
  1271           [SOME (cterm_of thy a), NONE, SOME (cterm_of thy b)] th
  1249       end;
  1272       end;
  1250 
  1273 
       
  1274     val fs_cp_sort =
       
  1275       map (fs_class_of thy9) dt_atoms @
       
  1276       maps (fn s => map (cp_class_of thy9 s) (dt_atoms \ s)) dt_atoms;
       
  1277 
  1251     (**********************************************************************
  1278     (**********************************************************************
  1252       The subgoals occurring in the proof of induct_aux have the
  1279       The subgoals occurring in the proof of induct_aux have the
  1253       following parameters:
  1280       following parameters:
  1254 
  1281 
  1255         x_1 ... x_k p_1 ... p_m z
  1282         x_1 ... x_k p_1 ... p_m z
  1261         z   : freshness context
  1288         z   : freshness context
  1262     ***********************************************************************)
  1289     ***********************************************************************)
  1263 
  1290 
  1264     val _ = warning "proving strong induction theorem ...";
  1291     val _ = warning "proving strong induction theorem ...";
  1265 
  1292 
  1266     val induct_aux = Goal.prove_global thy9 [] ind_prems' ind_concl' (fn {prems, context} =>
  1293     val induct_aux = Goal.prove_global thy9 []
       
  1294         (map (augment_sort thy9 fs_cp_sort) ind_prems')
       
  1295         (augment_sort thy9 fs_cp_sort ind_concl') (fn {prems, context} =>
  1267       let
  1296       let
  1268         val (prems1, prems2) = chop (length dt_atomTs) prems;
  1297         val (prems1, prems2) = chop (length dt_atomTs) prems;
  1269         val ind_ss2 = HOL_ss addsimps
  1298         val ind_ss2 = HOL_ss addsimps
  1270           finite_Diff :: abs_fresh @ abs_supp @ fs_atoms;
  1299           finite_Diff :: abs_fresh @ abs_supp @ fs_atoms;
  1271         val ind_ss1 = ind_ss2 addsimps fresh_left @ calc_atm @
  1300         val ind_ss1 = ind_ss2 addsimps fresh_left @ calc_atm @
  1274           abs_perm @ calc_atm @ perm_swap;
  1303           abs_perm @ calc_atm @ perm_swap;
  1275         val ind_ss4 = HOL_basic_ss addsimps fresh_left @ prems1 @
  1304         val ind_ss4 = HOL_basic_ss addsimps fresh_left @ prems1 @
  1276           fin_set_fresh @ calc_atm;
  1305           fin_set_fresh @ calc_atm;
  1277         val ind_ss5 = HOL_basic_ss addsimps pt1_atoms;
  1306         val ind_ss5 = HOL_basic_ss addsimps pt1_atoms;
  1278         val ind_ss6 = HOL_basic_ss addsimps flat perm_simps';
  1307         val ind_ss6 = HOL_basic_ss addsimps flat perm_simps';
  1279         val th = Goal.prove context [] [] aux_ind_concl
  1308         val th = Goal.prove context [] []
       
  1309           (augment_sort thy9 fs_cp_sort aux_ind_concl)
  1280           (fn {context = context1, ...} =>
  1310           (fn {context = context1, ...} =>
  1281              EVERY (indtac dt_induct tnames 1 ::
  1311              EVERY (indtac dt_induct tnames 1 ::
  1282                maps (fn ((_, (_, _, constrs)), (_, constrs')) =>
  1312                maps (fn ((_, (_, _, constrs)), (_, constrs')) =>
  1283                  map (fn ((cname, cargs), is) =>
  1313                  map (fn ((cname, cargs), is) =>
  1284                    REPEAT (rtac allI 1) THEN
  1314                    REPEAT (rtac allI 1) THEN
  1349            REPEAT (etac allE 1),
  1379            REPEAT (etac allE 1),
  1350            REPEAT (TRY (rtac conjI 1) THEN asm_full_simp_tac ind_ss5 1)]
  1380            REPEAT (TRY (rtac conjI 1) THEN asm_full_simp_tac ind_ss5 1)]
  1351       end);
  1381       end);
  1352 
  1382 
  1353     val induct_aux' = Thm.instantiate ([],
  1383     val induct_aux' = Thm.instantiate ([],
  1354       map (fn (s, T) =>
  1384       map (fn (s, v as Var (_, T)) =>
  1355         let val pT = TVar (("'n", 0), HOLogic.typeS) --> T --> HOLogic.boolT
  1385         (cterm_of thy9 v, cterm_of thy9 (Free (s, T))))
  1356         in (cterm_of thy9 (Var ((s, 0), pT)), cterm_of thy9 (Free (s, pT))) end)
  1386           (pnames ~~ map head_of (HOLogic.dest_conj
  1357           (pnames ~~ recTs) @
  1387              (HOLogic.dest_Trueprop (concl_of induct_aux)))) @
  1358       map (fn (_, f) =>
  1388       map (fn (_, f) =>
  1359         let val f' = Logic.varify f
  1389         let val f' = Logic.varify f
  1360         in (cterm_of thy9 f',
  1390         in (cterm_of thy9 f',
  1361           cterm_of thy9 (Const ("Nominal.supp", fastype_of f')))
  1391           cterm_of thy9 (Const ("Nominal.supp", fastype_of f')))
  1362         end) fresh_fs) induct_aux;
  1392         end) fresh_fs) induct_aux;
  1363 
  1393 
  1364     val induct = Goal.prove_global thy9 [] ind_prems ind_concl
  1394     val induct = Goal.prove_global thy9 []
       
  1395       (map (augment_sort thy9 fs_cp_sort) ind_prems)
       
  1396       (augment_sort thy9 fs_cp_sort ind_concl)
  1365       (fn {prems, ...} => EVERY
  1397       (fn {prems, ...} => EVERY
  1366          [rtac induct_aux' 1,
  1398          [rtac induct_aux' 1,
  1367           REPEAT (resolve_tac fs_atoms 1),
  1399           REPEAT (resolve_tac fs_atoms 1),
  1368           REPEAT ((resolve_tac prems THEN_ALL_NEW
  1400           REPEAT ((resolve_tac prems THEN_ALL_NEW
  1369             (etac meta_spec ORELSE' full_simp_tac (HOL_basic_ss addsimps [fresh_def]))) 1)])
  1401             (etac meta_spec ORELSE' full_simp_tac (HOL_basic_ss addsimps [fresh_def]))) 1)])
  1378 
  1410 
  1379     val _ = warning "defining recursion combinator ...";
  1411     val _ = warning "defining recursion combinator ...";
  1380 
  1412 
  1381     val used = foldr add_typ_tfree_names [] recTs;
  1413     val used = foldr add_typ_tfree_names [] recTs;
  1382 
  1414 
  1383     val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts' used;
  1415     val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts used;
  1384 
  1416 
  1385     val rec_sort = if null dt_atomTs then HOLogic.typeS else
  1417     val rec_sort = if null dt_atomTs then HOLogic.typeS else
  1386       let val names = map (Sign.base_name o fst o dest_Type) dt_atomTs
  1418       Sign.certify_sort thy10 pt_cp_sort;
  1387       in Sign.certify_sort thy10 (map (Sign.intern_class thy10)
       
  1388         (map (fn s => "pt_" ^ s) names @
       
  1389          List.concat (map (fn s => List.mapPartial (fn s' =>
       
  1390            if s = s' then NONE
       
  1391            else SOME ("cp_" ^ s ^ "_" ^ s')) names) names)))
       
  1392       end;
       
  1393 
  1419 
  1394     val rec_result_Ts = map (fn TFree (s, _) => TFree (s, rec_sort)) rec_result_Ts';
  1420     val rec_result_Ts = map (fn TFree (s, _) => TFree (s, rec_sort)) rec_result_Ts';
  1395     val rec_fn_Ts = map (typ_subst_atomic (rec_result_Ts' ~~ rec_result_Ts)) rec_fn_Ts';
  1421     val rec_fn_Ts = map (typ_subst_atomic (rec_result_Ts' ~~ rec_result_Ts)) rec_fn_Ts';
  1396 
  1422 
  1397     val rec_set_Ts = map (fn (T1, T2) =>
  1423     val rec_set_Ts = map (fn (T1, T2) =>
  1427     val rec_ctxt = Free ("z", fsT');
  1453     val rec_ctxt = Free ("z", fsT');
  1428 
  1454 
  1429     fun make_rec_intr T p rec_set ((rec_intr_ts, rec_prems, rec_prems',
  1455     fun make_rec_intr T p rec_set ((rec_intr_ts, rec_prems, rec_prems',
  1430           rec_eq_prems, l), ((cname, cargs), idxs)) =
  1456           rec_eq_prems, l), ((cname, cargs), idxs)) =
  1431       let
  1457       let
  1432         val Ts = map (typ_of_dtyp descr'' sorts') cargs;
  1458         val Ts = map (typ_of_dtyp descr'' sorts) cargs;
  1433         val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts;
  1459         val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts;
  1434         val frees' = partition_cargs idxs frees;
  1460         val frees' = partition_cargs idxs frees;
  1435         val binders = List.concat (map fst frees');
  1461         val binders = List.concat (map fst frees');
  1436         val atomTs = distinct op = (maps (map snd o fst) frees');
  1462         val atomTs = distinct op = (maps (map snd o fst) frees');
  1437         val recs = List.mapPartial
  1463         val recs = List.mapPartial
  1506           in
  1532           in
  1507             (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y)
  1533             (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y)
  1508           end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs));
  1534           end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs));
  1509         val ths = map (fn th => standard (th RS mp)) (split_conj_thm
  1535         val ths = map (fn th => standard (th RS mp)) (split_conj_thm
  1510           (Goal.prove_global thy11 [] []
  1536           (Goal.prove_global thy11 [] []
  1511             (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps)))
  1537             (augment_sort thy1 pt_cp_sort
       
  1538               (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps))))
  1512             (fn _ => rtac rec_induct 1 THEN REPEAT
  1539             (fn _ => rtac rec_induct 1 THEN REPEAT
  1513                (simp_tac (Simplifier.theory_context thy11 HOL_basic_ss
  1540                (simp_tac (Simplifier.theory_context thy11 HOL_basic_ss
  1514                   addsimps flat perm_simps'
  1541                   addsimps flat perm_simps'
  1515                   addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
  1542                   addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
  1516                 (resolve_tac rec_intrs THEN_ALL_NEW
  1543                 (resolve_tac rec_intrs THEN_ALL_NEW
  1517                  asm_simp_tac (HOL_ss addsimps (fresh_bij @ perm_bij))) 1))))
  1544                  asm_simp_tac (HOL_ss addsimps (fresh_bij @ perm_bij))) 1))))
  1518         val ths' = map (fn ((P, Q), th) =>
  1545         val ths' = map (fn ((P, Q), th) =>
  1519           Goal.prove_global thy11 [] []
  1546           Goal.prove_global thy11 [] []
  1520             (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P))
  1547             (augment_sort thy1 pt_cp_sort
       
  1548               (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P)))
  1521             (fn _ => dtac (Thm.instantiate ([],
  1549             (fn _ => dtac (Thm.instantiate ([],
  1522                  [(cterm_of thy11 (Var (("pi", 0), permT)),
  1550                  [(cterm_of thy11 (Var (("pi", 0), permT)),
  1523                    cterm_of thy11 (Const ("List.rev", permT --> permT) $ pi))]) th) 1 THEN
  1551                    cterm_of thy11 (Const ("List.rev", permT --> permT) $ pi))]) th) 1 THEN
  1524                NominalPermeq.perm_simp_tac HOL_ss 1)) (ps ~~ ths)
  1552                NominalPermeq.perm_simp_tac HOL_ss 1)) (ps ~~ ths)
  1525       in (ths, ths') end) dt_atomTs);
  1553       in (ths, ths') end) dt_atomTs);
  1535         val fins = map (fn (f, T) => HOLogic.mk_Trueprop
  1563         val fins = map (fn (f, T) => HOLogic.mk_Trueprop
  1536           (finite $ (Const ("Nominal.supp", T --> aset) $ f)))
  1564           (finite $ (Const ("Nominal.supp", T --> aset) $ f)))
  1537             (rec_fns ~~ rec_fn_Ts)
  1565             (rec_fns ~~ rec_fn_Ts)
  1538       in
  1566       in
  1539         map (fn th => standard (th RS mp)) (split_conj_thm
  1567         map (fn th => standard (th RS mp)) (split_conj_thm
  1540           (Goal.prove_global thy11 [] fins
  1568           (Goal.prove_global thy11 []
  1541             (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
  1569             (map (augment_sort thy11 fs_cp_sort) fins)
  1542               (map (fn (((T, U), R), i) =>
  1570             (augment_sort thy11 fs_cp_sort
  1543                  let
  1571               (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
  1544                    val x = Free ("x" ^ string_of_int i, T);
  1572                 (map (fn (((T, U), R), i) =>
  1545                    val y = Free ("y" ^ string_of_int i, U)
  1573                    let
  1546                  in
  1574                      val x = Free ("x" ^ string_of_int i, T);
  1547                    HOLogic.mk_imp (R $ x $ y,
  1575                      val y = Free ("y" ^ string_of_int i, U)
  1548                      finite $ (Const ("Nominal.supp", U --> aset) $ y))
  1576                    in
  1549                  end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ (1 upto length recTs)))))
  1577                      HOLogic.mk_imp (R $ x $ y,
       
  1578                        finite $ (Const ("Nominal.supp", U --> aset) $ y))
       
  1579                    end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~
       
  1580                      (1 upto length recTs))))))
  1550             (fn {prems = fins, ...} =>
  1581             (fn {prems = fins, ...} =>
  1551               (rtac rec_induct THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT
  1582               (rtac rec_induct THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT
  1552                (NominalPermeq.finite_guess_tac (HOL_ss addsimps [fs_name]) 1))))
  1583                (NominalPermeq.finite_guess_tac (HOL_ss addsimps [fs_name]) 1))))
  1553       end) dt_atomTs;
  1584       end) dt_atomTs;
  1554 
  1585 
  1557     val finite_premss = map (fn aT =>
  1588     val finite_premss = map (fn aT =>
  1558       map (fn (f, T) => HOLogic.mk_Trueprop
  1589       map (fn (f, T) => HOLogic.mk_Trueprop
  1559         (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
  1590         (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
  1560            (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ f)))
  1591            (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ f)))
  1561            (rec_fns ~~ rec_fn_Ts)) dt_atomTs;
  1592            (rec_fns ~~ rec_fn_Ts)) dt_atomTs;
       
  1593 
       
  1594     val rec_fns' = map (augment_sort thy11 fs_cp_sort) rec_fns;
  1562 
  1595 
  1563     val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) =>
  1596     val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) =>
  1564       let
  1597       let
  1565         val name = Sign.base_name (fst (dest_Type aT));
  1598         val name = Sign.base_name (fst (dest_Type aT));
  1566         val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1");
  1599         val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1");
  1568         val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop
  1601         val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop
  1569           (fresh_const aT fT $ a $ f)) (rec_fns ~~ rec_fn_Ts)
  1602           (fresh_const aT fT $ a $ f)) (rec_fns ~~ rec_fn_Ts)
  1570       in
  1603       in
  1571         map (fn (((T, U), R), eqvt_th) =>
  1604         map (fn (((T, U), R), eqvt_th) =>
  1572           let
  1605           let
  1573             val x = Free ("x", T);
  1606             val x = Free ("x", augment_sort_typ thy11 fs_cp_sort T);
  1574             val y = Free ("y", U);
  1607             val y = Free ("y", U);
  1575             val y' = Free ("y'", U)
  1608             val y' = Free ("y'", U)
  1576           in
  1609           in
  1577             standard (Goal.prove (ProofContext.init thy11) [] (finite_prems @
  1610             standard (Goal.prove (ProofContext.init thy11) []
  1578                 [HOLogic.mk_Trueprop (R $ x $ y),
  1611               (map (augment_sort thy11 fs_cp_sort)
  1579                  HOLogic.mk_Trueprop (HOLogic.mk_all ("y'", U,
  1612                 (finite_prems @
  1580                    HOLogic.mk_imp (R $ x $ y', HOLogic.mk_eq (y', y)))),
  1613                    [HOLogic.mk_Trueprop (R $ x $ y),
  1581                  HOLogic.mk_Trueprop (fresh_const aT T $ a $ x)] @
  1614                     HOLogic.mk_Trueprop (HOLogic.mk_all ("y'", U,
  1582               freshs)
  1615                       HOLogic.mk_imp (R $ x $ y', HOLogic.mk_eq (y', y)))),
       
  1616                     HOLogic.mk_Trueprop (fresh_const aT T $ a $ x)] @
       
  1617                  freshs))
  1583               (HOLogic.mk_Trueprop (fresh_const aT U $ a $ y))
  1618               (HOLogic.mk_Trueprop (fresh_const aT U $ a $ y))
  1584               (fn {prems, context} =>
  1619               (fn {prems, context} =>
  1585                  let
  1620                  let
  1586                    val (finite_prems, rec_prem :: unique_prem ::
  1621                    val (finite_prems, rec_prem :: unique_prem ::
  1587                      fresh_prems) = chop (length finite_prems) prems;
  1622                      fresh_prems) = chop (length finite_prems) prems;
  1588                    val unique_prem' = unique_prem RS spec RS mp;
  1623                    val unique_prem' = unique_prem RS spec RS mp;
  1589                    val unique = [unique_prem', unique_prem' RS sym] MRS trans;
  1624                    val unique = [unique_prem', unique_prem' RS sym] MRS trans;
  1590                    val _ $ (_ $ (_ $ S $ _)) $ _ = prop_of supports_fresh;
  1625                    val _ $ (_ $ (_ $ S $ _)) $ _ = prop_of supports_fresh;
  1591                    val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns)
  1626                    val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns')
  1592                  in EVERY
  1627                  in EVERY
  1593                    [rtac (Drule.cterm_instantiate
  1628                    [rtac (Drule.cterm_instantiate
  1594                       [(cterm_of thy11 S,
  1629                       [(cterm_of thy11 S,
  1595                         cterm_of thy11 (Const ("Nominal.supp",
  1630                         cterm_of thy11 (Const ("Nominal.supp",
  1596                           fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))]
  1631                           fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))]
  1629         Const ("Ex1", (U --> HOLogic.boolT) --> HOLogic.boolT) $
  1664         Const ("Ex1", (U --> HOLogic.boolT) --> HOLogic.boolT) $
  1630           Abs ("y", U, R $ Free x $ Bound 0))
  1665           Abs ("y", U, R $ Free x $ Bound 0))
  1631       (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets);
  1666       (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets);
  1632 
  1667 
  1633     val induct_aux_rec = Drule.cterm_instantiate
  1668     val induct_aux_rec = Drule.cterm_instantiate
  1634       (map (pairself (cterm_of thy11))
  1669       (map (pairself (cterm_of thy11) o apsnd (augment_sort thy11 fs_cp_sort))
  1635          (map (fn (aT, f) => (Logic.varify f, Abs ("z", HOLogic.unitT,
  1670          (map (fn (aT, f) => (Logic.varify f, Abs ("z", HOLogic.unitT,
  1636             Const ("Nominal.supp", fun_tupleT --> HOLogic.mk_setT aT) $ fun_tuple)))
  1671             Const ("Nominal.supp", fun_tupleT --> HOLogic.mk_setT aT) $ fun_tuple)))
  1637               fresh_fs @
  1672               fresh_fs @
  1638           map (fn (((P, T), (x, U)), Q) =>
  1673           map (fn (((P, T), (x, U)), Q) =>
  1639            (Var ((P, 0), Logic.varifyT (fsT' --> T --> HOLogic.boolT)),
  1674            (Var ((P, 0), Logic.varifyT (fsT' --> T --> HOLogic.boolT)),
  1666         (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
  1701         (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
  1667            (Const ("Nominal.supp", fsT' --> HOLogic.mk_setT aT) $ rec_ctxt))) dt_atomTs;
  1702            (Const ("Nominal.supp", fsT' --> HOLogic.mk_setT aT) $ rec_ctxt))) dt_atomTs;
  1668 
  1703 
  1669     val rec_unique_thms = split_conj_thm (Goal.prove
  1704     val rec_unique_thms = split_conj_thm (Goal.prove
  1670       (ProofContext.init thy11) (map fst rec_unique_frees)
  1705       (ProofContext.init thy11) (map fst rec_unique_frees)
  1671       (List.concat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems')
  1706       (map (augment_sort thy11 fs_cp_sort)
  1672       (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj rec_unique_concls))
  1707         (List.concat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems'))
       
  1708       (augment_sort thy11 fs_cp_sort
       
  1709         (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj rec_unique_concls)))
  1673       (fn {prems, context} =>
  1710       (fn {prems, context} =>
  1674          let
  1711          let
  1675            val k = length rec_fns;
  1712            val k = length rec_fns;
  1676            val (finite_thss, ths1) = fold_map (fn T => fn xs =>
  1713            val (finite_thss, ths1) = fold_map (fn T => fn xs =>
  1677              apfst (pair T) (chop k xs)) dt_atomTs prems;
  1714              apfst (pair T) (chop k xs)) dt_atomTs prems;
  1678            val (finite_ctxt_ths, ths2) = chop (length dt_atomTs) ths1;
  1715            val (finite_ctxt_ths, ths2) = chop (length dt_atomTs) ths1;
  1679            val (P_ind_ths, fcbs) = chop k ths2;
  1716            val (P_ind_ths, fcbs) = chop k ths2;
  1680            val P_ths = map (fn th => th RS mp) (split_conj_thm
  1717            val P_ths = map (fn th => th RS mp) (split_conj_thm
  1681              (Goal.prove context
  1718              (Goal.prove context
  1682                (map fst (rec_unique_frees'' @ rec_unique_frees')) []
  1719                (map fst (rec_unique_frees'' @ rec_unique_frees')) []
  1683                (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
  1720                (augment_sort thy11 fs_cp_sort
  1684                   (map (fn (((x, y), S), P) => HOLogic.mk_imp
  1721                  (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
  1685                     (S $ Free x $ Free y, P $ (Free y)))
  1722                     (map (fn (((x, y), S), P) => HOLogic.mk_imp
  1686                       (rec_unique_frees'' ~~ rec_unique_frees' ~~ rec_sets ~~ rec_preds))))
  1723                       (S $ Free x $ Free y, P $ (Free y)))
       
  1724                         (rec_unique_frees'' ~~ rec_unique_frees' ~~
       
  1725                            rec_sets ~~ rec_preds)))))
  1687                (fn _ =>
  1726                (fn _ =>
  1688                   rtac rec_induct 1 THEN
  1727                   rtac rec_induct 1 THEN
  1689                   REPEAT ((resolve_tac P_ind_ths THEN_ALL_NEW assume_tac) 1))));
  1728                   REPEAT ((resolve_tac P_ind_ths THEN_ALL_NEW assume_tac) 1))));
  1690            val rec_fin_supp_thms' = map
  1729            val rec_fin_supp_thms' = map
  1691              (fn (ths, (T, fin_ths)) => (T, map (curry op MRS fin_ths) ths))
  1730              (fn (ths, (T, fin_ths)) => (T, map (curry op MRS fin_ths) ths))
  1720                     val boundsr = List.concat (map fst cargsr');
  1759                     val boundsr = List.concat (map fst cargsr');
  1721                     val (params1, _ :: params2) =
  1760                     val (params1, _ :: params2) =
  1722                       chop (length params div 2) (map term_of params);
  1761                       chop (length params div 2) (map term_of params);
  1723                     val params' = params1 @ params2;
  1762                     val params' = params1 @ params2;
  1724                     val rec_prems = filter (fn th => case prop_of th of
  1763                     val rec_prems = filter (fn th => case prop_of th of
  1725                       _ $ (S $ _ $ _) => S mem rec_sets | _ => false) prems';
  1764                         _ $ p => (case head_of p of
       
  1765                           Const (s, _) => s mem rec_set_names
       
  1766                         | _ => false)
       
  1767                       | _ => false) prems';
  1726                     val fresh_prems = filter (fn th => case prop_of th of
  1768                     val fresh_prems = filter (fn th => case prop_of th of
  1727                         _ $ (Const ("Nominal.fresh", _) $ _ $ _) => true
  1769                         _ $ (Const ("Nominal.fresh", _) $ _ $ _) => true
  1728                       | _ $ (Const ("Not", _) $ _) => true
  1770                       | _ $ (Const ("Not", _) $ _) => true
  1729                       | _ => false) prems';
  1771                       | _ => false) prems';
  1730                     val Ts = map fastype_of boundsl;
  1772                     val Ts = map fastype_of boundsl;
  1731 
  1773 
  1732                     val _ = warning "step 1: obtaining fresh names";
  1774                     val _ = warning "step 1: obtaining fresh names";
  1733                     val (freshs1, freshs2, context'') = fold
  1775                     val (freshs1, freshs2, context'') = fold
  1734                       (obtain_fresh_name (rec_ctxt :: rec_fns @ params')
  1776                       (obtain_fresh_name (rec_ctxt :: rec_fns' @ params')
  1735                          (List.concat (map snd finite_thss) @
  1777                          (List.concat (map snd finite_thss) @
  1736                             finite_ctxt_ths @ rec_prems)
  1778                             finite_ctxt_ths @ rec_prems)
  1737                          rec_fin_supp_thms')
  1779                          rec_fin_supp_thms')
  1738                       Ts ([], [], context');
  1780                       Ts ([], [], context');
  1739                     val pi1 = map perm_of_pair (boundsl ~~ freshs1);
  1781                     val pi1 = map perm_of_pair (boundsl ~~ freshs1);
  1800                     (** (ts, pi1^-1 o pi2 o vs) in rec_set **)
  1842                     (** (ts, pi1^-1 o pi2 o vs) in rec_set **)
  1801                     val _ = warning "step 6: (ts, pi1^-1 o pi2 o vs) in rec_set";
  1843                     val _ = warning "step 6: (ts, pi1^-1 o pi2 o vs) in rec_set";
  1802                     val rec_prems' = map (fn th =>
  1844                     val rec_prems' = map (fn th =>
  1803                       let
  1845                       let
  1804                         val _ $ (S $ x $ y) = prop_of th;
  1846                         val _ $ (S $ x $ y) = prop_of th;
  1805                         val k = find_index (equal S) rec_sets;
  1847                         val Const (s, _) = head_of S;
       
  1848                         val k = find_index (equal s) rec_set_names;
  1806                         val pi = rpi1 @ pi2;
  1849                         val pi = rpi1 @ pi2;
  1807                         fun mk_pi z = fold_rev (mk_perm []) pi z;
  1850                         fun mk_pi z = fold_rev (mk_perm []) pi z;
  1808                         fun eqvt_tac p =
  1851                         fun eqvt_tac p =
  1809                           let
  1852                           let
  1810                             val U as Type (_, [Type (_, [T, _])]) = fastype_of p;
  1853                             val U as Type (_, [Type (_, [T, _])]) = fastype_of p;
  1986         val prems' = List.concat finite_premss @ finite_ctxt_prems @
  2029         val prems' = List.concat finite_premss @ finite_ctxt_prems @
  1987           rec_prems @ rec_prems' @ map (subst_atomic ps) prems;
  2030           rec_prems @ rec_prems' @ map (subst_atomic ps) prems;
  1988         fun solve rules prems = resolve_tac rules THEN_ALL_NEW
  2031         fun solve rules prems = resolve_tac rules THEN_ALL_NEW
  1989           (resolve_tac prems THEN_ALL_NEW atac)
  2032           (resolve_tac prems THEN_ALL_NEW atac)
  1990       in
  2033       in
  1991         Goal.prove_global thy12 [] prems' concl'
  2034         Goal.prove_global thy12 []
       
  2035           (map (augment_sort thy12 fs_cp_sort) prems')
       
  2036           (augment_sort thy12 fs_cp_sort concl')
  1992           (fn {prems, ...} => EVERY
  2037           (fn {prems, ...} => EVERY
  1993             [rewrite_goals_tac reccomb_defs,
  2038             [rewrite_goals_tac reccomb_defs,
  1994              rtac the1_equality 1,
  2039              rtac the1_equality 1,
  1995              solve rec_unique_thms prems 1,
  2040              solve rec_unique_thms prems 1,
  1996              resolve_tac rec_intrs 1,
  2041              resolve_tac rec_intrs 1,
  1997              REPEAT (solve (prems @ rec_total_thms) prems 1)])
  2042              REPEAT (solve (prems @ rec_total_thms) prems 1)])
  1998       end) (rec_eq_prems ~~
  2043       end) (rec_eq_prems ~~
  1999         DatatypeProp.make_primrecs new_type_names descr' sorts' thy12);
  2044         DatatypeProp.make_primrecs new_type_names descr' sorts thy12);
  2000 
  2045 
  2001     val dt_infos = map (make_dt_info pdescr sorts induct reccomb_names rec_thms)
  2046     val dt_infos = map (make_dt_info pdescr sorts induct reccomb_names rec_thms)
  2002       ((0 upto length descr1 - 1) ~~ descr1 ~~ distinct_thms ~~ inject_thms);
  2047       ((0 upto length descr1 - 1) ~~ descr1 ~~ distinct_thms ~~ inject_thms);
  2003 
  2048 
  2004     (* FIXME: theorems are stored in database for testing only *)
  2049     (* FIXME: theorems are stored in database for testing only *)